src/Tools/isac/Build_Isac.thy
author wneuper <Walther.Neuper@jku.at>
Mon, 26 Sep 2022 10:57:53 +0200
changeset 60556 486223010ea8
parent 60551 3b5be6fae2f0
child 60557 0be383bdb883
permissions -rw-r--r--
follow up 2: Problem.adapt_to_typ on loading by CalcTree, CalcTreeTEST
     1 (*  Title:  build the Isac-Kernel: MathEngine & Knowledge
     2     Author: Walther Neuper, TU Graz, 100808
     3    (c) due to copyright terms
     4 
     5 For creating a heap image of isac see ~~/ROOT.
     6 For debugging see text at begin below, e.g. theory dependencies:
     7   ~$ evince file:///home/wneuper/.isabelle/isabisac/browser_info/Unsorted/Isac/session_graph.pdf &
     8 
     9 ATTENTION: no errors in this theory do not mean that there are no errors in Isac ..
    10 .. open theories collecting files from folders: BaseDefinitions.thy, ProgLang.thy etc.
    11 Errors are rigorously detected by isabelle build.
    12 *)
    13 
    14 theory Build_Isac
    15 imports
    16 (*  theory Know_Store imports Complex_Main
    17     $ISABELLE_ISAC/BaseDefinitions
    18       ML_file libraryC.sml
    19       ML_file theoryC.sml
    20       ML_file unparseC.sml
    21       ML_file "rule-def.sml"
    22       ML_file "thmC-def.sml"
    23       ML_file "eval-def.sml"
    24       ML_file "rewrite-order.sml"
    25       ML_file rule.sml
    26       ML_file "error-pattern-def.sml"
    27       ML_file "rule-set.sml"
    28       
    29       ML_file "store.sml"
    30       ML_file "check-unique.sml"
    31       ML_file "specification.sml"
    32       ML_file "model-pattern.sml"
    33       ML_file "problem-def.sml"
    34       ML_file "method-def.sml"
    35       ML_file "cas-def.sml"
    36       ML_file "thy-write.sml"
    37   theory BaseDefinitions imports Know_Store
    38   $ISABELLE_ISAC/BaseDefinitions
    39     ML_file termC.sml
    40     ML_file substitution.sml
    41     ML_file contextC.sml
    42     ML_file environment.sml
    43 ( ** )    "BaseDefinitions/BaseDefinitions"( **)
    44 (*
    45       theory Calc_Binop imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    46         at $ISABELLE_ISAC/ProgLang
    47         ML_file evaluate.sml
    48       theory ListC imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    49       theory Program imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    50     theory Prog_Expr imports Calc_Binop ListC Program
    51      theory Prog_Tac imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    52       theory Tactical imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    53     theory Auto_Prog imports Prog_Tac Tactical
    54   theory ProgLang imports Prog_Expr Auto_Prog
    55     at $ISABELLE_ISAC/ProgLang
    56 ( ** )    "ProgLang/ProgLang"( **)
    57 (*
    58   theory MathEngBasic imports
    59     "$ISABELLE_ISAC/ProgLang/ProgLang" "$ISABELLE_ISAC/Specify/Input_Descript"
    60     at $ISABELLE_ISAC/MathEngBasic
    61     ML_file thmC.sml
    62     ML_file problem.sml
    63     ML_file method.sml
    64     ML_file "cas-command.sml"
    65   
    66     ML_file rewrite.sml
    67   
    68     ML_file "model-def.sml"
    69     ML_file "istate-def.sml"
    70     ML_file "calc-tree-elem.sml"
    71     ML_file "pre-conds-def.sml"
    72   
    73     ML_file tactic.sml
    74     ML_file applicable.sml
    75   
    76     ML_file position.sml
    77     ML_file "ctree-basic.sml"
    78     ML_file "ctree-access.sml"
    79     ML_file "ctree-navi.sml"
    80     ML_file ctree.sml
    81   
    82     ML_file "state-steps.sml"
    83     ML_file calculation.sml
    84 ( ** )    "MathEngBasic/MathEngBasic"( **)
    85 (*
    86     theory Input_Descript imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    87     $ISABELLE_ISAC/Specify
    88   theory Specify imports "$ISABELLE_ISAC/ProgLang/ProgLang" Input_Descript
    89   $ISABELLE_ISAC/Specify
    90     ML_file formalise.sml
    91     ML_file "o-model.sml"
    92     ML_file "i-model.sml"
    93     ML_file "p-model.sml"
    94     ML_file model.sml
    95     ML_file "pre-conditions.sml"
    96     ML_file refine.sml
    97     ML_file "mstools.sml" (*..TODO review*)
    98     ML_file ptyps.sml     (*..TODO review*)
    99     ML_file "test-out.sml"
   100     ML_file "specify-step.sml"
   101     ML_file calchead.sml  (*..TODO review*)
   102     ML_file "input-calchead.sml"
   103     ML_file "step-specify.sml"
   104     ML_file specify.sml
   105 ( ** )    "Specify/Specify"( **)
   106 (*
   107   theory Interpret imports "$ISABELLE_ISAC/Specify/Specify"
   108   $ISABELLE_ISAC/Interpret
   109     ML_file istate.sml
   110     ML_file "sub-problem.sml"
   111     ML_file "thy-read.sml"
   112     ML_file "solve-step.sml"
   113     ML_file "error-pattern.sml"
   114     ML_file derive.sml
   115     ML_file "li-tool.sml"
   116     ML_file "lucas-interpreter.sml"
   117     ML_file "step-solve.sml"
   118 ( ** )    "Interpret/Interpret"( **)
   119 (*
   120   theory MathEngine imports Interpret.Interpret
   121   $ISABELLE_ISAC/MathEngine
   122     ML_file "fetch-tactics.sml"
   123     ML_file solve.sml
   124     ML_file step.sml
   125     ML_file "detail-step.sml"
   126     ML_file "mathengine-stateless.sml"
   127     ML_file messages.sml
   128     ML_file states.sml
   129 ( ** )    "MathEngine/MathEngine"( **)
   130 (*
   131   theory Test_Code imports "$ISABELLE_ISAC/MathEngine/MathEngine"
   132   $ISABELLE_ISAC/Test_Code
   133     ML_file "test-code.sml"
   134 ( ** )    "Test_Code/Test_Code"( **)
   135 (*
   136   theory BridgeLibisabelle imports "$ISABELLE_ISAC/MathEngine/MathEngine"
   137   $ISABELLE_ISAC/BridgeLibisabelle
   138     ML_file "thy-present.sml"
   139     ML_file mathml.sml   
   140     ML_file datatypes.sml
   141     ML_file "pbl-met-hierarchy.sml"
   142     ML_file "thy-hierarchy.sml" 
   143     ML_file "interface-xml.sml"
   144     ML_file interface.sml
   145 ( ** )    "BridgeLibisabelle/BridgeLibisabelle"( **)
   146 (*
   147     theory Calculation imports Calculation (*preliminary for starting Isabelle/Isac*)
   148     $ISABELLE_ISAC/BridgeJEdit
   149       ML_file parseC.sml
   150       ML_file preliminary.sml
   151   theory BridgeJEdit imports Calculation (*preliminary for starting Isabelle/Isac*)
   152   $ISABELLE_ISAC/BridgeJEdit
   153 ( ** )    "BridgeLibisabelle/BridgeLibisabelle"( **)
   154 (*
   155     theory Isac imports  "$ISABELLE_ISAC/MathEngine/MathEngine"
   156       ML_file parseC.sml
   157   theory BridgeJEdit imports Isac    
   158 ( **)    "BridgeJEdit/BridgeJEdit"  (*DEactivate after devel.of BridgeJEdit*)
   159 
   160           "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
   161 
   162 (*//-----------------------------------------------------------------------------------------\\*)
   163 (*\\-----------------------------------------------------------------------------------------//*)
   164 begin
   165 
   166 text \<open>
   167   show theory dependencies using the graph browser, 
   168   open "browser_info/HOL/Isac/session.graph"
   169   and proceed from the ancestors towards the siblings.
   170 \<close>
   171 
   172 section \<open>check presence of definitions from directories\<close>
   173 
   174 (*declare [[ML_print_depth = 999]]*)
   175 ML \<open>
   176 \<close> ML \<open>
   177 \<close> ML \<open>
   178 \<close> ML \<open> (* \<rightarrow> termC.sml*)
   179 \<close> ML \<open> (* \<rightarrow> model-pattern.sml*)
   180 \<close> ML \<open>
   181 Model_Pattern.adapt_term_to_type: Proof.context -> term -> term ;
   182 Model_Pattern.adapt_to_type: Proof.context -> Model_Pattern.single -> Model_Pattern.single
   183 \<close> ML \<open>
   184 \<close> ML \<open> (* \<rightarrow> problem.sml*)
   185 \<close> ML \<open>
   186 val id = ["univariate", "equation", "test"]
   187 \<close> ML \<open>
   188 Problem.from_store_PIDE: Proof.context -> Problem.id -> Problem.T
   189 \<close> ML \<open>
   190 \<close> ML \<open> (* \<rightarrow> refine.sml*)
   191 \<close> ML \<open>
   192 \<close> text \<open>  local
   193 refin_PIDE
   194 \<close> ML \<open>
   195 \<close> text \<open>  \<isac_test>
   196 refine_PIDE
   197 \<close> ML \<open>
   198 \<close> ML \<open>
   199 \<close> ML \<open>(*---------------------------- why "real" in pbl? ----------------------------*)
   200 Test_Tool.show_ptyps ();
   201 \<close> ML \<open>
   202 KEStore_Elems.get_pbls @{theory Poly}; (*! real ! due to Simplify :: "real => real" etc*)
   203 \<close> ML \<open>
   204 Problem.from_store_PIDE @{context} ["polynomial", "simplification"]
   205 \<close> ML \<open>
   206 val input = (["polynomial", "simplification"], 
   207   [("#Given", ["Simplify t_t"]), ("#Find", ["normalform n_n"])], 
   208   Rule_Set.empty, NONE (*cas*), 
   209   [["simplification","for_polynomials"]]) : Problem.input
   210 \<close> ML \<open>
   211 Problem.prep_input_PIDE @{theory Poly} "guh" ["math-author-1"] ["polynomial", "simplification"]
   212   input; (*! real !*)
   213 \<close> ML \<open>
   214 \<close> ML \<open>
   215 \<close> ML \<open>
   216 "~~~~~ fun prep_input_PIDE , args:"; val (thy, guh, maa, init, (pblID, dsc_dats, ev, ca, metIDs)) =
   217   (@{theory Poly}, "guh", ["math-author-1"], ["polynomial", "simplification"], input);
   218 \<close> ML \<open>
   219       fun eq f (f', _) = f = f';
   220 \<close> ML \<open>
   221       val gi = filter (eq "#Given") dsc_dats;
   222 \<close> ML \<open>
   223         val (_, gi') :: [] = (*case*) gi (*of*);
   224 \<close> ML \<open>
   225       map (Problem.split_did_PIDE o (Syntax.read_term_global thy)) gi'
   226 \<close> ML \<open>
   227 \<close> ML \<open>
   228 (*+*)Syntax.read_term_global thy "Simplify t_t" (*Simplify :: "real => real"*)
   229 \<close> ML \<open>
   230 (*+*)TermC.parse_patt thy "matches (?a = 0) e_e"
   231 (* = Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ 
   232   
   233 Var (("a", 0), "real") $ Const ("Groups.zero_class.zero", "real")) $ Free ("e_e", "bool")*)
   234 \<close> ML \<open>
   235 (*+*)TermC.parse_patt_PIDE thy "matches (?a = 0) e_e"
   236 (*t = Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ 
   237   Var (("a", 0), "?'a1") $ Const ("Groups.zero_class.zero", "?'a1")) $ Free ("e_e", "bool")*)
   238 \<close> ML \<open>
   239 \<close> ML \<open>
   240 \<close> ML \<open>
   241 \<close> ML \<open>
   242 KEStore_Elems.get_pbls @{theory Isac_Knowledge}; (*! real ! due to Simplify :: "real => real" etc*)
   243 (*val it =
   244    [Node ("empty_probl_id", [{cas = NONE, guh = "pbl_empty", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls = Empty, thy = {Pure}, where_ = []}], []),
   245     Node
   246      ("simplification",
   247       [{cas = SOME (Const ("Simplify.Simplify", "real \<Rightarrow> real") $ Free ("t_t", "real")), guh = "pbl_simp", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
   248         [("#Given", (Const ("Simplify.Term", "real \<Rightarrow> una"), Free ("t_t", "real"))), ("#Find", (Const ("Simplify.normalform", "real \<Rightarrow> una"), Free ("n_n", "real")))], prls =
   249         Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
   250         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   251           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   252           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   253           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   254           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   255           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
   256           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
   257           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
   258           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
   259           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
   260           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
   261           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify:30},
   262         where_ = []}],
   263       [Node
   264         ("polynomial",
   265          [{cas = SOME (Const ("Simplify.Simplify", "real \<Rightarrow> real") $ Free ("t_t", "real")), guh = "pbl_simp_poly", init = ["empty_probl_id"], mathauthors = [], met = [["simplification", "for_polynomials"]],
   266            ppc = [("#Given", (Const ("Simplify.Term", "real \<Rightarrow> una"), Free ("t_t", "real"))), ("#Find", (Const ("Simplify.normalform", "real \<Rightarrow> una"), Free ("n_n", "real")))], prls =
   267            Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [Eval ("Poly.is_polyexp", fn)], scr = Empty_Prog, srls = Empty}, thy =
   268            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   269              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   270              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   271              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   272              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   273              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   274              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   275              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   276              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   277              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   278              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   279              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   280              Isac.Poly:734},
   281            where_ = [Const ("Poly.is_polyexp", "real \<Rightarrow> bool") $ Free ("t_t", "real")]}],
   282          []),
   283        Node
   284         ("rational",
   285          [{cas = SOME (Const ("Simplify.Simplify", "real \<Rightarrow> real") $ Free ("t_t", "real")), guh = "pbl_simp_rat", init = ["empty_probl_id"], mathauthors = [], met = [["simplification", "of_rationals"]], ppc =
   286            [("#Given", (Const ("Simplify.Term", "real \<Rightarrow> una"), Free ("t_t", "real"))), ("#Find", (Const ("Simplify.normalform", "real \<Rightarrow> una"), Free ("n_n", "real")))], prls =
   287            Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
   288            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   289              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   290              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   291              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   292              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   293              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   294              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   295              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   296              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   297              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   298              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   299              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   300              Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational:315},
   301            where_ = [Const ("Rational.is_ratpolyexp", "real \<Rightarrow> bool") $ Free ("t_t", "real")]}],
   302          [Node
   303            ("partial_fraction",
   304             [{cas = NONE, guh = "pbl_simp_rat_partfrac", init = ["empty_probl_id"], mathauthors = [], met = [["simplification", "of_rationals", "to_partial_fraction"]], ppc =
   305               [("#Given", (Const ("Input_Descript.functionTerm", "real \<Rightarrow> una"), Free ("t_t", "real"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   306                ("#Find", (Const ("Partial_Fractions.decomposedFunction", "real \<Rightarrow> una"), Free ("p_p'''", "real")))],
   307               prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
   308               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   309                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   310                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   311                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   312                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   313                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   314                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   315                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   316                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   317                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   318                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   319                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   320                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.Partial_Fractions:136},
   321               where_ = []}],
   322             [])])]),
   323     Node
   324      ("vereinfachen",
   325       [{cas = SOME (Const ("Simplify.Vereinfache", "real \<Rightarrow> real") $ Free ("t_t", "real")), guh = "pbl_vereinfache", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
   326         [("#Given", (Const ("Simplify.Term", "real \<Rightarrow> una"), Free ("t_t", "real"))), ("#Find", (Const ("Simplify.normalform", "real \<Rightarrow> una"), Free ("n_n", "real")))], prls =
   327         Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
   328         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   329           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   330           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   331           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   332           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   333           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
   334           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
   335           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
   336           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
   337           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
   338           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
   339           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify:40},
   340         where_ = []}],
   341       [Node
   342         ("polynom",
   343          [{cas = NONE, guh = "pbl_vereinf_poly", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls = Empty, thy =
   344            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   345              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   346              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   347              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   348              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   349              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   350              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   351              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   352              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   353              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   354              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   355              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   356              Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational, Isac.PolyMinus:199},
   357            where_ = []}],
   358          [Node
   359            ("plus_minus",
   360             [{cas = SOME (Const ("Simplify.Vereinfache", "real \<Rightarrow> real") $ Free ("t_t", "real")), guh = "pbl_vereinf_poly_minus", init = ["empty_probl_id"], mathauthors = [], met =
   361               [["simplification", "for_polynomials", "with_minus"]], ppc =
   362               [("#Given", (Const ("Simplify.Term", "real \<Rightarrow> una"), Free ("t_t", "real"))), ("#Find", (Const ("Simplify.normalform", "real \<Rightarrow> una"), Free ("n_n", "real")))], prls =
   363               Repeat
   364                {calc = [], erls = Empty, errpatts = [], id = "prls_pbl_vereinf_poly", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   365                 [Eval ("Poly.is_polyexp", fn), Eval ("Prog_Expr.matchsub", fn), Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a"), Thm ("not_true", "(\<not> True) = False"),
   366                  Thm ("not_false", "(\<not> False) = True")],
   367                 scr = Empty_Prog, srls = Empty},
   368               thy =
   369               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   370                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   371                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   372                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   373                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   374                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   375                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   376                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   377                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   378                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   379                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   380                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   381                 Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational, Isac.PolyMinus:209},
   382               where_ =
   383               [Const ("Poly.is_polyexp", "real \<Rightarrow> bool") $ Free ("t_t", "real"),
   384                Const ("HOL.Not", "bool \<Rightarrow> bool") $
   385                  (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   386                    (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   387                      (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("a", 0), "?'a1") $
   388                        (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1"))) $
   389                      Free ("t_t", "?'a1")) $
   390                    (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   391                      (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   392                        (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("a", 0), "?'a1") $
   393                          (Const ("Groups.minus_class.minus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1"))) $
   394                        Free ("t_t", "?'a1")) $
   395                      (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   396                        (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   397                          (Const ("Groups.minus_class.minus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("a", 0), "?'a1") $
   398                            (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1"))) $
   399                          Free ("t_t", "?'a1")) $
   400                        (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   401                          (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("a", 0), "?'a1") $
   402                            (Const ("Groups.minus_class.minus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1"))) $
   403                          Free ("t_t", "?'a1"))))),
   404                Const ("HOL.Not", "bool \<Rightarrow> bool") $
   405                  (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   406                    (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   407                      (Const ("Groups.times_class.times", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("a", 0), "?'a1") $
   408                        (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1"))) $
   409                      Free ("t_t", "?'a1")) $
   410                    (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   411                      (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   412                        (Const ("Groups.times_class.times", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("a", 0), "?'a1") $
   413                          (Const ("Groups.minus_class.minus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1"))) $
   414                        Free ("t_t", "?'a1")) $
   415                      (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   416                        (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   417                          (Const ("Groups.times_class.times", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1")) $
   418                            Var (("a", 0), "?'a1")) $
   419                          Free ("t_t", "?'a1")) $
   420                        (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   421                          (Const ("Groups.times_class.times", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ (Const ("Groups.minus_class.minus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1")) $
   422                            Var (("a", 0), "?'a1")) $
   423                          Free ("t_t", "?'a1")))))]}],
   424             []),
   425           Node
   426            ("klammer",
   427             [{cas = SOME (Const ("Simplify.Vereinfache", "real \<Rightarrow> real") $ Free ("t_t", "real")), guh = "pbl_vereinf_poly_klammer", init = ["empty_probl_id"], mathauthors = [], met =
   428               [["simplification", "for_polynomials", "with_parentheses"]], ppc =
   429               [("#Given", (Const ("Simplify.Term", "real \<Rightarrow> una"), Free ("t_t", "real"))), ("#Find", (Const ("Simplify.normalform", "real \<Rightarrow> una"), Free ("n_n", "real")))], prls =
   430               Repeat
   431                {calc = [], erls = Empty, errpatts = [], id = "prls_pbl_vereinf_poly_klammer", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   432                 [Eval ("Poly.is_polyexp", fn), Eval ("Prog_Expr.matchsub", fn), Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a"), Thm ("not_true", "(\<not> True) = False"),
   433                  Thm ("not_false", "(\<not> False) = True")],
   434                 scr = Empty_Prog, srls = Empty},
   435               thy =
   436               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   437                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   438                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   439                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   440                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   441                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   442                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   443                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   444                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   445                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   446                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   447                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   448                 Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational, Isac.PolyMinus:219},
   449               where_ =
   450               [Const ("Poly.is_polyexp", "real \<Rightarrow> bool") $ Free ("t_t", "real"),
   451                Const ("HOL.Not", "bool \<Rightarrow> bool") $
   452                  (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   453                    (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   454                      (Const ("Groups.times_class.times", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("a", 0), "?'a1") $
   455                        (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1"))) $
   456                      Free ("t_t", "?'a1")) $
   457                    (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   458                      (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   459                        (Const ("Groups.times_class.times", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("a", 0), "?'a1") $
   460                          (Const ("Groups.minus_class.minus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1"))) $
   461                        Free ("t_t", "?'a1")) $
   462                      (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   463                        (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   464                          (Const ("Groups.times_class.times", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1")) $
   465                            Var (("a", 0), "?'a1")) $
   466                          Free ("t_t", "?'a1")) $
   467                        (Const ("Prog_Expr.matchsub", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
   468                          (Const ("Groups.times_class.times", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ (Const ("Groups.minus_class.minus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Var (("c", 0), "?'a1")) $
   469                            Var (("a", 0), "?'a1")) $
   470                          Free ("t_t", "?'a1")))))]}],
   471             []),
   472           Node
   473            ("binom_klammer",
   474             [{cas = SOME (Const ("Simplify.Vereinfache", "real \<Rightarrow> real") $ Free ("t_t", "real")), guh = "pbl_vereinf_poly_klammer_mal", init = ["empty_probl_id"], mathauthors = [], met =
   475               [["simplification", "for_polynomials", "with_parentheses_mult"]], ppc =
   476               [("#Given", (Const ("Simplify.Term", "real \<Rightarrow> una"), Free ("t_t", "real"))), ("#Find", (Const ("Simplify.normalform", "real \<Rightarrow> una"), Free ("n_n", "real")))], prls =
   477               Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [Eval ("Poly.is_polyexp", fn)], scr = Empty_Prog, srls = Empty}, thy =
   478               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   479                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   480                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   481                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   482                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   483                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   484                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   485                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   486                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   487                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   488                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   489                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   490                 Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational, Isac.PolyMinus:229},
   491               where_ = [Const ("Poly.is_polyexp", "real \<Rightarrow> bool") $ Free ("t_t", "real")]}],
   492             [])])]),
   493     Node
   494      ("probe",
   495       [{cas = NONE, guh = "pbl_probe", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls = Empty, thy =
   496         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   497           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   498           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   499           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   500           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   501           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
   502           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
   503           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
   504           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
   505           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
   506           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
   507           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational,
   508           Isac.PolyMinus:239},
   509         where_ = []}],
   510       [Node
   511         ("polynom",
   512          [{cas = SOME (Const ("PolyMinus.Probe", "bool \<Rightarrow> bool list \<Rightarrow> bool") $ Free ("e_e", "bool") $ Free ("w_w", "bool list")), guh = "pbl_probe_poly", init = ["empty_probl_id"], mathauthors = [], met =
   513            [["probe", "fuer_polynom"]], ppc =
   514            [("#Given", (Const ("PolyMinus.Pruefe", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("PolyMinus.mitWert", "bool list \<Rightarrow> tobooll"), Free ("w_w", "bool list"))),
   515             ("#Find", (Const ("PolyMinus.Geprueft", "bool \<Rightarrow> una"), Free ("p_p", "bool")))],
   516            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "prls_pbl_probe_poly", preconds = [], rew_ord = ("dummy_ord", fn), rules = [Eval ("Poly.is_polyexp", fn)], scr = Empty_Prog, srls = Empty},
   517            thy =
   518            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   519              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   520              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   521              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   522              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   523              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   524              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   525              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   526              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   527              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   528              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   529              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   530              Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational, Isac.PolyMinus:249},
   531            where_ = [Const ("Poly.is_polyexp", "real \<Rightarrow> bool") $ Free ("e_e", "real")]}],
   532          []),
   533        Node
   534         ("bruch",
   535          [{cas = SOME (Const ("PolyMinus.Probe", "bool \<Rightarrow> bool list \<Rightarrow> bool") $ Free ("e_e", "bool") $ Free ("w_w", "bool list")), guh = "pbl_probe_bruch", init = ["empty_probl_id"], mathauthors = [], met =
   536            [["probe", "fuer_bruch"]], ppc =
   537            [("#Given", (Const ("PolyMinus.Pruefe", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("PolyMinus.mitWert", "bool list \<Rightarrow> tobooll"), Free ("w_w", "bool list"))),
   538             ("#Find", (Const ("PolyMinus.Geprueft", "bool \<Rightarrow> una"), Free ("p_p", "bool")))],
   539            prls =
   540            Repeat {calc = [], erls = Empty, errpatts = [], id = "prls_pbl_probe_bruch", preconds = [], rew_ord = ("dummy_ord", fn), rules = [Eval ("Rational.is_ratpolyexp", fn)], scr = Empty_Prog, srls = Empty},
   541            thy =
   542            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   543              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   544              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   545              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   546              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   547              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   548              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   549              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   550              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   551              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   552              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   553              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   554              Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational, Isac.PolyMinus:259},
   555            where_ = [Const ("Rational.is_ratpolyexp", "real \<Rightarrow> bool") $ Free ("e_e", "real")]}],
   556          [])]),
   557     Node
   558      ("equation",
   559       [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh = "pbl_equ", init =
   560         ["empty_probl_id"], mathauthors = [], met = [], ppc =
   561         [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   562          ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   563         prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "equation_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules = [Eval ("Prog_Expr.matches", fn)], scr = Empty_Prog, srls = Empty}, thy =
   564         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   565           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   566           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   567           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   568           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   569           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
   570           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
   571           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
   572           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
   573           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
   574           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
   575           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Equation:50},
   576         where_ = [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Var (("b", 0), "?'a1")) $ Free ("e_e", "bool")]}],
   577       [Node
   578         ("univariate",
   579          [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh = "pbl_equ_univ",
   580            init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
   581            [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   582             ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   583            prls =
   584            Repeat {calc = [], erls = Empty, errpatts = [], id = "univariate_equation_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules = [Eval ("Prog_Expr.matches", fn)], scr = Empty_Prog, srls = Empty},
   585            thy =
   586            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   587              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   588              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   589              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   590              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   591              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   592              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   593              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   594              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   595              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   596              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   597              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Equation:60},
   598            where_ = [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Var (("b", 0), "?'a1")) $ Free ("e_e", "bool")]}],
   599          [Node
   600            ("rootX",
   601             [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   602               "pbl_equ_univ_root", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
   603               [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   604                ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   605               prls =
   606               Repeat
   607                {calc = [], erls = Empty, errpatts = [], id = "RootEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   608                 [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("RootEq.is_sqrtTerm_in", fn), Eval ("RootEq.is_rootTerm_in", fn),
   609                  Eval ("RootEq.is_normSqrtTerm_in", fn), Eval ("HOL.eq", fn), Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"),
   610                  Thm ("and_false", "(?a \<and> False) = False"), Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   611                 scr = Empty_Prog, srls = Empty},
   612               thy =
   613               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   614                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   615                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   616                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   617                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   618                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   619                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   620                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   621                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   622                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   623                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   624                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   625                 Isac.Poly, Isac.Root, Isac.Equation, Isac.RootEq:252},
   626               where_ =
   627               [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("RootEq.is_rootTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   628                  (Const ("RootEq.is_rootTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.rhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real"))]}],
   629             [Node
   630               ("sq",
   631                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   632                  "pbl_equ_univ_root_sq", init = ["empty_probl_id"], mathauthors = [], met = [["RootEq", "solve_sq_root_equation"]], ppc =
   633                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   634                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   635                  prls =
   636                  Repeat
   637                   {calc = [], erls = Empty, errpatts = [], id = "RootEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   638                    [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("RootEq.is_sqrtTerm_in", fn), Eval ("RootEq.is_rootTerm_in", fn),
   639                     Eval ("RootEq.is_normSqrtTerm_in", fn), Eval ("HOL.eq", fn), Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"),
   640                     Thm ("and_false", "(?a \<and> False) = False"), Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   641                    scr = Empty_Prog, srls = Empty},
   642                  thy =
   643                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
   644                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   645                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   646                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   647                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   648                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   649                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   650                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   651                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
   652                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
   653                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
   654                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
   655                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.RootEq:262},
   656                  where_ =
   657                  [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   658                     (Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   659                       (Const ("RootEq.is_sqrtTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   660                       (Const ("RootEq.is_normSqrtTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real"))) $
   661                     (Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   662                       (Const ("RootEq.is_sqrtTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.rhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   663                       (Const ("RootEq.is_normSqrtTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.rhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")))]}],
   664                [Node
   665                  ("rat",
   666                   [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   667                     "pbl_equ_univ_root_sq_rat", init = ["empty_probl_id"], mathauthors = [], met = [["RootRatEq", "elim_rootrat_equation"]], ppc =
   668                     [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   669                      ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   670                     prls =
   671                     Repeat
   672                      {calc = [], erls = Empty, errpatts = [], id = "RootRatEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   673                       [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("RootEq.is_rootTerm_in", fn),
   674                        Eval ("RootRatEq.is_rootRatAddTerm_in", fn), Eval ("HOL.eq", fn), Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"),
   675                        Thm ("and_false", "(?a \<and> False) = False"), Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   676                       scr = Empty_Prog, srls = Empty},
   677                     thy =
   678                     {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
   679                       HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   680                       HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base,
   681                       HOL.BNF_Def, HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power,
   682                       HOL.Groups_Big, HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big,
   683                       HOL.Euclidean_Division, HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer,
   684                       HOL.Lifting_Set, HOL.List, HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence,
   685                       HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick,
   686                       HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main,
   687                       HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series,
   688                       HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac,
   689                       Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine,
   690                       Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq,
   691                       Isac.RootRat, Isac.RootRatEq:108},
   692                     where_ =
   693                     [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   694                        (Const ("RootRatEq.is_rootRatAddTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   695                        (Const ("RootRatEq.is_rootRatAddTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.rhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real"))]}],
   696                   [])]),
   697              Node
   698               ("normalise",
   699                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   700                  "pbl_equ_univ_root_norm", init = ["empty_probl_id"], mathauthors = [], met = [["RootEq", "norm_sq_root_equation"]], ppc =
   701                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   702                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   703                  prls =
   704                  Repeat
   705                   {calc = [], erls = Empty, errpatts = [], id = "RootEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   706                    [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("RootEq.is_sqrtTerm_in", fn), Eval ("RootEq.is_rootTerm_in", fn),
   707                     Eval ("RootEq.is_normSqrtTerm_in", fn), Eval ("HOL.eq", fn), Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"),
   708                     Thm ("and_false", "(?a \<and> False) = False"), Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   709                    scr = Empty_Prog, srls = Empty},
   710                  thy =
   711                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
   712                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   713                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   714                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   715                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   716                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   717                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   718                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   719                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
   720                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
   721                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
   722                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
   723                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.RootEq:272},
   724                  where_ =
   725                  [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   726                     (Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   727                       (Const ("RootEq.is_sqrtTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   728                       (Const ("HOL.Not", "bool \<Rightarrow> bool") $
   729                         (Const ("RootEq.is_normSqrtTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")))) $
   730                     (Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   731                       (Const ("RootEq.is_sqrtTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.rhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   732                       (Const ("HOL.Not", "bool \<Rightarrow> bool") $
   733                         (Const ("RootEq.is_normSqrtTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.rhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real"))))]}],
   734                [])]),
   735           Node
   736            ("LINEAR",
   737             [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   738               "pbl_equ_univ_lin", init = ["empty_probl_id"], mathauthors = [], met = [["LinEq", "solve_lineq_equation"]], ppc =
   739               [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   740                ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   741               prls =
   742               Repeat
   743                {calc = [], erls = Empty, errpatts = [], id = "LinEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   744                 [Eval ("HOL.eq", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn),
   745                  Eval ("Prog_Expr.occurs_in", fn), Eval ("Prog_Expr.ident", fn), Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"),
   746                  Thm ("and_false", "(?a \<and> False) = False"), Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   747                 scr = Empty_Prog, srls = Empty},
   748               thy =
   749               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   750                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   751                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   752                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   753                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   754                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   755                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   756                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   757                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   758                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   759                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   760                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   761                 Isac.Poly, Isac.Equation, Isac.LinEq:107},
   762               where_ =
   763               [Const ("HOL.False", "bool"),
   764                Const ("HOL.Not", "bool \<Rightarrow> bool") $ (Const ("Poly.is_polyrat_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")),
   765                Const ("HOL.Not", "bool \<Rightarrow> bool") $ (Const ("Poly.is_polyrat_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.rhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")),
   766                Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Poly.has_degree_in", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   767                  Const ("Groups.one_class.one", "real"),
   768                Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Poly.has_degree_in", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Prog_Expr.rhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   769                  Const ("Groups.one_class.one", "real")]}],
   770             []),
   771           Node
   772            ("rational",
   773             [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   774               "pbl_equ_univ_rat", init = ["empty_probl_id"], mathauthors = [], met = [["RatEq", "solve_rat_equation"]], ppc =
   775               [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   776                ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   777               prls =
   778               Repeat
   779                {calc = [], erls = Empty, errpatts = [], id = "RatEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   780                 [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("RatEq.is_ratequation_in", fn), Eval ("HOL.eq", fn),
   781                  Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
   782                  Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   783                 scr = Empty_Prog, srls = Empty},
   784               thy =
   785               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   786                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   787                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   788                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   789                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   790                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   791                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   792                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   793                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   794                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   795                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   796                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   797                 Isac.Poly, Isac.GCD_Poly_ML, Isac.Equation, Isac.Rational, Isac.LinEq, Isac.RatEq:163},
   798               where_ = [Const ("RatEq.is_ratequation_in", "bool \<Rightarrow> real \<Rightarrow> bool") $ Free ("e_e", "bool") $ Free ("v_v", "real")]}],
   799             []),
   800           Node
   801            ("polynomial",
   802             [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   803               "pbl_equ_univ_poly", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
   804               [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   805                ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   806               prls =
   807               Repeat
   808                {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   809                 [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
   810                  Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
   811                  Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
   812                  Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   813                 scr = Empty_Prog, srls = Empty},
   814               thy =
   815               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
   816                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   817                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   818                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   819                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   820                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   821                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   822                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   823                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
   824                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
   825                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
   826                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
   827                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:502},
   828               where_ =
   829               [Const ("HOL.Not", "bool \<Rightarrow> bool") $ (Const ("RatEq.is_ratequation_in", "bool \<Rightarrow> real \<Rightarrow> bool") $ Free ("e_e", "bool") $ Free ("v_v", "real")),
   830                Const ("HOL.Not", "bool \<Rightarrow> bool") $ (Const ("RootEq.is_rootTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")),
   831                Const ("HOL.Not", "bool \<Rightarrow> bool") $ (Const ("RootEq.is_rootTerm_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.rhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real"))]}],
   832             [Node
   833               ("degree_0",
   834                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   835                  "pbl_equ_univ_poly_deg0", init = ["empty_probl_id"], mathauthors = [], met = [["PolyEq", "solve_d0_polyeq_equation"]], ppc =
   836                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   837                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   838                  prls =
   839                  Repeat
   840                   {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   841                    [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
   842                     Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
   843                     Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
   844                     Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   845                    scr = Empty_Prog, srls = Empty},
   846                  thy =
   847                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
   848                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   849                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   850                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   851                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   852                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   853                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   854                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   855                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
   856                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
   857                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
   858                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
   859                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:512},
   860                  where_ =
   861                  [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Const ("Groups.zero_class.zero", "?'a1")) $ Free ("e_e", "bool"),
   862                   Const ("Poly.is_poly_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real"),
   863                   Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Poly.has_degree_in", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   864                     Const ("Groups.zero_class.zero", "real")]}],
   865                []),
   866              Node
   867               ("degree_1",
   868                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   869                  "pbl_equ_univ_poly_deg1", init = ["empty_probl_id"], mathauthors = [], met = [["PolyEq", "solve_d1_polyeq_equation"]], ppc =
   870                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   871                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   872                  prls =
   873                  Repeat
   874                   {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   875                    [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
   876                     Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
   877                     Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
   878                     Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   879                    scr = Empty_Prog, srls = Empty},
   880                  thy =
   881                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
   882                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   883                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   884                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   885                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   886                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   887                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   888                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   889                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
   890                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
   891                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
   892                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
   893                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:522},
   894                  where_ =
   895                  [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Const ("Groups.zero_class.zero", "?'a1")) $ Free ("e_e", "bool"),
   896                   Const ("Poly.is_poly_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real"),
   897                   Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Poly.has_degree_in", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   898                     Const ("Groups.one_class.one", "real")]}],
   899                []),
   900              Node
   901               ("degree_2",
   902                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   903                  "pbl_equ_univ_poly_deg2", init = ["empty_probl_id"], mathauthors = [], met = [["PolyEq", "solve_d2_polyeq_equation"]], ppc =
   904                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   905                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   906                  prls =
   907                  Repeat
   908                   {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   909                    [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
   910                     Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
   911                     Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
   912                     Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   913                    scr = Empty_Prog, srls = Empty},
   914                  thy =
   915                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
   916                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   917                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
   918                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
   919                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
   920                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
   921                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
   922                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
   923                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
   924                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
   925                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
   926                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
   927                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:532},
   928                  where_ =
   929                  [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Const ("Groups.zero_class.zero", "?'a1")) $ Free ("e_e", "bool"),
   930                   Const ("Poly.is_poly_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real"),
   931                   Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Poly.has_degree_in", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
   932                     (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))]}],
   933                [Node
   934                  ("sq_only",
   935                   [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
   936                     "pbl_equ_univ_poly_deg2_sqonly", init = ["empty_probl_id"], mathauthors = [], met = [["PolyEq", "solve_d2_polyeq_sqonly_equation"]], ppc =
   937                     [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
   938                      ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
   939                     prls =
   940                     Repeat
   941                      {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
   942                       [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
   943                        Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
   944                        Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
   945                        Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
   946                       scr = Empty_Prog, srls = Empty},
   947                     thy =
   948                     {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
   949                       HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
   950                       HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base,
   951                       HOL.BNF_Def, HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power,
   952                       HOL.Groups_Big, HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big,
   953                       HOL.Euclidean_Division, HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer,
   954                       HOL.Lifting_Set, HOL.List, HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence,
   955                       HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick,
   956                       HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main,
   957                       HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series,
   958                       HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac,
   959                       Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine,
   960                       Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq,
   961                       Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:542},
   962                     where_ =
   963                     [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   964                        (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   965                          (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
   966                            (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $
   967                              (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
   968                                (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
   969                            Const ("Groups.zero_class.zero", "real")) $
   970                          Free ("e_e", "bool")) $
   971                        (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   972                          (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   973                            (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
   974                              (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $
   975                                (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $
   976                                  (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
   977                                    (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))))) $
   978                              Const ("Groups.zero_class.zero", "real")) $
   979                            Free ("e_e", "bool")) $
   980                          (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   981                            (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   982                              (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
   983                                (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
   984                                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))) $
   985                                Const ("Groups.zero_class.zero", "real")) $
   986                              Free ("e_e", "bool")) $
   987                            (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   988                              (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
   989                                (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $
   990                                  (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
   991                                    (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
   992                                Const ("Groups.zero_class.zero", "real")) $
   993                              Free ("e_e", "bool")))),
   994                      Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   995                        (Const ("HOL.Not", "bool \<Rightarrow> bool") $
   996                          (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
   997                            (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
   998                              (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $ Var (("v_", 0), "real")) $
   999                                (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1000                                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  1001                              Const ("Groups.zero_class.zero", "real")) $
  1002                            Free ("e_e", "bool"))) $
  1003                        (Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1004                          (Const ("HOL.Not", "bool \<Rightarrow> bool") $
  1005                            (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1006                              (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1007                                (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  1008                                  (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $
  1009                                    (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $ Var (("v_", 0), "real"))) $
  1010                                  (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1011                                    (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  1012                                Const ("Groups.zero_class.zero", "real")) $
  1013                              Free ("e_e", "bool"))) $
  1014                          (Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1015                            (Const ("HOL.Not", "bool \<Rightarrow> bool") $
  1016                              (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1017                                (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1018                                  (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $ Var (("v_", 0), "real")) $
  1019                                    (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("c", 0), "real") $
  1020                                      (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1021                                        (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))))) $
  1022                                  Const ("Groups.zero_class.zero", "real")) $
  1023                                Free ("e_e", "bool"))) $
  1024                            (Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1025                              (Const ("HOL.Not", "bool \<Rightarrow> bool") $
  1026                                (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1027                                  (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1028                                    (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  1029                                      (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $
  1030                                        (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $ Var (("v_", 0), "real"))) $
  1031                                      (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("c", 0), "real") $
  1032                                        (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1033                                          (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))))) $
  1034                                    Const ("Groups.zero_class.zero", "real")) $
  1035                                  Free ("e_e", "bool"))) $
  1036                              (Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1037                                (Const ("HOL.Not", "bool \<Rightarrow> bool") $
  1038                                  (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1039                                    (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1040                                      (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1041                                        (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1042                                          (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  1043                                      Const ("Groups.zero_class.zero", "real")) $
  1044                                    Free ("e_e", "bool"))) $
  1045                                (Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1046                                  (Const ("HOL.Not", "bool \<Rightarrow> bool") $
  1047                                    (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1048                                      (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1049                                        (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $ Var (("v_", 0), "real")) $
  1050                                          (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1051                                            (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  1052                                        Const ("Groups.zero_class.zero", "real")) $
  1053                                      Free ("e_e", "bool"))) $
  1054                                  (Const ("HOL.conj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1055                                    (Const ("HOL.Not", "bool \<Rightarrow> bool") $
  1056                                      (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1057                                        (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1058                                          (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1059                                            (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("c", 0), "real") $
  1060                                              (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1061                                                (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))))) $
  1062                                          Const ("Groups.zero_class.zero", "real")) $
  1063                                        Free ("e_e", "bool"))) $
  1064                                    (Const ("HOL.Not", "bool \<Rightarrow> bool") $
  1065                                      (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1066                                        (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1067                                          (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  1068                                            (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $ Var (("v_", 0), "real")) $
  1069                                            (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("c", 0), "real") $
  1070                                              (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1071                                                (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))))) $
  1072                                          Const ("Groups.zero_class.zero", "real")) $
  1073                                        Free ("e_e", "bool")))))))))]}],
  1074                   []),
  1075                 Node
  1076                  ("bdv_only",
  1077                   [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  1078                     "pbl_equ_univ_poly_deg2_bdvonly", init = ["empty_probl_id"], mathauthors = [], met = [["PolyEq", "solve_d2_polyeq_bdvonly_equation"]], ppc =
  1079                     [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1080                      ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  1081                     prls =
  1082                     Repeat
  1083                      {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1084                       [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
  1085                        Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
  1086                        Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
  1087                        Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
  1088                       scr = Empty_Prog, srls = Empty},
  1089                     thy =
  1090                     {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1091                       HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1092                       HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base,
  1093                       HOL.BNF_Def, HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power,
  1094                       HOL.Groups_Big, HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big,
  1095                       HOL.Euclidean_Division, HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer,
  1096                       HOL.Lifting_Set, HOL.List, HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence,
  1097                       HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick,
  1098                       HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main,
  1099                       HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series,
  1100                       HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac,
  1101                       Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine,
  1102                       Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq,
  1103                       Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:552},
  1104                     where_ =
  1105                     [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1106                        (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1107                          (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1108                            (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $ Var (("v_", 0), "real")) $
  1109                              (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1110                                (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  1111                            Const ("Groups.zero_class.zero", "real")) $
  1112                          Free ("e_e", "bool")) $
  1113                        (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1114                          (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1115                            (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1116                              (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1117                                (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1118                                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  1119                              Const ("Groups.zero_class.zero", "real")) $
  1120                            Free ("e_e", "bool")) $
  1121                          (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1122                            (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1123                              (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1124                                (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1125                                  (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $
  1126                                    (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1127                                      (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))))) $
  1128                                Const ("Groups.zero_class.zero", "real")) $
  1129                              Free ("e_e", "bool")) $
  1130                            (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1131                              (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1132                                (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1133                                  (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $ Var (("v_", 0), "real")) $
  1134                                    (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $
  1135                                      (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1136                                        (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))))) $
  1137                                  Const ("Groups.zero_class.zero", "real")) $
  1138                                Free ("e_e", "bool")) $
  1139                              (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1140                                (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1141                                  (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1142                                    (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1143                                      (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))) $
  1144                                    Const ("Groups.zero_class.zero", "real")) $
  1145                                  Free ("e_e", "bool")) $
  1146                                (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1147                                  (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1148                                    (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $
  1149                                      (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1150                                        (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  1151                                    Const ("Groups.zero_class.zero", "real")) $
  1152                                  Free ("e_e", "bool"))))))]}],
  1153                   []),
  1154                 Node
  1155                  ("pqFormula",
  1156                   [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  1157                     "pbl_equ_univ_poly_deg2_pq", init = ["empty_probl_id"], mathauthors = [], met = [["PolyEq", "solve_d2_polyeq_pq_equation"]], ppc =
  1158                     [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1159                      ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  1160                     prls =
  1161                     Repeat
  1162                      {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1163                       [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
  1164                        Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
  1165                        Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
  1166                        Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
  1167                       scr = Empty_Prog, srls = Empty},
  1168                     thy =
  1169                     {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1170                       HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1171                       HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base,
  1172                       HOL.BNF_Def, HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power,
  1173                       HOL.Groups_Big, HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big,
  1174                       HOL.Euclidean_Division, HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer,
  1175                       HOL.Lifting_Set, HOL.List, HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence,
  1176                       HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick,
  1177                       HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main,
  1178                       HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series,
  1179                       HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac,
  1180                       Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine,
  1181                       Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq,
  1182                       Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:562},
  1183                     where_ =
  1184                     [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1185                        (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1186                          (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1187                            (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $
  1188                              (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Const ("Groups.one_class.one", "real") $
  1189                                (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1190                                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))))) $
  1191                            Const ("Groups.zero_class.zero", "real")) $
  1192                          Free ("e_e", "bool")) $
  1193                        (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1194                          (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1195                            (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $
  1196                              (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1197                                (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  1198                            Const ("Groups.zero_class.zero", "real")) $
  1199                          Free ("e_e", "bool"))]}],
  1200                   []),
  1201                 Node
  1202                  ("abcFormula",
  1203                   [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  1204                     "pbl_equ_univ_poly_deg2_abc", init = ["empty_probl_id"], mathauthors = [], met = [["PolyEq", "solve_d2_polyeq_abc_equation"]], ppc =
  1205                     [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1206                      ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  1207                     prls =
  1208                     Repeat
  1209                      {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1210                       [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
  1211                        Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
  1212                        Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
  1213                        Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
  1214                       scr = Empty_Prog, srls = Empty},
  1215                     thy =
  1216                     {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1217                       HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1218                       HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base,
  1219                       HOL.BNF_Def, HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power,
  1220                       HOL.Groups_Big, HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big,
  1221                       HOL.Euclidean_Division, HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer,
  1222                       HOL.Lifting_Set, HOL.List, HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence,
  1223                       HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick,
  1224                       HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main,
  1225                       HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series,
  1226                       HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac,
  1227                       Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine,
  1228                       Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq,
  1229                       Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:572},
  1230                     where_ =
  1231                     [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1232                        (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1233                          (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1234                            (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $
  1235                              (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1236                                (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  1237                            Const ("Groups.zero_class.zero", "real")) $
  1238                          Free ("e_e", "bool")) $
  1239                        (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1240                          (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  1241                            (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $
  1242                              (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $
  1243                                (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("v_", 0), "real") $
  1244                                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))))) $
  1245                            Const ("Groups.zero_class.zero", "real")) $
  1246                          Free ("e_e", "bool"))]}],
  1247                   [])]),
  1248              Node
  1249               ("degree_3",
  1250                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  1251                  "pbl_equ_univ_poly_deg3", init = ["empty_probl_id"], mathauthors = [], met = [["PolyEq", "solve_d3_polyeq_equation"]], ppc =
  1252                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1253                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  1254                  prls =
  1255                  Repeat
  1256                   {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1257                    [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
  1258                     Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
  1259                     Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
  1260                     Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
  1261                    scr = Empty_Prog, srls = Empty},
  1262                  thy =
  1263                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1264                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1265                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1266                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1267                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1268                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1269                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1270                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1271                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  1272                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  1273                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  1274                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  1275                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:582},
  1276                  where_ =
  1277                  [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Const ("Groups.zero_class.zero", "?'a1")) $ Free ("e_e", "bool"),
  1278                   Const ("Poly.is_poly_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real"),
  1279                   Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Poly.has_degree_in", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
  1280                     (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit1", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))]}],
  1281                []),
  1282              Node
  1283               ("degree_4",
  1284                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  1285                  "pbl_equ_univ_poly_deg4", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  1286                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1287                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  1288                  prls =
  1289                  Repeat
  1290                   {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1291                    [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
  1292                     Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
  1293                     Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
  1294                     Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
  1295                    scr = Empty_Prog, srls = Empty},
  1296                  thy =
  1297                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1298                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1299                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1300                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1301                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1302                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1303                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1304                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1305                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  1306                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  1307                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  1308                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  1309                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:592},
  1310                  where_ =
  1311                  [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Const ("Groups.zero_class.zero", "?'a1")) $ Free ("e_e", "bool"),
  1312                   Const ("Poly.is_poly_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real"),
  1313                   Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Poly.has_degree_in", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
  1314                     (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))]}],
  1315                []),
  1316              Node
  1317               ("normalise",
  1318                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  1319                  "pbl_equ_univ_poly_norm", init = ["empty_probl_id"], mathauthors = [], met = [["PolyEq", "normalise_poly"]], ppc =
  1320                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1321                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  1322                  prls =
  1323                  Repeat
  1324                   {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1325                    [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
  1326                     Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
  1327                     Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
  1328                     Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
  1329                    scr = Empty_Prog, srls = Empty},
  1330                  thy =
  1331                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1332                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1333                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1334                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1335                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1336                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1337                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1338                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1339                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  1340                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  1341                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  1342                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  1343                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:602},
  1344                  where_ =
  1345                  [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1346                     (Const ("HOL.Not", "bool \<Rightarrow> bool") $
  1347                       (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Const ("Groups.zero_class.zero", "?'a1")) $
  1348                         Free ("e_e", "bool"))) $
  1349                     (Const ("HOL.Not", "bool \<Rightarrow> bool") $ (Const ("Poly.is_poly_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")))]}],
  1350                [])]),
  1351           Node
  1352            ("expanded",
  1353             [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  1354               "pbl_equ_univ_expand", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  1355               [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1356                ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  1357               prls =
  1358               Repeat
  1359                {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1360                 [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
  1361                  Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
  1362                  Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
  1363                  Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
  1364                 scr = Empty_Prog, srls = Empty},
  1365               thy =
  1366               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1367                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1368                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1369                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1370                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1371                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1372                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1373                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1374                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1375                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1376                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1377                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1378                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:612},
  1379               where_ =
  1380               [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Const ("Groups.zero_class.zero", "?'a1")) $ Free ("e_e", "bool"),
  1381                Const ("Poly.is_expanded_in", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")]}],
  1382             [Node
  1383               ("degree_2",
  1384                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  1385                  "pbl_equ_univ_expand_deg2", init = ["empty_probl_id"], mathauthors = [], met = [["PolyEq", "complete_square"]], ppc =
  1386                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1387                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  1388                  prls =
  1389                  Repeat
  1390                   {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1391                    [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
  1392                     Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
  1393                     Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
  1394                     Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
  1395                    scr = Empty_Prog, srls = Empty},
  1396                  thy =
  1397                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1398                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1399                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1400                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1401                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1402                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1403                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1404                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1405                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  1406                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  1407                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  1408                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  1409                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq:622},
  1410                  where_ =
  1411                  [Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("Poly.has_degree_in", "real \<Rightarrow> real \<Rightarrow> real") $ (Const ("Prog_Expr.lhs", "bool \<Rightarrow> real") $ Free ("e_e", "bool")) $ Free ("v_v", "real")) $
  1412                     (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))]}],
  1413                [])]),
  1414           Node
  1415            ("logarithmic",
  1416             [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  1417               "pbl_test_equ_univ_log", init = ["empty_probl_id"], mathauthors = [], met = [["Equation", "solve_log"]], ppc =
  1418               [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1419                ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  1420               prls =
  1421               Repeat
  1422                {calc = [], erls = Empty, errpatts = [], id = "PolyEq_prls", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1423                 [Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn), Eval ("Prog_Expr.lhs", fn), Eval ("Prog_Expr.rhs", fn), Eval ("Poly.is_expanded_in", fn), Eval ("Poly.is_poly_in", fn),
  1424                  Eval ("Poly.has_degree_in", fn), Eval ("Poly.is_polyrat_in", fn), Eval ("HOL.eq", fn), Eval ("RootEq.is_rootTerm_in", fn), Eval ("RatEq.is_ratequation_in", fn),
  1425                  Thm ("not_true", "(\<not> True) = False"), Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"),
  1426                  Thm ("or_true", "(?a \<or> True) = True"), Thm ("or_false", "(?a \<or> False) = ?a")],
  1427                 scr = Empty_Prog, srls = Empty},
  1428               thy =
  1429               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1430                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1431                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1432                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1433                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1434                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1435                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1436                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1437                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1438                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1439                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1440                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1441                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.LogExp:48},
  1442               where_ =
  1443               [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  1444                  (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("LogExp.alog", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $ Var (("v_v", 0), "real")) $ Var (("b", 0), "real")) $
  1445                  Free ("e_e", "bool")]}],
  1446             [])]),
  1447        Node
  1448         ("makeFunctionTo",
  1449          [{cas = NONE, guh = "pbl_equ_fromfun", init = ["empty_probl_id"], mathauthors = [], met = [["Equation", "fromFunction"]], ppc =
  1450            [("#Given", (Const ("Input_Descript.functionEq", "bool \<Rightarrow> una"), Free ("fu_n", "bool"))), ("#Given", (Const ("Equation.substitution", "bool \<Rightarrow> una"), Free ("su_b", "bool"))),
  1451             ("#Find", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("equ'''", "bool")))],
  1452            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1453            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1454              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1455              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1456              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1457              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1458              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1459              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1460              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1461              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1462              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1463              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1464              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1465              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1466              Isac.Diff, Isac.Integrate, Isac.EqSystem, Isac.Biegelinie:144},
  1467            where_ = []}],
  1468          []),
  1469        Node
  1470         ("diophantine",
  1471          [{cas =
  1472            SOME
  1473             (
  1474                Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $
  1475                  (Const ("Product_Type.prod.case_prod", "(bool \<Rightarrow> int \<Rightarrow> bool \<times> real) \<Rightarrow> bool \<times> int \<Rightarrow> bool \<times> real") $
  1476                    Abs ("x", "bool", Abs ("y", "int", Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Bound 1 $ (Const ("Int.ring_1_class.of_int", "int \<Rightarrow> real") $ Bound 0))) $
  1477                    (Const ("Product_Type.Pair", "bool \<Rightarrow> int \<Rightarrow> bool \<times> int") $ Free ("e_e", "bool") $ Free ("v_v", "int")))
  1478                ),
  1479            guh = "pbl_equ_dio", init = ["empty_probl_id"], mathauthors = [], met = [["LinEq", "solve_lineq_equation"]], ppc =
  1480            [("#Given", (Const ("Input_Descript.boolTestGiven", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.intTestGiven", "int \<Rightarrow> una"), Free ("v_v", "int"))),
  1481             ("#Find", (Const ("Input_Descript.boolTestFind", "bool \<Rightarrow> una"), Free ("s_s", "bool")))],
  1482            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1483            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1484              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1485              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1486              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1487              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1488              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1489              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1490              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1491              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1492              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1493              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1494              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1495              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1496              Isac.Diff, Isac.Test, Isac.DiophantEq:34},
  1497            where_ = []}],
  1498          [])]),
  1499     Node
  1500      ("function",
  1501       [{cas = NONE, guh = "pbl_fun", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls =
  1502         Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1503         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1504           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1505           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1506           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1507           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1508           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
  1509           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
  1510           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
  1511           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
  1512           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
  1513           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
  1514           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation,
  1515           Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff:184},
  1516         where_ = []}],
  1517       [Node
  1518         ("derivative_of",
  1519          [{cas = SOME (Const ("Diff.Diff", "real \<times> real \<Rightarrow> real") $ (Const ("Product_Type.Pair", "real \<Rightarrow> real \<Rightarrow> real \<times> real") $ Free ("f_f", "real") $ Free ("v_v", "real"))), guh = "pbl_fun_deriv", init =
  1520            ["empty_probl_id"], mathauthors = [], met = [["diff", "differentiate_on_R"], ["diff", "after_simplification"]], ppc =
  1521            [("#Given", (Const ("Input_Descript.functionTerm", "real \<Rightarrow> una"), Free ("f_f", "real"))), ("#Given", (Const ("Input_Descript.differentiateFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1522             ("#Find", (Const ("Input_Descript.derivative", "real \<Rightarrow> una"), Free ("f_f'", "real")))],
  1523            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1524            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1525              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1526              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1527              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1528              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1529              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1530              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1531              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1532              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1533              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1534              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1535              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1536              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1537              Isac.Diff:194},
  1538            where_ = []}],
  1539          [Node
  1540            ("named",
  1541             [{cas = SOME (Const ("Diff.Differentiate", "bool \<times> real \<Rightarrow> bool") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("f_f", "bool") $ Free ("v_v", "real"))), guh =
  1542               "pbl_fun_deriv_nam", init = ["empty_probl_id"], mathauthors = [], met = [["diff", "differentiate_equality"]], ppc =
  1543               [("#Given", (Const ("Input_Descript.functionEq", "bool \<Rightarrow> una"), Free ("f_f", "bool"))), ("#Given", (Const ("Input_Descript.differentiateFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1544                ("#Find", (Const ("Diff.derivativeEq", "bool \<Rightarrow> una"), Free ("f_f'", "bool")))],
  1545               prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1546               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1547                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1548                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1549                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1550                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1551                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1552                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1553                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1554                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1555                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1556                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1557                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1558                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1559                 Isac.Diff:204},
  1560               where_ = []}],
  1561             [])]),
  1562        Node
  1563         ("integrate",
  1564          [{cas = SOME (Const ("Integrate.Integrate", "real \<times> real \<Rightarrow> real") $ (Const ("Product_Type.Pair", "real \<Rightarrow> real \<Rightarrow> real \<times> real") $ Free ("f_f", "real") $ Free ("v_v", "real"))), guh = "pbl_fun_integ",
  1565            init = ["empty_probl_id"], mathauthors = [], met = [["diff", "integration"]], ppc =
  1566            [("#Given", (Const ("Input_Descript.functionTerm", "real \<Rightarrow> una"), Free ("f_f", "real"))), ("#Given", (Const ("Integrate.integrateBy", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1567             ("#Find", (Const ("Integrate.antiDerivative", "real \<Rightarrow> una"), Free ("F_F", "real")))],
  1568            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1569            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1570              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1571              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1572              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1573              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1574              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1575              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1576              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1577              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1578              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1579              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1580              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1581              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1582              Isac.Diff, Isac.Integrate:127},
  1583            where_ = []}],
  1584          [Node
  1585            ("named",
  1586             [{cas = SOME (Const ("Integrate.Integrate", "real \<times> real \<Rightarrow> real") $ (Const ("Product_Type.Pair", "real \<Rightarrow> real \<Rightarrow> real \<times> real") $ Free ("f_f", "real") $ Free ("v_v", "real"))), guh =
  1587               "pbl_fun_integ_nam", init = ["empty_probl_id"], mathauthors = [], met = [["diff", "integration", "named"]], ppc =
  1588               [("#Given", (Const ("Input_Descript.functionTerm", "real \<Rightarrow> una"), Free ("f_f", "real"))), ("#Given", (Const ("Integrate.integrateBy", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1589                ("#Find", (Const ("Integrate.antiDerivativeName", "(real \<Rightarrow> real) \<Rightarrow> una"), Free ("F_F", "real \<Rightarrow> real")))],
  1590               prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1591               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1592                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1593                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1594                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1595                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1596                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1597                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1598                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1599                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1600                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1601                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1602                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1603                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1604                 Isac.Diff, Isac.Integrate:137},
  1605               where_ = []}],
  1606             [])]),
  1607        Node
  1608         ("maximum_of",
  1609          [{cas = NONE, guh = "pbl_fun_max", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  1610            [("#Given", (Const ("Input_Descript.fixedValues", "bool list \<Rightarrow> nam"), Free ("f_ix", "bool list"))), ("#Find", (Const ("Input_Descript.maximum", "real \<Rightarrow> toreal"), Free ("m_m", "real"))),
  1611             ("#Find", (Const ("Input_Descript.valuesFor", "real list \<Rightarrow> toreall"), Free ("v_s", "real list"))), ("#Relate", (Const ("Input_Descript.relations", "bool list \<Rightarrow> una"), Free ("r_s", "bool list")))],
  1612            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1613            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1614              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1615              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1616              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1617              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1618              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1619              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1620              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1621              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1622              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1623              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1624              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1625              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1626              Isac.Diff, Isac.Diff_App:255},
  1627            where_ = []}],
  1628          [Node
  1629            ("on_interval",
  1630             [{cas = NONE, guh = "pbl_fun_max_interv", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  1631               [("#Given", (Const ("Input_Descript.functionEq", "bool \<Rightarrow> una"), Free ("t_t", "bool"))), ("#Given", (Const ("Input_Descript.boundVariable", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1632                ("#Given", (Const ("Input_Descript.interval", "real set \<Rightarrow> una"), Free ("i_tv", "real set"))), ("#Find", (Const ("Input_Descript.maxArgument", "bool \<Rightarrow> toreal"), Free ("v_0", "bool")))],
  1633               prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1634               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1635                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1636                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1637                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1638                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1639                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1640                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1641                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1642                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1643                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1644                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1645                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1646                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1647                 Isac.Diff, Isac.Diff_App:295},
  1648               where_ = []}],
  1649             [])]),
  1650        Node
  1651         ("make",
  1652          [{cas = NONE, guh = "pbl_fun_make", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  1653            [("#Given", (Const ("Input_Descript.functionOf", "real \<Rightarrow> una"), Free ("f_f", "real"))), ("#Given", (Const ("Input_Descript.boundVariable", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1654             ("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("eqs", "bool list"))), ("#Find", (Const ("Input_Descript.functionEq", "bool \<Rightarrow> una"), Free ("f_1", "bool")))],
  1655            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1656            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1657              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1658              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1659              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1660              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1661              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1662              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1663              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1664              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1665              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1666              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1667              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1668              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1669              Isac.Diff, Isac.Diff_App:265},
  1670            where_ = []}],
  1671          [Node
  1672            ("by_explicit",
  1673             [{cas = NONE, guh = "pbl_fun_max_expl", init = ["empty_probl_id"], mathauthors = [], met = [["Diff_App", "make_fun_by_explicit"]], ppc =
  1674               [("#Given", (Const ("Input_Descript.functionOf", "real \<Rightarrow> una"), Free ("f_f", "real"))), ("#Given", (Const ("Input_Descript.boundVariable", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1675                ("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("eqs", "bool list"))), ("#Find", (Const ("Input_Descript.functionEq", "bool \<Rightarrow> una"), Free ("f_1", "bool")))],
  1676               prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1677               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1678                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1679                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1680                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1681                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1682                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1683                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1684                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1685                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1686                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1687                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1688                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1689                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1690                 Isac.Diff, Isac.Diff_App:275},
  1691               where_ = []}],
  1692             []),
  1693           Node
  1694            ("by_new_variable",
  1695             [{cas = NONE, guh = "pbl_fun_max_newvar", init = ["empty_probl_id"], mathauthors = [], met = [["Diff_App", "make_fun_by_new_variable"]], ppc =
  1696               [("#Given", (Const ("Input_Descript.functionOf", "real \<Rightarrow> una"), Free ("f_f", "real"))), ("#Given", (Const ("Input_Descript.boundVariable", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  1697                ("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("eqs", "bool list"))), ("#Find", (Const ("Input_Descript.functionEq", "bool \<Rightarrow> una"), Free ("f_1", "bool")))],
  1698               prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1699               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1700                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1701                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1702                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1703                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1704                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1705                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1706                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1707                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1708                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1709                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1710                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1711                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1712                 Isac.Diff, Isac.Diff_App:285},
  1713               where_ = []}],
  1714             [])])]),
  1715     Node
  1716      ("system",
  1717       [{cas = SOME (Const ("EqSystem.solveSystem", "bool list \<Rightarrow> real list \<Rightarrow> bool list") $ Free ("e_s", "bool list") $ Free ("v_s", "real list")), guh = "pbl_equsys", init = ["empty_probl_id"], mathauthors =
  1718         [], met = [], ppc =
  1719         [("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("e_s", "bool list"))), ("#Given", (Const ("EqSystem.solveForVars", "real list \<Rightarrow> toreall"), Free ("v_s", "real list"))),
  1720          ("#Find", (Const ("EqSystem.solution", "bool list \<Rightarrow> toreall"), Free ("ss'''", "bool list")))],
  1721         prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1722         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1723           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1724           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1725           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1726           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1727           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
  1728           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
  1729           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
  1730           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
  1731           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
  1732           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
  1733           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation,
  1734           Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Integrate, Isac.EqSystem:188},
  1735         where_ = []}],
  1736       [Node
  1737         ("LINEAR",
  1738          [{cas = SOME (Const ("EqSystem.solveSystem", "bool list \<Rightarrow> real list \<Rightarrow> bool list") $ Free ("e_s", "bool list") $ Free ("v_s", "real list")), guh = "pbl_equsys_lin", init = ["empty_probl_id"],
  1739            mathauthors = [], met = [], ppc =
  1740            [("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("e_s", "bool list"))), ("#Given", (Const ("EqSystem.solveForVars", "real list \<Rightarrow> toreall"), Free ("v_s", "real list"))),
  1741             ("#Find", (Const ("EqSystem.solution", "bool list \<Rightarrow> toreall"), Free ("ss'''", "bool list")))],
  1742            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1743            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1744              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1745              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1746              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1747              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1748              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1749              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1750              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1751              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1752              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1753              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1754              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1755              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1756              Isac.Diff, Isac.Integrate, Isac.EqSystem:198},
  1757            where_ = []}],
  1758          [Node
  1759            ("2x2",
  1760             [{cas = SOME (Const ("EqSystem.solveSystem", "bool list \<Rightarrow> real list \<Rightarrow> bool list") $ Free ("e_s", "bool list") $ Free ("v_s", "real list")), guh = "pbl_equsys_lin_2x2", init = ["empty_probl_id"],
  1761               mathauthors = [], met = [], ppc =
  1762               [("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("e_s", "bool list"))),
  1763                ("#Given", (Const ("EqSystem.solveForVars", "real list \<Rightarrow> toreall"), Free ("v_s", "real list"))), ("#Find", (Const ("EqSystem.solution", "bool list \<Rightarrow> toreall"), Free ("ss'''", "bool list")))],
  1764               prls =
  1765               Repeat
  1766                {calc = [], erls = Empty, errpatts = [], id = "prls_2x2_linear_system", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1767                 [Thm ("LENGTH_CONS", "Length (?x # ?xs) = 1 + Length ?xs"), Thm ("LENGTH_NIL", "Length [] = 0"), Eval ("Groups.plus_class.plus", fn), Eval ("HOL.eq", fn)], scr = Empty_Prog, srls = Empty},
  1768               thy =
  1769               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1770                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1771                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1772                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1773                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1774                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1775                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1776                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1777                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1778                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1779                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1780                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1781                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1782                 Isac.Diff, Isac.Integrate, Isac.EqSystem:208},
  1783               where_ =
  1784               [Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("ListC.Length", "bool list \<Rightarrow> real") $ Free ("e_s", "bool list")) $
  1785                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))),
  1786                Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("ListC.Length", "?'a list \<Rightarrow> real") $ Free ("v_s", "?'a list")) $
  1787                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))]}],
  1788             [Node
  1789               ("triangular",
  1790                [{cas = SOME (Const ("EqSystem.solveSystem", "bool list \<Rightarrow> real list \<Rightarrow> bool list") $ Free ("e_s", "bool list") $ Free ("v_s", "real list")), guh = "pbl_equsys_lin_2x2_tri", init =
  1791                  ["empty_probl_id"], mathauthors = [], met = [["EqSystem", "top_down_substitution", "2x2"]], ppc =
  1792                  [("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("e_s", "bool list"))),
  1793                   ("#Given", (Const ("EqSystem.solveForVars", "real list \<Rightarrow> toreall"), Free ("v_s", "real list"))), ("#Find", (Const ("EqSystem.solution", "bool list \<Rightarrow> toreall"), Free ("ss'''", "bool list")))],
  1794                  prls =
  1795                  Repeat
  1796                   {calc = [], erls =
  1797                    Repeat
  1798                     {calc = [], erls = Empty, errpatts = [], id = "erls_prls_triangular", preconds = [], rew_ord = ("Rewrite_Ord.id_empty", fn), rules =
  1799                      [Eval ("Orderings.ord_class.less", fn), Eval ("Groups.plus_class.plus", fn), Eval ("EqSystem.occur_exactly_in", fn)], scr = Empty_Prog, srls = Empty},
  1800                    errpatts = [], id = "prls_triangular", preconds = [], rew_ord = ("Rewrite_Ord.id_empty", fn), rules =
  1801                    [Thm ("NTH_CONS", "1 < ?n \<Longrightarrow> NTH ?n (?x # ?xs) = NTH (?n + - 1) ?xs"), Eval ("Groups.plus_class.plus", fn), Thm ("NTH_NIL", "NTH 1 (?x # ?xs) = ?x"), Thm ("tl_Cons", "tl (?x # ?xs) = ?xs"),
  1802                     Thm ("tl_Nil", "tl [] = []"), Eval ("EqSystem.occur_exactly_in", fn)],
  1803                    scr = Empty_Prog, srls = Empty},
  1804                  thy =
  1805                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1806                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1807                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1808                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1809                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1810                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1811                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1812                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1813                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  1814                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  1815                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  1816                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  1817                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq,
  1818                    Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Integrate, Isac.EqSystem:218},
  1819                  where_ =
  1820                  [Const ("EqSystem.occur_exactly_in", "real list \<Rightarrow> real list \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("List.list.tl", "real list \<Rightarrow> real list") $ Free ("v_s", "real list")) $ Free ("v_s", "real list") $
  1821                     (Const ("ListC.NTH", "real \<Rightarrow> bool list \<Rightarrow> bool") $ Const ("Groups.one_class.one", "real") $ Free ("e_s", "bool list")),
  1822                   Const ("EqSystem.occur_exactly_in", "real list \<Rightarrow> real list \<Rightarrow> bool \<Rightarrow> bool") $ Free ("v_s", "real list") $ Free ("v_s", "real list") $
  1823                     (Const ("ListC.NTH", "real \<Rightarrow> bool list \<Rightarrow> bool") $ (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))) $
  1824                       Free ("e_s", "bool list"))]}],
  1825                []),
  1826              Node
  1827               ("normalise",
  1828                [{cas = SOME (Const ("EqSystem.solveSystem", "bool list \<Rightarrow> real list \<Rightarrow> bool list") $ Free ("e_s", "bool list") $ Free ("v_s", "real list")), guh = "pbl_equsys_lin_2x2_norm", init =
  1829                  ["empty_probl_id"], mathauthors = [], met = [["EqSystem", "normalise", "2x2"]], ppc =
  1830                  [("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("e_s", "bool list"))),
  1831                   ("#Given", (Const ("EqSystem.solveForVars", "real list \<Rightarrow> toreall"), Free ("v_s", "real list"))), ("#Find", (Const ("EqSystem.solution", "bool list \<Rightarrow> toreall"), Free ("ss'''", "bool list")))],
  1832                  prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1833                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1834                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1835                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1836                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1837                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1838                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1839                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1840                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1841                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  1842                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  1843                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  1844                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  1845                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq,
  1846                    Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Integrate, Isac.EqSystem:228},
  1847                  where_ = []}],
  1848                [])]),
  1849           Node
  1850            ("3x3",
  1851             [{cas = SOME (Const ("EqSystem.solveSystem", "bool list \<Rightarrow> real list \<Rightarrow> bool list") $ Free ("e_s", "bool list") $ Free ("v_s", "real list")), guh = "pbl_equsys_lin_3x3", init = ["empty_probl_id"],
  1852               mathauthors = [], met = [], ppc =
  1853               [("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("e_s", "bool list"))),
  1854                ("#Given", (Const ("EqSystem.solveForVars", "real list \<Rightarrow> toreall"), Free ("v_s", "real list"))), ("#Find", (Const ("EqSystem.solution", "bool list \<Rightarrow> toreall"), Free ("ss'''", "bool list")))],
  1855               prls =
  1856               Repeat
  1857                {calc = [], erls = Empty, errpatts = [], id = "prls_3x3_linear_system", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1858                 [Thm ("LENGTH_CONS", "Length (?x # ?xs) = 1 + Length ?xs"), Thm ("LENGTH_NIL", "Length [] = 0"), Eval ("Groups.plus_class.plus", fn), Eval ("HOL.eq", fn)], scr = Empty_Prog, srls = Empty},
  1859               thy =
  1860               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1861                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1862                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1863                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1864                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1865                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1866                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1867                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1868                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1869                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1870                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1871                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1872                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1873                 Isac.Diff, Isac.Integrate, Isac.EqSystem:238},
  1874               where_ =
  1875               [Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("ListC.Length", "bool list \<Rightarrow> real") $ Free ("e_s", "bool list")) $
  1876                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit1", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))),
  1877                Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("ListC.Length", "?'a list \<Rightarrow> real") $ Free ("v_s", "?'a list")) $
  1878                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit1", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))]}],
  1879             []),
  1880           Node
  1881            ("4x4",
  1882             [{cas = SOME (Const ("EqSystem.solveSystem", "bool list \<Rightarrow> real list \<Rightarrow> bool list") $ Free ("e_s", "bool list") $ Free ("v_s", "real list")), guh = "pbl_equsys_lin_4x4", init = ["empty_probl_id"],
  1883               mathauthors = [], met = [], ppc =
  1884               [("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("e_s", "bool list"))),
  1885                ("#Given", (Const ("EqSystem.solveForVars", "real list \<Rightarrow> toreall"), Free ("v_s", "real list"))), ("#Find", (Const ("EqSystem.solution", "bool list \<Rightarrow> toreall"), Free ("ss'''", "bool list")))],
  1886               prls =
  1887               Repeat
  1888                {calc = [], erls = Empty, errpatts = [], id = "prls_4x4_linear_system", preconds = [], rew_ord = ("dummy_ord", fn), rules =
  1889                 [Thm ("LENGTH_CONS", "Length (?x # ?xs) = 1 + Length ?xs"), Thm ("LENGTH_NIL", "Length [] = 0"), Eval ("Groups.plus_class.plus", fn), Eval ("HOL.eq", fn)], scr = Empty_Prog, srls = Empty},
  1890               thy =
  1891               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1892                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1893                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1894                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1895                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1896                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1897                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1898                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1899                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  1900                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  1901                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  1902                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  1903                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  1904                 Isac.Diff, Isac.Integrate, Isac.EqSystem:248},
  1905               where_ =
  1906               [Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("ListC.Length", "bool list \<Rightarrow> real") $ Free ("e_s", "bool list")) $
  1907                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))),
  1908                Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $ (Const ("ListC.Length", "?'a list \<Rightarrow> real") $ Free ("v_s", "?'a list")) $
  1909                  (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))]}],
  1910             [Node
  1911               ("triangular",
  1912                [{cas = SOME (Const ("EqSystem.solveSystem", "bool list \<Rightarrow> real list \<Rightarrow> bool list") $ Free ("e_s", "bool list") $ Free ("v_s", "real list")), guh = "pbl_equsys_lin_4x4_tri", init =
  1913                  ["empty_probl_id"], mathauthors = [], met = [["EqSystem", "top_down_substitution", "4x4"]], ppc =
  1914                  [("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("e_s", "bool list"))),
  1915                   ("#Given", (Const ("EqSystem.solveForVars", "real list \<Rightarrow> toreall"), Free ("v_s", "real list"))), ("#Find", (Const ("EqSystem.solution", "bool list \<Rightarrow> toreall"), Free ("ss'''", "bool list")))],
  1916                  prls =
  1917                  Repeat
  1918                   {calc = [], erls =
  1919                    Repeat
  1920                     {calc = [], erls = Empty, errpatts = [], id = "erls_prls_triangular", preconds = [], rew_ord = ("Rewrite_Ord.id_empty", fn), rules =
  1921                      [Eval ("Orderings.ord_class.less", fn), Eval ("Groups.plus_class.plus", fn), Eval ("EqSystem.occur_exactly_in", fn)], scr = Empty_Prog, srls = Empty},
  1922                    errpatts = [], id = "prls_tri_4x4_lin_sys", preconds = [], rew_ord = ("Rewrite_Ord.id_empty", fn), rules =
  1923                    [Thm ("NTH_CONS", "1 < ?n \<Longrightarrow> NTH ?n (?x # ?xs) = NTH (?n + - 1) ?xs"), Eval ("Groups.plus_class.plus", fn), Thm ("NTH_NIL", "NTH 1 (?x # ?xs) = ?x"), Thm ("tl_Cons", "tl (?x # ?xs) = ?xs"),
  1924                     Thm ("tl_Nil", "tl [] = []"), Eval ("EqSystem.occur_exactly_in", fn), Eval ("Prog_Expr.occurs_in", fn)],
  1925                    scr = Empty_Prog, srls = Empty},
  1926                  thy =
  1927                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1928                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1929                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1930                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1931                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1932                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1933                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1934                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1935                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  1936                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  1937                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  1938                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  1939                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq,
  1940                    Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Integrate, Isac.EqSystem:258},
  1941                  where_ =
  1942                  [Const ("Prog_Expr.occurs_in", "real \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("ListC.NTH", "real \<Rightarrow> real list \<Rightarrow> real") $ Const ("Groups.one_class.one", "real") $ Free ("v_s", "real list")) $
  1943                     (Const ("ListC.NTH", "real \<Rightarrow> bool list \<Rightarrow> bool") $ Const ("Groups.one_class.one", "real") $ Free ("e_s", "bool list")),
  1944                   Const ("Prog_Expr.occurs_in", "real \<Rightarrow> bool \<Rightarrow> bool") $
  1945                     (Const ("ListC.NTH", "real \<Rightarrow> real list \<Rightarrow> real") $ (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))) $
  1946                       Free ("v_s", "real list")) $
  1947                     (Const ("ListC.NTH", "real \<Rightarrow> bool list \<Rightarrow> bool") $ (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))) $
  1948                       Free ("e_s", "bool list")),
  1949                   Const ("Prog_Expr.occurs_in", "real \<Rightarrow> bool \<Rightarrow> bool") $
  1950                     (Const ("ListC.NTH", "real \<Rightarrow> real list \<Rightarrow> real") $ (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit1", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))) $
  1951                       Free ("v_s", "real list")) $
  1952                     (Const ("ListC.NTH", "real \<Rightarrow> bool list \<Rightarrow> bool") $ (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit1", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))) $
  1953                       Free ("e_s", "bool list")),
  1954                   Const ("Prog_Expr.occurs_in", "real \<Rightarrow> bool \<Rightarrow> bool") $
  1955                     (Const ("ListC.NTH", "real \<Rightarrow> real list \<Rightarrow> real") $
  1956                       (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))) $
  1957                       Free ("v_s", "real list")) $
  1958                     (Const ("ListC.NTH", "real \<Rightarrow> bool list \<Rightarrow> bool") $
  1959                       (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))) $
  1960                       Free ("e_s", "bool list"))]}],
  1961                []),
  1962              Node
  1963               ("normalise",
  1964                [{cas = SOME (Const ("EqSystem.solveSystem", "bool list \<Rightarrow> real list \<Rightarrow> bool list") $ Free ("e_s", "bool list") $ Free ("v_s", "real list")), guh = "pbl_equsys_lin_4x4_norm", init =
  1965                  ["empty_probl_id"], mathauthors = [], met = [["EqSystem", "normalise", "4x4"]], ppc =
  1966                  [("#Given", (Const ("Input_Descript.equalities", "bool list \<Rightarrow> tobooll"), Free ("e_s", "bool list"))),
  1967                   ("#Given", (Const ("EqSystem.solveForVars", "real list \<Rightarrow> toreall"), Free ("v_s", "real list"))), ("#Find", (Const ("EqSystem.solution", "bool list \<Rightarrow> toreall"), Free ("ss'''", "bool list")))],
  1968                  prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1969                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  1970                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1971                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1972                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1973                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1974                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  1975                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  1976                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  1977                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  1978                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  1979                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  1980                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  1981                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq,
  1982                    Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Integrate, Isac.EqSystem:268},
  1983                  where_ = []}],
  1984                [])])])]),
  1985     Node
  1986      ("Biegelinien",
  1987       [{cas = NONE, guh = "pbl_bieg", init = ["empty_probl_id"], mathauthors = [], met = [["IntegrierenUndKonstanteBestimmen2"]], ppc =
  1988         [("#Given", (Const ("Biegelinie.Traegerlaenge", "real \<Rightarrow> una"), Free ("l_l", "real"))), ("#Given", (Const ("Biegelinie.Streckenlast", "real \<Rightarrow> una"), Free ("q_q", "real"))),
  1989          ("#Find", (Const ("Biegelinie.Biegelinie", "(real \<Rightarrow> real) \<Rightarrow> una"), Free ("b_b", "real \<Rightarrow> real"))), ("#Relate", (Const ("Biegelinie.Randbedingungen", "bool list \<Rightarrow> una"), Free ("r_b", "bool list")))],
  1990         prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  1991         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  1992           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  1993           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  1994           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  1995           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  1996           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
  1997           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
  1998           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
  1999           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
  2000           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
  2001           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
  2002           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation,
  2003           Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Integrate, Isac.EqSystem,
  2004           Isac.Biegelinie:74},
  2005         where_ = []}],
  2006       [Node
  2007         ("MomentBestimmte",
  2008          [{cas = NONE, guh = "pbl_bieg_mom", init = ["empty_probl_id"], mathauthors = [], met = [["IntegrierenUndKonstanteBestimmen"]], ppc =
  2009            [("#Given", (Const ("Biegelinie.Traegerlaenge", "real \<Rightarrow> una"), Free ("l_l", "real"))), ("#Given", (Const ("Biegelinie.Streckenlast", "real \<Rightarrow> una"), Free ("q_q", "real"))),
  2010             ("#Find", (Const ("Biegelinie.Biegelinie", "(real \<Rightarrow> real) \<Rightarrow> una"), Free ("b_b", "real \<Rightarrow> real"))),
  2011             ("#Relate", (Const ("Biegelinie.RandbedingungenBiegung", "bool list \<Rightarrow> una"), Free ("r_b", "bool list"))),
  2012             ("#Relate", (Const ("Biegelinie.RandbedingungenMoment", "bool list \<Rightarrow> una"), Free ("r_m", "bool list")))],
  2013            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2014            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2015              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2016              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2017              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2018              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2019              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2020              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2021              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2022              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2023              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2024              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2025              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2026              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2027              Isac.Diff, Isac.Integrate, Isac.EqSystem, Isac.Biegelinie:84},
  2028            where_ = []}],
  2029          []),
  2030        Node
  2031         ("MomentGegebene",
  2032          [{cas = NONE, guh = "pbl_bieg_momg", init = ["empty_probl_id"], mathauthors = [], met = [["IntegrierenUndKonstanteBestimmen", "2xIntegrieren"]], ppc = [], prls =
  2033            Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2034            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2035              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2036              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2037              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2038              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2039              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2040              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2041              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2042              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2043              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2044              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2045              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2046              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2047              Isac.Diff, Isac.Integrate, Isac.EqSystem, Isac.Biegelinie:94},
  2048            where_ = []}],
  2049          []),
  2050        Node
  2051         ("einfache",
  2052          [{cas = NONE, guh = "pbl_bieg_einf", init = ["empty_probl_id"], mathauthors = [], met = [["IntegrierenUndKonstanteBestimmen", "4x4System"]], ppc = [], prls =
  2053            Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2054            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2055              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2056              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2057              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2058              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2059              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2060              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2061              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2062              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2063              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2064              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2065              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2066              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2067              Isac.Diff, Isac.Integrate, Isac.EqSystem, Isac.Biegelinie:104},
  2068            where_ = []}],
  2069          []),
  2070        Node
  2071         ("QuerkraftUndMomentBestimmte",
  2072          [{cas = NONE, guh = "pbl_bieg_momquer", init = ["empty_probl_id"], mathauthors = [], met = [["IntegrierenUndKonstanteBestimmen", "1xIntegrieren"]], ppc = [], prls =
  2073            Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2074            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2075              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2076              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2077              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2078              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2079              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2080              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2081              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2082              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2083              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2084              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2085              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2086              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2087              Isac.Diff, Isac.Integrate, Isac.EqSystem, Isac.Biegelinie:114},
  2088            where_ = []}],
  2089          []),
  2090        Node
  2091         ("vonBelastungZu",
  2092          [{cas = NONE, guh = "pbl_bieg_vonq", init = ["empty_probl_id"], mathauthors = [], met = [["Biegelinien", "ausBelastung"]], ppc =
  2093            [("#Given", (Const ("Biegelinie.Streckenlast", "real \<Rightarrow> una"), Free ("q_q", "real"))), ("#Given", (Const ("Biegelinie.FunktionsVariable", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  2094             ("#Find", (Const ("Biegelinie.Funktionen", "bool list \<Rightarrow> una"), Free ("funs'''", "bool list")))],
  2095            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2096            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2097              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2098              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2099              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2100              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2101              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2102              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2103              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2104              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2105              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2106              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2107              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2108              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2109              Isac.Diff, Isac.Integrate, Isac.EqSystem, Isac.Biegelinie:124},
  2110            where_ = []}],
  2111          []),
  2112        Node
  2113         ("setzeRandbedingungen",
  2114          [{cas = NONE, guh = "pbl_bieg_randbed", init = ["empty_probl_id"], mathauthors = [], met = [["Biegelinien", "setzeRandbedingungenEin"]], ppc =
  2115            [("#Given", (Const ("Biegelinie.Funktionen", "bool list \<Rightarrow> una"), Free ("fun_s", "bool list"))), ("#Given", (Const ("Biegelinie.Randbedingungen", "bool list \<Rightarrow> una"), Free ("r_b", "bool list"))),
  2116             ("#Find", (Const ("Biegelinie.Gleichungen", "bool list \<Rightarrow> una"), Free ("equs'''", "bool list")))],
  2117            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2118            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2119              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2120              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2121              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2122              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2123              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2124              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2125              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2126              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2127              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2128              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2129              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2130              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2131              Isac.Diff, Isac.Integrate, Isac.EqSystem, Isac.Biegelinie:134},
  2132            where_ = []}],
  2133          [])]),
  2134     Node
  2135      ("Berechnung",
  2136       [{cas = NONE, guh = "pbl_algein", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls =
  2137         Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2138         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2139           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2140           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2141           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2142           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2143           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
  2144           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
  2145           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
  2146           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
  2147           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
  2148           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
  2149           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational,
  2150           Isac.AlgEin:30},
  2151         where_ = []}],
  2152       [Node
  2153         ("numerischSymbolische",
  2154          [{cas = NONE, guh = "pbl_algein_numsym", init = ["empty_probl_id"], mathauthors = [], met = [["Berechnung", "erstNumerisch"], ["Berechnung", "erstSymbolisch"]], ppc =
  2155            [("#Given", (Const ("AlgEin.KantenLaenge", "bool \<Rightarrow> una"), Free ("k_k", "bool"))), ("#Given", (Const ("AlgEin.Querschnitt", "bool \<Rightarrow> una"), Free ("q__q", "bool"))),
  2156             ("#Given", (Const ("AlgEin.KantenUnten", "bool list \<Rightarrow> una"), Free ("u_u", "bool list"))), ("#Given", (Const ("AlgEin.KantenSenkrecht", "bool list \<Rightarrow> una"), Free ("s_s", "bool list"))),
  2157             ("#Given", (Const ("AlgEin.KantenOben", "bool list \<Rightarrow> una"), Free ("o_o", "bool list"))), ("#Find", (Const ("AlgEin.GesamtLaenge", "real \<Rightarrow> una"), Free ("l_l", "real")))],
  2158            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2159            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2160              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2161              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2162              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2163              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2164              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2165              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2166              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2167              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2168              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2169              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2170              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2171              Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational, Isac.AlgEin:40},
  2172            where_ = []}],
  2173          [])]),
  2174     Node
  2175      ("Programming",
  2176       [{cas = NONE, guh = "pbl_Programming", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls =
  2177         Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2178         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2179           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2180           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2181           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2182           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2183           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
  2184           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
  2185           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
  2186           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
  2187           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
  2188           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
  2189           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational,
  2190           Isac.InsSort:124},
  2191         where_ = []}],
  2192       [Node
  2193         ("SORT",
  2194          [{cas = NONE, guh = "pbl_Prog_sort", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls =
  2195            Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2196            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2197              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2198              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2199              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2200              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2201              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2202              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2203              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2204              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2205              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2206              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2207              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2208              Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational, Isac.InsSort:134},
  2209            where_ = []}],
  2210          [Node
  2211            ("insertion",
  2212             [{cas = SOME (Const ("InsSort.Sort", "int xlist \<Rightarrow> int xlist") $ Free ("u_u", "int xlist")), guh = "pbl_Prog_sort_ins", init = ["empty_probl_id"], mathauthors = [], met =
  2213               [["Programming", "SORT", "insertion_steps"]], ppc =
  2214               [("#Given", (Const ("InsSort.unsorted", "int xlist \<Rightarrow> unl"), Free ("u_u", "int xlist"))), ("#Find", (Const ("InsSort.sorted", "int xlist \<Rightarrow> unl"), Free ("s_s", "int xlist")))], prls =
  2215               Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2216               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2217                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2218                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2219                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2220                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2221                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2222                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2223                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2224                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2225                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2226                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2227                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2228                 Isac.Poly, Isac.GCD_Poly_ML, Isac.Rational, Isac.InsSort:144},
  2229               where_ = []}],
  2230             [])])]),
  2231     Node
  2232      ("test",
  2233       [{cas = NONE, guh = "pbl_test", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls =
  2234         Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2235         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2236           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2237           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2238           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2239           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2240           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
  2241           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
  2242           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
  2243           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
  2244           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
  2245           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
  2246           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation,
  2247           Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Test:397},
  2248         where_ = []}],
  2249       [Node
  2250         ("equation",
  2251          [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh = "pbl_test_equ",
  2252            init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  2253            [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  2254             ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2255            prls =
  2256            Repeat
  2257             {calc =
  2258              [("is_num", ("Prog_Expr.is_num", fn)), ("matches", ("Prog_Expr.matches", fn)), ("PLUS", ("Groups.plus_class.plus", fn)), ("TIMES", ("Groups.times_class.times", fn)),
  2259               ("POWER", ("BaseDefinitions.realpow", fn)), ("le", ("Orderings.ord_class.less", fn)), ("leq", ("Orderings.ord_class.less_eq", fn)), ("ident", ("Prog_Expr.ident", fn))],
  2260              erls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, errpatts = [], id = "matches",
  2261              preconds = [], rew_ord = ("termlessI", fn), rules =
  2262              [Thm ("refl", "?t = ?t"), Thm ("order_refl", "?x \<le> ?x"), Thm ("radd_left_cancel_le", "(?k + ?m \<le> ?k + ?n) = (?m \<le> ?n)"), Thm ("not_true", "(\<not> True) = False"),
  2263               Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"), Thm ("or_true", "(?a \<or> True) = True"),
  2264               Thm ("or_false", "(?a \<or> False) = ?a"), Thm ("and_commute", "(?a \<and> ?b) = (?b \<and> ?a)"), Thm ("or_commute", "(?a \<or> ?b) = (?b \<or> ?a)"), Eval ("Prog_Expr.is_num", fn), Eval ("Prog_Expr.matches", fn),
  2265               Eval ("Groups.plus_class.plus", fn), Eval ("Groups.times_class.times", fn), Eval ("BaseDefinitions.realpow", fn), Eval ("Orderings.ord_class.less", fn), Eval ("Orderings.ord_class.less_eq", fn),
  2266               Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn)],
  2267              scr = Empty_Prog, srls = Empty},
  2268            thy =
  2269            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2270              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2271              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2272              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2273              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2274              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2275              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2276              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2277              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2278              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2279              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2280              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2281              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2282              Isac.Diff, Isac.Test:407},
  2283            where_ = [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Var (("b", 0), "?'a1")) $ Free ("e_e", "bool")]}],
  2284          [Node
  2285            ("univariate",
  2286             [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh = "pbl_test_uni",
  2287               init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  2288               [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  2289                ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2290               prls =
  2291               Repeat
  2292                {calc =
  2293                 [("is_num", ("Prog_Expr.is_num", fn)), ("matches", ("Prog_Expr.matches", fn)), ("PLUS", ("Groups.plus_class.plus", fn)), ("TIMES", ("Groups.times_class.times", fn)),
  2294                  ("POWER", ("BaseDefinitions.realpow", fn)), ("le", ("Orderings.ord_class.less", fn)), ("leq", ("Orderings.ord_class.less_eq", fn)), ("ident", ("Prog_Expr.ident", fn))],
  2295                 erls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, errpatts = [], id = "matches",
  2296                 preconds = [], rew_ord = ("termlessI", fn), rules =
  2297                 [Thm ("refl", "?t = ?t"), Thm ("order_refl", "?x \<le> ?x"), Thm ("radd_left_cancel_le", "(?k + ?m \<le> ?k + ?n) = (?m \<le> ?n)"), Thm ("not_true", "(\<not> True) = False"),
  2298                  Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"), Thm ("or_true", "(?a \<or> True) = True"),
  2299                  Thm ("or_false", "(?a \<or> False) = ?a"), Thm ("and_commute", "(?a \<and> ?b) = (?b \<and> ?a)"), Thm ("or_commute", "(?a \<or> ?b) = (?b \<or> ?a)"), Eval ("Prog_Expr.is_num", fn), Eval ("Prog_Expr.matches", fn),
  2300                  Eval ("Groups.plus_class.plus", fn), Eval ("Groups.times_class.times", fn), Eval ("BaseDefinitions.realpow", fn), Eval ("Orderings.ord_class.less", fn), Eval ("Orderings.ord_class.less_eq", fn),
  2301                  Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn)],
  2302                 scr = Empty_Prog, srls = Empty},
  2303               thy =
  2304               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2305                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2306                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2307                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2308                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2309                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2310                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2311                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2312                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2313                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2314                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2315                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2316                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2317                 Isac.Diff, Isac.Test:417},
  2318               where_ = [Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Var (("a", 0), "?'a1") $ Var (("b", 0), "?'a1")) $ Free ("e_e", "bool")]}],
  2319             [Node
  2320               ("LINEAR",
  2321                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  2322                  "pbl_test_uni_lin", init = ["empty_probl_id"], mathauthors = [], met = [["Test", "solve_linear"]], ppc =
  2323                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  2324                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2325                  prls =
  2326                  Repeat
  2327                   {calc =
  2328                    [("is_num", ("Prog_Expr.is_num", fn)), ("matches", ("Prog_Expr.matches", fn)), ("PLUS", ("Groups.plus_class.plus", fn)), ("TIMES", ("Groups.times_class.times", fn)),
  2329                     ("POWER", ("BaseDefinitions.realpow", fn)), ("le", ("Orderings.ord_class.less", fn)), ("leq", ("Orderings.ord_class.less_eq", fn)), ("ident", ("Prog_Expr.ident", fn))],
  2330                    erls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, errpatts = [], id = "matches",
  2331                    preconds = [], rew_ord = ("termlessI", fn), rules =
  2332                    [Thm ("refl", "?t = ?t"), Thm ("order_refl", "?x \<le> ?x"), Thm ("radd_left_cancel_le", "(?k + ?m \<le> ?k + ?n) = (?m \<le> ?n)"), Thm ("not_true", "(\<not> True) = False"),
  2333                     Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"), Thm ("or_true", "(?a \<or> True) = True"),
  2334                     Thm ("or_false", "(?a \<or> False) = ?a"), Thm ("and_commute", "(?a \<and> ?b) = (?b \<and> ?a)"), Thm ("or_commute", "(?a \<or> ?b) = (?b \<or> ?a)"), Eval ("Prog_Expr.is_num", fn),
  2335                     Eval ("Prog_Expr.matches", fn), Eval ("Groups.plus_class.plus", fn), Eval ("Groups.times_class.times", fn), Eval ("BaseDefinitions.realpow", fn), Eval ("Orderings.ord_class.less", fn),
  2336                     Eval ("Orderings.ord_class.less_eq", fn), Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn)],
  2337                    scr = Empty_Prog, srls = Empty},
  2338                  thy =
  2339                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  2340                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2341                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2342                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2343                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2344                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2345                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2346                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2347                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  2348                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  2349                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  2350                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  2351                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq,
  2352                    Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Test:427},
  2353                  where_ =
  2354                  [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2355                     (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $ (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ Free ("v_v", "?'a1") $ Const ("Groups.zero_class.zero", "?'a1")) $ Free ("e_e", "bool")) $
  2356                     (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2357                       (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2358                         (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ (Const ("Groups.times_class.times", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Free ("v_v", "?'a1")) $
  2359                           Const ("Groups.zero_class.zero", "?'a1")) $
  2360                         Free ("e_e", "bool")) $
  2361                       (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2362                         (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2363                           (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $ (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("a", 0), "?'a1") $ Free ("v_v", "?'a1")) $
  2364                             Const ("Groups.zero_class.zero", "?'a1")) $
  2365                           Free ("e_e", "bool")) $
  2366                         (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2367                           (Const ("HOL.eq", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> bool") $
  2368                             (Const ("Groups.plus_class.plus", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("a", 0), "?'a1") $
  2369                               (Const ("Groups.times_class.times", "?'a1 \<Rightarrow> ?'a1 \<Rightarrow> ?'a1") $ Var (("b", 0), "?'a1") $ Free ("v_v", "?'a1"))) $
  2370                             Const ("Groups.zero_class.zero", "?'a1")) $
  2371                           Free ("e_e", "bool"))))]}],
  2372                []),
  2373              Node
  2374               ("plain_square",
  2375                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  2376                  "pbl_test_uni_plain2", init = ["empty_probl_id"], mathauthors = [], met = [["Test", "solve_plain_square"]], ppc =
  2377                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  2378                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2379                  prls =
  2380                  Repeat
  2381                   {calc =
  2382                    [("is_num", ("Prog_Expr.is_num", fn)), ("matches", ("Prog_Expr.matches", fn)), ("PLUS", ("Groups.plus_class.plus", fn)), ("TIMES", ("Groups.times_class.times", fn)),
  2383                     ("POWER", ("BaseDefinitions.realpow", fn)), ("le", ("Orderings.ord_class.less", fn)), ("leq", ("Orderings.ord_class.less_eq", fn)), ("ident", ("Prog_Expr.ident", fn))],
  2384                    erls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, errpatts = [], id = "matches",
  2385                    preconds = [], rew_ord = ("termlessI", fn), rules =
  2386                    [Thm ("refl", "?t = ?t"), Thm ("order_refl", "?x \<le> ?x"), Thm ("radd_left_cancel_le", "(?k + ?m \<le> ?k + ?n) = (?m \<le> ?n)"), Thm ("not_true", "(\<not> True) = False"),
  2387                     Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"), Thm ("or_true", "(?a \<or> True) = True"),
  2388                     Thm ("or_false", "(?a \<or> False) = ?a"), Thm ("and_commute", "(?a \<and> ?b) = (?b \<and> ?a)"), Thm ("or_commute", "(?a \<or> ?b) = (?b \<or> ?a)"), Eval ("Prog_Expr.is_num", fn),
  2389                     Eval ("Prog_Expr.matches", fn), Eval ("Groups.plus_class.plus", fn), Eval ("Groups.times_class.times", fn), Eval ("BaseDefinitions.realpow", fn), Eval ("Orderings.ord_class.less", fn),
  2390                     Eval ("Orderings.ord_class.less_eq", fn), Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.matches", fn)],
  2391                    scr = Empty_Prog, srls = Empty},
  2392                  thy =
  2393                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  2394                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2395                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2396                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2397                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2398                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2399                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2400                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2401                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  2402                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  2403                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  2404                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  2405                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq,
  2406                    Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Test:475},
  2407                  where_ =
  2408                  [Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2409                     (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2410                       (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2411                         (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $
  2412                           (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $
  2413                             (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("v_v", "real") $
  2414                               (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))))) $
  2415                         Const ("Groups.zero_class.zero", "real")) $
  2416                       Free ("e_e", "bool")) $
  2417                     (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2418                       (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2419                         (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2420                           (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("b", 0), "real") $
  2421                             (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("v_v", "real") $
  2422                               (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  2423                           Const ("Groups.zero_class.zero", "real")) $
  2424                         Free ("e_e", "bool")) $
  2425                       (Const ("HOL.disj", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2426                         (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2427                           (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2428                             (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $ Var (("a", 0), "real") $
  2429                               (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("v_v", "real") $
  2430                                 (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  2431                             Const ("Groups.zero_class.zero", "real")) $
  2432                           Free ("e_e", "bool")) $
  2433                         (Const ("Prog_Expr.matches", "bool \<Rightarrow> bool \<Rightarrow> bool") $
  2434                           (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2435                             (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("v_v", "real") $
  2436                               (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))) $
  2437                             Const ("Groups.zero_class.zero", "real")) $
  2438                           Free ("e_e", "bool"))))]}],
  2439                []),
  2440              Node
  2441               ("polynomial",
  2442                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  2443                  "pbl_test_uni_poly", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  2444                  [("#Given",
  2445                    (Const ("Input_Descript.equality", "bool \<Rightarrow> una"),
  2446                     Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2447                       (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2448                         (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2449                           (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("v_v", "real") $
  2450                             (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))) $
  2451                           (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("p_p", "real") $ Free ("v_v", "real"))) $
  2452                         Free ("q__q", "real")) $
  2453                       Const ("Groups.zero_class.zero", "real"))),
  2454                   ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))), ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2455                  prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2456                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  2457                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2458                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2459                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2460                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2461                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2462                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2463                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2464                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  2465                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  2466                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  2467                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  2468                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq,
  2469                    Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Test:485},
  2470                  where_ = [Const ("HOL.False", "bool")]}],
  2471                [Node
  2472                  ("degree_two",
  2473                   [{cas =
  2474                     SOME
  2475                      (
  2476                         Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $
  2477                           (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $
  2478                             (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2479                               (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2480                                 (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2481                                   (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("v_v", "real") $
  2482                                     (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))) $
  2483                                   (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("p_p", "real") $ Free ("v_v", "real"))) $
  2484                                 Free ("q__q", "real")) $
  2485                               Const ("Groups.zero_class.zero", "real")) $
  2486                             Free ("v_v", "real"))
  2487                         ),
  2488                     guh = "pbl_test_uni_poly_deg2", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  2489                     [("#Given",
  2490                       (Const ("Input_Descript.equality", "bool \<Rightarrow> una"),
  2491                        Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2492                          (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2493                            (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2494                              (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("v_v", "real") $
  2495                                (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))) $
  2496                              (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("p_p", "real") $ Free ("v_v", "real"))) $
  2497                            Free ("q__q", "real")) $
  2498                          Const ("Groups.zero_class.zero", "real"))),
  2499                      ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))), ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2500                     prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2501                     {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  2502                       HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2503                       HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base,
  2504                       HOL.BNF_Def, HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power,
  2505                       HOL.Groups_Big, HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big,
  2506                       HOL.Euclidean_Division, HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer,
  2507                       HOL.Lifting_Set, HOL.List, HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence,
  2508                       HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick,
  2509                       HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main,
  2510                       HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series,
  2511                       HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac,
  2512                       Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine,
  2513                       Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq,
  2514                       Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Test:495},
  2515                     where_ = []}],
  2516                   [Node
  2517                     ("pq_formula",
  2518                      [{cas =
  2519                        SOME
  2520                         (
  2521                            Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $
  2522                              (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $
  2523                                (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2524                                  (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2525                                    (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2526                                      (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("v_v", "real") $
  2527                                        (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))) $
  2528                                      (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("p_p", "real") $ Free ("v_v", "real"))) $
  2529                                    Free ("q__q", "real")) $
  2530                                  Const ("Groups.zero_class.zero", "real")) $
  2531                                Free ("v_v", "real"))
  2532                            ),
  2533                        guh = "pbl_test_uni_poly_deg2_pq", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  2534                        [("#Given",
  2535                          (Const ("Input_Descript.equality", "bool \<Rightarrow> una"),
  2536                           Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2537                             (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2538                               (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2539                                 (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("v_v", "real") $
  2540                                   (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num")))) $
  2541                                 (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("p_p", "real") $ Free ("v_v", "real"))) $
  2542                               Free ("q__q", "real")) $
  2543                             Const ("Groups.zero_class.zero", "real"))),
  2544                         ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  2545                         ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2546                        prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2547                        {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  2548                          HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2549                          HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base,
  2550                          HOL.BNF_Def, HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power,
  2551                          HOL.Groups_Big, HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big,
  2552                          HOL.Euclidean_Division, HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer,
  2553                          HOL.Lifting_Set, HOL.List, HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence,
  2554                          HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick,
  2555                          HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main,
  2556                          HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series,
  2557                          HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program,
  2558                          Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret,
  2559                          Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq,
  2560                          Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Test:505},
  2561                        where_ = []}],
  2562                      []),
  2563                    Node
  2564                     ("abc_formula",
  2565                      [{cas =
  2566                        SOME
  2567                         (
  2568                            Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $
  2569                              (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $
  2570                                (Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2571                                  (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2572                                    (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2573                                      (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("a_a", "real") $
  2574                                        (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("x", "real") $
  2575                                          (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  2576                                      (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("b_b", "real") $ Free ("x", "real"))) $
  2577                                    Free ("c_c", "real")) $
  2578                                  Const ("Groups.zero_class.zero", "real")) $
  2579                                Free ("v_v", "real"))
  2580                            ),
  2581                        guh = "pbl_test_uni_poly_deg2_abc", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  2582                        [("#Given",
  2583                          (Const ("Input_Descript.equality", "bool \<Rightarrow> una"),
  2584                           Const ("HOL.eq", "real \<Rightarrow> real \<Rightarrow> bool") $
  2585                             (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2586                               (Const ("Groups.plus_class.plus", "real \<Rightarrow> real \<Rightarrow> real") $
  2587                                 (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("a_a", "real") $
  2588                                   (Const ("BaseDefinitions.realpow", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("x", "real") $
  2589                                     (Const ("Num.numeral_class.numeral", "num \<Rightarrow> real") $ (Const ("Num.num.Bit0", "num \<Rightarrow> num") $ Const ("Num.num.One", "num"))))) $
  2590                                 (Const ("Groups.times_class.times", "real \<Rightarrow> real \<Rightarrow> real") $ Free ("b_b", "real") $ Free ("x", "real"))) $
  2591                               Free ("c_c", "real")) $
  2592                             Const ("Groups.zero_class.zero", "real"))),
  2593                         ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  2594                         ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2595                        prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2596                        {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  2597                          HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2598                          HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base,
  2599                          HOL.BNF_Def, HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power,
  2600                          HOL.Groups_Big, HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big,
  2601                          HOL.Euclidean_Division, HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer,
  2602                          HOL.Lifting_Set, HOL.List, HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence,
  2603                          HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick,
  2604                          HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main,
  2605                          HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series,
  2606                          HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program,
  2607                          Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret,
  2608                          Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq,
  2609                          Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Test:515},
  2610                        where_ = []}],
  2611                      [])])]),
  2612              Node
  2613               ("squareroot",
  2614                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  2615                  "pbl_test_uni_root", init = ["empty_probl_id"], mathauthors = [], met = [["Test", "square_equation"]], ppc =
  2616                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  2617                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2618                  prls =
  2619                  Repeat {calc = [], erls = Empty, errpatts = [], id = "contains_root", preconds = [], rew_ord = ("dummy_ord", fn), rules = [Eval ("Test.contains_root", fn)], scr = Empty_Prog, srls = Empty},
  2620                  thy =
  2621                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  2622                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2623                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2624                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2625                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2626                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2627                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2628                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2629                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  2630                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  2631                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  2632                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  2633                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq,
  2634                    Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Test:525},
  2635                  where_ = [Const ("Test.precond_rootpbl", "?'a \<Rightarrow> bool") $ Free ("v_v", "?'a")]}],
  2636                []),
  2637              Node
  2638               ("normalise",
  2639                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  2640                  "pbl_test_uni_norm", init = ["empty_probl_id"], mathauthors = [], met = [["Test", "norm_univar_equation"]], ppc =
  2641                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  2642                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2643                  prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2644                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  2645                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2646                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2647                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2648                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2649                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2650                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2651                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2652                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  2653                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  2654                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  2655                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  2656                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq,
  2657                    Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Test:535},
  2658                  where_ = []}],
  2659                []),
  2660              Node
  2661               ("sqroot-test",
  2662                [{cas = SOME (Const ("Equation.solve", "bool \<times> real \<Rightarrow> bool list") $ (Const ("Product_Type.Pair", "bool \<Rightarrow> real \<Rightarrow> bool \<times> real") $ Free ("e_e", "bool") $ Free ("v_v", "real"))), guh =
  2663                  "pbl_test_uni_roottest", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  2664                  [("#Given", (Const ("Input_Descript.equality", "bool \<Rightarrow> una"), Free ("e_e", "bool"))), ("#Given", (Const ("Input_Descript.solveFor", "real \<Rightarrow> una"), Free ("v_v", "real"))),
  2665                   ("#Find", (Const ("Input_Descript.solutions", "bool list \<Rightarrow> toreall"), Free ("v_v'i'", "bool list")))],
  2666                  prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2667                  {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive,
  2668                    HOL.Rings, HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2669                    HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2670                    HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2671                    HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2672                    HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2673                    HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2674                    HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2675                    HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull,
  2676                    HOL.Modules, HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex,
  2677                    HOL.MacLaurin, Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr,
  2678                    Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle,
  2679                    Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq,
  2680                    Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Test:545},
  2681                  where_ = [Const ("Test.precond_rootpbl", "?'a \<Rightarrow> bool") $ Free ("v_v", "?'a")]}],
  2682                [])])]),
  2683        Node
  2684         ("inttype",
  2685          [{cas = NONE, guh = "pbl_test_intsimp", init = ["empty_probl_id"], mathauthors = [], met = [["Test", "intsimp"]], ppc =
  2686            [("#Given", (Const ("Input_Descript.intTestGiven", "int \<Rightarrow> una"), Free ("t_t", "int"))), ("#Find", (Free ("intTestFind", "'a \<Rightarrow> 'b"), Free ("s_s", "'a")))], prls =
  2687            Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2688            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2689              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2690              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2691              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2692              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2693              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2694              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2695              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2696              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2697              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2698              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2699              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2700              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2701              Isac.Diff, Isac.Test:555},
  2702            where_ = []}],
  2703          [])]),
  2704     Node
  2705      ("tool",
  2706       [{cas = NONE, guh = "pbl_tool", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls =
  2707         Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2708         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2709           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2710           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2711           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2712           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2713           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
  2714           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
  2715           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
  2716           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
  2717           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
  2718           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
  2719           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation,
  2720           Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Diff_App:305},
  2721         where_ = []}],
  2722       [Node
  2723         ("find_values",
  2724          [{cas = NONE, guh = "pbl_tool_findvals", init = ["empty_probl_id"], mathauthors = [], met = [], ppc =
  2725            [("#Given", (Const ("Input_Descript.maxArgument", "bool \<Rightarrow> toreal"), Free ("m_ax", "bool"))), ("#Given", (Const ("Input_Descript.functionEq", "bool \<Rightarrow> una"), Free ("f_f", "bool"))),
  2726             ("#Given", (Const ("Input_Descript.boundVariable", "real \<Rightarrow> una"), Free ("v_v", "real"))), ("#Find", (Const ("Input_Descript.valuesFor", "real list \<Rightarrow> toreall"), Free ("v_ls", "real list"))),
  2727             ("#Relate", (Const ("Input_Descript.additionalRels", "bool list \<Rightarrow> una"), Free ("r_s", "bool list")))],
  2728            prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2729            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2730              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2731              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2732              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2733              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2734              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2735              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2736              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2737              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2738              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2739              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2740              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2741              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2742              Isac.Diff, Isac.Diff_App:315},
  2743            where_ = []}],
  2744          [])]),
  2745     Node
  2746      ("Optimisation",
  2747       [{cas = NONE, guh = "pbl_opti", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls =
  2748         Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2749         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2750           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2751           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2752           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2753           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2754           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
  2755           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
  2756           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
  2757           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
  2758           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
  2759           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
  2760           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation,
  2761           Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Diff_App:325},
  2762         where_ = []}],
  2763       [Node
  2764         ("univariate_calculus",
  2765          [{cas = NONE, guh = "pbl_opti_univ", init = ["empty_probl_id"], mathauthors = [], met = [["Optimisation", "by_univariate_calculus"]], ppc =
  2766            [("#Given", (Const ("Input_Descript.Constants", "bool list \<Rightarrow> nam"), Free ("fixes", "bool list"))), ("#Find", (Const ("Input_Descript.Maximum", "real \<Rightarrow> toreal"), Free ("maxx", "real"))),
  2767             ("#Find", (Const ("Input_Descript.AdditionalValues", "real list \<Rightarrow> toreall"), Free ("vals", "real list"))), ("#Relate", (Const ("Input_Descript.Extremum", "bool \<Rightarrow> toreal"), Free ("extr", "bool"))),
  2768             ("#Relate", (Const ("Input_Descript.SideConditions", "bool list \<Rightarrow> una"), Free ("sideconds", "bool list")))],
  2769            prls =
  2770            Repeat
  2771             {calc =
  2772              [("le", ("Orderings.ord_class.less", fn)), ("leq", ("Orderings.ord_class.less_eq", fn)), ("ident", ("Prog_Expr.ident", fn)), ("is_num", ("Prog_Expr.is_num", fn)),
  2773               ("occurs_in", ("Prog_Expr.occurs_in", fn)), ("matches", ("Prog_Expr.matches", fn))],
  2774              erls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, errpatts = [], id = "eval_rls",
  2775              preconds = [], rew_ord = ("termlessI", fn), rules =
  2776              [Thm ("refl", "?t = ?t"), Thm ("order_refl", "?x \<le> ?x"), Thm ("radd_left_cancel_le", "(?k + ?m \<le> ?k + ?n) = (?m \<le> ?n)"), Thm ("not_true", "(\<not> True) = False"),
  2777               Thm ("not_false", "(\<not> False) = True"), Thm ("and_true", "(?a \<and> True) = ?a"), Thm ("and_false", "(?a \<and> False) = False"), Thm ("or_true", "(?a \<or> True) = True"),
  2778               Thm ("or_false", "(?a \<or> False) = ?a"), Thm ("and_commute", "(?a \<and> ?b) = (?b \<and> ?a)"), Thm ("or_commute", "(?a \<or> ?b) = (?b \<or> ?a)"), Eval ("Orderings.ord_class.less", fn),
  2779               Eval ("Orderings.ord_class.less_eq", fn), Eval ("Prog_Expr.ident", fn), Eval ("Prog_Expr.is_num", fn), Eval ("Prog_Expr.occurs_in", fn), Eval ("Prog_Expr.matches", fn)],
  2780              scr = Empty_Prog, srls = Empty},
  2781            thy =
  2782            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2783              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2784              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2785              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2786              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2787              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2788              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2789              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2790              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2791              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2792              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2793              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2794              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2795              Isac.Diff, Isac.Diff_App:335},
  2796            where_ = []}],
  2797          [])]),
  2798     Node
  2799      ("SignalProcessing",
  2800       [{cas = NONE, guh = "pbl_SP", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls =
  2801         Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2802         {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2803           HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2804           HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2805           HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2806           HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2807           HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List, HOL.Groups_List,
  2808           HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation, HOL.Quickcheck_Random,
  2809           HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing, HOL.Mirabelle, HOL.Extraction,
  2810           HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules, HOL.Real, HOL.Topological_Spaces,
  2811           HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin, Complex_Main, Specify.Know_Store,
  2812           Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog, Specify.ProgLang, Specify.Input_Descript,
  2813           Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify, Isac.Poly, Isac.Root, Isac.Equation,
  2814           Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp, Isac.Diff, Isac.Diff_App, Isac.Partial_Fractions,
  2815           Isac.Inverse_Z_Transform:96},
  2816         where_ = []}],
  2817       [Node
  2818         ("Z_Transform",
  2819          [{cas = NONE, guh = "pbl_SP_Ztrans", init = ["empty_probl_id"], mathauthors = [], met = [], ppc = [], prls =
  2820            Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2821            {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2822              HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2823              HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2824              HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2825              HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2826              HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2827              HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2828              HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2829              HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2830              HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2831              Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2832              Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2833              Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2834              Isac.Diff, Isac.Diff_App, Isac.Partial_Fractions, Isac.Inverse_Z_Transform:106},
  2835            where_ = []}],
  2836          [Node
  2837            ("Inverse",
  2838             [{cas = NONE, guh = "pbl_SP_Ztrans_inv", init = ["empty_probl_id"], mathauthors = [], met = [["SignalProcessing", "Z_Transform", "Inverse"]], ppc =
  2839               [("#Given", (Const ("Inverse_Z_Transform.filterExpression", "bool \<Rightarrow> una"), Free ("X_eq", "bool"))), ("#Find", (Const ("Inverse_Z_Transform.stepResponse", "bool \<Rightarrow> una"), Free ("n_eq", "bool")))],
  2840               prls = Repeat {calc = [], erls = Empty, errpatts = [], id = "empty", preconds = [], rew_ord = ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Empty}, thy =
  2841               {Pure, Tools.Code_Generator, HOL.HOL, HOL.Orderings, HOL.Groups, HOL.Lattices, HOL.Boolean_Algebras, HOL.Set, HOL.Fun, HOL.Complete_Lattices, HOL.Ctr_Sugar, HOL.Typedef, HOL.Inductive, HOL.Rings,
  2842                 HOL.Nat, HOL.Product_Type, HOL.Sum_Type, HOL.Fields, HOL.Finite_Set, HOL.Relation, HOL.Transitive_Closure, HOL.Wellfounded, HOL.Wfrec, HOL.Order_Relation, HOL.Hilbert_Choice,
  2843                 HOL.BNF_Wellorder_Relation, HOL.BNF_Wellorder_Embedding, HOL.Zorn, HOL.BNF_Wellorder_Constructions, HOL.BNF_Cardinal_Order_Relation, HOL.BNF_Cardinal_Arithmetic, HOL.Fun_Def_Base, HOL.BNF_Def,
  2844                 HOL.BNF_Composition, HOL.Basic_BNFs, HOL.BNF_Fixpoint_Base, HOL.BNF_Least_Fixpoint, HOL.Meson, HOL.ATP, HOL.Basic_BNF_LFPs, HOL.Metis, HOL.Transfer, HOL.Num, HOL.Power, HOL.Groups_Big,
  2845                 HOL.Equiv_Relations, HOL.Lifting, HOL.Complete_Partial_Order, HOL.Option, HOL.Argo, HOL.Partial_Function, HOL.SAT, HOL.Quotient, HOL.Fun_Def, HOL.Int, HOL.Lattices_Big, HOL.Euclidean_Division,
  2846                 HOL.Parity, HOL.Divides, HOL.Numeral_Simprocs, HOL.Semiring_Normalization, HOL.Groebner_Basis, HOL.Set_Interval, HOL.Presburger, HOL.SMT, HOL.Sledgehammer, HOL.Lifting_Set, HOL.List,
  2847                 HOL.Groups_List, HOL.Bit_Operations, HOL.Code_Numeral, HOL.Map, HOL.Enum, HOL.String, HOL.Predicate, HOL.Lazy_Sequence, HOL.Typerep, HOL.Limited_Sequence, HOL.Random, HOL.Code_Evaluation,
  2848                 HOL.Quickcheck_Random, HOL.Random_Pred, HOL.Random_Sequence, HOL.Quickcheck_Exhaustive, HOL.Predicate_Compile, HOL.Record, HOL.GCD, HOL.Nitpick, HOL.Factorial, HOL.Quickcheck_Narrowing,
  2849                 HOL.Mirabelle, HOL.Extraction, HOL.Nunchaku, HOL.BNF_Greatest_Fixpoint, HOL.Filter, HOL.Conditionally_Complete_Lattices, HOL.Binomial, Main, HOL.Archimedean_Field, HOL.Rat, HOL.Hull, HOL.Modules,
  2850                 HOL.Real, HOL.Topological_Spaces, HOL.Vector_Spaces, HOL.Real_Vector_Spaces, HOL.Limits, HOL.Inequalities, HOL.Deriv, HOL.Series, HOL.NthRoot, HOL.Transcendental, HOL.Complex, HOL.MacLaurin,
  2851                 Complex_Main, Specify.Know_Store, Specify.BaseDefinitions, Specify.Calc_Binop, Specify.ListC, Specify.Program, Specify.Prog_Tac, Specify.Tactical, Specify.Prog_Expr, Specify.Auto_Prog,
  2852                 Specify.ProgLang, Specify.Input_Descript, Specify.MathEngBasic, Specify.Specify, Interpret.Interpret, Isac.MathEngine, Isac.Test_Code, Isac.BridgeLibisabelle, Isac.Base_Tools, Isac.Simplify,
  2853                 Isac.Poly, Isac.Root, Isac.Equation, Isac.GCD_Poly_ML, Isac.Rational, Isac.LinEq, Isac.RootEq, Isac.RatEq, Isac.RootRat, Isac.RootRatEq, Isac.PolyEq, Isac.Calculus, Isac.Trig, Isac.LogExp,
  2854                 Isac.Diff, Isac.Diff_App, Isac.Partial_Fractions, Isac.Inverse_Z_Transform:116},
  2855               where_ = []}],
  2856             [])])])]:
  2857    Probl_Def.store*)
  2858 \<close> text \<open>
  2859 ("thy1", ["aaa", "bbb"], ["ccc", "ddd"]) |> References.select_input References.empty |> #1 |> ThyC.get_theory |> Proof_Context.init_global
  2860 \<close> text \<open>
  2861 Proof_Context.init_global (References.select_input speco spec |> #1 |> ThyC.get_theory)
  2862 \<close> ML \<open>
  2863 \<close> ML \<open>
  2864 \<close> ML \<open>
  2865 \<close> ML \<open>
  2866 \<close> ML \<open>
  2867 \<close> ML \<open>
  2868 \<close> ML \<open>
  2869 \<close>
  2870 (*\----- end delete -----------------------------------------------------------------------/*)
  2871 ML \<open>Eval.adhoc_thm; (*from "ProgLang/evaluate.sml" *)\<close>
  2872 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
  2873 ML \<open>Input_Descript.for_real_list; (*from "Input_Descript.thy" *)\<close>
  2874 ML \<open>Test_Code.me;\<close>
  2875 text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
  2876 ML \<open>prog_expr\<close>
  2877                                                          
  2878 ML \<open>Prog_Expr.eval_occurs_in\<close>
  2879 ML \<open>@{thm last_thmI}\<close>
  2880 (** )ML \<open>@{thm Querkraft_Belastung}\<close>( *exception FAIL NONE raised (line 161 of "General/scan.ML")*)
  2881 
  2882 declare [[check_unique = false]]
  2883 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
  2884 ML \<open>@{theory "Isac_Knowledge"}\<close>
  2885 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
  2886   ERROR: app_py: not found: ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]*)\<close>
  2887 
  2888 section \<open>State of approaching Isabelle by Isac\<close>
  2889 text \<open>
  2890   Mathias Lehnfeld gives the following list in his thesis in section 
  2891   4.2.3 Relation to Ongoing Isabelle Development.
  2892 \<close>
  2893 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
  2894 text \<open>
  2895   DONE
  2896 \<close>
  2897 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
  2898 subsection \<open>(2) Make Isac’s programming language usable\<close>
  2899 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
  2900 text \<open>
  2901   In 2002 isac already strived for floating point numbers. Since that time
  2902   isac represents numerals as "Free", see below (*1*). These experiments are
  2903   unsatisfactory with respect to logical soundness.
  2904   Since Isabelle now has started to care about floating point numbers, it is high 
  2905   time to adopt these together with the other numerals. Isabelle2012/13's numerals
  2906   are different from Isabelle2011, see "$ISABELLE_ISAC_TEST/Tools/isac/ProgLang/termC.sml".
  2907   
  2908   The transition from "Free" to standard numerals is a task to be scheduled for 
  2909   several weeks. The urgency of this task follows from the experience, 
  2910   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
  2911   some of the long identifiers of theorems which would tremendously simplify
  2912   building a hierarchy of theorems according to (1.2), see (*2*) below.
  2913 \<close>
  2914 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
  2915 
  2916 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
  2917 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
  2918 
  2919 end