src/Tools/isac/Knowledge/Equation.thy
author wenzelm
Mon, 21 Jun 2021 15:36:09 +0200
changeset 60309 70a1d102660d
parent 60306 51ec2e101e9f
child 60314 1cf9c607fa6a
permissions -rw-r--r--
more antiquotations for Isabelle/HOL consts/types, without change of semantics;
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@37950
    36
val univariate_equation_prls = 
walther@59852
    37
    Rule_Set.append_rules "univariate_equation_prls" Rule_Set.empty 
wenzelm@60294
    38
	       [\<^rule_eval>\<open>Prog_Expr.matches\<close> (Prog_Expr.eval_matches "")];
wneuper@59472
    39
\<close>
wenzelm@60289
    40
rule_set_knowledge univariate_equation_prls = \<open>Auto_Prog.prep_rls @{theory} univariate_equation_prls\<close>
neuper@37950
    41
wenzelm@60306
    42
problem pbl_equ : "equation" =
wenzelm@60306
    43
  \<open>Rule_Set.append_rules "equation_prls" Rule_Set.empty
wenzelm@60306
    44
    [\<^rule_eval>\<open>Prog_Expr.matches\<close> (Prog_Expr.eval_matches "")]\<close>
wenzelm@60306
    45
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
    46
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
    47
  Where: "matches (?a = ?b) e_e"
wenzelm@60306
    48
  Find: "solutions v_v'i'"
wenzelm@60306
    49
wenzelm@60306
    50
problem pbl_equ_univ : "univariate/equation" =
wenzelm@60306
    51
  \<open>univariate_equation_prls\<close>
wenzelm@60306
    52
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
    53
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
    54
  Where: "matches (?a = ?b) e_e"
wenzelm@60306
    55
  Find: "solutions v_v'i'"
neuper@37950
    56
wneuper@59472
    57
ML\<open>
neuper@37950
    58
(*.function for handling the cas-input "solve (x+1=2, x)":
wneuper@59279
    59
   make a model which is already in ctree-internal format.*)
neuper@37950
    60
(* val (h,argl) = strip_comb (str2term "solve (x+1=2, x)");
wneuper@59186
    61
   val (h,argl) = strip_comb ((Thm.term_of o the o (parse thy)) 
neuper@37950
    62
				  "solveTest (x+1=2, x)");
neuper@37950
    63
   *)
wenzelm@60309
    64
fun argl2dtss [Const (\<^const_name>\<open>Pair\<close>, _) $ eq $ bdv] =
wenzelm@60291
    65
    [((the o (TermC.parseNEW \<^context>)) "equality", [eq]),
wenzelm@60291
    66
     ((the o (TermC.parseNEW \<^context>)) "solveFor", [bdv]),
wenzelm@60291
    67
     ((the o (TermC.parseNEW \<^context>)) "solutions", 
wenzelm@60291
    68
      [(the o (TermC.parseNEW \<^context>)) "L"])
neuper@37950
    69
     ]
walther@59962
    70
  | argl2dtss _ = raise ERROR "Equation.ML: wrong argument for argl2dtss";
wneuper@59472
    71
\<close>
wneuper@59472
    72
setup \<open>KEStore_Elems.add_cas
wenzelm@60290
    73
  [((Thm.term_of o the o (TermC.parse @{theory})) "solveTest", 
walther@59997
    74
      (("Test", ["univariate", "equation", "test"], ["no_met"]), argl2dtss)),
wenzelm@60290
    75
    ((Thm.term_of o the o (TermC.parse @{theory})) "solve",
walther@59997
    76
      (("Isac_Knowledge", ["univariate", "equation"], ["no_met"]), argl2dtss))]\<close>
neuper@37950
    77
wenzelm@60303
    78
method met_equ : "Equation" =
wenzelm@60303
    79
  \<open>{rew_ord'="tless_true", rls'=Rule_Set.Empty, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty, crls = Atools_erls,
wenzelm@60303
    80
          errpats = [], nrls = Rule_Set.empty}\<close>
neuper@37950
    81
wenzelm@60303
    82
ML \<open>
walther@60278
    83
\<close> ML \<open>
wneuper@59472
    84
\<close>
neuper@37906
    85
end