src/Tools/isac/Knowledge/Equation.thy
author Walther Neuper <neuper@ist.tugraz.at>
Wed, 08 Sep 2010 16:47:22 +0200
branchisac-update-Isa09-2
changeset 37991 028442673981
parent 37981 b2877b9d455a
child 38009 b49723351533
permissions -rw-r--r--
tuned src + test

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