src/Tools/isac/Knowledge/Equation.thy
author Walther Neuper <neuper@ist.tugraz.at>
Wed, 04 May 2011 09:01:10 +0200
branchdecompose-isar
changeset 41972 22c5483e9bfb
parent 41952 0e76e17a430a
child 42398 04d3f0133827
permissions -rw-r--r--
update all "Pair" to "Product_Type.Pair"

regression test --- x+1=2 start SubProblem 'stac2tac_ TODO: no match for SubProblem
     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 consts
     9 
    10   (*descriptions in the related problems TODOshift here from Descriptions.thy*)
    11   substitution      :: "bool => una"
    12 
    13   (*the CAS-commands*)
    14   solve             :: "[bool * 'a] => bool list" (* solve (x+1=2, x) *)
    15   solveTest         :: "[bool * 'a] => bool list" (* for test collection *)
    16   
    17   (*Script-names*)
    18   Function2Equality :: "[bool, bool,       bool] 
    19 					 => bool"
    20                   ("((Script Function2Equality (_ _ =))// (_))" 9)
    21 
    22 text {* defines equation and univariate-equation
    23         created by: rlang 
    24               date: 02.09
    25         changed by: rlang
    26         last change by: rlang
    27                   date: 02.11.29
    28         (c) by Richard Lang, 2003 *}
    29 
    30 ML {*
    31 val thy = @{theory};
    32 val ctxt = thy2ctxt thy;
    33 
    34 val univariate_equation_prls = 
    35     append_rls "univariate_equation_prls" e_rls 
    36 	       [Calc ("Tools.matches",eval_matches "")];
    37 ruleset' := 
    38 overwritelthy @{theory} (!ruleset',
    39 		   [("univariate_equation_prls",
    40 		     prep_rls univariate_equation_prls)]);
    41 
    42 
    43 store_pbt 
    44  (prep_pbt thy "pbl_equ" [] e_pblID
    45  (["equation"],
    46   [("#Given" ,["equality e_e","solveFor v_v"]),
    47    ("#Where" ,["matches (?a = ?b) e_e"]),
    48    ("#Find"  ,["solutions v_v'i'"])
    49   ],
    50   append_rls "equation_prls" e_rls 
    51 	     [Calc ("Tools.matches",eval_matches "")],
    52   SOME "solve (e_e::bool, v_v)",
    53   []));
    54 
    55 store_pbt
    56  (prep_pbt thy "pbl_equ_univ" [] e_pblID
    57  (["univariate","equation"],
    58   [("#Given" ,["equality e_e","solveFor v_v"]),
    59    ("#Where" ,["matches (?a = ?b) e_e"]),
    60    ("#Find"  ,["solutions v_v'i'"])
    61   ],
    62   univariate_equation_prls,SOME "solve (e_e::bool, v_v)",[]));
    63 
    64 
    65 (*.function for handling the cas-input "solve (x+1=2, x)":
    66    make a model which is already in ptree-internal format.*)
    67 (* val (h,argl) = strip_comb (str2term "solve (x+1=2, x)");
    68    val (h,argl) = strip_comb ((term_of o the o (parse thy)) 
    69 				  "solveTest (x+1=2, x)");
    70    *)
    71 fun argl2dtss [Const ("Product_Type.Pair", _) $ eq $ bdv] =
    72     [((the o (parseNEW ctxt)) "equality", [eq]),
    73      ((the o (parseNEW ctxt)) "solveFor", [bdv]),
    74      ((the o (parseNEW ctxt)) "solutions", 
    75       [(the o (parseNEW ctxt)) "L"])
    76      ]
    77   | argl2dtss _ = error "Equation.ML: wrong argument for argl2dtss";
    78 
    79 castab := 
    80 overwritel (!castab, 
    81 	    [((term_of o the o (parse thy)) "solveTest", 
    82 	      (("Test", ["univariate","equation","test"], ["no_met"]), 
    83 	       argl2dtss)),
    84 	     ((term_of o the o (parse thy)) "solve",  
    85 	      (("Isac", ["univariate","equation"], ["no_met"]), 
    86 	       argl2dtss))
    87 	     ]);
    88 
    89 
    90 
    91 store_met
    92     (prep_met thy "met_equ" [] e_metID
    93 	      (["Equation"],
    94 	       [],
    95 	       {rew_ord'="tless_true", rls'=Erls, calc = [], 
    96 		srls = e_rls, 
    97 		prls=e_rls,
    98 	     crls = Atools_erls, nrls = e_rls},
    99 "empty_script"
   100 ));
   101 *}
   102 
   103 end