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