src/Tools/isac/Knowledge/Simplify.thy
author wenzelm
Mon, 21 Jun 2021 21:53:23 +0200
changeset 60314 1cf9c607fa6a
parent 60306 51ec2e101e9f
child 60331 40eb8aa2b0d6
permissions -rw-r--r--
Isar command 'cas' as front-end for KEStore_Elems.add_cas, without change of semantics;
more uniform set_data;
neuper@37906
     1
(* simplification of terms
neuper@37906
     2
   author: Walther Neuper 050912
neuper@37906
     3
   (c) due to copyright terms
neuper@37906
     4
*)
neuper@37906
     5
wneuper@59424
     6
theory Simplify imports Base_Tools begin
neuper@37906
     7
neuper@37906
     8
consts
neuper@37906
     9
neuper@37906
    10
  (*descriptions in the related problem*)
neuper@38083
    11
  Term        :: "real => una"
neuper@38083
    12
(*TERM -->   Const ("Pure.term", "RealDef.real => prop")  (*!!!*) $ 
neuper@38083
    13
               Free ("ttt", "RealDef.real")
neuper@38083
    14
  term -->   Free ("term", "RealDef.real => RealDef.real") $
neuper@38083
    15
               Free ("ttt", "RealDef.real")
neuper@38083
    16
    but 'term' is a keyword in *.thy*)
neuper@37950
    17
  normalform  :: "real => una"
neuper@37906
    18
neuper@37906
    19
  (*the CAS-command*)
neuper@37906
    20
  Simplify    :: "real => real"  (*"Simplify (1+2a+3+4a)*)
neuper@37906
    21
  Vereinfache :: "real => real"  (*"Vereinfache (1+2a+3+4a)*)
neuper@37906
    22
neuper@37950
    23
(** problems **)
wenzelm@60306
    24
wenzelm@60306
    25
problem pbl_simp : "simplification" =
wenzelm@60306
    26
  \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)]\<close>
wenzelm@60306
    27
  CAS: "Simplify t_t"
wenzelm@60306
    28
  Given: "Term t_t"
wenzelm@60306
    29
  Find: "normalform n_n"
wenzelm@60306
    30
wenzelm@60306
    31
problem pbl_vereinfache : "vereinfachen" =
wenzelm@60306
    32
  \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)]\<close>
wenzelm@60306
    33
  CAS: "Vereinfache t_t"
wenzelm@60306
    34
  Given: "Term t_t"
wenzelm@60306
    35
  Find: "normalform n_n"
neuper@37950
    36
neuper@37950
    37
(** methods **)
wenzelm@60303
    38
wenzelm@60303
    39
method met_tsimp : "simplification" =
wenzelm@60303
    40
  \<open>{rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty, crls = Rule_Set.empty,
wenzelm@60303
    41
    errpats = [], nrls = Rule_Set.empty}\<close>
wenzelm@60303
    42
  Given: "Term t_t"
wenzelm@60303
    43
  Find: "normalform n_n"
s1210629013@55373
    44
wneuper@59472
    45
ML \<open>
neuper@37950
    46
neuper@37950
    47
(** CAS-command **)
neuper@37950
    48
neuper@37950
    49
(*.function for handling the cas-input "Simplify (2*a + 3*a)":
wneuper@59279
    50
   make a model which is already in ctree-internal format.*)
neuper@37950
    51
(* val (h,argl) = strip_comb (str2term "Simplify (2*a + 3*a)");
wneuper@59352
    52
   val (h,argl) = strip_comb ((Thm.term_of o the o (parse (Thy_Info_get_theory "Simplify"))) 
neuper@37950
    53
				  "Simplify (2*a + 3*a)");
neuper@37950
    54
   *)
neuper@37950
    55
fun argl2dtss t =
wenzelm@60290
    56
    [((Thm.term_of o the o (TermC.parse @{theory})) "Term", t),
wenzelm@60290
    57
     ((Thm.term_of o the o (TermC.parse @{theory})) "normalform", 
wenzelm@60290
    58
      [(Thm.term_of o the o (TermC.parse @{theory})) "N"])
neuper@37950
    59
     ]
walther@59962
    60
  (*| argl2dtss _ = raise ERROR "Simplify.ML: wrong argument for argl2dtss"*);
wneuper@59472
    61
\<close>
wenzelm@60314
    62
cas Simplify = \<open>argl2dtss\<close>
wenzelm@60314
    63
  Problem: "simplification"
wenzelm@60314
    64
cas Vereinfache = \<open>argl2dtss\<close>
wenzelm@60314
    65
  Problem: "vereinfachen"
neuper@37906
    66
walther@60269
    67
ML \<open>
walther@60269
    68
\<close> ML \<open>
walther@60269
    69
\<close> ML \<open>
walther@60269
    70
\<close>
neuper@37906
    71
end