src/Tools/isac/Knowledge/Equation.thy
author wenzelm
Wed, 26 May 2021 16:24:05 +0200
changeset 60286 31efa1b39a20
parent 60278 343efa173023
child 60289 a7b88fc19a92
permissions -rw-r--r--
command 'setup_rule' semantic equivalent for KEStore_Elems.add_rlss;
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
  
wneuper@59472
    27
text \<open>defines equation and univariate-equation
neuper@37950
    28
        created by: rlang 
neuper@37950
    29
              date: 02.09
neuper@37950
    30
        changed by: rlang
neuper@37950
    31
        last change by: rlang
neuper@37950
    32
                  date: 02.11.29
wneuper@59472
    33
        (c) by Richard Lang, 2003\<close>
neuper@37950
    34
wneuper@59472
    35
ML \<open>
neuper@37972
    36
val thy = @{theory};
walther@59879
    37
val ctxt = ThyC.to_ctxt thy;
neuper@37972
    38
neuper@37950
    39
val univariate_equation_prls = 
walther@59852
    40
    Rule_Set.append_rules "univariate_equation_prls" Rule_Set.empty 
walther@59878
    41
	       [Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches "")];
wneuper@59472
    42
\<close>
wenzelm@60286
    43
setup_rule univariate_equation_prls = \<open>Auto_Prog.prep_rls @{theory} univariate_equation_prls\<close>
wneuper@59472
    44
setup \<open>KEStore_Elems.add_pbts
walther@59973
    45
  [(Problem.prep_input thy "pbl_equ" [] Problem.id_empty
s1210629013@55339
    46
      (["equation"],
walther@59997
    47
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55339
    48
          ("#Where" ,["matches (?a = ?b) e_e"]),
s1210629013@55339
    49
          ("#Find"  ,["solutions v_v'i'"])],
walther@59878
    50
        Rule_Set.append_rules "equation_prls" Rule_Set.empty  [Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches "")],
s1210629013@55339
    51
        SOME "solve (e_e::bool, v_v)", [])),
walther@59973
    52
    (Problem.prep_input thy "pbl_equ_univ" [] Problem.id_empty
walther@59997
    53
      (["univariate", "equation"],
walther@59997
    54
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55339
    55
          ("#Where" ,["matches (?a = ?b) e_e"]),
s1210629013@55339
    56
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59472
    57
        univariate_equation_prls, SOME "solve (e_e::bool, v_v)", []))]\<close>
neuper@37950
    58
neuper@37950
    59
wneuper@59472
    60
ML\<open>
neuper@37950
    61
(*.function for handling the cas-input "solve (x+1=2, x)":
wneuper@59279
    62
   make a model which is already in ctree-internal format.*)
neuper@37950
    63
(* val (h,argl) = strip_comb (str2term "solve (x+1=2, x)");
wneuper@59186
    64
   val (h,argl) = strip_comb ((Thm.term_of o the o (parse thy)) 
neuper@37950
    65
				  "solveTest (x+1=2, x)");
neuper@37950
    66
   *)
neuper@41972
    67
fun argl2dtss [Const ("Product_Type.Pair", _) $ eq $ bdv] =
wneuper@59389
    68
    [((the o (TermC.parseNEW ctxt)) "equality", [eq]),
wneuper@59389
    69
     ((the o (TermC.parseNEW ctxt)) "solveFor", [bdv]),
wneuper@59389
    70
     ((the o (TermC.parseNEW ctxt)) "solutions", 
wneuper@59389
    71
      [(the o (TermC.parseNEW ctxt)) "L"])
neuper@37950
    72
     ]
walther@59962
    73
  | argl2dtss _ = raise ERROR "Equation.ML: wrong argument for argl2dtss";
wneuper@59472
    74
\<close>
wneuper@59472
    75
setup \<open>KEStore_Elems.add_cas
wneuper@59389
    76
  [((Thm.term_of o the o (TermC.parse thy)) "solveTest", 
walther@59997
    77
      (("Test", ["univariate", "equation", "test"], ["no_met"]), argl2dtss)),
wneuper@59389
    78
    ((Thm.term_of o the o (TermC.parse thy)) "solve",
walther@59997
    79
      (("Isac_Knowledge", ["univariate", "equation"], ["no_met"]), argl2dtss))]\<close>
neuper@37950
    80
neuper@37950
    81
wneuper@59472
    82
setup \<open>KEStore_Elems.add_mets
walther@60154
    83
    [MethodC.prep_input thy "met_equ" [] MethodC.id_empty
s1210629013@55373
    84
	    (["Equation"], [],
walther@59852
    85
	      {rew_ord'="tless_true", rls'=Rule_Set.Empty, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty, crls = Atools_erls,
walther@59852
    86
          errpats = [], nrls = Rule_Set.empty},
wneuper@59545
    87
        @{thm refl})]
walther@60278
    88
\<close> ML \<open>
walther@60278
    89
\<close> ML \<open>
wneuper@59472
    90
\<close>
neuper@37906
    91
end