src/Tools/isac/Knowledge/Equation.thy
author Mathias Lehnfeld <s1210629013@students.fh-hagenberg.at>
Mon, 27 Jan 2014 21:49:27 +0100
changeset 55359 73dc85c025ab
parent 55339 cccd24e959ba
child 55363 d78bc1342183
permissions -rw-r--r--
cleanup, naming: 'KEStore_Elems' in Tests now 'Test_KEStore_Elems', 'store_pbts' now 'add_pbts'
     1 (* equations and functions; functions NOT as lambda-terms
     2    author: Walther Neuper 2005, 2006
     3    (c) due to copyright terms
     4 *)
     5 
     6 theory Equation imports Atools begin
     7 
     8 text {* univariate equations over terms :
     9   In 2003 Richard Lang prototyped ISAC's equation solver while also alpha-testing
    10   the Lucas-Interpreter's Subproblem mechanism.
    11   This prototype suffers from the drop-out of Matthias Goldgruber for a year,
    12   who had started to work on simplification in parallel. So this implementation
    13   replaced simplifiers by many specific theorems for specific terms in specific
    14   phases of equation solving; these specific solutions wait for generalisation
    15   in future improvements of ISAC's equation solver.
    16 *}
    17 
    18 consts
    19 
    20   (*descriptions in the related problems TODOshift here from Descriptions.thy*)
    21   substitution      :: "bool => una"
    22 
    23   (*the CAS-commands*)
    24   solve             :: "[bool * 'a] => bool list" (* solve (x+1=2, x) *)
    25   solveTest         :: "[bool * 'a] => bool list" (* for test collection *)
    26   
    27   (*Script-names*)
    28   Function2Equality :: "[bool, bool,       bool] 
    29 					 => bool"
    30                   ("((Script Function2Equality (_ _ =))// (_))" 9)
    31 
    32 text {* defines equation and univariate-equation
    33         created by: rlang 
    34               date: 02.09
    35         changed by: rlang
    36         last change by: rlang
    37                   date: 02.11.29
    38         (c) by Richard Lang, 2003 *}
    39 
    40 ML {*
    41 val thy = @{theory};
    42 val ctxt = thy2ctxt thy;
    43 
    44 val univariate_equation_prls = 
    45     append_rls "univariate_equation_prls" e_rls 
    46 	       [Calc ("Tools.matches",eval_matches "")];
    47 *}
    48 setup {* KEStore_Elems.add_rlss [("univariate_equation_prls",
    49   (Context.theory_name @{theory}, prep_rls univariate_equation_prls))] *}
    50 ML {*
    51 store_pbt 
    52  (prep_pbt thy "pbl_equ" [] e_pblID
    53  (["equation"],
    54   [("#Given" ,["equality e_e","solveFor v_v"]),
    55    ("#Where" ,["matches (?a = ?b) e_e"]),
    56    ("#Find"  ,["solutions v_v'i'"])
    57   ],
    58   append_rls "equation_prls" e_rls 
    59 	     [Calc ("Tools.matches",eval_matches "")],
    60   SOME "solve (e_e::bool, v_v)",
    61   []));
    62 
    63 store_pbt
    64  (prep_pbt thy "pbl_equ_univ" [] e_pblID
    65  (["univariate","equation"],
    66   [("#Given" ,["equality e_e","solveFor v_v"]),
    67    ("#Where" ,["matches (?a = ?b) e_e"]),
    68    ("#Find"  ,["solutions v_v'i'"])
    69   ],
    70   univariate_equation_prls,SOME "solve (e_e::bool, v_v)",[]));
    71 *}
    72 setup {* KEStore_Elems.add_pbts
    73   [(prep_pbt thy "pbl_equ" [] e_pblID
    74       (["equation"],
    75         [("#Given" ,["equality e_e","solveFor v_v"]),
    76           ("#Where" ,["matches (?a = ?b) e_e"]),
    77           ("#Find"  ,["solutions v_v'i'"])],
    78         append_rls "equation_prls" e_rls  [Calc ("Tools.matches",eval_matches "")],
    79         SOME "solve (e_e::bool, v_v)", [])),
    80     (prep_pbt thy "pbl_equ_univ" [] e_pblID
    81       (["univariate","equation"],
    82         [("#Given" ,["equality e_e","solveFor v_v"]),
    83           ("#Where" ,["matches (?a = ?b) e_e"]),
    84           ("#Find"  ,["solutions v_v'i'"])],
    85         univariate_equation_prls, SOME "solve (e_e::bool, v_v)", []))] *}
    86 
    87 
    88 ML{*
    89 (*.function for handling the cas-input "solve (x+1=2, x)":
    90    make a model which is already in ptree-internal format.*)
    91 (* val (h,argl) = strip_comb (str2term "solve (x+1=2, x)");
    92    val (h,argl) = strip_comb ((term_of o the o (parse thy)) 
    93 				  "solveTest (x+1=2, x)");
    94    *)
    95 fun argl2dtss [Const ("Product_Type.Pair", _) $ eq $ bdv] =
    96     [((the o (parseNEW ctxt)) "equality", [eq]),
    97      ((the o (parseNEW ctxt)) "solveFor", [bdv]),
    98      ((the o (parseNEW ctxt)) "solutions", 
    99       [(the o (parseNEW ctxt)) "L"])
   100      ]
   101   | argl2dtss _ = error "Equation.ML: wrong argument for argl2dtss";
   102 *}
   103 setup {* KEStore_Elems.add_cas
   104   [((term_of o the o (parse thy)) "solveTest", 
   105       (("Test", ["univariate","equation","test"], ["no_met"]), argl2dtss)),
   106     ((term_of o the o (parse thy)) "solve",
   107       (("Isac", ["univariate","equation"], ["no_met"]), argl2dtss))]*}
   108 
   109 
   110 ML {*
   111 store_met
   112     (prep_met thy "met_equ" [] e_metID
   113 	      (["Equation"],
   114 	       [],
   115 	       {rew_ord'="tless_true", rls'=Erls, calc = [], 
   116 		srls = e_rls,
   117 		prls=e_rls,
   118 	     crls = Atools_erls, errpats = [], nrls = e_rls},
   119 "empty_script"
   120 ));
   121 *}
   122 
   123 end