src/Tools/isac/Knowledge/Simplify.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Thu, 22 Dec 2016 11:36:20 +0100
changeset 59279 255c853ea2f0
parent 59269 1da53d1540fe
child 59352 172b53399454
permissions -rw-r--r--
renamed Ctree.ptree --> Ctree.ctree
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
neuper@37950
     6
theory Simplify imports Atools 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@37906
    23
  (*Script-name*)
neuper@37906
    24
  SimplifyScript      :: "[real,  real] => real"
neuper@37906
    25
                  ("((Script SimplifyScript (_ =))// (_))" 9)
neuper@37906
    26
neuper@37950
    27
ML {*
neuper@37972
    28
val thy = @{theory};
s1210629013@55363
    29
*}
neuper@37950
    30
(** problems **)
s1210629013@55359
    31
setup {* KEStore_Elems.add_pbts
wneuper@59269
    32
  [(Specify.prep_pbt thy "pbl_simp" [] e_pblID
s1210629013@55339
    33
      (["simplification"],
s1210629013@55339
    34
        [("#Given" ,["Term t_t"]),
s1210629013@55339
    35
          ("#Find"  ,["normalform n_n"])],
s1210629013@55339
    36
        append_rls "e_rls" e_rls [(*for preds in where_*)], SOME "Simplify t_t", [])),
wneuper@59269
    37
    (Specify.prep_pbt thy "pbl_vereinfache" [] e_pblID
s1210629013@55339
    38
      (["vereinfachen"],
s1210629013@55339
    39
        [("#Given", ["Term t_t"]),
s1210629013@55339
    40
          ("#Find", ["normalform n_n"])],
s1210629013@55339
    41
        append_rls "e_rls" e_rls [(*for preds in where_*)], SOME "Vereinfache t_t", []))] *}
neuper@37950
    42
neuper@37950
    43
(** methods **)
s1210629013@55373
    44
setup {* KEStore_Elems.add_mets
wneuper@59269
    45
  [Specify.prep_met thy "met_tsimp" [] e_metID
s1210629013@55373
    46
	    (["simplification"],
s1210629013@55373
    47
	      [("#Given" ,["Term t_t"]),
s1210629013@55373
    48
		      ("#Find"  ,["normalform n_n"])],
s1210629013@55373
    49
		    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls=e_rls, crls = e_rls,
s1210629013@55373
    50
		      errpats = [], nrls = e_rls},
s1210629013@55373
    51
	      "empty_script")]
s1210629013@55373
    52
*}
s1210629013@55373
    53
neuper@42425
    54
ML {*
neuper@37950
    55
neuper@37950
    56
(** CAS-command **)
neuper@37950
    57
neuper@37950
    58
(*.function for handling the cas-input "Simplify (2*a + 3*a)":
wneuper@59279
    59
   make a model which is already in ctree-internal format.*)
neuper@37950
    60
(* val (h,argl) = strip_comb (str2term "Simplify (2*a + 3*a)");
wneuper@59186
    61
   val (h,argl) = strip_comb ((Thm.term_of o the o (parse (Thy_Info.get_theory "Simplify"))) 
neuper@37950
    62
				  "Simplify (2*a + 3*a)");
neuper@37950
    63
   *)
neuper@37950
    64
fun argl2dtss t =
wneuper@59186
    65
    [((Thm.term_of o the o (parse thy)) "Term", t),
wneuper@59186
    66
     ((Thm.term_of o the o (parse thy)) "normalform", 
wneuper@59186
    67
      [(Thm.term_of o the o (parse thy)) "N"])
neuper@37950
    68
     ]
s1210629013@52170
    69
  (*| argl2dtss _ = error "Simplify.ML: wrong argument for argl2dtss"*);
neuper@37950
    70
*}
s1210629013@52170
    71
setup {* KEStore_Elems.add_cas
wneuper@59186
    72
  [((Thm.term_of o the o (parse thy)) "Simplify",  
s1210629013@52170
    73
    (("Isac", ["simplification"], ["no_met"]), argl2dtss)),
wneuper@59186
    74
   ((Thm.term_of o the o (parse thy)) "Vereinfache",  
s1210629013@52170
    75
     (("Isac", ["vereinfachen"], ["no_met"]), argl2dtss))]*}
neuper@37906
    76
neuper@37906
    77
end