src/Tools/isac/Interpret/script.sml
author Walther Neuper <wneuper@ist.tugraz.at>
Thu, 07 Mar 2019 17:29:30 +0100
changeset 59513 deb1efba3119
parent 59503 3f58f3a2c3e6
child 59515 59b219e92469
permissions -rw-r--r--
[-Test_Isac] funpack: Const AA enforces care with thys for parsing
     1 (* Title:  interpreter for scripts
     2    Author: Walther Neuper 2000
     3    (c) due to copyright terms
     4 *)
     5 
     6 signature LUCAS_INTERPRETER =
     7 sig
     8 
     9   type step = Tac.tac_ * Generate.mout * Ctree.ctree * Ctree.pos' * Ctree.pos' list
    10   datatype locate = NotLocatable | Steps of Selem.istate * step list
    11   
    12   val next_tac : (*diss: next-tactic-function*)
    13     Rule.theory' * Rule.rls -> Ctree.state -> Rule.scr -> Selem.istate * 'a -> Tac.tac_ * (Selem.istate * 'a) * (term * Selem.safe)
    14   val locate_gen : (*diss: locate-function*)
    15     Rule.theory' * Rule.rls -> Tac.tac_ -> Ctree.state -> Rule.scr * 'a -> Selem.istate * Proof.context -> locate
    16                          
    17 (* can these functions be local to Lucin or part of LItools ? *)
    18   val sel_rules : Ctree.ctree -> Ctree.pos' -> Tac.tac list 
    19   val init_form : 'a -> Rule.scr -> (term * term) list -> term option
    20   val tac_2tac : Tac.tac_ -> Tac.tac
    21   val init_scrstate : theory -> Model.itm list -> Celem.metID -> Selem.istate * Proof.context * Rule.scr
    22   val from_pblobj' : Rule.theory' -> Ctree.pos' -> Ctree.ctree -> Rule.rls * (Selem.istate * Proof.context) * Rule.scr
    23   val from_pblobj_or_detail' : Rule.theory' -> Ctree.pos' -> Ctree.ctree -> 
    24     Rule.rls * (Selem.istate * Proof.context) * Rule.scr
    25   val rule2thm'' : Rule.rule -> Celem.thm''
    26   val rule2rls' : Rule.rule -> string
    27 (* ---- for tests only: shifted from below to remove the Warning "unused" at fun.def. --------- *)
    28   val sel_appl_atomic_tacs : Ctree.ctree -> Ctree.pos' -> Tac.tac list
    29 (*/-------------------------------------------------------- ! aktivate for Test_Isac BEGIN ---\* )
    30   datatype asap = Aundef | AssOnly | AssGen
    31   datatype appy = Appy of Tac.tac_ * Selem.scrstate | Napp of LTool.env | Skip of term * LTool.env
    32   datatype appy_ = Napp_ | Skip_
    33   val appy : Rule.theory' * Rule.rls -> Ctree.state -> LTool.env -> Celem.lrd list -> term ->
    34     term option -> term -> appy
    35   val formal_args : term -> term list
    36   val body_of : term -> term
    37   val get_stac : 'a -> term -> term option 
    38   val go : Celem.loc_ -> term -> term
    39   val handle_leaf : string -> Rule.theory' -> Rule.rls -> LTool.env -> term option -> term -> term -> 
    40     term option * LTool.stacexpr
    41   val id_of_scr : term -> string
    42    val is_spec_pos : Ctree.pos_ -> bool
    43   val itms2args : 'a -> Celem.metID -> Model.itm list -> term list
    44   val nstep_up : Rule.theory' * Rule.rls -> Ctree.state -> Rule.scr -> LTool.env ->
    45     Celem.lrd list -> appy_ -> term option -> term -> appy
    46   val nxt_up: Rule.theory' * Rule.rls -> Ctree.state -> Rule.scr -> LTool.env ->
    47     Celem.lrd list -> appy_ -> term -> term option -> term -> appy
    48   val stac2tac : Ctree.ctree -> theory -> term -> Tac.tac
    49   val stac2tac_ : Ctree.ctree -> theory -> term -> Tac.tac * Tac.tac_
    50   val upd_env_opt : LTool.env -> term option * term -> LTool.env
    51 ( *\--- ! aktivate for Test_Isac END ----------------------------------------------------------/*)
    52 
    53 (*----- unused code, kept as hints to design ideas ---------------------------------------------*)
    54   type assoc
    55   val assoc2str : assoc -> string
    56 end 
    57 
    58 (* traces the leaves (ie. non-tactical nodes) of Prog found by next_tac, see "and scr" *)   
    59 val trace_script = Unsynchronized.ref false; (* TODO: how are traces done in Isabelle? *)
    60 
    61 structure Lucin(*: LUCAS_INTERPRETER*) =
    62 struct
    63 open Ctree
    64 (*TODO open Celem for L,R,D;
    65   the few other Celem.items are acceptable: metID, e_metID, assoc_thy, metID2str, maxthy, thm''*)
    66 
    67 (**)
    68 (* data for creating a new node in ctree; designed for use as:
    69    fun ass* scrstate steps = / ... case ass* scrstate steps of /
    70    Assoc (scrstate, steps) => ... ass* scrstate steps *)
    71 type step =
    72     Tac.tac_        (*transformed from associated tac                   *)
    73     * Generate.mout (*result with indentation etc.                      *)
    74     * ctree         (*containing node created by tac_ + resp. scrstate  *)
    75     * pos'          (*position in ctree; ctree * pos' is the proofstate *)
    76     * pos' list;    (*of ctree-nodes probably cut (by fst tac_)         *)
    77 
    78 fun rule2thm'' (Rule.Thm (id, thm)) = (id, thm)
    79   | rule2thm'' r = error ("rule2thm': not defined for " ^ Rule.rule2str r);
    80 fun rule2rls' (Rule.Rls_ rls) = Rule.id_rls rls
    81   | rule2rls' r = error ("rule2rls': not defined for " ^ Rule.rule2str r);
    82 
    83 (*.makes a (rule,term) list to a Step (m, mout, pt', p', cid) for solve;
    84    complicated with current t in rrlsstate.*)
    85 fun rts2steps steps ((pt, p), (f, f'', rss, rts), (thy', ro, er, pa)) [(r, (f', am))] =
    86       let
    87         val thy = Celem.assoc_thy thy'
    88         val ctxt = get_ctxt pt p |> Stool.insert_assumptions am
    89 	      val m = Tac.Rewrite' (thy', ro, er, pa, rule2thm'' r, f, (f', am))
    90 	      val is = Selem.RrlsState (f', f'', rss, rts)
    91 	      val p = case p of (_, Frm) => p | (p', Res) => (lev_on p', Res) | _ => error "rts2steps: p1"
    92 	      val (p', cid, mout, pt') = Generate.generate1 thy m (is, ctxt) p pt
    93       in (is, (m, mout, pt', p', cid) :: steps) end
    94   | rts2steps steps ((pt, p) ,(f, f'', rss, rts), (thy', ro, er, pa)) ((r, (f', am)) :: rts') =
    95       let
    96         val thy = Celem.assoc_thy thy'
    97         val ctxt = get_ctxt pt p |> Stool.insert_assumptions am
    98 	      val m = Tac.Rewrite' (thy', ro, er, pa, rule2thm'' r, f, (f', am))
    99 	      val is = Selem.RrlsState (f', f'', rss, rts)
   100 	      val p = case p of (_, Frm) => p | (p', Res) => (lev_on p', Res) | _ => error "rts2steps: p1"
   101 	      val (p', cid, mout, pt') = Generate.generate1 thy m (is, ctxt) p pt
   102       in rts2steps ((m, mout, pt', p', cid)::steps) 
   103 		    ((pt', p'), (f', f'', rss, rts), (thy', ro, er, pa)) rts'
   104 		  end
   105   | rts2steps _ _ _ = error "rts2steps: uncovered fun-def"
   106 
   107 (* functions for the environment stack: NOT YET IMPLEMENTED
   108 fun accessenv id es = the (assoc ((top es) : LTool.env, id))
   109     handle _ => error ("accessenv: " ^ free2str id ^ " not in LTool.env");
   110 fun updateenv id vl (es : LTool.env stack) = 
   111     (push (overwrite(top es, (id, vl))) (pop es)) : LTool.env stack;
   112 fun pushenv id vl (es : LTool.env stack) = 
   113     (push (overwrite(top es, (id, vl))) es) : LTool.env stack;
   114 val popenv = pop : LTool.env stack -> LTool.env stack;
   115 *)
   116 
   117 fun de_esc_underscore str =
   118   let
   119     fun scan [] = []
   120     | scan (s :: ss) = if s = "'" then (scan ss) else (s :: (scan ss))
   121   in (implode o scan o Symbol.explode) str end;
   122 
   123 (*go at a location in a script and fetch the contents*)
   124 fun go [] t = t
   125   | go (Celem.D :: p) (Abs(_, _, t0)) = go (p : Celem.loc_) t0
   126   | go (Celem.L :: p) (t1 $ _) = go p t1
   127   | go (Celem.R :: p) (_ $ t2) = go p t2
   128   | go l _ = error ("go: no " ^ Celem.loc_2str l);
   129 
   130 (*.get argument of first stactic in a script for init_form.*)
   131 fun get_stac thy (_ $ body) =
   132   let
   133     fun get_t y (Const ("Script.Seq",_) $ e1 $ e2) a = 
   134     	  (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   135       | get_t y (Const ("Script.Seq",_) $ e1 $ e2 $ a) _ = 
   136     	  (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   137       | get_t y (Const ("Script.Try",_) $ e) a = get_t y e a
   138       | get_t y (Const ("Script.Try",_) $ e $ a) _ = get_t y e a
   139       | get_t y (Const ("Script.Repeat",_) $ e) a = get_t y e a
   140       | get_t y (Const ("Script.Repeat",_) $ e $ a) _ = get_t y e a
   141       | get_t y (Const ("Script.Or",_) $e1 $ e2) a =
   142     	  (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   143       | get_t y (Const ("Script.Or",_) $e1 $ e2 $ a) _ =
   144     	  (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   145       | get_t y (Const ("Script.While",_) $ _ $ e) a = get_t y e a
   146       | get_t y (Const ("Script.While",_) $ _ $ e $ a) _ = get_t y e a
   147       | get_t y (Const ("Script.Letpar",_) $ e1 $ Abs (_, _, e2)) a = 
   148     	  (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   149     (*| get_t y (Const ("HOL.Let",_) $ e1 $ Abs (_,_,e2)) a =
   150 	      (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   151       | get_t y (Abs (_,_,e)) a = get_t y e a*)
   152       | get_t y (Const ("HOL.Let",_) $ e1 $ Abs (_, _, _)) a =
   153     	get_t y e1 a (*don't go deeper without evaluation !*)
   154       | get_t _ (Const ("If", _) $ _ $ _ $ _) _ = NONE
   155     	(*(case get_t y e1 a of NONE => get_t y e2 a | la => la)*)
   156     
   157       | get_t _ (Const ("Script.Rewrite",_) $ _ $ _ $ a) _ = SOME a
   158       | get_t _ (Const ("Script.Rewrite",_) $ _ $ _    ) a = SOME a
   159       | get_t _ (Const ("Script.Rewrite'_Inst",_) $ _ $ _ $ _ $ a) _ = SOME a
   160       | get_t _ (Const ("Script.Rewrite'_Inst",_) $ _ $ _ $ _ )    a = SOME a
   161       | get_t _ (Const ("Script.Rewrite'_Set",_) $ _ $ _ $ a) _ = SOME a
   162       | get_t _ (Const ("Script.Rewrite'_Set",_) $ _ $ _ )    a = SOME a
   163       | get_t _ (Const ("Script.Rewrite'_Set'_Inst",_) $ _ $ _ $ _ $a)_ =SOME a
   164       | get_t _ (Const ("Script.Rewrite'_Set'_Inst",_) $ _ $ _ $ _ )  a =SOME a
   165       | get_t _ (Const ("Script.Calculate",_) $ _ $ a) _ = SOME a
   166       | get_t _ (Const ("Script.Calculate",_) $ _ )    a = SOME a
   167     
   168       | get_t _ (Const ("Script.Substitute",_) $ _ $ a) _ = SOME a
   169       | get_t _ (Const ("Script.Substitute",_) $ _ )    a = SOME a
   170     
   171       | get_t _ (Const ("Script.SubProblem",_) $ _ $ _) _ = NONE
   172 
   173       | get_t _ _ _ = ((*tracing ("### get_t yac: list-expr "^(term2str x));*) NONE)
   174     in get_t thy body Rule.e_term end
   175   | get_stac _ t = error ("get_stac: no fun-def. for " ^ Rule.term2str t);
   176     
   177 fun init_form thy (Rule.Prog sc) env =
   178     (case get_stac thy sc of NONE => NONE | SOME stac => SOME (subst_atomic env stac))
   179   | init_form _ _ _ = error "init_form: no match";
   180 
   181 (* get the arguments of the script out of the scripts parsetree *)
   182 (* version for later switch to partial_function *)
   183 fun formal_args tm = (tm
   184   |> HOLogic.dest_eq
   185   |> fst
   186   |> strip_comb
   187   |> snd)
   188   handle TERM _ => raise TERM ("formal_args", [tm])
   189 (* version to be replaced during switch to partial_function *)
   190 fun formal_args scr = (fst o split_last o snd o strip_comb) scr;
   191 
   192 (* get the body of a program *)
   193 (* version for later switch to partial_function *)
   194 fun body_of tm = (tm
   195   |> HOLogic.dest_eq
   196   |> snd)
   197   handle TERM _ => raise TERM ("body_of", [tm])
   198 (* version introduced BEFORE switch to partial_function *)
   199 fun body_of (_ $ body) = body
   200   | body_of t = raise TERM ("body_of", [t])
   201 
   202 
   203 (* get the identifier of the script out of the scripts parsetree *)
   204 fun id_of_scr sc = (id_of o fst o strip_comb) sc;
   205 
   206 (*WN020526: not clear, when a is available in ass_up for eval_true*)
   207 (*WN060906: in "fun handle_leaf" eg. uses "SOME M__"(from some PREVIOUS
   208   curried Rewrite) for CURRENT value (which may be different from PREVIOUS);
   209   thus "NONE" must be set at the end of currying (ill designed anyway)*)
   210 fun upd_env_opt env (SOME a, v) = LTool.upd_env env (a, v)
   211   | upd_env_opt env (NONE, _) = 
   212       ((*tracing ("*** upd_env_opt: (NONE," ^ term2str v ^ ")");*) env);
   213 
   214 type dsc = typ; (* <-> nam..unknow in Descript.thy *)
   215 
   216 (*.create the actual parameters (args) of script: their order 
   217   is given by the order in met.pat .*)
   218 (*WN.5.5.03: ?: does this allow for different descriptions ???
   219              ?: why not taken from formal args of script ???
   220 !: FIXXXME penv: push it here in itms2args into script-evaluation*)
   221 (* val (thy, mI, itms) = (thy, metID, itms);
   222    *)
   223 val errmsg = "ERROR: the guard is missing (#ppc in 'type met' added in prep_met)."
   224 fun itms2args _ mI itms =
   225   let
   226     val mvat = Model.max_vt itms
   227     fun okv mvat (_, vats, b, _, _) = member op = vats mvat andalso b
   228     val itms = filter (okv mvat) itms
   229     fun test_dsc d (_, _, _, _, itm_) = (d = Model.d_in itm_)
   230     fun itm2arg itms (_,(d,_)) =
   231         case find_first (test_dsc d) itms of
   232           NONE => error ("itms2args: '" ^ Rule.term2str d ^ "' not in itms")
   233         | SOME (_, _, _, _, itm_) => Model.penvval_in itm_
   234       (*| SOME (_,_,_,_,itm_) => mk_arg thy (Model.d_in itm_) (ts_in itm_);
   235             penv postponed; presently penv holds already LTool.env for script*)
   236     val pats = (#ppc o Specify.get_met) mI
   237     val _ = if pats = [] then raise ERROR errmsg else ()
   238   in (flat o (map (itm2arg itms))) pats end;
   239 
   240 (* convert a script-tac 'stac' to 'tac' for users;
   241    for "Script.SubProblem" also create a 'tac_' for internal use. FIXME separate?
   242    if stac is an initac, then convert to a 'tac_' (as required in appy).
   243    arg ctree for pushing the thy specified in rootpbl into subpbls    *)
   244 fun stac2tac_ _ thy (Const ("Script.Rewrite", _) $ thmID $ _ $ _) =
   245     let
   246       val tid = HOLogic.dest_string thmID
   247     in (Tac.Rewrite (tid, Rewrite.assoc_thm'' thy tid), Tac.Empty_Tac_) end
   248   | stac2tac_ _ thy (Const ("Script.Rewrite'_Inst", _) $ sub $ thmID $ _ $ _) =
   249     let
   250       val tid = HOLogic.dest_string thmID
   251     in (Tac.Rewrite_Inst (Selem.subst'_to_sube sub, (tid, Rewrite.assoc_thm'' thy tid)), Tac.Empty_Tac_) end
   252   | stac2tac_ _ _ (Const ("Script.Rewrite'_Set",_) $ rls $ _ $ _) =
   253      (Tac.Rewrite_Set (HOLogic.dest_string rls), Tac.Empty_Tac_)
   254   | stac2tac_ _ _ (Const ("Script.Rewrite'_Set'_Inst", _) $ sub $ rls $ _ $ _) =
   255       (Tac.Rewrite_Set_Inst (Selem.subst'_to_sube sub, HOLogic.dest_string rls), Tac.Empty_Tac_)
   256   | stac2tac_ _ _ (Const ("Script.Calculate", _) $ op_ $ _) =
   257       (Tac.Calculate (HOLogic.dest_string op_), Tac.Empty_Tac_)
   258   | stac2tac_ _ _ (Const ("Script.Take", _) $ t) = (Tac.Take (Rule.term2str t), Tac.Empty_Tac_)
   259   | stac2tac_ _ _ (Const ("Script.Substitute", _) $ isasub $ _) =
   260     (Tac.Substitute ((Selem.subte2sube o TermC.isalist2list) isasub), Tac.Empty_Tac_)
   261   | stac2tac_ _ thy (Const("Script.Check'_elementwise", _) $ _ $ 
   262     (Const ("Set.Collect", _) $ Abs (_, _, pred))) =
   263       (Tac.Check_elementwise (Rule.term_to_string''' thy pred), Tac.Empty_Tac_)
   264   | stac2tac_ _ _ (Const("Script.Or'_to'_List", _) $ _ ) = (Tac.Or_to_List, Tac.Empty_Tac_)
   265 
   266     (*compare "| assod _ (Subproblem'"*)
   267   | stac2tac_ pt _ (stac as Const ("Script.SubProblem", _) $
   268      (Const ("Product_Type.Pair", _) $ dI' $ (Const ("Product_Type.Pair", _) $ pI' $ mI')) $ ags') =
   269     let
   270       val dI = HOLogic.dest_string dI';
   271       val thy = Celem.maxthy (Celem.assoc_thy dI) (rootthy pt);
   272 	    val pI = pI' |> HOLogic.dest_list |> map HOLogic.dest_string;
   273 	    val mI = mI' |> HOLogic.dest_list |> map HOLogic.dest_string;
   274 	    val ags = TermC.isalist2list ags';
   275 	    val (pI, pors, mI) = 
   276 	      if mI = ["no_met"] 
   277 	      then
   278           let
   279             val pors = (Chead.match_ags thy ((#ppc o Specify.get_pbt) pI) ags)
   280 		          handle ERROR "actual args do not match formal args" 
   281 			        => (Chead.match_ags_msg pI stac ags(*raise exn*); [])
   282 		        val pI' = Specify.refine_ori' pors pI;
   283 		      in (pI', pors (* refinement over models with diff.prec only *), 
   284 		          (hd o #met o Specify.get_pbt) pI') end
   285 	      else (pI, (Chead.match_ags thy ((#ppc o Specify.get_pbt) pI) ags)
   286 		      handle ERROR "actual args do not match formal args"
   287 		      => (Chead.match_ags_msg pI stac ags(*raise exn*); []), mI);
   288       val (fmz_, vals) = Chead.oris2fmz_vals pors;
   289 	    val {cas,ppc,thy,...} = Specify.get_pbt pI
   290 	    val dI = Rule.theory2theory' thy (*.take dI from _refined_ pbl.*)
   291 	    val dI = Rule.theory2theory' (Celem.maxthy (Celem.assoc_thy dI) (rootthy pt));
   292       val ctxt = dI |> Rule.Thy_Info_get_theory |> Proof_Context.init_global |> Stool.declare_constraints' vals
   293 	    val hdl =
   294         case cas of
   295 		      NONE => LTool.pblterm dI pI
   296 		    | SOME t => subst_atomic ((Chead.vars_of_pbl_' ppc) ~~~ vals) t
   297       val f = LTool.subpbl (strip_thy dI) pI
   298     in (Tac.Subproblem (dI, pI),	Tac.Subproblem' ((dI, pI, mI), pors, hdl, fmz_, ctxt, f))
   299     end
   300   | stac2tac_ _ thy t = error ("stac2tac_ TODO: no match for " ^ Rule.term_to_string''' thy t);
   301 
   302 fun stac2tac pt thy t = (fst o stac2tac_ pt thy) t;
   303 
   304 datatype ass = 
   305     Ass of
   306       Tac.tac_ *  (* SubProblem gets args instantiated in assod *)
   307   	  term        (* for itr_arg, result in ets *)
   308   | AssWeak of
   309       Tac.tac_ *
   310   	  term     (*for itr_arg,result in ets*)
   311   | NotAss;
   312 
   313 (* check if tac_ is associated with stac.
   314    Additional task: check if term t (the result has been calculated from) in tac_
   315    has been changed (see "datatype tac_"); if yes, recalculate result
   316    TODO.WN120106 recalculate impl.only for Substitute'
   317 args
   318   pt     : ctree for pushing the thy specified in rootpbl into subpbls
   319   d      : unused (planned for data for comparison)
   320   tac_   : from user (via applicable_in); to be compared with ...
   321   stac   : found in Script
   322 returns
   323   Ass    : associated: e.g. thmID in stac = thmID in m
   324                        +++ arg   in stac = arg   in m
   325   AssWeak: weakly ass.:e.g. thmID in stac = thmID in m, //arg//
   326   NotAss :             e.g. thmID in stac/=/thmID in m (not =)
   327 *)
   328 fun assod _ _ (m as Tac.Rewrite_Inst' (_, _, _, _, _, thm'' as (thmID, _), f, (f', _))) stac =
   329     (case stac of
   330 	    (Const ("Script.Rewrite'_Inst", _) $ _ $ thmID_ $ _ $ f_) =>
   331 	      if thmID = HOLogic.dest_string thmID_
   332         then 
   333 	        if f = f_ 
   334           then ((*tracing"3### assod ..Ass";*) Ass (m,f')) 
   335 	        else ((*tracing"3### assod ..AssWeak";*) AssWeak(m, f'))
   336 	      else ((*tracing"3### assod ..NotAss";*) NotAss)
   337     | (Const ("Script.Rewrite'_Set'_Inst",_) $ _ $ rls_ $ _ $ f_) =>
   338 	      if Rtools.contains_rule (Rule.Thm thm'') (assoc_rls (HOLogic.dest_string rls_))
   339         then if f = f_ then Ass (m,f') else AssWeak (m,f')
   340 	      else NotAss
   341     | _ => NotAss)
   342   | assod _ _ (m as Tac.Rewrite' (_, _, _, _, thm'' as (thmID, _), f, (f', _))) stac =
   343     (case stac of
   344 	    (Const ("Script.Rewrite", _) $ thmID_ $ _ $ f_) =>
   345 	      ((*tracing ("3### assod: stac = " ^ ter2str t);
   346 	       tracing ("3### assod: f(m)= " ^ term2str f);*)
   347 	      if thmID = HOLogic.dest_string thmID_
   348         then 
   349 	        if f = f_
   350           then ((*tracing"3### assod ..Ass";*) Ass (m,f')) 
   351 	        else 
   352             ((*tracing"### assod ..AssWeak";
   353 		         tracing("### assod: f(m)  = " ^ term2str f);
   354 		         tracing("### assod: f(stac)= " ^ term2str f_)*)
   355 		         AssWeak (m,f'))
   356 	      else ((*tracing"3### assod ..NotAss";*) NotAss))
   357     | (Const ("Script.Rewrite'_Set", _) $ rls_ $ _ $ f_) =>
   358 	       if Rtools.contains_rule (Rule.Thm thm'') (assoc_rls (HOLogic.dest_string rls_))
   359          then if f = f_ then Ass (m, f') else AssWeak (m, f')
   360 	       else NotAss
   361     | _ => NotAss)
   362   | assod _ _ (m as Tac.Rewrite_Set_Inst' (_, _, _, rls, f, (f', _))) 
   363       (Const ("Script.Rewrite'_Set'_Inst", _) $ _ $ rls_ $ _ $ f_) = 
   364     if Rule.id_rls rls = HOLogic.dest_string rls_ 
   365     then if f = f_ then Ass (m, f') else AssWeak (m ,f')
   366     else NotAss
   367   | assod _ _ (m as Tac.Detail_Set_Inst' (_, _, _, rls, f, (f',_))) 
   368       (Const ("Script.Rewrite'_Set'_Inst", _) $ _ $ rls_ $ _ $ f_) = 
   369     if Rule.id_rls rls = HOLogic.dest_string rls_
   370     then if f = f_ then Ass (m, f') else AssWeak (m, f')
   371     else NotAss
   372   | assod _ _ (m as Tac.Rewrite_Set' (_, _, rls, f, (f', _))) 
   373       (Const ("Script.Rewrite'_Set", _) $ rls_ $ _ $ f_) = 
   374     if Rule.id_rls rls = HOLogic.dest_string rls_
   375     then if f = f_ then Ass (m, f') else AssWeak (m, f')
   376     else NotAss
   377   | assod _ _ (m as Tac.Detail_Set' (_, _, rls, f, (f', _))) 
   378       (Const ("Script.Rewrite'_Set", _) $ rls_ $ _ $ f_) = 
   379     if Rule.id_rls rls = HOLogic.dest_string rls_
   380     then if f = f_ then Ass (m, f') else AssWeak (m, f')
   381     else NotAss
   382   | assod _ _ (m as Tac.Calculate' (_, op_, f, (f', _))) stac =
   383     (case stac of
   384 	    (Const ("Script.Calculate",_) $ op__ $ f_) =>
   385 	      if op_ = HOLogic.dest_string op__
   386         then if f = f_ then Ass (m, f') else AssWeak (m, f')
   387 	      else NotAss
   388     | (Const ("Script.Rewrite'_Set'_Inst", _) $ _ $ rls_ $ _ $ f_)  =>
   389         let val thy = Celem.assoc_thy "Isac";
   390         in
   391           if Rtools.contains_rule (Rule.Calc (assoc_calc' thy op_ |> snd)) 
   392             (assoc_rls (HOLogic.dest_string rls_))
   393           then if f = f_ then Ass (m, f') else AssWeak (m, f')
   394           else NotAss
   395         end
   396     | (Const ("Script.Rewrite'_Set",_) $ rls_ $ _ $ f_) =>
   397         let val thy = Celem.assoc_thy "Isac";
   398         in
   399           if Rtools.contains_rule (Rule.Calc (assoc_calc' thy op_ |> snd))
   400             (assoc_rls (HOLogic.dest_string rls_))
   401           then if f = f_ then Ass (m,f') else AssWeak (m,f')
   402           else NotAss
   403         end
   404     | _ => NotAss)
   405   | assod _ _ (m as Tac.Check_elementwise' (consts, _, (consts_chkd, _)))
   406       (Const ("Script.Check'_elementwise",_) $ consts' $ _) =
   407     if consts = consts'
   408     then Ass (m, consts_chkd)
   409     else NotAss
   410   | assod _ _ (m as Tac.Or_to_List' (_, list)) (Const ("Script.Or'_to'_List", _) $ _) = Ass (m, list) 
   411   | assod _ _ (m as Tac.Take' term) (Const ("Script.Take", _) $ _) = Ass (m, term)
   412   | assod _ _ (m as Tac.Substitute' (ro, erls, subte, f, f')) (Const ("Script.Substitute", _) $ _ $ t) =
   413 	  if f = t then Ass (m, f')
   414 	  else (*compare | applicable_in (p,p_) pt (m as Substitute sube)*)
   415 		  if foldl and_ (true, map TermC.contains_Var subte)
   416 		  then
   417 		    let val t' = subst_atomic (map HOLogic.dest_eq subte (*TODO subte2subst*)) t
   418 		    in if t = t' then error "assod: Substitute' not applicable to val of Expr"
   419 		       else Ass (Tac.Substitute' (ro, erls, subte, t, t'), t')
   420 		    end
   421 		  else (case Rewrite.rewrite_terms_ (Rule.Isac ()) ro erls subte t of
   422 		         SOME (t', _) =>  Ass (Tac.Substitute' (ro, erls, subte, t, t'), t')
   423 		       | NONE => error "assod: Substitute' not applicable to val of Expr")
   424 
   425     (*compare "| stac2tac_ thy (Const ("Script.SubProblem",_)"*)
   426   | assod pt _ (Tac.Subproblem' ((domID, pblID, _), _, _, _, _, _))
   427       (stac as Const ("Script.SubProblem", _) $ (Const ("Product_Type.Pair", _) $ 
   428         dI' $ (Const ("Product_Type.Pair", _) $ pI' $ mI')) $ ags') =
   429     let 
   430       val dI = HOLogic.dest_string dI';
   431       val thy = Celem.maxthy (Celem.assoc_thy dI) (rootthy pt);
   432 	    val pI = pI' |> HOLogic.dest_list |> map HOLogic.dest_string;
   433 	    val mI = mI' |> HOLogic.dest_list |> map HOLogic.dest_string;
   434 	    val ags = TermC.isalist2list ags';
   435 	    val (pI, pors, mI) = 
   436 	      if mI = ["no_met"] 
   437 	      then
   438           let
   439             val pors = (Chead.match_ags thy ((#ppc o Specify.get_pbt) pI) ags)
   440 		          handle ERROR "actual args do not match formal args"
   441 			          => (Chead.match_ags_msg pI stac ags(*raise exn*);[]);
   442 		        val pI' = Specify.refine_ori' pors pI;
   443 		      in (pI', pors (*refinement over models with diff.prec only*), (hd o #met o Specify.get_pbt) pI')
   444           end
   445 	      else (pI, (Chead.match_ags thy ((#ppc o Specify.get_pbt) pI) ags)
   446 		      handle ERROR "actual args do not match formal args"
   447 		      => (Chead.match_ags_msg pI stac ags(*raise exn*); []), mI);
   448       val (fmz_, vals) = Chead.oris2fmz_vals pors;
   449 	    val {cas, ppc, thy, ...} = Specify.get_pbt pI
   450 	    val dI = Rule.theory2theory' thy (*take dI from _refined_ pbl*)
   451 	    val dI = Rule.theory2theory' (Celem.maxthy (Celem.assoc_thy dI) (rootthy pt))
   452 	    val ctxt = dI |> Rule.Thy_Info_get_theory |> Proof_Context.init_global |> Stool.declare_constraints' vals
   453 	    val hdl = 
   454         case cas of
   455 		      NONE => LTool.pblterm dI pI
   456 		    | SOME t => subst_atomic ((Chead.vars_of_pbl_' ppc) ~~~ vals) t
   457       val f = LTool.subpbl (strip_thy dI) pI
   458     in 
   459       if domID = dI andalso pblID = pI
   460       then Ass (Tac.Subproblem' ((dI, pI, mI), pors, hdl, fmz_, ctxt, f), f) 
   461       else NotAss
   462     end
   463   | assod _ _ m _ = 
   464     (if (!trace_script) 
   465      then tracing("@@@ the 'tac_' proposed to apply does NOT match the leaf found in the script:\n"
   466 		   ^ "@@@ tac_ = " ^ Tac.tac_2str m)
   467      else ();
   468     NotAss);
   469 
   470 fun tac_2tac (Tac.Refine_Tacitly' (pI, _, _, _, _)) = Tac.Refine_Tacitly pI
   471   | tac_2tac (Tac.Model_Problem' (_, _, _)) = Tac.Model_Problem
   472   | tac_2tac (Tac.Add_Given' (t, _)) = Tac.Add_Given t
   473   | tac_2tac (Tac.Add_Find' (t, _)) = Tac.Add_Find t
   474   | tac_2tac (Tac.Add_Relation' (t, _)) = Tac.Add_Relation t
   475  
   476   | tac_2tac (Tac.Specify_Theory' dI) = Tac.Specify_Theory dI
   477   | tac_2tac (Tac.Specify_Problem' (dI, _)) = Tac.Specify_Problem dI
   478   | tac_2tac (Tac.Specify_Method' (dI, _, _)) = Tac.Specify_Method dI
   479   
   480   | tac_2tac (Tac.Rewrite' (_, _, _, _, thm, _, _)) = Tac.Rewrite thm
   481   | tac_2tac (Tac.Rewrite_Inst' (_, _, _, _, sub, thm, _, _)) = Tac.Rewrite_Inst (Selem.subst2subs sub, thm)
   482 
   483   | tac_2tac (Tac.Rewrite_Set' (_, _, rls, _, _)) = Tac.Rewrite_Set (Rule.id_rls rls)
   484   | tac_2tac (Tac.Detail_Set' (_, _, rls, _, _)) = Tac.Detail_Set (Rule.id_rls rls)
   485 
   486   | tac_2tac (Tac.Rewrite_Set_Inst' (_, _, sub, rls, _, _)) = 
   487     Tac.Rewrite_Set_Inst (Selem.subst2subs sub, Rule.id_rls rls)
   488   | tac_2tac (Tac.Detail_Set_Inst' (_, _, sub, rls, _, _)) = 
   489     Tac.Detail_Set_Inst (Selem.subst2subs sub, Rule.id_rls rls)
   490 
   491   | tac_2tac (Tac.Calculate' (_, op_, _, _)) = Tac.Calculate (op_)
   492   | tac_2tac (Tac.Check_elementwise' (_, pred, _)) = Tac.Check_elementwise pred
   493 
   494   | tac_2tac (Tac.Or_to_List' _) = Tac.Or_to_List
   495   | tac_2tac (Tac.Take' term) = Tac.Take (Rule.term2str term)
   496   | tac_2tac (Tac.Substitute' (_, _, subte, _, _)) = Tac.Substitute (Selem.subte2sube subte) 
   497   | tac_2tac (Tac.Tac_ (_, _, id, _)) = Tac.Tac id
   498 
   499   | tac_2tac (Tac.Subproblem' ((domID, pblID, _), _, _, _,_ ,_)) = Tac.Subproblem (domID, pblID)
   500   | tac_2tac (Tac.Check_Postcond' (pblID, _)) = Tac.Check_Postcond pblID
   501   | tac_2tac Tac.Empty_Tac_ = Tac.Empty_Tac
   502   | tac_2tac m = error ("tac_2tac: not impl. for "^(Tac.tac_2str m));
   503 
   504 fun make_rule thy t =
   505   let val ct = Thm.global_cterm_of thy (HOLogic.Trueprop $ t)
   506   in Rule.Thm (Rule.term_to_string''' thy (Thm.term_of ct), Thm.make_thm ct) end;
   507 
   508 fun rep_tac_ (Tac.Rewrite_Inst' (thy', _, _, _, subs, (thmID, _), f, (f', _))) = 
   509     let
   510       val b = @{term False};
   511       val subs' = Selem.subst_to_subst' subs;
   512       val sT' = type_of subs';
   513       val fT = type_of f;
   514       val lhs = Const ("Script.Rewrite'_Inst", [sT', HOLogic.stringT, HOLogic.boolT, fT] ---> fT) 
   515         $ subs' $ HOLogic.mk_string thmID $ b $ f;
   516     in (((make_rule (Celem.assoc_thy thy')) o HOLogic.mk_eq) (lhs, f'), (lhs, f')) end
   517   | rep_tac_ (Tac.Rewrite' (thy', _, _, put, (thmID, _), f, (f', _)))=
   518     let 
   519       val fT = type_of f;
   520       val b = if put then @{term True} else @{term False};
   521       val lhs = Const ("Script.Rewrite", [HOLogic.stringT, HOLogic.boolT, fT] ---> fT)
   522         $ HOLogic.mk_string thmID $ b $ f;
   523     in (((make_rule (Celem.assoc_thy thy')) o HOLogic.mk_eq) (lhs, f'), (lhs, f')) end
   524   | rep_tac_ (Tac.Rewrite_Set_Inst' (thy', _, subs, rls, f, (f', _))) =
   525     let
   526       val b = @{term False};
   527       val subs' = Selem.subst_to_subst' subs;
   528       val sT' = type_of subs';
   529       val fT = type_of f;
   530       val lhs = Const ("Script.Rewrite'_Set'_Inst", [sT', HOLogic.stringT, HOLogic.boolT, fT] ---> fT) 
   531         $ subs' $ HOLogic.mk_string (Rule.id_rls rls) $ b $ f;
   532     in (((make_rule (Celem.assoc_thy thy')) o HOLogic.mk_eq) (lhs, f'), (lhs, f')) end
   533   | rep_tac_ (Tac.Rewrite_Set' (thy', put, rls, f, (f', _))) =
   534     let 
   535       val fT = type_of f;
   536       val b = if put then @{term True} else @{term False};
   537       val lhs = Const ("Script.Rewrite'_Set", [HOLogic.stringT, HOLogic.boolT, fT] ---> fT) 
   538         $ HOLogic.mk_string (Rule.id_rls rls) $ b $ f;
   539     in (((make_rule (Celem.assoc_thy thy')) o HOLogic.mk_eq) (lhs,f'),(lhs,f')) end
   540   | rep_tac_ (Tac.Calculate' (thy', op_, f, (f', _)))=
   541     let
   542       val fT = type_of f;
   543       val lhs = Const ("Script.Calculate",[HOLogic.stringT,fT] ---> fT) $ HOLogic.mk_string op_ $ f
   544     in (((make_rule (Celem.assoc_thy thy')) o HOLogic.mk_eq) (lhs,f'),(lhs,f')) end
   545   | rep_tac_ (Tac.Check_elementwise' (_, _, (t', _))) = (Rule.Erule, (Rule.e_term, t'))
   546   | rep_tac_ (Tac.Subproblem' (_, _, _, _, _, t')) = (Rule.Erule, (Rule.e_term, t'))
   547   | rep_tac_ (Tac.Take' t') = (Rule.Erule, (Rule.e_term, t'))
   548   | rep_tac_ (Tac.Substitute' (_, _, _, t, t')) = (Rule.Erule, (t, t'))
   549   | rep_tac_ (Tac.Or_to_List' (t, t')) = (Rule.Erule, (t, t'))
   550   | rep_tac_ m = error ("rep_tac_: not impl.for " ^ Tac.tac_2str m)
   551 
   552 fun tac_2res m = (snd o snd o rep_tac_) m;
   553 
   554 (* handle a leaf at the end of recursive descent:
   555    a leaf is either a tactic or an 'expr' in "let v = expr"
   556    where "expr" does not contain a tactic.
   557    Handling a leaf comprises
   558    (1) 'subst_stacexpr' substitute LTool.env and complete curried tactic
   559    (2) rewrite the leaf by 'srls'
   560 *)
   561 fun handle_leaf call thy srls E a v t =
   562       (*WN050916 'upd_env_opt' is a blind copy from previous version*)
   563     case LTool.subst_stacexpr E a v t of
   564 	    (a', LTool.STac stac) => (*script-tactic*)
   565 	      let val stac' =
   566             Rewrite.eval_listexpr_ (Celem.assoc_thy thy) srls (subst_atomic (upd_env_opt E (a,v)) stac)
   567 	      in
   568           (if (! trace_script) 
   569 	         then tracing ("@@@ "^call^" leaf '" ^ Rule.term2str t^"' ---> STac '" ^ Rule.term2str stac ^"'")
   570 	         else ();
   571 	         (a', LTool.STac stac'))
   572 	      end
   573     | (a', LTool.Expr lexpr) => (*leaf-expression*)
   574 	      let val lexpr' =
   575             Rewrite.eval_listexpr_ (Celem.assoc_thy thy) srls (subst_atomic (upd_env_opt E (a,v)) lexpr)
   576 	      in
   577           (if (! trace_script) 
   578 	         then tracing("@@@ "^call^" leaf '" ^ Rule.term2str t^"' ---> Expr '" ^ Rule.term2str lexpr'^"'")
   579 	         else ();
   580 	         (a', LTool.Expr lexpr')) (*lexpr' is the value of the Expr*)
   581 	      end;
   582 
   583 (** locate an applicable stac in a script **)
   584 datatype assoc =   (* ExprVal in the sense of denotational semantics               *)
   585   Assoc of         (* the stac is associated, strongly or weakly                   *)
   586   Selem.scrstate * (* the current; returned for next_tac etc. outside ass*         *)  
   587   (step list)      (* list of steps done until associated stac found;
   588 	                    initiated with the data for doing the 1st step,
   589                       thus the head holds these data further on,
   590 		                  while the tail holds steps finished (incl.scrstate in ctree) *)
   591 | NasApp of        (* stac not associated, but applicable, ctree-node generated    *)
   592   Selem.scrstate * (step list)
   593 | NasNap of        (* stac not associated, not applicable, nothing generated;
   594 	                    for distinction in Or, for leaving iterations, leaving Seq,
   595 		                  evaluate scriptexpressions                                   *)
   596   term * LTool.env;
   597 fun assoc2str (Assoc _) = "Assoc"
   598   | assoc2str (NasNap _) = "NasNap"
   599   | assoc2str (NasApp _) = "NasApp";
   600 
   601 datatype asap = (* arg. of assy _only_ for distinction w.r.t. Or                 *)
   602   Aundef        (* undefined: set only by (topmost) Or                           *)
   603 | AssOnly       (* do not execute appl stacs - there could be an associated
   604 	                 in parallel Or-branch                                         *)
   605 | AssGen;       (* no Ass(Weak) found within Or, thus 
   606                    search for _applicable_ stacs, execute and generate pt        *)
   607 (*this constructions doesnt allow arbitrary nesting of Or !!!                    *)
   608 
   609 (* assy, ass_up, astep_up scan for locate_gen in a script.
   610    search is clearly separated into (1)-(2):
   611    (1) assy is recursive descent;
   612    (2) ass_up resumes interpretation at a location somewhere in the script;
   613        astep_up does only get to the parentnode of the scriptexpr.
   614    consequence:
   615    * call of (2) means _always_ that in this branch below
   616      there was an appl.stac (Repeat, Or e1, ...) found by the previous step.
   617 *)
   618 (*WN161112 blanks between list elements left as is until istate is introduced here*)
   619 fun assy ya ((E,l,a,v,S,b),ss) (Const ("HOL.Let",_) $ e $ (Abs (id,T,body))) =
   620     (case assy ya ((E , l @ [Celem.L, Celem.R], a,v,S,b), ss) e of
   621        NasApp ((E',l,a,v,S,_),ss) => 
   622          let
   623            val id' = TermC.mk_Free (id, T);
   624            val E' = LTool.upd_env E' (id', v);
   625          in assy ya ((E', l @ [Celem.R, Celem.D], a,v,S,b),ss) body end
   626      | NasNap (v,E) =>
   627          let
   628            val id' = TermC.mk_Free (id, T);
   629            val E' = LTool.upd_env E (id', v);
   630          in assy ya ((E', l @ [Celem.R, Celem.D], a,v,S,b),ss) body end
   631      | ay => ay)
   632   | assy (ya as (thy,_,srls,_,_)) ((E,l,_,v,S,b),ss) (Const ("Script.While",_) $ c $ e $ a) =
   633     if Rewrite.eval_true_ thy srls (subst_atomic (LTool.upd_env E (a,v)) c) 
   634     then assy ya ((E, l @ [Celem.L, Celem.R], SOME a,v,S,b),ss)  e
   635     else NasNap (v, E)
   636   | assy (ya as (thy,_,srls,_,_)) ((E,l,a,v,S,b),ss) (Const ("Script.While",_) $ c $ e) =
   637     if Rewrite.eval_true_ thy srls (subst_atomic (upd_env_opt E (a,v)) c) 
   638     then assy ya ((E, l @ [Celem.R], a,v,S,b),ss) e
   639     else NasNap (v, E)
   640   | assy (ya as (thy,_,srls,_,_)) ((E,l,a,v,S,b),ss) (Const ("If",_) $ c $ e1 $ e2) =
   641     if Rewrite.eval_true_ thy srls (subst_atomic (upd_env_opt E (a,v)) c) 
   642     then assy ya ((E, l @ [Celem.L, Celem.R], a,v,S,b),ss) e1
   643     else assy ya ((E, l @ [Celem.R], a,v,S,b),ss) e2 
   644   | assy ya ((E,l,_,v,S,b),ss) (Const ("Script.Try"(*2*),_) $ e $ a) =
   645     (case assy ya ((E, l @ [Celem.L, Celem.R], SOME a,v,S,b),ss) e of ay => ay) 
   646   | assy ya ((E,l,a,v,S,b),ss) (Const ("Script.Try"(*1*),_) $ e) =
   647     (case assy ya ((E, l @ [Celem.R], a,v,S,b),ss) e of ay => ay)
   648   | assy ya ((E,l,_,v,S,b),ss) (Const ("Script.Seq"(*2*),_) $e1 $ e2 $ a) =
   649     (case assy ya ((E, l @ [Celem.L, Celem.L, Celem.R], SOME a,v,S,b),ss) e1 of
   650 	     NasNap (v, E) => assy ya ((E, l @ [Celem.L, Celem.R], SOME a,v,S,b),ss) e2
   651      | NasApp ((E,_,_,v,_,_),ss) => assy ya ((E, l @ [Celem.L, Celem.R], SOME a,v,S,b),ss) e2
   652      | ay => ay)
   653   | assy ya ((E,l,a,v,S,b),ss) (Const ("Script.Seq"(*1*),_) $e1 $ e2) =
   654     (case assy ya ((E, l @ [Celem.L, Celem.R], a,v,S,b),ss) e1 of
   655 	     NasNap (v, E) => assy ya ((E, l @ [Celem.R], a,v,S,b),ss) e2
   656      | NasApp ((E,_,_,v,_,_),ss) => assy ya ((E, l @ [Celem.R], a,v,S,b),ss) e2
   657      | ay => ay)
   658   | assy ya ((E,l,_,v,S,b),ss) (Const ("Script.Repeat",_) $ e $ a) =
   659     assy ya ((E,(l @ [Celem.L, Celem.R]),SOME a,v,S,b),ss) e
   660   | assy ya ((E,l,a,v,S,b),ss) (Const ("Script.Repeat",_) $ e) =
   661     assy ya ((E,(l @ [Celem.R]),a,v,S,b),ss) e
   662   | assy (y,x,s,sc,Aundef) ((E,l,_,v,S,b),ss) (Const ("Script.Or",_) $e1 $ e2 $ a) =
   663     (case assy (y,x,s,sc,AssOnly) ((E,(l @ [Celem.L, Celem.L, Celem.R]),SOME a,v,S,b),ss) e1 of
   664 	     NasNap (v, E) => 
   665 	       (case assy (y,x,s,sc,AssOnly) ((E,(l @ [Celem.L, Celem.R]),SOME a,v,S,b),ss) e2 of
   666 	          NasNap (v, E) => 
   667 	            (case assy (y,x,s,sc,AssGen) ((E,(l @ [Celem.L, Celem.L, Celem.R]),SOME a,v,S,b),ss) e1 of
   668 	               NasNap (v, E) => 
   669 	                 assy (y,x,s,sc,AssGen) ((E, (l @ [Celem.L, Celem.R]), SOME a,v,S,b),ss) e2
   670 	             | ay => ay)
   671 	        | ay =>ay)
   672      | NasApp _ => error ("assy: FIXXXME ///must not return NasApp///")
   673      | ay => (ay))
   674   | assy ya ((E,l,a,v,S,b),ss) (Const ("Script.Or",_) $e1 $ e2) =
   675     (case assy ya ((E, (l @ [Celem.L, Celem.R]),a,v,S,b), ss) e1 of
   676 	     NasNap (v, E) => assy ya ((E,(l @ [Celem.R]),a,v,S,b), ss) e2
   677      | ay => (ay))
   678     (*here is not a tactical like TRY etc, but a tactic creating a step in calculation*)
   679   | assy (thy',ctxt,sr,d,ap) ((E,l,a,v,S,_), (m,_,pt,(p,p_),c)::ss) t =
   680     (case handle_leaf "locate" thy' sr E a v t of
   681 	     (a', LTool.Expr _) => 
   682 	        (NasNap (Rewrite.eval_listexpr_ (Celem.assoc_thy thy') sr
   683 		     (subst_atomic (upd_env_opt E (a',v)) t), E))
   684      | (a', LTool.STac stac) =>
   685 	       let
   686            val p' = 
   687              case p_ of Frm => p 
   688              | Res => lev_on p
   689 		         | _ => error ("assy: call by " ^ pos'2str (p,p_));
   690 	       in
   691            case assod pt d m stac of
   692 	         Ass (m,v') =>
   693 	           let val (p'',c',f',pt') =
   694                  Generate.generate1 (Celem.assoc_thy thy') m (Selem.ScrState (E,l,a',v',S,true), ctxt) (p',p_) pt;
   695 	           in Assoc ((E,l,a',v',S,true), (m,f',pt',p'',c @ c')::ss) end
   696            | AssWeak (m,v') => 
   697 	           let val (p'',c',f',pt') =
   698                Generate.generate1 (Celem.assoc_thy thy') m (Selem.ScrState (E,l,a',v',S,false), ctxt) (p',p_) pt;
   699 	           in Assoc ((E,l,a',v',S,false), (m,f',pt',p'',c @ c')::ss) end
   700            | NotAss =>
   701              (case ap of   (*switch for Or: 1st AssOnly, 2nd AssGen*)
   702                 AssOnly => (NasNap (v, E))
   703               | _ =>
   704                   (case Applicable.applicable_in (p,p_) pt (stac2tac pt (Celem.assoc_thy thy') stac) of
   705 		                 Chead.Appl m' =>
   706 		                   let
   707                          val is = (E,l,a',tac_2res m',S,false(*FIXXXME.WN0?*))
   708 		                     val (p'',c',f',pt') =
   709 		                       Generate.generate1 (Celem.assoc_thy thy') m' (Selem.ScrState is, ctxt) (p', p_) pt;
   710 		                   in NasApp (is,(m,f',pt',p'',c @ c')::ss) end
   711 		               | Chead.Notappl _ => (NasNap (v, E))
   712 		              )
   713 		         )
   714          end)
   715   | assy _ (_, []) t = error ("assy: uncovered fun-def with " ^ Rule.term2str t);
   716 
   717 (*WN161112 blanks between list elements left as is until istate is introduced here*)
   718 fun ass_up (ys as (y,ctxt,s,Rule.Prog sc,d)) ((E,l,a,v,S,b),ss) (Const ("HOL.Let",_) $ _) =
   719     let 
   720 	    val l = drop_last l; (*comes from e, goes to Abs*)
   721       val (i, T, body) =
   722         (case go l sc of
   723            Const ("HOL.Let",_) $ _ $ (Abs (i, T, body)) => (i, T, body)
   724          | t => error ("ass_up..HOL.Let $ _ with " ^ Rule.term2str t))
   725       val i = TermC.mk_Free (i, T);
   726       val E = LTool.upd_env E (i, v);
   727     in case assy (y,ctxt,s,d,Aundef) ((E, l @ [Celem.R, Celem.D], a,v,S,b),ss) body of
   728 	       Assoc iss => Assoc iss
   729 	     | NasApp iss => astep_up ys iss 
   730 	     | NasNap (v, E) => astep_up ys ((E,l,a,v,S,b),ss) 
   731 	  end
   732   | ass_up ys iss (Abs (_,_,_)) = astep_up ys iss (*TODO 5.9.00: LTool.env ?*)
   733   | ass_up ys iss (Const ("HOL.Let",_) $ _ $ (Abs _)) = astep_up ys iss (*TODO 5.9.00: LTool.env ?*)
   734   | ass_up ysa iss (Const ("Script.Seq",_) $ _ $ _ $ _) =
   735     astep_up ysa iss (*all has been done in (*2*) below*)
   736   | ass_up ysa iss (Const ("Script.Seq",_) $ _ $ _) =
   737     astep_up ysa iss (*2*: comes from e2*)
   738 
   739   | ass_up (ysa as (y,ctxt,s,Rule.Prog sc,d)) ((E,l,a,v,S,b),ss)
   740 	  (Const ("Script.Seq",_) $ _ ) = (*2*: comes from e1, goes to e2*)
   741      let 
   742        val up = drop_last l;
   743      val e2 =
   744        (case go up sc of
   745           Const ("Script.Seq",_) $ _ $ e2 => e2
   746         | t => error ("ass_up..Script.Seq $ _ with " ^ Rule.term2str t))
   747      in case assy (y,ctxt,s,d,Aundef) ((E, up @ [Celem.R], a,v,S,b),ss) e2 of
   748          NasNap (v,E) => astep_up ysa ((E,up,a,v,S,b),ss)
   749        | NasApp iss => astep_up ysa iss
   750        | ay => ay 
   751      end
   752   | ass_up ysa iss (Const ("Script.Try"(*2*),_) $ _ $ _) = astep_up ysa iss
   753   | ass_up ysa iss (Const ("Script.Try"(*1*),_) $ _) = astep_up ysa iss
   754   | ass_up (ys as (y,ctxt,s,_,d)) ((E,l,_,v,S,b),ss)
   755 	   (*(Const ("Script.While",_) $ c $ e $ a) = WN050930 blind fix*)
   756 	    (t as Const ("Script.While",_) $ c $ e $ a) =
   757     if Rewrite.eval_true_ y s (subst_atomic (LTool.upd_env E (a,v)) c)
   758     then case assy (y,ctxt,s,d,Aundef) ((E, l @ [Celem.L, Celem.R], SOME a,v,S,b),ss) e of 
   759       NasNap (v,E') => astep_up ys ((E',l, SOME a,v,S,b),ss)
   760     | NasApp ((E',l,a,v,S,b),ss) =>
   761       ass_up ys ((E',l,a,v,S,b),ss) t (*WN050930 't' was not assigned*)
   762     | ay => ay
   763     else astep_up ys ((E,l, SOME a,v,S,b),ss)
   764   | ass_up (ys as (y,ctxt,s,_,d)) ((E,l,a,v,S,b),ss)
   765 	   (*(Const ("Script.While",_) $ c $ e) = WN050930 blind fix*)
   766 	     (t as Const ("Script.While",_) $ c $ e) =
   767     if Rewrite.eval_true_ y s (subst_atomic (upd_env_opt E (a,v)) c)
   768     then case assy (y,ctxt,s,d,Aundef) ((E, l @ [Celem.R], a,v,S,b),ss) e of 
   769        NasNap (v,E') => astep_up ys ((E',l, a,v,S,b),ss)
   770      | NasApp ((E',l,a,v,S,b),ss) =>
   771        ass_up ys ((E',l,a,v,S,b),ss) t (*WN050930 't' was not assigned*)
   772      | ay => ay
   773     else astep_up ys ((E,l, a,v,S,b),ss)
   774   | ass_up y iss (Const ("If",_) $ _ $ _ $ _) = astep_up y iss
   775   | ass_up (ys as (y,ctxt,s,_,d)) ((E,l,_,v,S,b),ss)
   776 	    (t as Const ("Script.Repeat",_) $ e $ a) =
   777     (case assy (y,ctxt,s,d, Aundef) ((E, (l @ [Celem.L, Celem.R]), SOME a,v,S,b),ss) e of 
   778       NasNap (v,E') => astep_up ys ((E',l, SOME a,v,S,b),ss)
   779     | NasApp ((E',l,a,v,S,b),ss) =>
   780       ass_up ys ((E',l,a,v,S,b),ss) t
   781     | ay => ay)
   782   | ass_up (ys as (y,ctxt,s,_,d)) ((E,l,a,v,S,b),ss)
   783 	    (t as Const ("Script.Repeat",_) $ e) =
   784     (case assy (y,ctxt,s,d,Aundef) ((E, (l @ [Celem.R]), a,v,S,b),ss) e of 
   785        NasNap (v', E') => astep_up ys ((E',l,a,v',S,b),ss)
   786      | NasApp ((E',l,a,v',S,_),ss) => ass_up ys ((E',l,a,v',S,b),ss) t
   787      | ay => ay)
   788   | ass_up y iss (Const ("Script.Or",_) $ _ $ _ $ _) = astep_up y iss
   789   | ass_up y iss (Const ("Script.Or",_) $ _ $ _) = astep_up y iss
   790   | ass_up y ((E,l,a,v,S,b),ss) (Const ("Script.Or",_) $ _ ) = 
   791     astep_up y ((E, (drop_last l), a,v,S,b),ss)
   792   | ass_up _ _ t =
   793     error ("ass_up not impl for t= " ^ Rule.term2str t)
   794 and astep_up (ys as (_,_,_,Rule.Prog sc,_)) ((E,l,a,v,S,b),ss) =
   795   if 1 < length l
   796   then 
   797     let val up = drop_last l;
   798     in ass_up ys ((E,up,a,v,S,b),ss) (go up sc) end
   799   else (NasNap (v, E))
   800   | astep_up _ ((_,l,_,_,_,_),_) = error ("astep_up: uncovered fun-def with " ^ Celem.loc_2str l)
   801 
   802 (*check if there are tacs for rewriting only*)
   803 fun rew_only ([]:step list) = true
   804   | rew_only (((Tac.Rewrite' _          ,_,_,_,_))::ss) = rew_only ss
   805   | rew_only (((Tac.Rewrite_Inst' _     ,_,_,_,_))::ss) = rew_only ss
   806   | rew_only (((Tac.Rewrite_Set' _      ,_,_,_,_))::ss) = rew_only ss
   807   | rew_only (((Tac.Rewrite_Set_Inst' _ ,_,_,_,_))::ss) = rew_only ss
   808   | rew_only (((Tac.Calculate' _        ,_,_,_,_))::ss) = rew_only ss
   809   | rew_only (((Tac.Begin_Trans' _      ,_,_,_,_))::ss) = rew_only ss
   810   | rew_only (((Tac.End_Trans' _        ,_,_,_,_))::ss) = rew_only ss
   811   | rew_only _ = false; 
   812 
   813 datatype locate =
   814   Steps of Selem.istate (* producing hd of step list (which was latest)
   815 	                         for next_tac, for reporting Safe|Unsafe to DG  *)
   816 	   * step             (* (scrstate producing this step is in ctree !)   *) 
   817 		 list               (* locate_gen may produce intermediate steps      *)
   818 | NotLocatable;         (* no (m Ass m') or (m AssWeak m') found          *)
   819 
   820 (* locate_gen tries to locate an input tac m in the script. 
   821    pursuing this goal the script is executed until an (m' equiv m) is found,
   822    or the end of the script
   823 args
   824    m   : input by the user, already checked by applicable_in,
   825          (to be searched within Or; and _not_ an m doing the step on ctree !)
   826    p,pt: (incl ets) at the time of input
   827    scr : the script
   828    d   : canonical simplifier for locating Take, Substitute, Subproblems etc.
   829    ets : ets at the time of input
   830    l   : the location (in scr) of the stac which generated the current formula
   831 returns
   832    Steps: pt,p (incl. ets) with m done
   833           pos' list of proofobjs cut (from generate)
   834           safe: implied from last proofobj
   835 	  ets:
   836    ///ToDo : ets contains a list of tacs to be done before m can be done
   837           NOT IMPL. -- "error: do other step before"
   838    NotLocatable: thus generate_hard
   839 *)
   840 (*WN161112 blanks between list elements left as is until istate is introduced here*)
   841 fun locate_gen (thy', _) (Tac.Rewrite' (_, ro, er, pa, thm, f, _)) (pt, p) 
   842 	    (Rule.Rfuns {locate_rule=lo,...}, _) (Selem.RrlsState (_,f'',rss,rts), _) = 
   843     (case lo rss f (Rule.Thm thm) of
   844 	    [] => NotLocatable
   845     | rts' => Steps (rts2steps [] ((pt,p),(f,f'',rss,rts),(thy',ro,er,pa)) rts'))
   846   | locate_gen (thy', srls) m (pt, p)
   847 	    (scr as Rule.Prog sc, d) (Selem.ScrState (E,l,a,v,S,b), ctxt)  = 
   848     let val thy = Celem.assoc_thy thy';
   849     in case if l = [] orelse (
   850 		       (*init.in solve..Apply_Method...*)(last_elem o fst) p = 0 andalso snd p = Res)
   851 	       then (assy (thy',ctxt,srls,d,Aundef) ((E,[Celem.R],a,v,S,b), [(m,Generate.EmptyMout,pt,p,[])]) (body_of sc))
   852 	       else (astep_up (thy',ctxt,srls,scr,d) ((E,l,a,v,S,b), [(m,Generate.EmptyMout,pt,p,[])]) ) of
   853 	    Assoc ((is as (_,_,_,_,_,strong_ass), ss as (_ :: _))) =>
   854 	      (if strong_ass
   855          then (Steps (Selem.ScrState is, ss))
   856 	       else
   857            if rew_only ss (*andalso 'not strong_ass'= associated weakly*)
   858 	         then
   859              let
   860                val (po,p_) = p
   861                val po' = case p_ of Frm => po | Res => lev_on po | _ => error ("locate_gen " ^ pos_2str p_)
   862                val (p'',c'',f'',pt'') = Generate.generate1 thy m (Selem.ScrState is, ctxt) (po',p_) pt
   863 	           in Steps (Selem.ScrState is, [(m, f'',pt'',p'',c'')]) end
   864 	         else Steps (Selem.ScrState is, ss))
   865 	  
   866     | NasApp _ => NotLocatable
   867     | err => error ("not-found-in-script: NotLocatable from " ^ @{make_string} err)
   868     end
   869   | locate_gen _ m _ (sc,_) (is, _) = 
   870     error ("locate_gen: wrong arguments,\n tac= " ^ Tac.tac_2str m ^ ",\n " ^
   871       "scr= " ^ Rule.scr2str sc ^ ",\n istate= " ^ Selem.istate2str is);
   872 
   873 (** find the next stactic in a script **)
   874 
   875 (*appy, nxt_up, nstep_up scanning for next_tac.
   876   search is clearly separated into (1)-(2):
   877   (1) appy is recursive descent;
   878   (2) nxt_up resumes interpretation at a location somewhere in the script;
   879       nstep_up does only get to the parentnode of the scriptexpr.
   880   consequence:
   881   * call of (2) means _always_ that in this branch below
   882     there was an applicable stac (Repeat, Or e1, ...)
   883 *)
   884 datatype appy =    (* ExprVal in the sense of denotational semantics  *)
   885     Appy of        (* applicable stac found, search stalled           *)
   886     Tac.tac_ *     (* tac_ associated (fun assod) with stac           *)
   887     Selem.scrstate (* after determination of stac WN.18.8.03          *)
   888   | Napp of        (* stac found was not applicable; 
   889 	                    this mode may become Skip in Repeat, Try and Or *)
   890     LTool.env      (* popped while nxt_up                             *)
   891   | Skip of        (* for restart after Appy, for leaving iterations,
   892 	                    for passing the value of scriptexpressions,
   893 		                  and for finishing the script successfully       *)
   894     term * LTool.env  (*a stack*);
   895 
   896 datatype appy_ = (* as argument in nxt_up, nstep_up, from appy               *)
   897 (*Appy              is only (final) returnvalue, not argument during search  *)
   898   Napp_          (* ev. detects 'script is not appropriate for this example' *)
   899 | Skip_;         (* detects 'script successfully finished'
   900 		                also used as init-value for resuming; this works,
   901 	                  because 'nxt_up Or e1' treats as Appy                    *)
   902 
   903 fun appy thy ptp E l (Const ("HOL.Let",_) $ e $ (Abs (i,T,b))) a v =
   904     (case appy thy ptp E (l @ [Celem.L, Celem.R]) e a v of
   905       Skip (res, E) => 
   906         let val E' = LTool.upd_env E (Free (i,T), res);
   907         in appy thy ptp E' (l @ [Celem.R, Celem.D]) b a v end
   908     | ay => ay)
   909   | appy (thy as (th,sr)) ptp E l (Const ("Script.While"(*1*),_) $ c $ e $ a) _ v =
   910     (if Rewrite.eval_true_ th sr (subst_atomic (LTool.upd_env E (a,v)) c)
   911      then appy thy ptp E (l @ [Celem.L, Celem.R]) e (SOME a) v
   912      else Skip (v, E))
   913   | appy (thy as (th,sr)) ptp E l (Const ("Script.While"(*2*),_) $ c $ e) a v =
   914     (if Rewrite.eval_true_ th sr (subst_atomic (upd_env_opt E (a,v)) c)
   915      then appy thy ptp E (l @ [Celem.R]) e a v
   916      else Skip (v, E))
   917   | appy (thy as (th,sr)) ptp E l (Const ("If",_) $ c $ e1 $ e2) a v =
   918     (if Rewrite.eval_true_ th sr (subst_atomic (upd_env_opt E (a,v)) c)
   919      then appy thy ptp E (l @ [Celem.L, Celem.R]) e1 a v
   920      else appy thy ptp E (l @ [Celem.R]) e2 a v)
   921   | appy thy ptp E l (Const ("Script.Repeat"(*1*),_) $ e $ a) _ v = 
   922     appy thy ptp E (l @ [Celem.L, Celem.R]) e (SOME a) v
   923   | appy thy ptp E l (Const ("Script.Repeat"(*2*),_) $ e) a v = appy thy ptp E (l @ [Celem.R]) e a v
   924   | appy thy ptp E l (Const ("Script.Try"(*2*),_) $ e $ a) _ v =
   925     (case appy thy ptp E (l @ [Celem.L, Celem.R]) e (SOME a) v of
   926       Napp E => (Skip (v, E))
   927     | ay => ay)
   928   | appy thy ptp E l(Const ("Script.Try"(*1*),_) $ e) a v =
   929     (case appy thy ptp E (l @ [Celem.R]) e a v of
   930       Napp E => (Skip (v, E))
   931     | ay => ay)
   932   | appy thy ptp E l (Const ("Script.Or"(*1*),_) $e1 $ e2 $ a) _ v =
   933     (case appy thy ptp E (l @ [Celem.L, Celem.L, Celem.R]) e1 (SOME a) v of
   934 	    Appy lme => Appy lme
   935     | _ => appy thy ptp E (*LTool.env*) (l @ [Celem.L, Celem.R]) e2 (SOME a) v)
   936   | appy thy ptp E l (Const ("Script.Or"(*2*),_) $e1 $ e2) a v =
   937     (case appy thy ptp E (l @ [Celem.L, Celem.R]) e1 a v of
   938 	    Appy lme => Appy lme
   939     | _ => appy thy ptp E (l @ [Celem.R]) e2 a v)
   940   | appy thy ptp E l (Const ("Script.Seq"(*2*),_) $ e1 $ e2 $ a) _ v =
   941     (case appy thy ptp E (l @ [Celem.L, Celem.L, Celem.R]) e1 (SOME a) v of
   942 	    Skip (v,E) => appy thy ptp E (l @ [Celem.L, Celem.R]) e2 (SOME a) v
   943     | ay => ay)
   944   | appy thy ptp E l (Const ("Script.Seq"(*1*),_) $ e1 $ e2) a v =
   945     (case appy thy ptp E (l @ [Celem.L, Celem.R]) e1 a v of
   946 	    Skip (v,E) => appy thy ptp E (l @ [Celem.R]) e2 a v
   947     | ay => ay)
   948   (* a leaf has been found *)   
   949   | appy ((th,sr)) (pt, p) E l t a v =
   950     case handle_leaf "next  " th sr E a v t of
   951 	    (_, LTool.Expr s) => Skip (s, E)
   952     | (a', LTool.STac stac) =>
   953 	    let val (m,m') = stac2tac_ pt (Celem.assoc_thy th) stac
   954       in case m of
   955 	      Tac.Subproblem _ => Appy (m', (E,l,a',tac_2res m',Selem.Sundef,false))
   956 	    | _ =>
   957         (case Applicable.applicable_in p pt m of
   958 		       Chead.Appl m' => (Appy (m', (E,l,a',tac_2res m',Selem.Sundef,false)))
   959 		     | _ => Napp E)
   960 	    end
   961 
   962 fun nxt_up thy ptp (scr as (Rule.Prog sc)) E l ay (Const ("HOL.Let", _) $ _) a v = (*comes from let=...*)
   963     if ay = Napp_
   964     then nstep_up thy ptp scr E (drop_last l) Napp_ a v
   965     else (*Skip_*)
   966 	    let
   967         val up = drop_last l
   968         val (i, T, body) =
   969           (case go up sc of
   970              Const ("HOL.Let",_) $ _ $ (Abs aa) => aa
   971            | t => error ("nxt_up..HOL.Let $ _ with " ^ Rule.term2str t))
   972         val i = TermC.mk_Free (i, T)
   973         val E = LTool.upd_env E (i, v)
   974       in
   975         case appy thy ptp E (up @ [Celem.R, Celem.D]) body a v  of
   976 	        Appy lre => Appy lre
   977 	      | Napp E => nstep_up thy ptp scr E up Napp_ a v
   978 	      | Skip (v,E) => nstep_up thy ptp scr E up Skip_ a v
   979 	    end
   980   | nxt_up thy ptp scr E l ay (Abs _) a v =  nstep_up thy ptp scr E l ay a v
   981   | nxt_up thy ptp scr E l ay (Const ("HOL.Let",_) $ _ $ (Abs _)) a v =
   982     nstep_up thy ptp scr E l ay a v
   983   (*no appy_: never causes Napp -> Helpless*)
   984   | nxt_up (thy as (th, sr)) ptp scr E l _ (Const ("Script.While"(*1*), _) $ c $ e $ _) a v = 
   985     if Rewrite.eval_true_ th sr (subst_atomic (upd_env_opt E (a, v)) c) 
   986     then case appy thy ptp E (l @ [Celem.L, Celem.R]) e a v of
   987 	     Appy lr => Appy lr
   988 	  | Napp E => nstep_up thy ptp scr E l Skip_ a v
   989 	  | Skip (v,E) => nstep_up thy ptp scr E l Skip_ a v
   990     else nstep_up thy ptp scr E l Skip_ a v
   991   (*no appy_: never causes Napp - Helpless*)
   992   | nxt_up (thy as (th, sr)) ptp scr E l _ (Const ("Script.While"(*2*), _) $ c $ e) a v = 
   993     if Rewrite.eval_true_ th sr (subst_atomic (upd_env_opt E (a, v)) c) 
   994     then case appy thy ptp E (l @ [Celem.R]) e a v of
   995 	    Appy lr => Appy lr
   996 	  | Napp E => nstep_up thy ptp scr E l Skip_ a v
   997 	  | Skip (v,E) => nstep_up thy ptp scr E l Skip_ a v
   998     else nstep_up thy ptp scr E l Skip_ a v
   999   | nxt_up thy ptp scr E l ay (Const ("If", _) $ _ $ _ $ _) a v = nstep_up thy ptp scr E l ay a v
  1000   | nxt_up thy ptp scr E l _ (*no appy_: there was already a stac below*)
  1001       (Const ("Script.Repeat"(*1*), _) $ e $ _) a v =
  1002     (case appy thy ptp (*upd_env*) E (*a,v)*) (l @ [Celem.L, Celem.R]) e a v  of
  1003       Appy lr => Appy lr
  1004     | Napp E =>  nstep_up thy ptp scr E l Skip_ a v
  1005     | Skip (v,E) =>  nstep_up thy ptp scr E l Skip_ a v)
  1006   | nxt_up thy ptp scr E l _ (*no appy_: there was already a stac below*)
  1007       (Const ("Script.Repeat"(*2*), _) $ e) a v =
  1008     (case appy thy ptp (*upd_env*) E (*a,v)*) (l @ [Celem.R]) e a v  of
  1009       Appy lr => Appy lr
  1010     | Napp E => nstep_up thy ptp scr E l Skip_ a v
  1011     | Skip (v,E) => nstep_up thy ptp scr E l Skip_ a v)
  1012   | nxt_up thy ptp scr E l _ (Const ("Script.Try"(*2*),_) $ _ $ _) a v = (*makes Napp to Skip*)
  1013     nstep_up thy ptp scr E l Skip_ a v
  1014 
  1015   | nxt_up thy ptp scr E l _ (Const ("Script.Try"(*1*), _) $ _) a v = (*makes Napp to Skip*)
  1016     nstep_up thy ptp scr E l Skip_ a v
  1017   | nxt_up thy ptp scr E l ay (Const ("Script.Or",_) $ _ $ _ $ _) a v =
  1018     nstep_up thy ptp scr E l ay a v
  1019   | nxt_up thy ptp scr E l ay (Const ("Script.Or",_) $ _ $ _) a v = nstep_up thy ptp scr E l ay a v
  1020   | nxt_up thy ptp scr E l ay (Const ("Script.Or",_) $ _ ) a v = 
  1021     nstep_up thy ptp scr E (drop_last l) ay a v
  1022   | nxt_up thy ptp scr E l ay (Const ("Script.Seq"(*1*),_) $ _ $ _ $ _) a v =
  1023     (*all has been done in (*2*) below*) nstep_up thy ptp scr E l ay a v
  1024   | nxt_up thy ptp scr E l ay (Const ("Script.Seq"(*2*),_) $ _ $ _) a v = (*comes from e2*)
  1025     nstep_up thy ptp scr E l ay a v
  1026   | nxt_up thy ptp (scr as Rule.Prog sc) E l ay (Const ("Script.Seq"(*3*),_) $ _) a v = (*comes from e1*)
  1027     if ay = Napp_
  1028     then nstep_up thy ptp scr E (drop_last l) Napp_ a v
  1029     else (*Skip_*)
  1030       let val up = drop_last l;
  1031           val e2 =
  1032             (case go up sc of
  1033                Const ("Script.Seq"(*2*), _) $ _ $ e2 => e2
  1034              | t => error ("nxt_up..Script.Seq $ _ with " ^ Rule.term2str t))
  1035       in case appy thy ptp E (up @ [Celem.R]) e2 a v  of
  1036         Appy lr => Appy lr
  1037       | Napp E => nstep_up thy ptp scr E up Napp_ a v
  1038       | Skip (v,E) => nstep_up thy ptp scr E up Skip_ a v end
  1039   | nxt_up _ _ _ _ _ _ t _ _ = error ("nxt_up not impl for " ^ Rule.term2str t)
  1040 and nstep_up thy ptp (Rule.Prog sc) E l ay a v = 
  1041     if 1 < length l
  1042     then 
  1043       let val up = drop_last l; 
  1044       in (nxt_up thy ptp (Rule.Prog sc) E up ay (go up sc) a v ) end
  1045     else (*interpreted to end*)
  1046       if ay = Skip_ then Skip (v, E) else Napp E 
  1047   | nstep_up _ _ _ _ l _ _ _ = error ("nstep_up: uncovered fun-def at " ^ Celem.loc_2str l)
  1048 
  1049 (* decide for the next applicable stac in the script;
  1050    returns (stactic, value) - the value in case the script is finished 
  1051    12.8.02:         ~~~~~ and no assumptions ??? FIXME ???
  1052    20.8.02: must return p in case of finished, because the next script
  1053             consulted need not be the calling script:
  1054             in case of detail ie. _inserted_ PrfObjs, the next stac
  1055             has to searched in a script with PblObj.status<>Complete !
  1056             (.. not true for other details ..PrfObj ??????????????????
  1057    20.8.02: do NOT return safe (is only changed in locate !!!)
  1058 *)
  1059 fun next_tac (thy,_) _ (Rule.Rfuns {next_rule, ...}) (Selem.RrlsState(f, f', rss, _), ctxt) =
  1060     if f = f'
  1061     then (Tac.End_Detail' (f',[])(*8.6.03*), (Selem.Uistate, ctxt), 
  1062     	(f', Selem.Sundef(*FIXME is no value of next_tac! vor 8.6.03*)))                (*finished*)
  1063     else
  1064       (case next_rule rss f of
  1065     	  NONE => (Tac.Empty_Tac_, (Selem.Uistate, ctxt), (Rule.e_term, Selem.Sundef))       (*helpless*)
  1066     	| SOME (Rule.Thm thm'')(*8.6.03: muss auch f' liefern ?!!*) => 
  1067     	    (Tac.Rewrite' (thy, "e_rew_ord", Rule.e_rls, false, thm'', f, (Rule.e_term, [(*!?!8.6.03*)])),
  1068   	         (Selem.Uistate, ctxt), (Rule.e_term, Selem.Sundef))                          (*next stac*)
  1069       | _ => error "next_tac: uncovered case next_rule")
  1070   | next_tac thy (ptp as (pt, (p, _))) (sc as Rule.Prog prog) 
  1071 	    (Selem.ScrState (E,l,a,v,s,_), ctxt) =
  1072     (case if l = [] then appy thy ptp E [Celem.R] (body_of prog) NONE v
  1073           else nstep_up thy ptp sc E l Skip_ a v of
  1074        Skip (v, _) =>                                                                 (*finished*)
  1075          (case par_pbl_det pt p of
  1076 	          (true, p', _) => 
  1077 	             let
  1078 	               val (_,pblID,_) = get_obj g_spec pt p';
  1079 	              in (Tac.Check_Postcond' (pblID, (v, [(*assigned in next step*)])), 
  1080 	                   (Selem.e_istate, ctxt), (v,s)) 
  1081                 end
  1082 	        | _ => (Tac.End_Detail' (Rule.e_term,[])(*8.6.03*), (Selem.e_istate, ctxt), (v,s)))
  1083      | Napp _ => (Tac.Empty_Tac_, (Selem.e_istate, ctxt), (Rule.e_term, Selem.Sundef))     (*helpless*)
  1084      | Appy (m', scrst as (_,_,_,v,_,_)) =>
  1085          (m', (Selem.ScrState scrst, ctxt), (v, Selem.Sundef)))                      (*next stac*)
  1086   | next_tac _ _ _ (is, _) = error ("next_tac: not impl for " ^ (Selem.istate2str is));
  1087 
  1088 (*.create the initial interpreter state from the items of the guard.*)
  1089 local
  1090 val errmsg = "ERROR: found no actual arguments for prog. of "
  1091 fun msg_miss (sc, metID, formals, actuals) =
  1092   "ERROR in creating the environment for '" ^ id_of_scr sc ^ 
  1093 	"' from \nthe items of the guard of " ^ Celem.metID2str metID ^ ",\n" ^
  1094 	"formal arg(s), from the script, miss actual arg(s), from the guards LTool.env:\n" ^
  1095 	(string_of_int o length) formals ^ " formals: " ^ Rule.terms2str formals ^ "\n" ^
  1096 	(string_of_int o length) actuals ^ " actuals: " ^ Rule.terms2str actuals
  1097 fun msg_type (sc, metID, a, f, formals, actuals) =
  1098   "ERROR in creating the environment for '" ^
  1099 	id_of_scr sc ^ "' from \nthe items of the guard of " ^
  1100 	Celem.metID2str metID ^ ",\n" ^
  1101 	"different types of formal arg, from the script, " ^
  1102 	"and actual arg, from the guards LTool.env:'\n" ^
  1103 	"formal: '" ^ Rule.term2str a ^ "::" ^ (Rule.type2str o type_of) a ^ "'\n" ^
  1104 	"actual: '" ^ Rule.term2str f ^ "::" ^ (Rule.type2str o type_of) f ^ "'\n" ^
  1105 	"in\n" ^
  1106 	"formals: " ^ Rule.terms2str formals ^ "\n" ^
  1107 	"actuals: " ^ Rule.terms2str actuals
  1108 in
  1109 fun init_scrstate thy itms metID =
  1110   let
  1111     val actuals = itms2args thy metID itms
  1112     val _ = if actuals <> [] then () else raise ERROR (errmsg ^ strs2str' metID)
  1113     val (scr, sc) = (case (#scr o Specify.get_met) metID of
  1114        scr as Rule.Prog sc => (scr, sc) | _ => raise ERROR ("init_scrstate with " ^ Celem.metID2str metID))
  1115     val formals = formal_args sc    
  1116 	  (*expects same sequence of (actual) args in itms and (formal) args in met*)
  1117 	  fun relate_args env [] [] = env
  1118 	    | relate_args _ _ [] = error (msg_miss (sc, metID, formals, actuals))
  1119 	    | relate_args env [] _ = env (*may drop Find!*)
  1120 	    | relate_args env (a::aa) (f::ff) = 
  1121 	      if type_of a = type_of f 
  1122 	      then relate_args (env @ [(a, f)]) aa ff
  1123         else error (msg_type (sc, metID, a, f, formals, actuals))
  1124     val env = relate_args [] formals actuals;
  1125     val ctxt = Proof_Context.init_global thy |> Stool.declare_constraints' actuals
  1126     val {pre, prls, ...} = Specify.get_met metID;
  1127     val pres = Stool.check_preconds thy prls pre itms |> map snd;
  1128     val ctxt = ctxt |> Stool.insert_assumptions pres;
  1129   in (Selem.ScrState (env, [], NONE, Rule.e_term, Selem.Safe, true), ctxt, scr) end;
  1130 end (*local*)
  1131 
  1132 (* decide, where to get script/istate from:
  1133    (* 1 *) from PblObj.LTool.env: at begin of script if no init_form
  1134    (* 2 *) from PblObj/PrfObj: if stac is in the middle of the script
  1135    (* 3 *) from rls/PrfObj: in case of detail a ruleset *)
  1136 fun from_pblobj_or_detail' _ (p, p_) pt =
  1137   if member op = [Pbl, Met] p_
  1138   then case get_obj g_env pt p of
  1139     NONE => error "from_pblobj_or_detail': no istate"
  1140   | SOME is =>
  1141       let
  1142         val metID = get_obj g_metID pt p
  1143         val {srls, ...} = Specify.get_met metID
  1144       in (srls, is, (#scr o Specify.get_met) metID) end
  1145   else
  1146     let val (pbl, p', rls') = par_pbl_det pt p
  1147     in if pbl 
  1148        then (*if last_elem p = 0 nothing written to pt yet*)                                (* 2 *)
  1149          let
  1150 	         val metID = get_obj g_metID pt p'
  1151 	         val {srls,...} = Specify.get_met metID
  1152 	       in (srls, get_loc pt (p,p_), (#scr o Specify.get_met) metID) end
  1153        else (*FIXME.WN0?: get from pbl or met !!! unused for Rrls in locate_gen, next_tac*) (* 3 *)
  1154 	       (Rule.e_rls, get_loc pt (p,p_),
  1155 	          case rls' of
  1156 		          Rule.Rls {scr = scr,...} => scr
  1157 	          | Rule.Seq {scr = scr,...} => scr
  1158 	          | Rule.Rrls {scr=rfuns,...} => rfuns
  1159 	          | Rule.Erls => error "from_pblobj_or_detail' with Erls")
  1160     end
  1161 
  1162 (*.get script and istate from PblObj, see                                                  ( * 1 *)
  1163 fun from_pblobj' thy' (p,p_) pt =
  1164   let
  1165     val p' = par_pblobj pt p
  1166 	  val thy = Celem.assoc_thy thy'
  1167 	  val itms =
  1168 	    (case get_obj I pt p' of
  1169 	      PblObj {meth = itms, ...} => itms
  1170 	    | PrfObj _ => error "from_pblobj' NOT with PrfObj")
  1171 	  val metID = get_obj g_metID pt p'
  1172 	  val {srls, scr, ...} = Specify.get_met metID
  1173   in
  1174     if last_elem p = 0 (*nothing written to pt yet*)
  1175     then
  1176        let val (is, ctxt, scr) = init_scrstate thy itms metID
  1177 	     in (srls, (is, ctxt), scr) end
  1178     else (srls, get_loc pt (p,p_), scr)
  1179   end;
  1180     
  1181 (*.get the stactics and problems of a script as tacs
  1182   instantiated with the current environment;
  1183   l is the location which generated the given formula.*)
  1184 (*WN.12.5.03: quick-and-dirty repair for listexpressions*)
  1185 fun is_spec_pos Pbl = true
  1186   | is_spec_pos Met = true
  1187   | is_spec_pos _ = false;
  1188 
  1189 (*. fetch _all_ tactics from script .*)
  1190 fun sel_rules _ (([],Res):pos') = 
  1191     raise PTREE "no tactics applicable at the end of a calculation"
  1192   | sel_rules pt (p,p_) =
  1193     if is_spec_pos p_ 
  1194     then [get_obj g_tac pt p]
  1195     else
  1196       let
  1197         val pp = par_pblobj pt p;
  1198         val thy' = get_obj g_domID pt pp;
  1199         val thy = Celem.assoc_thy thy';
  1200         val metID = get_obj g_metID pt pp;
  1201         val metID' = if metID = Celem.e_metID then (thd3 o snd3) (get_obj g_origin pt pp) else metID
  1202         val (sc, srls) = (case Specify.get_met metID' of
  1203             {scr = Rule.Prog sc, srls, ...} => (sc, srls) | _ => error "sel_rules 1")
  1204         val (env, a, v) = (case get_istate pt (p, p_) of
  1205             Selem.ScrState (env, _, a, v, _, _) => (env, a, v) | _ => error "sel_rules 2")
  1206       in map ((stac2tac pt thy) o LTool.rep_stacexpr o #2 o
  1207   	    (handle_leaf "selrul" thy' srls env a v)) (LTool.stacpbls sc)
  1208   	  end;
  1209 
  1210 (* fetch tactics from script and filter _applicable_ tactics;
  1211    in case of Rewrite_Set* go down to _atomic_ rewrite-tactics *)
  1212 fun sel_appl_atomic_tacs _ (([], Res) : pos') = 
  1213     raise PTREE "no tactics applicable at the end of a calculation"
  1214   | sel_appl_atomic_tacs pt (p, p_) =
  1215     if is_spec_pos p_ 
  1216     then [get_obj g_tac pt p]
  1217     else
  1218       let
  1219         val pp = par_pblobj pt p
  1220         val thy' = get_obj g_domID pt pp
  1221         val thy = Celem.assoc_thy thy'
  1222         val metID = get_obj g_metID pt pp
  1223         val metID' =
  1224           if metID = Celem.e_metID 
  1225           then (thd3 o snd3) (get_obj g_origin pt pp)
  1226           else metID
  1227         val (sc, srls, erls, ro) = (case Specify.get_met metID' of
  1228             {scr = Rule.Prog sc, srls, erls, rew_ord' = ro, ...} => (sc, srls, erls, ro)
  1229           | _ => error "sel_appl_atomic_tacs 1")
  1230         val (env, a, v) = (case get_istate pt (p, p_) of
  1231             Selem.ScrState (env, _, a, v, _, _) => (env, a, v) | _ => error "sel_appl_atomic_tacs 2")
  1232         val alltacs = (*we expect at least 1 stac in a script*)
  1233           map ((stac2tac pt thy) o LTool.rep_stacexpr o #2 o
  1234            (handle_leaf "selrul" thy' srls env a v)) (LTool.stacpbls sc)
  1235         val f =
  1236           (case p_ of Frm => get_obj g_form pt p | Res => (fst o (get_obj g_result pt)) p
  1237           | _ => error "")
  1238           (*WN071231 ? replace atomic_appl_tacs with applicable_in (ineff!) ?*)
  1239       in ((gen_distinct Tac.eq_tac) o flat o (map (Rtools.atomic_appl_tacs thy ro erls f))) alltacs end;
  1240 (**)
  1241 end
  1242 (**)