src/Tools/isac/Knowledge/Equation.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Wed, 21 Nov 2018 12:32:54 +0100
changeset 59472 3e904f8ec16c
parent 59424 406681ebe781
child 59473 28b67cae58c3
permissions -rw-r--r--
update to new Isabelle conventions: {*...*} to \<open>...\<close>
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
wneuper@59424
     6
theory Equation imports Base_Tools begin
neuper@37906
     7
wneuper@59472
     8
text \<open>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.
wneuper@59472
    16
\<close>
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
wneuper@59472
    32
text \<open>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
wneuper@59472
    38
        (c) by Richard Lang, 2003\<close>
neuper@37950
    39
wneuper@59472
    40
ML \<open>
neuper@37972
    41
val thy = @{theory};
wneuper@59416
    42
val ctxt = Rule.thy2ctxt thy;
neuper@37972
    43
neuper@37950
    44
val univariate_equation_prls = 
wneuper@59416
    45
    Rule.append_rls "univariate_equation_prls" Rule.e_rls 
wneuper@59416
    46
	       [Rule.Calc ("Tools.matches",eval_matches "")];
wneuper@59472
    47
\<close>
wneuper@59472
    48
setup \<open>KEStore_Elems.add_rlss [("univariate_equation_prls",
wneuper@59472
    49
  (Context.theory_name @{theory}, LTool.prep_rls @{theory} univariate_equation_prls))]\<close>
wneuper@59472
    50
setup \<open>KEStore_Elems.add_pbts
wneuper@59406
    51
  [(Specify.prep_pbt thy "pbl_equ" [] Celem.e_pblID
s1210629013@55339
    52
      (["equation"],
s1210629013@55339
    53
        [("#Given" ,["equality e_e","solveFor v_v"]),
s1210629013@55339
    54
          ("#Where" ,["matches (?a = ?b) e_e"]),
s1210629013@55339
    55
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59416
    56
        Rule.append_rls "equation_prls" Rule.e_rls  [Rule.Calc ("Tools.matches",eval_matches "")],
s1210629013@55339
    57
        SOME "solve (e_e::bool, v_v)", [])),
wneuper@59406
    58
    (Specify.prep_pbt thy "pbl_equ_univ" [] Celem.e_pblID
s1210629013@55339
    59
      (["univariate","equation"],
s1210629013@55339
    60
        [("#Given" ,["equality e_e","solveFor v_v"]),
s1210629013@55339
    61
          ("#Where" ,["matches (?a = ?b) e_e"]),
s1210629013@55339
    62
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59472
    63
        univariate_equation_prls, SOME "solve (e_e::bool, v_v)", []))]\<close>
neuper@37950
    64
neuper@37950
    65
wneuper@59472
    66
ML\<open>
neuper@37950
    67
(*.function for handling the cas-input "solve (x+1=2, x)":
wneuper@59279
    68
   make a model which is already in ctree-internal format.*)
neuper@37950
    69
(* val (h,argl) = strip_comb (str2term "solve (x+1=2, x)");
wneuper@59186
    70
   val (h,argl) = strip_comb ((Thm.term_of o the o (parse thy)) 
neuper@37950
    71
				  "solveTest (x+1=2, x)");
neuper@37950
    72
   *)
neuper@41972
    73
fun argl2dtss [Const ("Product_Type.Pair", _) $ eq $ bdv] =
wneuper@59389
    74
    [((the o (TermC.parseNEW ctxt)) "equality", [eq]),
wneuper@59389
    75
     ((the o (TermC.parseNEW ctxt)) "solveFor", [bdv]),
wneuper@59389
    76
     ((the o (TermC.parseNEW ctxt)) "solutions", 
wneuper@59389
    77
      [(the o (TermC.parseNEW ctxt)) "L"])
neuper@37950
    78
     ]
neuper@38031
    79
  | argl2dtss _ = error "Equation.ML: wrong argument for argl2dtss";
wneuper@59472
    80
\<close>
wneuper@59472
    81
setup \<open>KEStore_Elems.add_cas
wneuper@59389
    82
  [((Thm.term_of o the o (TermC.parse thy)) "solveTest", 
s1210629013@52170
    83
      (("Test", ["univariate","equation","test"], ["no_met"]), argl2dtss)),
wneuper@59389
    84
    ((Thm.term_of o the o (TermC.parse thy)) "solve",
wneuper@59472
    85
      (("Isac", ["univariate","equation"], ["no_met"]), argl2dtss))]\<close>
neuper@37950
    86
neuper@37950
    87
wneuper@59472
    88
setup \<open>KEStore_Elems.add_mets
wneuper@59406
    89
  [Specify.prep_met thy "met_equ" [] Celem.e_metID
s1210629013@55373
    90
	    (["Equation"], [],
wneuper@59416
    91
	      {rew_ord'="tless_true", rls'=Rule.Erls, calc = [], srls = Rule.e_rls, prls=Rule.e_rls, crls = Atools_erls,
wneuper@59416
    92
          errpats = [], nrls = Rule.e_rls},
s1210629013@55373
    93
        "empty_script")]
wneuper@59472
    94
\<close>
neuper@37950
    95
neuper@37906
    96
end