src/Tools/isac/Interpret/lucas-interpreter.sml
author Walther Neuper <walther.neuper@jku.at>
Fri, 10 Apr 2020 18:32:36 +0200
changeset 59868 d77aa0992e0f
parent 59861 65ec9f679c3f
child 59881 bdced24f62bf
permissions -rw-r--r--
use "UnparseC" for renaming identifiers
     1 (* Title:  Interpret/lucas-interpreter.sml
     2    Author: Walther Neuper 2019
     3    (c) due to copyright terms
     4 *)
     5 
     6 signature LUCAS_INTERPRETER =
     7 sig
     8   datatype next_step_result =
     9       Next_Step of Istate.T * Proof.context * Tactic.T
    10     | Helpless | End_Program of Istate.T (*TODO ?needed when..*)* Tactic.T
    11   val find_next_step: Program.T -> Calc.T -> Istate.T -> Proof.context
    12     -> next_step_result
    13 
    14   datatype input_tactic_result =
    15       Safe_Step of Istate.T * Proof.context * Tactic.T
    16     | Unsafe_Step of Istate.T * Proof.context * Tactic.T
    17     | Not_Locatable (*TODO rm..*) of string
    18   val locate_input_tactic: Program.T -> Calc.T -> Istate.T -> Proof.context
    19       -> Tactic.T -> input_tactic_result
    20 
    21   datatype input_term_result = Found_Step of Calc.T | Not_Derivable (**)
    22   val locate_input_term: Calc.T -> term -> input_term_result (**)
    23                         
    24   (* must reside here:
    25      find_next_step calls do_next and is called by by_tactic;
    26      by_tactic and do_next are mutually recursive via by_tactic..Apply_Method'
    27    *)
    28   val by_tactic: Tactic.T -> Istate.T * Proof.context -> Calc.T -> string * Chead.calcstate'
    29   val do_next: Calc.T -> string * Chead.calcstate'
    30 
    31 (* ---- for tests only: shifted from below to remove the Warning "unused" at fun.def. --------- *)
    32   datatype expr_val1 =
    33       Accept_Tac1 of Istate.pstate * Proof.context * Tactic.T
    34     | Reject_Tac1 of Istate.pstate * Proof.context * Tactic.T
    35     | Term_Val1 of term;
    36   val assoc2str: expr_val1 -> string
    37 (* ---- for Doc/Lucas-Interpreter ------------------------------------------------------------- *)
    38   val check_Seq_up: Istate.pstate -> term -> term
    39   datatype expr_val = Accept_Tac of Tactic.T * Istate.pstate * Proof.context
    40     | Reject_Tac | Term_Val of term
    41 
    42   val scan_dn: Calc.T * Proof.context -> Istate.pstate -> term -> expr_val
    43   val scan_up: term * (Calc.T * Proof.context) -> Istate.pstate -> term -> expr_val
    44   val go_scan_up: term * (Calc.T * Proof.context) -> Istate.pstate -> expr_val
    45   val scan_to_tactic: term * (Calc.T * Proof.context) -> Istate.T -> expr_val
    46   val check_tac: Calc.T * Proof.context -> Istate.pstate -> term * term option -> expr_val
    47 (*/-------------------------------------------------------- ! aktivate for Test_Isac BEGIN ---\* )
    48   val check_Let_up: Istate.pstate -> term -> term * term
    49   val compare_step: Generate.taci list * Pos.pos' list * (Calc.T) -> term -> string * Chead.calcstate'
    50 
    51   val scan_dn1: (Calc.T * Proof.context * Tactic.T) -> Istate.pstate -> term -> expr_val1
    52   val scan_up1: term * (Calc.T * Proof.context * Tactic.T) -> Istate.pstate -> term -> expr_val1;
    53   val go_scan_up1: term * (Calc.T * Proof.context * Tactic.T) -> Istate.pstate -> expr_val1;
    54   val scan_to_tactic1: term * (Calc.T * Proof.context * Tactic.T) -> Istate.T -> expr_val1
    55 
    56   val check_tac1: Calc.T * Proof.context * Tactic.T -> Istate.pstate -> term * term option -> expr_val1
    57 ( *\--- ! aktivate for Test_Isac END ----------------------------------------------------------/*)
    58 end
    59 
    60 (**)
    61 structure LI(**): LUCAS_INTERPRETER(**) =
    62 struct
    63 (**)
    64 open Ctree
    65 open Pos
    66 open TermC
    67 open Istate
    68 
    69 (*** auxiliary functions ***)
    70 
    71 fun at_location [] t = t
    72   | at_location (D :: p) (Abs(_, _, body)) = at_location (p : TermC.path) body
    73   | at_location (L :: p) (t1 $ _) = at_location p t1
    74   | at_location (R :: p) (_ $ t2) = at_location p t2
    75   | at_location l _ = error ("at_location: no " ^ TermC.string_of_path l);
    76 
    77 fun check_Let_up ({path, ...}: pstate) prog =
    78   case at_location (drop_last path) prog of
    79     Const ("HOL.Let",_) $ _ $ (Abs (i, T, body)) => (TermC.mk_Free (i, T), body)
    80   | t => raise ERROR ("scan_up1..\"HOL.Let $ _\" with: \"" ^ UnparseC.term t ^ "\"")
    81 fun check_Seq_up  ({path, ...}: pstate) prog =
    82   case at_location (drop_last path) prog of
    83     Const ("Tactical.Chain",_) $ _ $ e2=> e2
    84   | t => raise ERROR ("scan_up1..\"Tactical.Chain $ _\" with: \"" ^ UnparseC.term t ^ "\"")
    85 
    86 
    87 (*** determine a next tactic within a program ***)
    88 
    89   datatype next_step_result = Next_Step of Istate.T * Proof.context * Tactic.T
    90     | Helpless | End_Program of Istate.T * Tactic.T (*contains prog_result, without asms*)
    91 
    92 (** scan to a Prog_Tac **)
    93 
    94 datatype expr_val = (* "ExprVal" in the sense of denotational semantics           *)
    95   Reject_Tac        (* tactic is found but NOT acknowledged, scan is continued    *)
    96 | Accept_Tac of     (* tactic is found and acknowledged, scan is stalled          *)
    97     Tactic.T *      (* Prog_Tac is applicable_in cstate                           *)
    98     Istate.pstate * (* created by application of Tactic.T, for resuming execution *)
    99     Proof.context   (* created by application of Tactic.T                         *)
   100 | Term_Val of       (* Prog_Expr is found and evaluated, scan is continued        *)
   101     term;           (* value of Prog_Expr, for updating environment               *)
   102 
   103 (* check if a prog_tac found in a program is applicable_in *)
   104 fun check_tac ((pt, p), ctxt) ist (prog_tac, form_arg) =
   105   let
   106     val m = LItool.tac_from_prog pt (Proof_Context.theory_of ctxt) prog_tac
   107   in
   108     case m of
   109       Tactic.Subproblem _ => (*might involve problem refinement etc*)
   110         let
   111           val m' = snd (Sub_Problem.tac_from_prog pt prog_tac)
   112         in
   113           Accept_Tac (m', ist |> set_subst_found (form_arg, Tactic.result m'), ctxt)
   114         end
   115     | _ =>
   116       (case Applicable.applicable_in p pt m of
   117         Applicable.Appl m' => 
   118           Accept_Tac (m', ist |> set_subst_found (form_arg, Tactic.result m'),
   119             Tactic.insert_assumptions m' ctxt)
   120       | _ => Reject_Tac)
   121   end
   122 
   123 
   124 (*
   125   scan_dn, go_scan_up, scan_up scan for find_next_step.
   126   (1) scan_dn is recursive descent depth first strictly from L to R;
   127   (2) go_scan_up goes to the parent node and calls (3);
   128   (3) scan_up resumes according to the interpreter-state.
   129   Call of (2) means that there was an applicable Prog_Tac below = before.
   130 *)
   131 fun scan_dn cc (ist as {act_arg, ...}) (Const ("Tactical.Try"(*1*), _) $ e $ a) =
   132     (case scan_dn cc (ist|> path_down_form ([L, R], a)) e of
   133       Reject_Tac => Term_Val act_arg
   134     | (*Accept_Tac*) goback => goback)
   135   | scan_dn cc (ist as {act_arg, ...}) (Const ("Tactical.Try"(*2*), _) $ e) =
   136     (case scan_dn cc (ist |> path_down [R]) e of
   137       Reject_Tac => Term_Val act_arg
   138     | (*Accept_Tac*) goback => goback)
   139 
   140   | scan_dn cc ist (Const ("Tactical.Repeat"(*1*), _) $ e $ a) = 
   141     scan_dn cc (ist |> path_down_form ([L, R], a)) e
   142   | scan_dn cc ist (Const ("Tactical.Repeat"(*2*), _) $ e) =
   143     scan_dn cc (ist |> path_down [R]) e
   144 
   145   | scan_dn cc ist (Const ("Tactical.Chain"(*1*), _) $ e1 $ e2 $ a) =
   146     (case scan_dn cc (ist |> path_down_form ([L, L, R], a)) e1 of
   147 	    Term_Val v => scan_dn cc (ist |> path_down_form ([L, R], a) |> set_act v) e2
   148     | (*Accept_Tac*) goback => goback)
   149   | scan_dn cc ist (Const ("Tactical.Chain"(*2*), _) $ e1 $ e2) =
   150     (case scan_dn cc (ist |> path_down [L, R]) e1 of
   151 	    Term_Val v => scan_dn cc (ist |> path_down [R] |> set_act v) e2
   152     | (*Accept_Tac*) goback => goback)
   153 
   154   | scan_dn cc ist (Const ("HOL.Let"(*1*), _) $ e $ (Abs (i, T, b))) =
   155      (case scan_dn cc (ist |> path_down [L, R]) e of
   156        Term_Val res => scan_dn cc (ist |> path_down [R, D] |> upd_env'' (Free (i, T), res)) b
   157     | (*Accept_Tac*) goback => goback)
   158 
   159   | scan_dn (cc as (_, ctxt)) (ist as {eval, act_arg, ...}) (Const ("Tactical.While"(*1*), _) $ c $ e $ a) =
   160     if Rewrite.eval_true_ (Proof_Context.theory_of ctxt) eval (subst_atomic (ist |> get_act_env |> Env.update' a) c)
   161     then scan_dn cc (ist |> path_down_form ([L, R], a)) e
   162     else Term_Val act_arg
   163   | scan_dn (cc as (_, ctxt)) (ist as {eval, act_arg, ...}) (Const ("Tactical.While"(*2*), _) $ c $ e) =
   164     if Rewrite.eval_true_ (Proof_Context.theory_of ctxt) eval (subst_atomic (Env.update_opt' (get_subst ist)) c)
   165     then scan_dn cc (ist |> path_down [R]) e
   166     else Term_Val act_arg
   167 
   168   |scan_dn cc ist (Const ("Tactical.Or"(*1*), _) $e1 $ e2 $ a) =
   169     (case scan_dn cc (ist |> path_down_form ([L, L, R], a)) e1 of
   170 	    Accept_Tac lme => Accept_Tac lme
   171     | _ => scan_dn cc (ist |> path_down_form ([L, R], a)) e2)
   172   | scan_dn cc ist (Const ("Tactical.Or"(*2*), _) $e1 $ e2) =
   173     (case scan_dn cc (ist |> path_down [L, R]) e1 of
   174 	    Accept_Tac lme => Accept_Tac lme
   175     | _ => scan_dn cc (ist |> path_down [R]) e2)
   176 
   177   |  scan_dn (cc as (_, ctxt)) (ist as {eval, ...}) (Const ("Tactical.If"(*1*), _) $ c $ e1 $ e2) =
   178     if Rewrite.eval_true_ (Proof_Context.theory_of ctxt) eval (subst_atomic (Env.update_opt' (get_subst ist)) c)
   179     then scan_dn cc (ist |> path_down [L, R]) e1
   180     else scan_dn cc (ist |> path_down [R]) e2
   181 
   182   | scan_dn (cc as (_, ctxt)) (ist as {eval, ...}) t =
   183     if Tactical.contained_in t
   184     then raise TERM ("scan_dn expects Prog_Tac or Prog_Expr", [t])
   185     else
   186       case LItool.check_leaf "next  " ctxt eval (get_subst ist) t of
   187   	    (Program.Expr s, _) => Term_Val s (*TODO?: include set_found here and drop those after call*)
   188       | (Program.Tac prog_tac, form_arg) =>
   189           check_tac cc ist (prog_tac, form_arg)
   190 
   191 fun go_scan_up (pcc as (sc, _)) (ist as {path, act_arg, found_accept, ...}) = 
   192     if path = [R] (*root of the program body*) then
   193       if found_accept then
   194         Term_Val act_arg
   195       else raise ERROR "LItool.find_next_step without result"
   196     else scan_up pcc (ist |> path_up) (go_up path sc)
   197 (* scan is strictly to R, because at L there was an \<open>expr_val\<close> *)
   198 and scan_up pcc ist (Const ("Tactical.Try"(*1*), _) $ _ $ _) = go_scan_up pcc ist
   199   | scan_up pcc ist (Const ("Tactical.Try"(*2*), _) $ _) = go_scan_up pcc ist
   200 
   201   | scan_up (pcc as (_, cc)) ist (Const ("Tactical.Repeat"(*1*), _) $ e $ _) =
   202     (case scan_dn cc (ist |> path_down [L, R]) e of
   203       Accept_Tac ict => Accept_Tac ict
   204     | Reject_Tac =>  go_scan_up pcc ist
   205     | Term_Val v =>  go_scan_up pcc (ist |> set_act v |> set_found))
   206   | scan_up (pcc as (_, cc)) ist (Const ("Tactical.Repeat"(*2*), _) $ e) =
   207     (case scan_dn cc (ist |> path_down [R]) e of
   208       Accept_Tac ict => Accept_Tac ict
   209     | Reject_Tac => go_scan_up pcc ist
   210     | Term_Val v => go_scan_up pcc (ist |> set_act v |> set_found))
   211 
   212   | scan_up pcc ist (Const ("Tactical.Chain"(*1*), _) $ _ $ _ $ _) = go_scan_up pcc ist
   213   | scan_up pcc ist (Const ("Tactical.Chain"(*2*), _) $ _ $ _) = go_scan_up pcc ist
   214   | scan_up (pcc as (sc, cc)) ist (Const ("Tactical.Chain"(*3*), _) $ _) =
   215       let 
   216         val e2 = check_Seq_up ist sc
   217       in
   218         case scan_dn cc (ist |> path_up_down [R]) e2 of
   219           Accept_Tac ict => Accept_Tac ict
   220         | Reject_Tac => go_scan_up pcc (ist |> path_up)
   221         | Term_Val v => go_scan_up pcc (ist |> path_up |> set_act v |> set_found)
   222       end
   223 
   224   | scan_up (pcc as (sc, cc)) ist (Const ("HOL.Let"(*1*), _) $ _) =
   225 	    let
   226         val (i, body) = check_Let_up ist sc
   227       in
   228         case scan_dn cc (ist |> path_up_down [R, D] |> upd_env i) body of
   229 	        Accept_Tac ict => Accept_Tac ict
   230 	      | Reject_Tac => go_scan_up pcc (ist |> path_up)
   231 	      | Term_Val v => go_scan_up pcc (ist |> path_up |> set_act v |> set_found)
   232 	    end
   233   | scan_up pcc ist (Abs _(*2*)) = go_scan_up pcc ist
   234   | scan_up pcc ist (Const ("HOL.Let"(*3*), _) $ _ $ (Abs _)) = go_scan_up pcc ist
   235 
   236   | scan_up (pcc as (_, cc as (_, ctxt)))  (ist as {eval, ...})
   237       (Const ("Tactical.While"(*1*), _) $ c $ e $ _) = 
   238     if Rewrite.eval_true_ (Proof_Context.theory_of ctxt) eval (subst_atomic (Env.update_opt' (get_subst ist)) c) 
   239     then
   240       case scan_dn cc (ist |> path_down [L, R]) e of
   241   	     Accept_Tac ict => Accept_Tac ict
   242   	  | Reject_Tac => go_scan_up pcc ist
   243   	  | Term_Val v => go_scan_up pcc (ist |> set_act v |> set_found)
   244     else
   245       go_scan_up pcc (ist (*|> set_found*))
   246   | scan_up  (pcc as (_, cc as (_, ctxt))) (ist as {eval, ...})
   247       (Const ("Tactical.While"(*2*), _) $ c $ e) = 
   248     if Rewrite.eval_true_ (Proof_Context.theory_of ctxt) eval (subst_atomic (Env.update_opt' (get_subst ist)) c) 
   249     then
   250       case scan_dn cc (ist |> path_down [R]) e of
   251   	    Accept_Tac ict => Accept_Tac ict
   252   	  | Reject_Tac => go_scan_up pcc ist
   253   	  | Term_Val v => go_scan_up pcc (ist |> set_act v |> set_found)
   254     else
   255       go_scan_up pcc ist
   256 
   257   | scan_up pcc ist (Const ("Tactical.Or"(*1*), _) $ _ $ _ $ _) = go_scan_up pcc ist
   258   | scan_up pcc ist (Const ("Tactical.Or"(*2*), _) $ _ $ _) = go_scan_up pcc ist
   259   | scan_up pcc ist (Const ("Tactical.Or"(*3*), _) $ _ ) = go_scan_up pcc ist
   260 
   261   | scan_up pcc ist (Const ("Tactical.If"(*1*), _) $ _ $ _ $ _) = go_scan_up pcc ist
   262 
   263   | scan_up _ _ t = error ("scan_up not impl for " ^ UnparseC.term t)
   264 
   265 (* scan program until an applicable tactic is found *)
   266 fun scan_to_tactic (prog, cc) (Pstate (ist as {path, ...})) =
   267   if path = [] then
   268     scan_dn cc (trans_scan_dn ist) (Program.body_of prog)
   269   else go_scan_up (prog, cc) (trans_scan_up ist)
   270 | scan_to_tactic _ _ = raise ERROR "scan_to_tactic: uncovered pattern";
   271 
   272 (* find the next applicable Prog_Tac in a prog *)
   273 fun find_next_step (Rule.Prog prog) (ptp as(pt, (p, _))) (Pstate ist) ctxt =
   274    (case scan_to_tactic (prog, (ptp, ctxt)) (Pstate ist) of
   275       Accept_Tac (tac, ist, ctxt) =>
   276         Next_Step (Pstate ist, Tactic.insert_assumptions tac ctxt, tac)
   277     | Term_Val prog_result =>
   278         (case parent_node pt p of
   279           (true, p', _) => 
   280             let
   281               val (_, pblID, _) = get_obj g_spec pt p';
   282             in
   283               End_Program (Pstate ist, Tactic.Check_Postcond' (pblID, prog_result))
   284             end
   285         | _ => End_Program (Pstate ist, Tactic.End_Detail' (TermC.empty,[])))
   286     | Reject_Tac => Helpless)
   287   | find_next_step _ _ ist _ =
   288     raise ERROR ("find_next_step: not impl for " ^ Istate.string_of ist);
   289 
   290 
   291 (*** locate an input tactic within a program ***)
   292 
   293 datatype input_tactic_result =
   294     Safe_Step of Istate.T * Proof.context * Tactic.T
   295   | Unsafe_Step of Istate.T * Proof.context * Tactic.T
   296   | Not_Locatable of string
   297 
   298 (*all functions ending with "1" are supposed to be replaced by those without "1"*)
   299 datatype expr_val1 = (* "ExprVal" in the sense of denotational semantics        *)
   300   Reject_Tac1 of     (* tactic is found but NOT acknowledged, scan is continued *)
   301     Istate.pstate * Proof.context * Tactic.T (*TODO: revise Pstate {or,...},no args as Reject_Tac*)
   302 | Accept_Tac1 of     (* tactic is found and acknowledged, scan is stalled       *)
   303     Istate.pstate *  (* the current state, returned for resuming execution      *)
   304     Proof.context *  (* updated according to evaluation of Prog_Tac             *)
   305     Tactic.T         (* Prog_Tac is associated to Tactic.input                  *)
   306 | Term_Val1 of       (* Prog_Expr is found and evaluated, scan is continued     *)
   307 	  term             (* value of Prog_Expr, for updating environment            *)
   308 fun assoc2str (Accept_Tac1 _) = "Accept_Tac1"                                         
   309   | assoc2str (Term_Val1 _) = "Term_Val1"
   310   | assoc2str (Reject_Tac1 _) = "Reject_Tac1";
   311 
   312 
   313 (** check a Prog_Tac: is it associated to Tactic ? **)
   314 
   315 fun check_tac1 ((pt, p), ctxt, tac) (ist as {act_arg, or, ...}) (prog_tac, form_arg) =
   316   case LItool.associate pt ctxt (tac, prog_tac) of
   317       LItool.Associated      (m, v', ctxt)
   318         => Accept_Tac1 (ist |> set_subst_true  (form_arg, v') |> set_found, ctxt, m)
   319     | LItool.Ass_Weak (m, v', ctxt) (*the ONLY ones in locate_tac ^v^v^v^v^ *)
   320         => Accept_Tac1 (ist |> set_subst_false (form_arg, v') |> set_found, ctxt, m)
   321     | LItool.Not_Associated =>
   322       (case or of (* switch for Tactical.Or: 1st AssOnly, 2nd AssGen *)
   323          AssOnly => Term_Val1 act_arg
   324        | _(*ORundef*) =>
   325           case Applicable.applicable_in p pt (LItool.tac_from_prog pt (Celem.assoc_thy "Isac_Knowledge") prog_tac) of
   326              Applicable.Appl m' => Reject_Tac1 (ist |> set_subst_false (form_arg, Tactic.result m'), ctxt, tac)
   327            | Applicable.Notappl _ => Term_Val1 act_arg)
   328 
   329 (** scan to a Prog_Tac **)
   330 
   331 (* scan is strictly first L, then R; tacticals have 2 args at maximum *)
   332 fun scan_dn1 cct ist (Const ("Tactical.Try"(*1*), _) $ e $ a) =
   333     (case scan_dn1 cct (ist |> path_down_form ([L, R], a)) e of goback => goback) 
   334   | scan_dn1 cct ist (Const ("Tactical.Try"(*2*), _) $ e) =
   335     (case scan_dn1 cct (ist |> path_down [R]) e of goback => goback)
   336 
   337   | scan_dn1 cct ist (Const ("Tactical.Repeat"(*1*), _) $ e $ a) =
   338     scan_dn1 cct (ist |> path_down_form ([L, R], a)) e
   339   | scan_dn1 cct ist (Const ("Tactical.Repeat"(*2*), _) $ e) =
   340     scan_dn1 cct (ist |> path_down [R]) e
   341 
   342   | scan_dn1 (cct as (cstate, _, _)) ist (Const ("Tactical.Chain"(*1*), _) $ e1 $ e2 $ a) =
   343     (case scan_dn1 cct (ist |> path_down_form ([L, L, R], a)) e1 of
   344 	     Term_Val1 v => scan_dn1 cct (ist |> path_down [L, R] |> set_subst' (a, v)) e2
   345      | Reject_Tac1 (ist', ctxt', tac') => scan_dn1 (cstate, ctxt', tac') (ist
   346         |> path_down_form ([L, R], a) |> trans_env_act ist') e2
   347      | goback => goback)
   348   | scan_dn1 (cct as (cstate, _, _)) ist (Const ("Tactical.Chain"(*2*), _) $e1 $ e2) =
   349     (case scan_dn1 cct (ist |> path_down [L, R]) e1 of
   350 	     Term_Val1 v => scan_dn1 cct (ist |> path_down [R] |> set_act v) e2
   351      | Reject_Tac1 (ist', ctxt', tac') =>
   352         scan_dn1 (cstate, ctxt', tac') (ist |> path_down [R] |> trans_env_act ist') e2
   353      | goback => goback)
   354 
   355   | scan_dn1 cct ist (Const ("HOL.Let", _) $ e $ (Abs (id, T, b))) =
   356     (case scan_dn1 cct (ist |> path_down [L, R]) e of
   357        Reject_Tac1 (ist', _, _) =>
   358          scan_dn1 cct (ist' |> path_down [R, D] |> upd_env (TermC.mk_Free (id, T)) |> trans_ass ist) b
   359      | Term_Val1 v =>
   360          scan_dn1 cct (ist |> path_down [R, D] |> upd_env'' (TermC.mk_Free (id, T), v)) b
   361      | goback => goback)
   362 
   363   | scan_dn1 (cct as (_, ctxt, _)) (ist as {eval, act_arg, ...})
   364       (Const ("Tactical.While"(*1*), _) $ c $ e $ a) =
   365     if Rewrite.eval_true_ (Proof_Context.theory_of ctxt) eval (subst_atomic (ist |> get_act_env |> Env.update' a) c)
   366     then scan_dn1 cct (ist |> path_down_form ([L, R], a)) e
   367     else Term_Val1 act_arg
   368   | scan_dn1 (cct as (_, ctxt, _)) (ist as {eval, act_arg, ...})
   369       (Const ("Tactical.While"(*2*),_) $ c $ e) =
   370     if Rewrite.eval_true_ (Proof_Context.theory_of ctxt) eval (subst_atomic (Env.update_opt' (get_subst ist)) c) 
   371     then scan_dn1 cct (ist |> path_down [R]) e
   372     else Term_Val1 act_arg
   373 
   374   | scan_dn1 cct (ist as {or = ORundef, ...}) (Const ("Tactical.Or"(*1*), _) $e1 $ e2 $ a) =
   375     (case scan_dn1 cct (ist |> path_down_form ([L, L, R], a) |> set_or AssOnly) e1 of
   376 	     Term_Val1 v =>
   377 	       (case scan_dn1 cct (ist |> path_down [L, R] |> set_subst' (a, v) |> set_or AssOnly) e2 of
   378 	          Term_Val1 v => 
   379 	            (case scan_dn1 cct (ist |> path_down [L, L, R] |> upd_env'' (a, v) |> set_or AssGen) e1 of
   380 	               Term_Val1 v => 
   381 	                 scan_dn1 cct (ist |> path_down [L, R] |> upd_env'' (a, v) |> set_or AssGen) e2
   382 	             | goback => goback)
   383 	        | goback => goback)
   384      | Reject_Tac1 _ => raise ERROR ("scan_dn1 Tactical.Or(*1*): must not return Reject_Tac1")
   385      | goback => goback)
   386   | scan_dn1 cct ist (Const ("Tactical.Or"(*2*), _) $e1 $ e2) =
   387     (case scan_dn1 cct (ist |> path_down [L, R]) e1 of
   388 	     Term_Val1 v => scan_dn1 cct (ist |> path_down [R] |> set_act v) e2
   389      | goback => goback)
   390 
   391   | scan_dn1 (cct as (_, ctxt, _)) (ist as {eval, ...}) (Const ("Tactical.If", _) $ c $ e1 $ e2) =
   392     if Rewrite.eval_true_ (Proof_Context.theory_of ctxt) eval (subst_atomic (Env.update_opt' (get_subst ist)) c)
   393     then scan_dn1 cct (ist |> path_down [L, R]) e1
   394     else scan_dn1 cct (ist |> path_down [R]) e2
   395 
   396   | scan_dn1 (cct as (_, ctxt, _)) (ist as {eval, ...}) t =
   397     if Tactical.contained_in t then raise TERM ("scan_dn1 expects Prog_Tac or Prog_Expr", [t])
   398     else
   399       case LItool.check_leaf "locate" ctxt eval (get_subst ist) t of
   400   	    (Program.Expr _, form_arg) => 
   401   	      Term_Val1 (Rewrite.eval_prog_expr (Proof_Context.theory_of ctxt) eval
   402   		      (subst_atomic (Env.update_opt'' (get_act_env ist, form_arg)) t))
   403       | (Program.Tac prog_tac, form_arg) =>
   404           check_tac1 cct ist (prog_tac, form_arg)
   405 
   406 fun go_scan_up1 (pcct as (prog, _)) (ist as {path, act_arg, ...}) =
   407     if 1 < length path then
   408       scan_up1 pcct (ist |> path_up) (at_location (path_up' path) prog)
   409     else Term_Val1 act_arg
   410 (* scan is strictly to R, because at L there was a expr_val *)
   411 and scan_up1 pcct ist (Const ("Tactical.Try"(*1*), _) $ _ $ _) = go_scan_up1 pcct ist
   412   | scan_up1 pcct ist (Const ("Tactical.Try"(*2*), _) $ _) = go_scan_up1 pcct ist
   413 
   414   | scan_up1 (pcct as (prog, cct as (cstate, _, _))) ist (t as Const ("Tactical.Repeat"(*1*), _) $ e $ a) =
   415     (case scan_dn1 cct (ist |> path_down_form ([L, R], a) |> set_or ORundef) e of 
   416       Term_Val1 v => go_scan_up1 pcct (ist |> set_act v |> set_form a)
   417     | Reject_Tac1 (ist', ctxt', tac') => scan_up1 (prog, (cstate, ctxt', tac')) ist' t
   418     | goback => goback)
   419   | scan_up1 (pcct as (prog, cct as (cstate, _, _))) ist (t as Const ("Tactical.Repeat"(*2*), _) $ e) =
   420     (case scan_dn1 cct (ist |> path_down [R] |> set_or ORundef) e of
   421        Term_Val1 v' => go_scan_up1 pcct (ist |> set_act v')
   422      | Reject_Tac1 (ist', ctxt', tac') => scan_up1 (prog, (cstate, ctxt', tac')) ist' t
   423      | goback => goback)
   424 
   425     (*all has been done in (*2*) below*)
   426   | scan_up1 pcct ist (Const ("Tactical.Chain"(*1*), _) $ _ $ _ $ _) = go_scan_up1 pcct ist
   427   | scan_up1 pcct ist (Const ("Tactical.Chain"(*2*), _) $ _ $ _) = go_scan_up1 pcct ist (*comes from e2*)    
   428   | scan_up1 (pcct as (prog, cct as (cstate, _, _))) ist (Const ("Tactical.Chain"(*3*), _) $ _ ) =
   429      let 
   430        val e2 = check_Seq_up ist prog (*comes from e1, goes to e2*)
   431      in
   432        case scan_dn1 cct (ist |> path_up_down [R] |> set_or ORundef) e2 of
   433          Term_Val1 v => go_scan_up1 pcct (ist |> path_up |> set_act v)
   434        | Reject_Tac1 (ist', ctxt', tac') => go_scan_up1 (prog, (cstate, ctxt', tac')) ist'
   435        | goback => goback 
   436      end
   437 
   438   | scan_up1 (pcct as (prog, cct as (cstate, _, _))) ist (Const ("HOL.Let"(*1*), _) $ _) =
   439     let 
   440       val (i, body) = check_Let_up ist prog
   441     in case scan_dn1 cct (ist |> path_up_down [R, D] |> upd_env i |> set_or ORundef) body of
   442 	       Accept_Tac1 iss => Accept_Tac1 iss
   443 	     | Reject_Tac1 (ist', ctxt', tac') => go_scan_up1 (prog, (cstate, ctxt', tac')) ist'
   444 	     | Term_Val1 v => go_scan_up1 pcct (ist |> path_up |> set_act v) 
   445 	  end
   446   | scan_up1 pcct ist (Abs(*2*) _) = go_scan_up1 pcct ist
   447   | scan_up1 pcct ist (Const ("HOL.Let"(*3*), _) $ _ $ (Abs _)) = go_scan_up1 pcct ist
   448 
   449   | scan_up1 (pcct as (prog, cct as (cstate, ctxt, _))) (ist as {eval, ...})
   450       (t as Const ("Tactical.While"(*1*),_) $ c $ e $ a) =
   451     if Rewrite.eval_true_ (Proof_Context.theory_of ctxt) eval (subst_atomic (Env.update' a (get_act_env ist)) c)
   452     then
   453       case scan_dn1 cct (ist |> path_down_form ([L, R], a) |> set_or ORundef) e of 
   454         Term_Val1 v => go_scan_up1 pcct (ist |> set_act v |> set_form a)
   455 
   456       | Reject_Tac1 (ist', ctxt', tac') => scan_up1 (prog, (cstate, ctxt', tac')) ist' t
   457       | goback => goback
   458     else go_scan_up1 pcct (ist |> set_form a)
   459   | scan_up1 (pcct as (prog, cct as (cstate, ctxt, _))) (ist as {eval, ...})
   460       (t as Const ("Tactical.While"(*2*), _) $ c $ e) =
   461     if Rewrite.eval_true_ (Proof_Context.theory_of ctxt) eval (subst_atomic (Env.update_opt' (get_subst ist)) c)
   462     then
   463       case scan_dn1 cct (ist |> path_down [R] |> set_or ORundef) e of 
   464         Term_Val1 v => go_scan_up1 pcct (ist |> set_act v)
   465        | Reject_Tac1 (ist', ctxt', tac') => scan_up1 (prog, (cstate, ctxt', tac')) ist' t
   466        | goback => goback
   467     else go_scan_up1 pcct ist
   468 
   469   | scan_up1 pcct ist (Const ("If", _) $ _ $ _ $ _) = go_scan_up1 pcct ist
   470 
   471   | scan_up1 pcct ist (Const ("Tactical.Or"(*1*), _) $ _ $ _ $ _) = go_scan_up1 pcct ist
   472   | scan_up1 pcct ist (Const ("Tactical.Or"(*2*), _) $ _ $ _) = go_scan_up1 pcct ist
   473   | scan_up1 pcct ist (Const ("Tactical.Or"(*3*), _) $ _ ) = go_scan_up1 pcct (ist |> path_up)
   474 
   475   | scan_up1 pcct ist (Const ("Tactical.If",_) $ _ $ _ $ _) = go_scan_up1 pcct ist
   476 
   477   | scan_up1 _ _ t = error ("scan_up1 not impl for t= " ^ UnparseC.term t)
   478 
   479 fun scan_to_tactic1 (prog, (cctt as ((_, p), _, _))) (Pstate (ist as {path, ...})) =
   480     if path = [] orelse Pos.at_first_tactic(*RM prog path RM WITH find_next_step IN solve*)p
   481     then scan_dn1 cctt (ist |> set_path [R] |> set_or ORundef) (Program.body_of prog)
   482     else go_scan_up1 (prog, cctt) ist
   483   | scan_to_tactic1 _ _ = raise ERROR "scan_to_tactic1: uncovered pattern in fun.def"
   484 
   485 (*locate an input tactic within a program*)
   486 fun locate_input_tactic (Rule.Prog prog) cstate istate ctxt tac =
   487     (case scan_to_tactic1 (prog, (cstate, ctxt, tac)) istate of
   488          Accept_Tac1 ((ist as {assoc, ...}), ctxt, tac') =>
   489           if assoc then
   490             Safe_Step (Pstate ist, ctxt, tac')
   491  	        else Unsafe_Step (Pstate ist, ctxt, tac')
   492       | Reject_Tac1 _ => Not_Locatable (Tactic.string_of tac ^ " not locatable")
   493       | err => raise ERROR ("not-found-in-program: NotLocatable from " ^ @{make_string} err))
   494   | locate_input_tactic _ _ _ _ _ = raise ERROR "locate_input_tactic: uncovered case in definition"
   495 
   496                                                                    
   497 (*** locate an input formula within a program ***)
   498 
   499 datatype input_term_result = Found_Step of Calc.T | Not_Derivable
   500 
   501 fun by_tactic (Tactic.Apply_Method' (mI, _, _, _)) (_, ctxt) (pt, pos as (p, _)) =
   502       let
   503          val (itms, (*l..*)oris, probl(*..l*)) = case get_obj I pt p of
   504            PblObj {meth = itms, origin = (oris, _, _), probl, ...} => (itms, oris, probl)
   505          | _ => error "LI.by_tactic Apply_Method': uncovered case get_obj"
   506          val {ppc, ...} = Specify.get_met mI;
   507          val itms = if itms <> [] then itms else Chead.complete_metitms oris probl [] ppc
   508          val itms = Specify.hack_until_review_Specify_1 mI itms
   509          val srls = LItool.get_simplifier (pt, pos)
   510          val (is, env, ctxt, prog) = case LItool.init_pstate srls ctxt itms mI of
   511            (is as Pstate {env, ...}, ctxt, scr) => (is, env, ctxt, scr)
   512          | _ => error "LI.by_tactic Apply_Method': uncovered case init_pstate"
   513         val ini = LItool.implicit_take prog env;
   514         val pos = start_new_level pos
   515       in 
   516         case ini of
   517           SOME t =>
   518           let                                                                   (* implicit Take *)
   519             val show_add = Tactic.Apply_Method' (mI, SOME t, is, ctxt);
   520             val (pos, c, _, pt) = Generate.generate1 show_add (is, ctxt) (pt, pos)
   521           in
   522            ("ok", ([(Tactic.Apply_Method mI, show_add, (pos, (is, ctxt)))], c, (pt, pos)))
   523           end
   524         | NONE =>
   525           let
   526             val (tac', ist', ctxt') = 
   527               case find_next_step prog (pt, (lev_dn p, Res)) is ctxt of
   528                 Next_Step (ist, ctxt, tac) => (tac, ist, ctxt)
   529               | _ => raise ERROR ("LI.by_tactic..Apply_Method find_next_step \<rightarrow> " ^ strs2str' mI)
   530           in
   531             case tac' of
   532               Tactic.Take' t =>
   533                 let                                                             (* explicit Take *)
   534                   val show_add = Tactic.Apply_Method' (mI, SOME t, ist', ctxt');
   535                   val (pos, c, _, pt) = Generate.generate1 show_add (ist', ctxt') (pt, pos)
   536                 in
   537                  ("ok", ([(Tactic.Apply_Method mI, show_add, (pos, (is, ctxt)))], c, (pt, pos)))
   538                 end
   539             | add as Tactic.Subproblem' (_, _, headline, _, _, _) =>
   540                 let                                                                (* Subproblem *)
   541                   val show = Tactic.Apply_Method' (mI, SOME headline, ist', ctxt');
   542                   val (pos, c, _, pt) = Generate.generate1 add (ist', ctxt') (pt, pos)
   543                 in
   544                  ("ok", ([(Tactic.Apply_Method mI, show, (pos, (ist', ctxt')))], c, (pt, pos)))
   545                 end
   546             | tac =>
   547                 raise ERROR ("LI.by_tactic..Apply_Method' does NOT expect " ^ Tactic.string_of tac)
   548           end
   549       end
   550   | by_tactic (Tactic.Check_Postcond' (pI, _)) (sub_ist, sub_ctxt) (pt, pos as (p, _)) =
   551       let
   552         val parent_pos = par_pblobj pt p
   553         val {scr, ...} = Specify.get_met (get_obj g_metID pt parent_pos);
   554         val prog_res =
   555            case find_next_step scr (pt, pos) sub_ist sub_ctxt of
   556   (*OLD*)    Next_Step (_, _, Tactic.Check_elementwise' (_, _, (prog_res, _))) => prog_res
   557     |(*OLD*) End_Program (_, Tactic.Check_Postcond' (_, prog_res)) => prog_res
   558            | _ => raise ERROR ("Step_Solve.by_tactic Check_Postcond " ^ strs2str' pI)
   559       in
   560         if parent_pos = [] then 
   561   	      let
   562             val tac = Tactic.Check_Postcond' (pI, prog_res)
   563             val is = Pstate (sub_ist |> the_pstate |> set_act prog_res |> set_found)
   564   	        val ((p, p_), ps, _, pt) = Generate.generate1 tac (is, sub_ctxt) (pt, (parent_pos, Res))
   565   	      in
   566   	        ("ok", ([(Tactic.Check_Postcond pI, tac, ((parent_pos, Res), (is, sub_ctxt)))], ps, (pt, (p, p_))))
   567   	      end
   568         else
   569           let (*resume program of parent PblObj, transfer result of Subproblem-program*)
   570             val (ist_parent, ctxt_parent) = case get_loc pt (parent_pos, Frm) of
   571               (Pstate i, c) => (i, c)
   572             | _ => error "LI.by_tactic Check_Postcond': uncovered case get_loc"
   573             val (prog_res', ctxt') = ContextC.subpbl_to_caller sub_ctxt prog_res ctxt_parent
   574             val tac = Tactic.Check_Postcond' (pI, prog_res')
   575             val ist' = Pstate (ist_parent |> set_act prog_res |> set_found)
   576             val ((p, p_), ps, _, pt) = Generate.generate1 tac (ist', ctxt') (pt, (parent_pos, Res))
   577           in
   578             ("ok", ([(Tactic.input_from_T tac, tac, ((parent_pos, Res), (ist', ctxt')))], ps, (pt, (p, p_))))
   579           end
   580       end
   581   | by_tactic (Tactic.End_Proof'') _ ptp = ("end-of-calculation", ([], [], ptp))
   582   | by_tactic tac_ ic (pt, pos) =
   583     let
   584       val pos = next_in_prog' pos
   585  	    val (pos', c, _, pt) = Generate.generate1 tac_ ic (pt, pos);
   586     in
   587       ("ok", ([(Tactic.input_from_T tac_, tac_, (pos, ic))], c, (pt, pos')))
   588     end
   589 (*find_next_step from program, by_tactic will update Ctree*)
   590 and do_next (ptp as (pt, pos as (p, p_))) =
   591     if Celem.e_metID = get_obj g_metID pt (par_pblobj pt p) then
   592       ("helpless", ([], [], (pt, (p, p_))))
   593     else 
   594       let 
   595         val thy' = get_obj g_domID pt (par_pblobj pt p);
   596 	      val ((ist, ctxt), sc) = LItool.resume_prog thy' (p,p_) pt;
   597       in
   598         case find_next_step sc (pt, pos) ist ctxt of
   599            Next_Step (ist, ctxt, tac) =>
   600              by_tactic tac (ist, Tactic.insert_assumptions tac ctxt) ptp
   601          | End_Program (ist, tac) => (*TODO RM ..*)
   602              (case tac of
   603                Tactic.End_Detail' res =>
   604                  ("ok", ([(Tactic.End_Detail, 
   605                    Tactic.End_Detail' res, (pos, (ist, ctxt)))], [], ptp))
   606              | _ => (*.. RM*) by_tactic tac (ist, ctxt) ptp
   607              )
   608          | Helpless => ("helpless", Chead.e_calcstate')
   609       end;
   610 
   611 (*
   612   compare inform with ctree.form at current pos by nrls;
   613   if found, embed the derivation generated during comparison
   614   if not, let the mat-engine compute the next ctree.form.
   615 
   616   Code's structure is copied from complete_solve
   617   CAUTION: tacis in returned calcstate' do NOT construct resulting ptp --
   618            all_modspec etc. has to be inserted at Subproblem'
   619 *)
   620 fun compare_step (tacis, c, ptp as (pt, pos as (p, _))) ifo =
   621   let
   622     val fo = Calc.get_current_formula ptp
   623 	  val {nrls, ...} = Specify.get_met (Ctree.get_obj Ctree.g_metID pt (Ctree.par_pblobj pt p))
   624 	  val {rew_ord, erls, rules, ...} = Rule_Set.rep nrls
   625 	  val (found, der) = Error_Fill_Pattern.concat_deriv rew_ord erls rules fo ifo; (*<---------------*)
   626   in
   627     if found
   628     then
   629        let
   630          val tacis' = map (Error_Fill_Pattern.mk_tacis rew_ord erls) der;
   631 		     val (c', ptp) = Generate.embed_deriv tacis' ptp;
   632 	     in ("ok", (tacis (*@ tacis'?WN050408*), c @ c', ptp)) end
   633      else 
   634 	     if pos = ([], Pos.Res) (*TODO: we should stop earlier with trying subproblems *)
   635 	     then ("no derivation found", (tacis, c, ptp): Chead.calcstate') 
   636 	     else
   637          let
   638            val msg_cs' as (_, (tacis, c', ptp)) = do_next ptp; (*<---------------------*)
   639 		       val (_, (tacis, c'', ptp)) = case tacis of
   640 			       ((Tactic.Subproblem _, _, _) :: _) => 
   641 			         let
   642                  val ptp as (pt, (p, _)) = Chead.all_modspec ptp (*<--------------------*)
   643 				         val mI = Ctree.get_obj Ctree.g_metID pt p
   644 			         in
   645 			           by_tactic (Tactic.Apply_Method' (mI, NONE, empty, ContextC.empty))
   646 			             (empty, ContextC.empty) ptp
   647                end
   648 			     | _ => msg_cs';
   649 		     in compare_step (tacis, c @ c' @ c'', ptp) ifo end
   650   end
   651 
   652 (* Locate a step in a program, which has been determined by input of a term *)
   653 fun locate_input_term (pt, pos) tm =
   654      let                                                          
   655    		val pos_pred = Pos.lev_back' pos (*f_pred ---"step pos cs"---> f_succ in appendFormula*)
   656    		val _(*f_pred*) = Ctree.get_curr_formula (pt, pos_pred)
   657      in
   658    		case compare_step ([], [], (pt, pos_pred)) tm of
   659    		   ("no derivation found", _) => Not_Derivable
   660        | ("ok", (_, _, cstate)) => 
   661            Found_Step cstate
   662        | _ => raise ERROR "compare_step: uncovered case"
   663      end
   664 
   665 (**)end(**)