src/Tools/isac/Knowledge/Simplify.thy
author Walther Neuper <neuper@ist.tugraz.at>
Fri, 31 Dec 2010 14:54:02 +0100
branchdecompose-isar
changeset 38083 a1d13f3de312
parent 38031 460c24a6a6ba
child 40836 69364e021751
permissions -rw-r--r--
removed error *** nxt_add: EX itm. not(dat(itm)<=dat(ori))

the reason was: 'TERM' and 'term' are occupied by Isabelle.
TODO: improve error diagnosis already in/near prep_ori.
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};
neuper@37950
    29
neuper@37950
    30
(** problems **)
neuper@37950
    31
store_pbt
neuper@37972
    32
 (prep_pbt thy "pbl_simp" [] e_pblID
neuper@37950
    33
 (["simplification"],
neuper@38083
    34
  [("#Given" ,["Term t_t"]),
neuper@37968
    35
   ("#Find"  ,["normalform n_n"])
neuper@37950
    36
  ],
neuper@37950
    37
  append_rls "e_rls" e_rls [(*for preds in where_*)], 
neuper@37968
    38
  SOME "Simplify t_t", 
neuper@37950
    39
  []));
neuper@37950
    40
neuper@37950
    41
store_pbt
neuper@37972
    42
 (prep_pbt thy "pbl_vereinfache" [] e_pblID
neuper@37950
    43
 (["vereinfachen"],
neuper@38083
    44
  [("#Given" ,["Term t_t"]),
neuper@37968
    45
   ("#Find"  ,["normalform n_n"])
neuper@37950
    46
  ],
neuper@37950
    47
  append_rls "e_rls" e_rls [(*for preds in where_*)], 
neuper@37968
    48
  SOME "Vereinfache t_t", 
neuper@37950
    49
  []));
neuper@37950
    50
neuper@37950
    51
(** methods **)
neuper@37950
    52
neuper@37950
    53
store_met
neuper@37972
    54
    (prep_met thy "met_tsimp" [] e_metID
neuper@37950
    55
	      (["simplification"],
neuper@38083
    56
	       [("#Given" ,["Term t_t"]),
neuper@37968
    57
		("#Find"  ,["normalform n_n"])
neuper@37950
    58
		],
neuper@37950
    59
	       {rew_ord'="tless_true",
neuper@37950
    60
		rls'= e_rls, 
neuper@37950
    61
		calc = [], 
neuper@37950
    62
		srls = e_rls, 
neuper@37950
    63
		prls=e_rls,
neuper@37950
    64
		crls = e_rls, nrls = e_rls},
neuper@37950
    65
	       "empty_script"
neuper@37950
    66
	       ));
neuper@37950
    67
neuper@37950
    68
(** CAS-command **)
neuper@37950
    69
neuper@37950
    70
(*.function for handling the cas-input "Simplify (2*a + 3*a)":
neuper@37950
    71
   make a model which is already in ptree-internal format.*)
neuper@37950
    72
(* val (h,argl) = strip_comb (str2term "Simplify (2*a + 3*a)");
neuper@37968
    73
   val (h,argl) = strip_comb ((term_of o the o (parse (theory "Simplify"))) 
neuper@37950
    74
				  "Simplify (2*a + 3*a)");
neuper@37950
    75
   *)
neuper@37950
    76
fun argl2dtss t =
neuper@38083
    77
    [((term_of o the o (parse thy)) "Term", t),
neuper@37972
    78
     ((term_of o the o (parse thy)) "normalform", 
neuper@37972
    79
      [(term_of o the o (parse thy)) "N"])
neuper@37950
    80
     ]
neuper@38031
    81
  | argl2dtss _ = error "Simplify.ML: wrong argument for argl2dtss";
neuper@37950
    82
neuper@37950
    83
castab := 
neuper@37950
    84
overwritel (!castab, 
neuper@37972
    85
	    [((term_of o the o (parse thy)) "Simplify",  
neuper@37968
    86
	      (("Isac", ["simplification"], ["no_met"]), 
neuper@37950
    87
	       argl2dtss)),
neuper@37972
    88
	     ((term_of o the o (parse thy)) "Vereinfache",  
neuper@37968
    89
	      (("Isac", ["vereinfachen"], ["no_met"]), 
neuper@37950
    90
	       argl2dtss))
neuper@37950
    91
	     ]);
neuper@37950
    92
*}
neuper@37906
    93
neuper@37906
    94
end