src/Tools/isac/ProgLang/Script.thy
author Walther Neuper <neuper@ist.tugraz.at>
Sun, 16 Jun 2013 12:31:41 +0200
changeset 48880 ea0c337066d9
parent 42400 dcacb8077a98
child 52139 511fc271f783
permissions -rw-r--r--
Isabelle2011 --> 2012 intermediate

Made paths absolute wrt ~~, which probably will be shortened later again.
'isabelle usedir -b HOL Isac' successfully creates a heap;
still no 'thehier' created in 'Knowledge/Build_Thydata.thy',
because these are not required by the math-engine.
     1 (* Title:  tactics, tacticals etc. for scripts
     2    Author: Walther Neuper 000224
     3    (c) due to copyright terms
     4  *)
     5 
     6 theory Script imports Tools begin
     7 
     8 typedecl
     9   ID	(* identifiers for thy, ruleset,... *)
    10 
    11 typedecl
    12   arg	(* argument of subproblem           *)
    13 
    14 consts
    15 
    16 (*types of subproblems' arguments*)
    17   REAL        :: "real => arg"
    18   REAL_LIST   :: "(real list) => arg"
    19   REAL_SET    :: "(real set) => arg"
    20   BOOL        :: "bool => arg"
    21   BOOL_LIST   :: "(bool list) => arg"
    22   REAL_REAL   :: "(real => real) => arg"
    23 
    24 (*tactics*)
    25   Rewrite      :: "[ID, bool, 'a] => 'a"
    26   Rewrite'_Inst:: "[(real * real) list, ID, bool, 'a] => 'a"
    27 			             ("(Rewrite'_Inst (_ _ _))" 11)
    28                                      (*without last argument ^^ for @@*)
    29   Rewrite'_Set :: "[ID, bool, 'a] => 'a" ("(Rewrite'_Set (_ _))" 11)
    30   Rewrite'_Set'_Inst
    31                :: "[(real * real) list, ID, bool, 'a] => 'a"
    32 		                     ("(Rewrite'_Set'_Inst (_ _ _))" 11)
    33                                      (*without last argument ^^ for @@*)
    34   Calculate    :: "[ID, 'a] => 'a" (*WN100816 PLUS, TIMES, POWER miss.in scr*)
    35   Calculate1   :: "[ID, 'a] => 'a" (*FIXXXME: unknown to script-interpreter*)
    36 
    37   (* WN0509 substitution now is rewriting by a list of terms (of type bool)
    38   Substitute   :: "[(real * real) list, 'a] => 'a"*)
    39   Substitute   :: "[bool list, 'a] => 'a"
    40 
    41   Map          :: "['a => 'b, 'a list] => 'b list"
    42   Tac          :: "ID => 'a"         (*deprecated; used in Test.ML*)
    43   Check'_elementwise ::
    44 		  "['a list, 'b set] => 'a list"
    45                                      ("Check'_elementwise (_ _)" 11)
    46   Take         :: "'a => 'a"         (*for non-var args as long as no 'o'*)
    47   SubProblem   :: "[ID * ID list * ID list, arg list] => 'a"
    48 
    49   Or'_to'_List :: "bool => 'a list"  ("Or'_to'_List (_)" 11)
    50   (*=========== record these ^^^ in 'tacs' in Script.ML =========*)
    51 
    52   Assumptions  :: bool
    53   Problem      :: "[ID * ID list] => 'a"
    54 
    55 (*special formulas for frontend 'CAS format'*)
    56   Subproblem   :: "(ID * ID list) => 'a" 
    57 
    58 (*script-expressions (tacticals)*)
    59   Seq      :: "['a => 'a, 'a => 'a, 'a] => 'a" (infixr "@@" 10) (*@ used*)
    60   Try      :: "['a => 'a, 'a] => 'a"
    61   Repeat   :: "['a => 'a, 'a] => 'a" 
    62   Or       :: "['a => 'a, 'a => 'a, 'a] => 'a" (infixr "Or" 10)
    63   While    :: "[bool, 'a => 'a, 'a] => 'a"     ("((While (_) Do)//(_))" 9)
    64 (*WN100723 because of "Error in syntax translation" below...
    65         (*'b => bool doesn't work with "contains_root _"*)
    66   Letpar   :: "['a, 'a => 'b] => 'b"
    67   (*--- defined in Isabelle/scr/HOL/HOL.thy:
    68   Let      :: "['a, 'a => 'b] => 'b"
    69   "_Let"   :: "[letbinds, 'a] => 'a"       ("(let (_)/ in (_))" 10)
    70   If       :: "[bool, 'a, 'a] => 'a"       ("(if (_)/ then (_)/ else (_))" 10)
    71   %x. P x  .. lambda is defined in Isabelles meta logic
    72   --- *)
    73 *)
    74   failtac :: 'a
    75   idletac :: 'a
    76   (*... + RECORD IN 'screxpr' in Script.ML *)
    77 
    78 (*for scripts generated automatically from rls*)
    79   Stepwise      :: "['z,     'z] => 'z" ("((Script Stepwise (_   =))// (_))" 9)
    80   Stepwise'_inst:: "['z,real,'z] => 'z" 
    81 	("((Script Stepwise'_inst (_ _ =))// (_))" 9)
    82 
    83 
    84 (*SHIFT -> resp.thys ----vvv---------------------------------------------*)
    85 (*script-names: initial capital letter,
    86 		type of last arg (=script-body) == result-type !
    87   Xxxx       :: script ids, duplicate result-type 'r in last argument:
    88              "['a, ... , \
    89 	       \         'r] => 'r
    90 *)
    91 			    
    92 (*make'_solution'_set :: "bool => bool list" 
    93 			("(make'_solution'_set (_))" 11)    
    94 					   
    95   max'_on'_interval
    96              :: "[ID * (ID list) * ID, bool,real,real set] => real"
    97                ("(max'_on'_interval (_)/ (_ _ _))" 9)
    98   find'_vals
    99              :: "[ID * (ID list) * ID,
   100 		  real,real,real,real,bool list] => bool list"
   101                ("(find'_vals (_)/ (_ _ _ _ _))" 9)
   102 
   103   make'_fun  :: "[ID * (ID list) * ID, real,real,bool list] => bool"
   104                ("(make'_fun (_)/ (_ _ _))" 9)
   105 
   106   solve'_univar
   107              :: "[ID * (ID list) * ID, bool,real] => bool list"
   108                ("(solve'_univar (_)/ (_ _ ))" 9)
   109   solve'_univar'_err
   110              :: "[ID * (ID list) * ID, bool,real,bool] => bool list"
   111                ("(solve'_univar (_)/ (_ _ _))" 9)
   112 ----------*)
   113 
   114   Testeq     :: "[bool, bool] => bool"
   115                ("((Script Testeq (_ =))// 
   116                   (_))" 9)
   117   
   118   Testeq2    :: "[bool, bool list] => bool list"
   119                ("((Script Testeq2 (_ =))// 
   120                   (_))" 9)
   121   
   122   Testterm   :: "[real, real] => real"
   123                ("((Script Testterm (_ =))// 
   124                   (_))" 9)
   125   
   126   Testchk    :: "[bool, real, real list] => real list"
   127                ("((Script Testchk (_ _ =))// 
   128                   (_))" 9)
   129   (*... + RECORD IN 'subpbls' in Script.ML *)
   130 (*SHIFT -> resp.thys ----^^^----------------------------*)
   131 
   132 (*Makarius 10.03
   133 syntax
   134 
   135   "_Letpar"     :: "[letbinds, 'a] => 'a"      ("(letpar (_)/ in (_))" 10)
   136 
   137 translations
   138 
   139   "_Letpar (_binds b bs) e"  == "_Letpar b (_Letpar bs e)"
   140   "letpar x = a in e"        == "Letpar a (%x. e)"
   141 *** Error in syntax translation rule: rhs contains extra variables
   142 *** ("_Letpar" ("_bind" x a) e)  ->  (Letpar a ("_abs" x e))
   143 *** At command "translations" (line 140 of "/usr/local/isabisac/src/Pure/isac/Scripts/Script").
   144 *)
   145 
   146 ML {* (*the former Script.ML*)
   147 
   148 (*--vvv----- SHIFT? or delete ?*)
   149 val IDTyp = Type("Script.ID",[]);
   150 
   151 val tacs = Unsynchronized.ref (distinct (remove op = ""
   152   ["Calculate",
   153    "Rewrite","Rewrite'_Inst","Rewrite'_Set","Rewrite'_Set'_Inst",
   154    "Substitute","Tac","Check'_elementswise",
   155    "Take","Subproblem","Or'_to'_List"]));
   156 
   157 val screxpr = Unsynchronized.ref (distinct (remove op = ""
   158   ["Let","If","Repeat","While","Try","Or"]));
   159 
   160 val listfuns = Unsynchronized.ref [(*_all_ functions in Isa99.List.thy *)
   161     "@","filter","concat","foldl","hd","last","set","list_all",
   162     "map","mem","nth","list_update","take","drop",	
   163     "takeWhile","dropWhile","tl","butlast",
   164     "rev","zip","upt","remdups","nodups","replicate",
   165 
   166     "Cons","Nil"];
   167 
   168 val scrfuns = Unsynchronized.ref ([]: string list);
   169 
   170 val listexpr = Unsynchronized.ref (union op = (!listfuns) (!scrfuns));
   171 
   172 val notsimp = Unsynchronized.ref 
   173   (distinct (remove op = "" 
   174              (!tacs @ !screxpr @ (*!subpbls @*) !scrfuns @ !listfuns)));
   175 
   176 val negotiable = Unsynchronized.ref ((!tacs (*@ !subpbls*)));
   177 
   178 val tacpbl = Unsynchronized.ref
   179   (distinct (remove op = "" (!tacs (*@ !subpbls*))));
   180 (*--^^^----- SHIFT? or delete ?*)
   181 
   182 *}
   183 
   184 end