src/Tools/isac/Knowledge/Equation.thy
author Walther Neuper <neuper@ist.tugraz.at>
Tue, 31 Aug 2010 16:38:22 +0200
branchisac-update-Isa09-2
changeset 37967 bd4f7a35e892
parent 37953 369b3012f6f6
child 37972 66fc615a1e89
permissions -rw-r--r--
updating Knowledge/Simplify, changes ahead + in test

# overwritelnthy thy --> overwritelnthy @{theory}
# test: thms-replace-Isa02-Isa09-2.sml @{thm ..}
     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 univariate_equation_prls = 
    32     append_rls "univariate_equation_prls" e_rls 
    33 	       [Calc ("Tools.matches",eval_matches "")];
    34 ruleset' := 
    35 overwritelthy @{theory} (!ruleset',
    36 		   [("univariate_equation_prls",
    37 		     prep_rls univariate_equation_prls)]);
    38 
    39 
    40 store_pbt 
    41  (prep_pbt (theory "Equation") "pbl_equ" [] e_pblID
    42  (["equation"],
    43   [("#Given" ,["equality e_","solveFor v_"]),
    44    ("#Where" ,["matches (?a = ?b) e_"]),
    45    ("#Find"  ,["solutions v_i_"])
    46   ],
    47   append_rls "equation_prls" e_rls 
    48 	     [Calc ("Tools.matches",eval_matches "")],
    49   SOME "solve (e_::bool, v_)",
    50   []));
    51 
    52 store_pbt
    53  (prep_pbt (theory "Equation") "pbl_equ_univ" [] e_pblID
    54  (["univariate","equation"],
    55   [("#Given" ,["equality e_","solveFor v_"]),
    56    ("#Where" ,["matches (?a = ?b) e_"]),
    57    ("#Find"  ,["solutions v_i_"])
    58   ],
    59   univariate_equation_prls,SOME "solve (e_::bool, v_)",[]));
    60 
    61 
    62 (*.function for handling the cas-input "solve (x+1=2, x)":
    63    make a model which is already in ptree-internal format.*)
    64 (* val (h,argl) = strip_comb (str2term "solve (x+1=2, x)");
    65    val (h,argl) = strip_comb ((term_of o the o (parse thy)) 
    66 				  "solveTest (x+1=2, x)");
    67    *)
    68 fun argl2dtss [Const ("Pair", _) $ eq $ bdv] =
    69     [((term_of o the o (parse thy)) "equality", [eq]),
    70      ((term_of o the o (parse thy)) "solveFor", [bdv]),
    71      ((term_of o the o (parse thy)) "solutions", 
    72       [(term_of o the o (parse thy)) "L"])
    73      ]
    74   | argl2dtss _ = raise error "Equation.ML: wrong argument for argl2dtss";
    75 
    76 castab := 
    77 overwritel (!castab, 
    78 	    [((term_of o the o (parse thy)) "solveTest", 
    79 	      (("Test.thy", ["univariate","equation","test"], ["no_met"]), 
    80 	       argl2dtss)),
    81 	     ((term_of o the o (parse thy)) "solve",  
    82 	      (("Isac.thy", ["univariate","equation"], ["no_met"]), 
    83 	       argl2dtss))
    84 	     ]);
    85 
    86 
    87 
    88 store_met
    89     (prep_met (theory "Equation") "met_equ" [] e_metID
    90 	      (["Equation"],
    91 	       [],
    92 	       {rew_ord'="tless_true", rls'=Erls, calc = [], 
    93 		srls = e_rls, 
    94 		prls=e_rls,
    95 	     crls = Atools_erls, nrls = e_rls},
    96 "empty_script"
    97 ));
    98 *}
    99 
   100 end