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