src/Tools/isac/Interpret/lucas-interpreter.sml
author Walther Neuper <walther.neuper@jku.at>
Sat, 16 Nov 2019 17:46:08 +0100
changeset 59705 be30fa5a7b76
parent 59704 178869f0f4ba
child 59706 a462f2e893c5
permissions -rw-r--r--
lucin: unify scanning to tactics
wneuper@59562
     1
(* Title:  Interpret/lucas-interpreter.sml
wneuper@59554
     2
   Author: Walther Neuper 2019
wneuper@59554
     3
   (c) due to copyright terms
wneuper@59554
     4
*)
wneuper@59554
     5
wneuper@59554
     6
signature LUCAS_INTERPRETER =
wneuper@59554
     7
sig
wneuper@59567
     8
  datatype input_tactic_result =
walther@59655
     9
      Safe_Step of Istate.T * Proof.context * Tactic.T
walther@59655
    10
    | Unsafe_Step of Istate.T * Proof.context * Tactic.T
wneuper@59569
    11
    | Not_Locatable of string
wneuper@59572
    12
  val locate_input_tactic : Rule.program -> Ctree.state -> Istate.T -> Proof.context -> Tactic.T
wneuper@59569
    13
    -> input_tactic_result
wneuper@59557
    14
wneuper@59564
    15
(*datatype input_formula_result = FStep of calcstate' | Not_Derivable *)
wneuper@59562
    16
  datatype input_formula_result =
wneuper@59572
    17
    Step of Ctree.state * Istate.T * Proof.context
wneuper@59562
    18
  | Not_Derivable of Chead.calcstate'
wneuper@59574
    19
(*val locate_input_formula : Ctree.state -> term -> input_formula_result *)
wneuper@59572
    20
  val locate_input_formula : Rule.program -> Ctree.state -> Istate.T -> Proof.context -> term
wneuper@59562
    21
    -> input_formula_result
wneuper@59564
    22
(*val begin_end_prog : ???*)
wneuper@59572
    23
  val begin_end_prog : Tactic.T -> Istate.T * Proof.context -> Ctree.state ->
wneuper@59572
    24
    (Tactic.input * Tactic.T * (Ctree.pos' * (Istate.T * Proof.context))) list * Ctree.pos' list * Ctree.state
wneuper@59574
    25
(*val do_solve_step : Ctree.state -> Ctree.state ???*)
wneuper@59574
    26
  val do_solve_step : Ctree.state ->
wneuper@59574
    27
    (Tactic.input * Tactic.T * (Ctree.pos' * (Istate.T * Proof.context))) list * Ctree.pos' list * Ctree.state
wneuper@59554
    28
wneuper@59569
    29
  datatype next_tactic_result =
wneuper@59572
    30
      NStep of Ctree.state * Istate.T * Proof.context * tactic
wneuper@59569
    31
    | Helpless | End_Program
wneuper@59569
    32
(*val determine_next_tactic :
wneuper@59572
    33
    Rule.program -> Ctree.state -> Istate.T -> Proof.context -> next_tactic_result *)
walther@59686
    34
  val determine_next_tactic: Rule.theory' * 'a -> Ctree.state -> Rule.program -> Istate.T * 'b
walther@59686
    35
    -> Tactic.T * (Istate.T * Proof.context) * (term * Telem.safe)
wneuper@59569
    36
wneuper@59557
    37
(* ---- for tests only: shifted from below to remove the Warning "unused" at fun.def. --------- *)
wneuper@59567
    38
  datatype assoc =
walther@59667
    39
      Assoc of Istate.pstate * Proof.context * Tactic.T
walther@59667
    40
    | NasApp of Istate.pstate * Proof.context * Tactic.T
walther@59705
    41
    | Skip1 of  term * Env.T;
wneuper@59557
    42
  val assoc2str : assoc -> string
walther@59702
    43
(*/-------------------------------------------------------- ! aktivate for Test_Isac BEGIN ---\* )
walther@59687
    44
  val go : Celem.path -> term -> term
walther@59687
    45
  val go_up: Celem.path -> term -> term
walther@59685
    46
  val check_Let_up : Istate.pstate -> term -> term * term
walther@59685
    47
  val check_Seq_up: Istate.pstate -> term -> term
walther@59691
    48
  val compare_step : Generate.taci list * Ctree.pos' list * (Ctree.state) -> term -> string * Chead.calcstate'
walther@59686
    49
walther@59691
    50
  val scan_dn1: (Ctree.state * Proof.context * Tactic.T) -> Istate.pstate -> term -> assoc
walther@59692
    51
  val scan_up1: term * (Ctree.state * Proof.context * Tactic.T) -> Istate.pstate -> term -> assoc;
walther@59692
    52
  val go_scan_up1: term * (Ctree.state * Proof.context * Tactic.T) -> Istate.pstate -> assoc;
walther@59692
    53
  val scan_to_tactic1: term * (Ctree.state * Proof.context * Tactic.T) -> Istate.T -> assoc
walther@59686
    54
walther@59705
    55
  datatype appy = Appy of Tactic.T * Istate.pstate | Napp of Env.T | Skip2 of term * Env.T
walther@59699
    56
  val scan_dn2: Ctree.state * Rule.theory'(*Proof.context*) -> Istate.pstate -> term -> appy
walther@59699
    57
  val scan_up2: term * (Ctree.state * Rule.theory') -> Istate.pstate -> term -> appy
walther@59699
    58
  val go_scan_up2: term * (Ctree.state * Rule.theory') -> Istate.pstate -> appy
walther@59699
    59
  val scan_to_tactic2: term * (Ctree.state * Rule.theory') -> Istate.T -> appy
walther@59702
    60
( *\--- ! aktivate for Test_Isac END ----------------------------------------------------------/*)
wneuper@59554
    61
end
wneuper@59554
    62
walther@59684
    63
(**)
wneuper@59557
    64
structure LucinNEW(**): LUCAS_INTERPRETER(**) = (*LucinNEW \<rightarrow> Lucin after code re-arrangement*)
wneuper@59554
    65
struct
walther@59684
    66
(**)
wneuper@59554
    67
open Ctree
walther@59695
    68
open Pos
walther@59659
    69
open Celem
walther@59661
    70
open Istate
wneuper@59554
    71
walther@59701
    72
(*** auxiliary functions ***)
wneuper@59567
    73
walther@59701
    74
(*go to a location in a program and fetch the resective sub-term*)
wneuper@59554
    75
fun go [] t = t
walther@59692
    76
  | go (D :: p) (Abs(_, _, body)) = go (p : Celem.path) body
walther@59661
    77
  | go (L :: p) (t1 $ _) = go p t1
walther@59661
    78
  | go (R :: p) (_ $ t2) = go p t2
walther@59687
    79
  | go l _ = error ("go: no " ^ Celem.path2str l);
walther@59668
    80
fun go_up l t =
walther@59668
    81
  if length l > 1 then go (drop_last l) t else raise ERROR ("go_up [] " ^ Rule.term2str t)
wneuper@59554
    82
walther@59685
    83
fun check_Let_up ({path, ...}: pstate) prog =
walther@59668
    84
  case go (drop_last path) prog of
walther@59664
    85
    Const ("HOL.Let",_) $ _ $ (Abs (i, T, body)) => (TermC.mk_Free (i, T), body)
walther@59691
    86
  | t => raise ERROR ("scan_up1..\"HOL.Let $ _\" with: \"" ^ Rule.term2str t ^ "\"")
walther@59685
    87
fun check_Seq_up  ({path, ...}: pstate) prog =
walther@59664
    88
  case go (drop_last path) prog of
walther@59688
    89
    Const ("Tactical.Chain",_) $ _ $ e2=> e2
walther@59691
    90
  | t => raise ERROR ("scan_up1..\"Tactical.Chain $ _\" with: \"" ^ Rule.term2str t ^ "\"")
walther@59691
    91
walther@59701
    92
walther@59701
    93
(*** locate an input tactic within a program ***)
walther@59691
    94
walther@59691
    95
datatype input_tactic_result =
walther@59691
    96
    Safe_Step of Istate.T * Proof.context * Tactic.T
walther@59691
    97
  | Unsafe_Step of Istate.T * Proof.context * Tactic.T
walther@59691
    98
  | Not_Locatable of string
walther@59691
    99
walther@59700
   100
datatype assoc =    (* "ExprVal" in the sense of denotational semantics *)
walther@59700
   101
  Assoc of          (* the Prog_Tac is associated, strongly or weakly   *)
walther@59700
   102
    Istate.pstate * (* the current, returned for resuming execution     *)
walther@59700
   103
    Proof.context * (* updated according to execution of Prog_Tac       *)
walther@59700
   104
    Tactic.T        (* the Prog_Tac found as accociated to the input    *)
walther@59700
   105
| NasApp of         (* Prog_Tac is NOT associated, but applicable       *)
walther@59700
   106
    Istate.pstate * Proof.context * Tactic.T
walther@59705
   107
| Skip1 of         (* Prog_Tac not associated, not applicable          *)              
walther@59703
   108
	   term *         (* in case of a Prog_Expr the respective value      *)
walther@59703
   109
     Env.T;         (* *)
walther@59691
   110
fun assoc2str (Assoc _) = "Assoc"
walther@59705
   111
  | assoc2str (Skip1 _) = "Skip1"
walther@59691
   112
  | assoc2str (NasApp _) = "NasApp";
walther@59691
   113
walther@59701
   114
walther@59701
   115
(** scan to a Prog_Tac **)
walther@59701
   116
walther@59700
   117
fun scan_dn1 cct ist (Const ("Tactical.Try"(*1*), _) $ e $ a) =
walther@59701
   118
    (case scan_dn1 cct (ist |> path_down_form ([L, R], a)) e of goback => goback) 
walther@59700
   119
  | scan_dn1 cct ist (Const ("Tactical.Try"(*2*), _) $ e) =
walther@59701
   120
    (case scan_dn1 cct (ist |> path_down [R]) e of goback => goback)
walther@59691
   121
walther@59700
   122
  | scan_dn1 cct ist (Const ("Tactical.Repeat"(*1*), _) $ e $ a) =
walther@59701
   123
    scan_dn1 cct (ist |> path_down_form ([L, R], a)) e
walther@59700
   124
  | scan_dn1 cct ist (Const ("Tactical.Repeat"(*2*), _) $ e) =
walther@59700
   125
    scan_dn1 cct (ist |> path_down [R]) e
walther@59691
   126
walther@59700
   127
  | scan_dn1 (cct as (cstate, _, _)) ist (Const ("Tactical.Chain"(*1*), _) $ e1 $ e2 $ a) =
walther@59700
   128
    (case scan_dn1 cct (ist |> path_down_form ([L, L, R], a)) e1 of
walther@59705
   129
	     Skip1 (v, E)
walther@59700
   130
        => scan_dn1 cct (ist |> path_down [L, R] |> set_subst E (a, v)) e2
walther@59691
   131
     | NasApp (ist', ctxt', tac')
walther@59700
   132
       => scan_dn1 (cstate, ctxt', tac')  (ist
walther@59701
   133
         |> path_down_form ([L, R], a) |> trans_env_act ist') e2
walther@59701
   134
     | goback => goback)
walther@59700
   135
  | scan_dn1 (cct as (cstate, _, _)) ist (Const ("Tactical.Chain"(*2*), _) $e1 $ e2) =
walther@59700
   136
    (case scan_dn1 cct (ist |> path_down [L, R]) e1 of
walther@59705
   137
	     Skip1 (v, E) => scan_dn1 cct (ist |> path_down [R] |> set_act_env (v, E)) e2
walther@59691
   138
     | NasApp (ist', ctxt', tac')
walther@59691
   139
        => scan_dn1 (cstate, ctxt', tac') (ist |> path_down [R] |> trans_env_act ist') e2
walther@59701
   140
     | goback => goback)
walther@59691
   141
walther@59700
   142
  | scan_dn1 cct ist (Const ("HOL.Let", _) $ e $ (Abs (id, T, b))) =
walther@59700
   143
    (case scan_dn1 cct (ist |> path_down [L, R]) e of
walther@59691
   144
       NasApp (ist', _, _) =>
walther@59700
   145
         scan_dn1 cct (ist' |> path_down [R, D] |> upd_env (TermC.mk_Free (id, T)) |> trans_ass ist) b
walther@59705
   146
     | Skip1 (v, E) =>
walther@59700
   147
         scan_dn1 cct (ist |> path_down [R, D] |> set_env (Env.update E (TermC.mk_Free (id, T), v))) b
walther@59703
   148
     | goback => goback)
walther@59691
   149
walther@59700
   150
  | scan_dn1 cct  (ist as {eval, ...}) (Const ("Tactical.While"(*1*), _) $ c $ e $ a) =
walther@59691
   151
    if Rewrite.eval_true_ "Isac_Knowledge" eval
walther@59697
   152
      (subst_atomic (ist |> get_act_env |> Env.update' a) c)
walther@59703
   153
    then scan_dn1 cct (ist |> path_down_form ([L, R], a)) e
walther@59705
   154
    else Skip1 (ist |> get_act_env)
walther@59700
   155
  | scan_dn1 cct (ist as {eval, ...}) (Const ("Tactical.While"(*2*),_) $ c $ e) =
walther@59697
   156
    if Rewrite.eval_true_ "Isac_Knowledge" eval (subst_atomic (Env.update_opt' (get_subst ist)) c) 
walther@59700
   157
    then scan_dn1 cct (ist |> path_down [R]) e
walther@59705
   158
    else Skip1 (ist |> get_act_env)
walther@59691
   159
walther@59700
   160
  | scan_dn1 cct (ist as {or = Aundef, ...}) (Const ("Tactical.Or"(*1*), _) $e1 $ e2 $ a) =
walther@59700
   161
    (case scan_dn1 cct (ist |> path_down_form ([L, L, R], a) |> set_or AssOnly) e1 of
walther@59705
   162
	     Skip1 (v, E) =>
walther@59700
   163
	       (case scan_dn1 cct (ist |> path_down [L, R] |> set_subst E (a, v) |> set_or AssOnly) e2 of
walther@59705
   164
	          Skip1 (v, E) => 
walther@59700
   165
	            (case scan_dn1 cct (ist |> path_down [L, L, R] |> set_subst E (a, v) |> set_or AssGen) e1 of
walther@59705
   166
	               Skip1 (v, E) => 
walther@59700
   167
	                 scan_dn1 cct (ist |> path_down [L, R] |> set_subst E (a, v) |> set_or AssGen) e2
walther@59703
   168
	             | goback => goback)
walther@59703
   169
	        | goback => goback)
walther@59691
   170
     | NasApp _ => raise ERROR ("scan_dn1 Tactical.Or(*1*): must not return NasApp")
walther@59703
   171
     | goback => goback)
walther@59700
   172
  | scan_dn1 cct ist (Const ("Tactical.Or"(*2*), _) $e1 $ e2) =
walther@59700
   173
    (case scan_dn1 cct (ist |> path_down [L, R]) e1 of
walther@59705
   174
	     Skip1 (v, E) => scan_dn1 cct (ist |> path_down [R] |> set_act_env (v, E)) e2
walther@59703
   175
     | goback => goback)
walther@59691
   176
walther@59700
   177
  | scan_dn1 cct  (ist as {eval, ...}) (Const ("Tactical.If", _) $ c $ e1 $ e2) =
walther@59697
   178
    if Rewrite.eval_true_ "Isac_Knowledge" eval (subst_atomic (Env.update_opt' (get_subst ist)) c)
walther@59700
   179
    then scan_dn1 cct (ist |> path_down [L, R]) e1
walther@59700
   180
    else scan_dn1 cct (ist |> path_down [R]) e2
walther@59691
   181
walther@59691
   182
    (*======= end of scanning tacticals, a Prog_Tac as leaf =======*)
walther@59691
   183
  | scan_dn1 ((pt, p), ctxt, tac) (ist as {env, eval, or, ...}) t =
walther@59691
   184
    (case Lucin.handle_leaf "locate" (Context.theory_name (Rule.Isac"")) eval (get_subst ist) t of
walther@59691
   185
	     (a', Celem.Expr _) => 
walther@59705
   186
	       (Skip1 (Rewrite.eval_listexpr_ (Celem.assoc_thy "Isac_Knowledge") eval
walther@59697
   187
		       (subst_atomic (Env.update_opt'' (get_act_env ist, a')) t), env))
walther@59691
   188
     | (a', Celem.STac stac) =>
walther@59691
   189
         (case Lucin.associate pt ctxt (tac, stac) of
walther@59703
   190
            (* HERE ----------------vvvv ContextC.insert_assumptions asm ctxt DONE -- 2nd in gen//*)
walther@59698
   191
	           Lucin.Ass      (m, v', ctxt) => Assoc (ist |> set_subst_true  (a', v'), ctxt, m)
walther@59698
   192
           | Lucin.Ass_Weak (m, v', ctxt) => Assoc (ist |> set_subst_false (a', v'), ctxt, m)
walther@59691
   193
walther@59691
   194
           | Lucin.NotAss =>
walther@59691
   195
             (case or of (* switch for Tactical.Or: 1st AssOnly, 2nd AssGen *)
walther@59705
   196
                AssOnly => (Skip1 (ist |> get_act_env))
walther@59691
   197
              | _(*Aundef*) =>
walther@59691
   198
                 case Applicable.applicable_in p pt (Lucin.stac2tac pt (Celem.assoc_thy "Isac_Knowledge") stac) of
walther@59698
   199
                    Chead.Appl m' => NasApp (ist |> set_subst_false (a', Lucin.tac_2res m'), ctxt, tac)
walther@59705
   200
                  | Chead.Notappl _ => (Skip1 (ist |> get_act_env)))
walther@59691
   201
    ));
walther@59691
   202
walther@59700
   203
fun scan_up1 pcct ist (Const ("Tactical.Try"(*1*), _) $ _ $ _) = go_scan_up1 pcct ist
walther@59700
   204
  | scan_up1 pcct ist (Const ("Tactical.Try"(*2*), _) $ _) = go_scan_up1 pcct ist
walther@59691
   205
walther@59700
   206
  | scan_up1 (pcct as (prog, cct as (cstate, _, _))) ist (t as Const ("Tactical.Repeat"(*1*), _) $ e $ a) =
walther@59700
   207
    (case scan_dn1 cct (ist |> path_down_form ([L, R], a) |> set_or Aundef) e of 
walther@59705
   208
      Skip1 (v, E') => go_scan_up1 pcct (ist |> set_form_env (v, E') |> set_form a)
walther@59691
   209
    | NasApp (ist', ctxt', tac') => scan_up1 (prog, (cstate, ctxt', tac')) ist' t
walther@59703
   210
    | goback => goback)
walther@59700
   211
  | scan_up1 (pcct as (prog, cct as (cstate, _, _))) ist (t as Const ("Tactical.Repeat"(*2*), _) $ e) =
walther@59700
   212
    (case scan_dn1 cct (ist |> path_down [R] |> set_or Aundef) e of
walther@59705
   213
       Skip1 (v', E') => go_scan_up1 pcct (ist |> set_act_env (v', E'))
walther@59691
   214
     | NasApp (ist', ctxt', tac') => scan_up1 (prog, (cstate, ctxt', tac')) ist' t
walther@59703
   215
     | goback => goback)
walther@59691
   216
walther@59703
   217
    (*all has been done in (*2*) below*)
walther@59700
   218
  | scan_up1 pcct ist (Const ("Tactical.Chain"(*1*), _) $ _ $ _ $ _) = go_scan_up1 pcct ist
walther@59703
   219
  | scan_up1 pcct ist (Const ("Tactical.Chain"(*2*), _) $ _ $ _) = go_scan_up1 pcct ist (*comes from e2*)
walther@59703
   220
    (*comes from e1, goes to e2*)
walther@59700
   221
  | scan_up1 (pcct as (prog, cct as (cstate, _, _))) ist (Const ("Tactical.Chain"(*3*), _) $ _ ) =
walther@59691
   222
     let 
walther@59692
   223
       val e2 = check_Seq_up ist prog
walther@59691
   224
     in
walther@59700
   225
       case scan_dn1 cct (ist |> path_up_down [R] |> set_or Aundef) e2 of
walther@59705
   226
         Skip1 (v, E) => go_scan_up1 pcct (ist |> path_up |> set_act_env (v, E))
walther@59691
   227
       | NasApp (ist', ctxt', tac') => go_scan_up1 (prog, (cstate, ctxt', tac')) ist'
walther@59703
   228
       | goback => goback 
walther@59691
   229
     end
walther@59691
   230
walther@59700
   231
  | scan_up1 (pcct as (prog, cct as (cstate, _, _))) ist (Const ("HOL.Let"(*1*), _) $ _) =
walther@59691
   232
    let 
walther@59692
   233
      val (i, body) = check_Let_up ist prog
walther@59700
   234
    in case scan_dn1 cct (ist |> path_up_down [R, D] |> upd_env i |> set_or Aundef) body of
walther@59691
   235
	       Assoc iss => Assoc iss
walther@59691
   236
	     | NasApp (ist', ctxt', tac') => go_scan_up1 (prog, (cstate, ctxt', tac')) ist'
walther@59705
   237
	     | Skip1 (v, E) => go_scan_up1 pcct (ist |> path_up |> set_act_env (v, E)) 
walther@59691
   238
	  end
walther@59700
   239
  | scan_up1 pcct ist (Abs(*2*) _) = go_scan_up1 pcct ist
walther@59700
   240
  | scan_up1 pcct ist (Const ("HOL.Let"(*3*), _) $ _ $ (Abs _)) = go_scan_up1 pcct ist
walther@59691
   241
walther@59703
   242
  | scan_up1 (pcct as (prog, cct as (cstate, _, _))) (ist as {eval, ...})
walther@59703
   243
      (t as Const ("Tactical.While"(*1*),_) $ c $ e $ a) =
walther@59697
   244
    if Rewrite.eval_true_ "Isac_Knowledge" eval (subst_atomic (Env.update' a (get_act_env ist)) c)
walther@59691
   245
    then
walther@59700
   246
      case scan_dn1 cct (ist |> path_down_form ([L, R], a) |> set_or Aundef) e of 
walther@59705
   247
        Skip1 (v, E') => go_scan_up1 pcct (ist |> set_act_env (v, E') |> set_form a)
walther@59691
   248
walther@59691
   249
      | NasApp (ist', ctxt', tac') => scan_up1 (prog, (cstate, ctxt', tac')) ist' t
walther@59703
   250
      | goback => goback
walther@59700
   251
    else go_scan_up1 pcct (ist |> set_form a)
walther@59703
   252
  | scan_up1 (pcct as (prog, cct as (cstate, _, _))) (ist as {eval, ...})
walther@59703
   253
      (t as Const ("Tactical.While"(*2*), _) $ c $ e) =
walther@59697
   254
    if Rewrite.eval_true_ "Isac_Knowledge" eval (subst_atomic (Env.update_opt' (get_subst ist)) c)
walther@59691
   255
    then
walther@59700
   256
      case scan_dn1 cct (ist |> path_down [R] |> set_or Aundef) e of 
walther@59705
   257
        Skip1 (v, E') => go_scan_up1 pcct (ist |> set_act_env (v, E'))
walther@59691
   258
       | NasApp (ist', ctxt', tac') => scan_up1 (prog, (cstate, ctxt', tac')) ist' t
walther@59703
   259
       | goback => goback
walther@59700
   260
    else go_scan_up1 pcct ist
walther@59691
   261
walther@59700
   262
  | scan_up1 pcct ist (Const ("If", _) $ _ $ _ $ _) = go_scan_up1 pcct ist
walther@59691
   263
walther@59700
   264
  | scan_up1 pcct ist (Const ("Tactical.Or"(*1*), _) $ _ $ _ $ _) = go_scan_up1 pcct ist
walther@59700
   265
  | scan_up1 pcct ist (Const ("Tactical.Or"(*2*), _) $ _ $ _) = go_scan_up1 pcct ist
walther@59700
   266
  | scan_up1 pcct ist (Const ("Tactical.Or"(*3*), _) $ _ ) = go_scan_up1 pcct (ist |> path_up)
walther@59691
   267
walther@59700
   268
  | scan_up1 pcct ist (Const ("Tactical.If",_) $ _ $ _ $ _) = go_scan_up1 pcct ist
walther@59691
   269
walther@59691
   270
  | scan_up1 _ _ t = error ("scan_up1 not impl for t= " ^ Rule.term2str t)
walther@59700
   271
and go_scan_up1 (pcct as (prog, _)) (ist as {path, ...}) =
walther@59691
   272
    if 1 < length path
walther@59700
   273
    then scan_up1 pcct (ist |> path_up) (go (path_up' path) prog)
walther@59705
   274
    else (Skip1 (get_act_env ist))
walther@59691
   275
walther@59692
   276
fun scan_to_tactic1 (prog, (cctt as ((_, p), _, _))) (Istate.Pstate (ist as {path, ...})) =
walther@59699
   277
    if path = [] orelse Pos.at_first_tactic(*RM prog path RM WITH determine_next_tactic IN solve*)p
walther@59698
   278
    then scan_dn1 cctt (ist |> set_path [R] |> set_or Aundef) (Program.body_of prog)
walther@59692
   279
    else go_scan_up1 (prog, cctt) ist
walther@59691
   280
  | scan_to_tactic1 _ _ = raise ERROR "scan_to_tactic1: uncovered pattern in fun.def"
walther@59691
   281
walther@59701
   282
(*locate an input tactic within a program*)
walther@59699
   283
fun locate_input_tactic (Rule.Prog prog) cstate istate ctxt tac =
walther@59699
   284
    (case scan_to_tactic1 (prog, (cstate, ctxt, tac)) istate of
walther@59691
   285
         Assoc ((ist as {assoc, ...}), ctxt, tac') =>
walther@59691
   286
          if assoc
walther@59691
   287
          then Safe_Step (Istate.Pstate ist, ctxt, tac')
walther@59691
   288
 	        else Unsafe_Step (Istate.Pstate ist, ctxt, tac')
walther@59691
   289
      | NasApp _ => Not_Locatable (Tactic.tac_2str tac ^ " not locatable")
walther@59691
   290
      | err => raise ERROR ("not-found-in-script: NotLocatable from " ^ @{make_string} err))
walther@59691
   291
  | locate_input_tactic _ _ _ _ _ = raise ERROR "locate_input_tactic: uncovered case in definition"
walther@59691
   292
wneuper@59562
   293
walther@59701
   294
(*** determine a next tactic by executing the program ***)
wneuper@59554
   295
wneuper@59572
   296
  datatype next_tactic_result = NStep of Ctree.state * Istate.T * Proof.context * tactic
wneuper@59567
   297
    | Helpless | End_Program
wneuper@59566
   298
walther@59701
   299
(*
walther@59701
   300
  scan_dn2, go_scan_up2, scan_up2 scan for determine_next_tactic.
walther@59691
   301
  (1) scan_dn2 is recursive descent;
walther@59700
   302
  (2) go_scan_up2 goes to the parent node and calls (3);
walther@59700
   303
  (3) scan_up2 resumes according to the interpreter-state.
walther@59701
   304
  Call of (2) means that there was an applicable Prog_Tac below
wneuper@59554
   305
*)
walther@59700
   306
datatype appy =    (* "ExprVal" in the sense of denotational semantics *)
walther@59700
   307
  Appy of        (* applicable Prog_Tac found, search stalled        *)
walther@59700
   308
    Tactic.T *     (* tactic_associated (fun associate) with Prog_Tac  *)
walther@59700
   309
    Istate.pstate  (* the current, returned for resuming execution     *)
walther@59700
   310
| Napp of        (* stac found was not applicable; 
walther@59705
   311
	                    this mode may become Skip2 in Repeat, Try and Or *)
walther@59691
   312
    Env.T      (* popped while scan_up2                             *)
walther@59705
   313
| Skip2 of        (* for restart after Appy, for leaving iterations,
wneuper@59554
   314
	                    for passing the value of scriptexpressions,
wneuper@59554
   315
		                  and for finishing the script successfully       *)
walther@59639
   316
    term *         (* ???*) 
walther@59659
   317
    Env.T      (* TODO: a stack of env for nested let?            *);
wneuper@59554
   318
walther@59701
   319
walther@59701
   320
(** scan to a Prog_Tac **)
walther@59701
   321
walther@59700
   322
fun scan_dn2 cc (ist as {act_arg, ...}) (Const ("Tactical.Try"(*1*), _) $ e $ a) =
walther@59700
   323
    (case scan_dn2 cc (ist|> path_down_form ([L, R], a)) e of
walther@59705
   324
      Napp E => (Skip2 (act_arg, E))
walther@59703
   325
    | goback => goback)
walther@59700
   326
  | scan_dn2 cc (ist as {act_arg, ...}) (Const ("Tactical.Try"(*2*), _) $ e) =
walther@59700
   327
    (case scan_dn2 cc (ist |> path_down [R]) e of
walther@59705
   328
      Napp E => (Skip2 (act_arg, E))
walther@59703
   329
    | goback => goback)
walther@59689
   330
walther@59700
   331
  | scan_dn2 cc ist (Const ("Tactical.Repeat"(*1*), _) $ e $ a) = 
walther@59700
   332
    scan_dn2 cc (ist |> path_down_form ([L, R], a)) e
walther@59700
   333
  | scan_dn2 cc ist (Const ("Tactical.Repeat"(*2*), _) $ e) =
walther@59700
   334
    scan_dn2 cc (ist |> path_down [R]) e
walther@59689
   335
walther@59700
   336
  | scan_dn2 cc ist (Const ("Tactical.Chain"(*1*), _) $ e1 $ e2 $ a) =
walther@59700
   337
    (case scan_dn2 cc (ist |> path_down_form ([L, L, R], a)) e1 of
walther@59705
   338
	    Skip2 (v, E) => scan_dn2 cc (ist |> path_down_form ([L, R], a) |> set_act_env (v, E)) e2 (*UPD*)
walther@59703
   339
    | goback => goback)
walther@59700
   340
  | scan_dn2 cc ist (Const ("Tactical.Chain"(*2*), _) $ e1 $ e2) =
walther@59700
   341
    (case scan_dn2 cc (ist |> path_down [L, R]) e1 of
walther@59705
   342
	    Skip2 (v, E) => scan_dn2 cc (ist |> path_down [R] |> set_act_env (v, E)) e2
walther@59703
   343
    | goback => goback)
walther@59689
   344
walther@59700
   345
  | scan_dn2 cc ist (Const ("HOL.Let"(*1*), _) $ e $ (Abs (i, T, b))) =
walther@59700
   346
     (case scan_dn2 cc (ist |> path_down [L, R]) e of
walther@59705
   347
       Skip2 (res, E) => scan_dn2 cc (ist |> path_down [R, D] |> upd_env' (E , (Free (i, T), res))) b
walther@59703
   348
     | goback => goback)
walther@59678
   349
walther@59700
   350
  | scan_dn2 (cc as (_, ctxt)) (ist as {eval, ...}) (Const ("Tactical.While"(*1*), _) $ c $ e $ a) =
walther@59697
   351
    if Rewrite.eval_true_ ctxt eval (subst_atomic (ist |> get_act_env |> Env.update' a) c)
walther@59700
   352
    then scan_dn2 cc (ist |> path_down_form ([L, R], a)) e
walther@59705
   353
    else Skip2 (get_act_env ist)
walther@59700
   354
  | scan_dn2 (cc as (_, ctxt)) (ist as {eval, ...}) (Const ("Tactical.While"(*2*), _) $ c $ e) =
walther@59697
   355
    if Rewrite.eval_true_ ctxt eval (subst_atomic (Env.update_opt' (get_subst ist)) c)
walther@59700
   356
    then scan_dn2 cc (ist |> path_down [R]) e
walther@59705
   357
    else Skip2 (get_act_env ist)
walther@59678
   358
walther@59700
   359
  |scan_dn2 cc ist (Const ("Tactical.Or"(*1*), _) $e1 $ e2 $ a) =
walther@59700
   360
    (case scan_dn2 cc (ist |> path_down_form ([L, L, R], a)) e1 of
walther@59689
   361
	    Appy lme => Appy lme
walther@59700
   362
    | _ => scan_dn2 cc (ist |> path_down_form ([L, R], a)) e2)
walther@59700
   363
  | scan_dn2 cc ist (Const ("Tactical.Or"(*2*), _) $e1 $ e2) =
walther@59700
   364
    (case scan_dn2 cc (ist |> path_down [L, R]) e1 of
walther@59689
   365
	    Appy lme => Appy lme
walther@59700
   366
    | _ => scan_dn2 cc (ist |> path_down [R]) e2)
walther@59689
   367
walther@59700
   368
  |  scan_dn2 (cc as (_, ctxt)) (ist as {eval, ...}) (Const ("Tactical.If"(*1*), _) $ c $ e1 $ e2) =
walther@59697
   369
    if Rewrite.eval_true_ ctxt eval (subst_atomic (Env.update_opt' (get_subst ist)) c)
walther@59700
   370
    then scan_dn2 cc (ist |> path_down [L, R]) e1
walther@59700
   371
    else scan_dn2 cc (ist |> path_down [R]) e2
walther@59678
   372
walther@59690
   373
    (*======= end of scanning tacticals, a Prog_Tac as leaf =======*)
walther@59691
   374
  | scan_dn2 ((pt, p), ctxt) (ist as {env, eval, ...}) t =
walther@59685
   375
    case Lucin.handle_leaf "next  " ctxt eval (get_subst ist) t of
walther@59705
   376
	    (_, Celem.Expr s) => Skip2 (s, env)
walther@59671
   377
    | (a', Celem.STac stac) =>
walther@59671
   378
	    let
walther@59684
   379
	      val (m, m') = Lucin.stac2tac_ pt (Celem.assoc_thy ctxt) stac
walther@59671
   380
      in
walther@59671
   381
        case m of (*TODO?: make applicable_in..Subproblem analogous to other tacs ?refine problem?*)
walther@59698
   382
  	      Tactic.Subproblem _ => Appy (m', ist |> set_subst_reset (a', Lucin.tac_2res m'))
walther@59671
   383
  	    | _ =>
walther@59671
   384
          (case Applicable.applicable_in p pt m of
walther@59698
   385
  		      Chead.Appl m' => (Appy (m', ist |> set_subst_reset (a', Lucin.tac_2res m')))
walther@59685
   386
  		    | _ => Napp env)
wneuper@59554
   387
	    end
wneuper@59554
   388
walther@59705
   389
    (*makes Napp to Skip2*)
walther@59700
   390
fun scan_up2 pcc ist (Const ("Tactical.Try"(*1*), _) $ _ $ _) = go_scan_up2 pcc (ist |> set_appy Skip_)
walther@59700
   391
  | scan_up2 pcc ist (Const ("Tactical.Try"(*2*), _) $ _) = go_scan_up2 pcc (ist |> set_appy Skip_)
walther@59689
   392
walther@59700
   393
  | scan_up2 (pcc as (_, cc)) (ist) (Const ("Tactical.Repeat"(*1*), _) $ e $ _) =
walther@59700
   394
    (case scan_dn2 cc (ist |> path_down [L, R]) e of (* no appy_: there was already a stac below *)
walther@59689
   395
      Appy lr => Appy lr
walther@59700
   396
    | Napp E =>  go_scan_up2 pcc (ist |> set_env E |> set_appy Skip_)
walther@59705
   397
    | Skip2 (v, E) =>  go_scan_up2 pcc (ist |> set_act_env (v, E) |> set_appy Skip_))
walther@59701
   398
    (*no appy_: there was already a stac below*)
walther@59700
   399
  | scan_up2  (pcc as (_, cc)) ist (Const ("Tactical.Repeat"(*2*), _) $ e) =
walther@59700
   400
    (case scan_dn2 cc (ist |> path_down [R]) e of
walther@59689
   401
      Appy lr => Appy lr
walther@59700
   402
    | Napp E => go_scan_up2 pcc (ist |> set_env E |> set_appy Skip_)
walther@59705
   403
    | Skip2 (v, E) => go_scan_up2 pcc (ist |> set_act_env (v, E) |> set_appy Skip_))
walther@59689
   404
walther@59700
   405
  | scan_up2 pcc ist (Const ("Tactical.Chain"(*1*), _) $ _ $ _ $ _) = go_scan_up2 pcc ist
walther@59700
   406
  | scan_up2 pcc ist (Const ("Tactical.Chain"(*2*), _) $ _ $ _) = go_scan_up2 pcc ist
walther@59700
   407
  | scan_up2 (pcc as (sc, cc)) (ist as {finish, ...}) (Const ("Tactical.Chain"(*3*), _) $ _) =
walther@59689
   408
    if finish = Napp_
walther@59700
   409
    then go_scan_up2 pcc (ist |> path_up)
walther@59689
   410
    else (*Skip_*)
walther@59689
   411
      let 
walther@59689
   412
        val e2 = check_Seq_up ist sc
walther@59689
   413
      in
walther@59700
   414
        case scan_dn2 cc (ist |> path_up_down [R]) e2 of
walther@59689
   415
          Appy lr => Appy lr
walther@59700
   416
        | Napp E => go_scan_up2 pcc (ist |> path_up |> set_env E)
walther@59705
   417
        | Skip2 (v, E) => go_scan_up2 pcc (ist |> path_up |> set_act_env (v, E) |> set_appy Skip_)
walther@59689
   418
      end
walther@59690
   419
walther@59700
   420
  | scan_up2 (pcc as (sc, cc)) (ist as {finish, ...}) (Const ("HOL.Let"(*1*), _) $ _) =
walther@59685
   421
    if finish = Napp_
walther@59700
   422
    then go_scan_up2 pcc (ist |> path_up)
walther@59671
   423
    else (*Skip_*)
walther@59671
   424
	    let
walther@59685
   425
        val (i, body) = check_Let_up ist sc
walther@59671
   426
      in
walther@59700
   427
        case scan_dn2 cc (ist |> path_up_down [R, D] |> upd_env i) body of
walther@59671
   428
	        Appy lre => Appy lre
walther@59700
   429
	      | Napp E => go_scan_up2 pcc (ist |> path_up |> set_env E)
walther@59705
   430
	      | Skip2 (v, E) => go_scan_up2 pcc (ist |> path_up |> set_act_env (v, E) |> set_appy Skip_)
walther@59671
   431
	    end
walther@59700
   432
  | scan_up2 pcc ist (Abs _(*2*)) =  go_scan_up2 pcc ist
walther@59700
   433
  | scan_up2 pcc ist (Const ("HOL.Let"(*3*),_) $ _ $ (Abs _)) =
walther@59700
   434
    go_scan_up2 pcc ist
walther@59678
   435
walther@59701
   436
    (*no appy_: never causes Napp -> Helpless*)
walther@59700
   437
  | scan_up2 (pcc as (_, cc as (_, ctxt)))  (ist as {eval, ...}) (Const ("Tactical.While"(*1*), _) $ c $ e $ _) = 
walther@59697
   438
    if Rewrite.eval_true_ ctxt eval (subst_atomic (Env.update_opt' (get_subst ist)) c) 
walther@59671
   439
    then
walther@59700
   440
      case scan_dn2 cc (ist |> path_down [L, R]) e of
walther@59671
   441
  	     Appy lr => Appy lr
walther@59700
   442
  	  | Napp E => go_scan_up2 pcc (ist |> set_env E |> set_appy Skip_)
walther@59705
   443
  	  | Skip2 (v, E) => go_scan_up2 pcc (ist |> set_act_env (v, E) |> set_appy Skip_)
walther@59671
   444
    else
walther@59700
   445
      go_scan_up2 pcc (ist |> set_appy Skip_)
walther@59701
   446
    (*no appy_: never causes Napp - Helpless*)
walther@59700
   447
  | scan_up2  (pcc as (_, cc as (_, ctxt))) (ist as {eval, ...}) (Const ("Tactical.While"(*2*), _) $ c $ e) = 
walther@59697
   448
    if Rewrite.eval_true_ ctxt eval (subst_atomic (Env.update_opt' (get_subst ist)) c) 
walther@59671
   449
    then
walther@59700
   450
      case scan_dn2 cc (ist |> path_down [R]) e of
walther@59671
   451
  	    Appy lr => Appy lr
walther@59700
   452
  	  | Napp E => go_scan_up2 pcc (ist |> set_env E |> set_appy Skip_)
walther@59705
   453
  	  | Skip2 (v, E) => go_scan_up2 pcc (ist |> set_act_env (v, E) |> set_appy Skip_)
walther@59671
   454
    else
walther@59700
   455
      go_scan_up2 pcc (ist |> set_appy Skip_)
walther@59678
   456
walther@59700
   457
  | scan_up2 pcc ist (Const ("Tactical.Or"(*1*), _) $ _ $ _ $ _) = go_scan_up2 pcc ist
walther@59700
   458
  | scan_up2 pcc ist (Const ("Tactical.Or"(*2*), _) $ _ $ _) = go_scan_up2 pcc ist
walther@59700
   459
  | scan_up2 pcc ist (Const ("Tactical.Or"(*3*), _) $ _ ) = go_scan_up2 pcc ist
walther@59678
   460
walther@59700
   461
  | scan_up2 pcc ist (Const ("Tactical.If"(*1*), _) $ _ $ _ $ _) = go_scan_up2 pcc ist
walther@59678
   462
walther@59691
   463
  | scan_up2 _ _ t = error ("scan_up2 not impl for " ^ Rule.term2str t)
walther@59700
   464
and go_scan_up2 (pcc as (sc, _)) (ist as {env, path, finish, ...}) = 
walther@59685
   465
    if 1 < length path
walther@59671
   466
    then 
walther@59700
   467
      scan_up2 pcc (ist |> path_up) (go_up path sc)
walther@59671
   468
    else (*interpreted to end*)
walther@59705
   469
      if finish = Skip_ then Skip2 (get_act_env ist) else Napp env
wneuper@59554
   470
walther@59700
   471
fun scan_to_tactic2 (prog, cc) (Istate.Pstate (ist as {path, ...})) =
walther@59699
   472
  if path = []
walther@59700
   473
  then scan_dn2 cc (trans_scan_down2 ist) (Program.body_of prog)
walther@59700
   474
  else go_scan_up2 (prog, cc) (trans_scan_up2 ist |> set_appy Skip_)
walther@59690
   475
| scan_to_tactic2 _ _ = raise ERROR "scan_to_tactic2: uncovered pattern";
wneuper@59565
   476
walther@59701
   477
(*decide for the next applicable Prog_Tac*)
walther@59699
   478
fun determine_next_tactic (thy, _) (ptp as (pt, (p, _))) (Rule.Prog prog) (Istate.Pstate ist, _(*ctxt*)) =
walther@59699
   479
   (case scan_to_tactic2 (prog, (ptp, thy)) (Istate.Pstate ist) of
walther@59705
   480
      Skip2 (v, _) =>                                                        (* program finished *)
walther@59675
   481
        (case par_pbl_det pt p of
walther@59675
   482
	        (true, p', _) => 
walther@59675
   483
	          let
walther@59680
   484
	            val (_, pblID, _) = get_obj g_spec pt p';
walther@59675
   485
	          in
walther@59675
   486
              (Tactic.Check_Postcond' (pblID, (v, [(*assigned in next step*)])), 
walther@59686
   487
	              (Istate.e_istate, ContextC.e_ctxt), (v, Telem.Safe)) 
walther@59675
   488
            end
walther@59675
   489
	      | _ =>
walther@59686
   490
          (Tactic.End_Detail' (Rule.e_term,[])(*8.6.03*), (Istate.e_istate, ContextC.e_ctxt), (v, Telem.Safe)))
walther@59675
   491
    | Napp _ =>
walther@59686
   492
        (Tactic.Empty_Tac_, (Istate.e_istate, ContextC.e_ctxt), (Rule.e_term, Telem.Helpless_))(* helpless *)
walther@59685
   493
    | Appy (m', (ist as {act_arg, ...})) =>
walther@59686
   494
        (m', (Pstate ist, ContextC.e_ctxt), (act_arg, Telem.Safe)))                  (* found next tac *)
wneuper@59560
   495
  | determine_next_tactic _ _ _ (is, _) =
wneuper@59572
   496
    error ("determine_next_tactic: not impl for " ^ (Istate.istate2str is));
wneuper@59554
   497
walther@59701
   498
walther@59701
   499
(*** locate an input formula in the program ***)
wneuper@59569
   500
wneuper@59562
   501
datatype input_formula_result =
wneuper@59572
   502
    Step of Ctree.state * Istate.T * Proof.context
wneuper@59562
   503
  | Not_Derivable of Chead.calcstate'
wneuper@59562
   504
wneuper@59562
   505
(* FIXME.WN050821 compare fun solve ... fun begin_end_prog
wneuper@59562
   506
   begin_end_prog (Apply_Method'     vvv FIXME: get args in applicable_in *)
wneuper@59582
   507
fun begin_end_prog (Tactic.Apply_Method' (mI, _, _, ctxt)) _ (pt, pos as (p, _)) =
wneuper@59562
   508
    let
wneuper@59562
   509
      val {ppc, ...} = Specify.get_met mI;
wneuper@59562
   510
      val (itms, oris, probl) = case get_obj I pt p of
wneuper@59562
   511
        PblObj {meth = itms, origin = (oris, _, _), probl, ...} => (itms, oris, probl)
wneuper@59562
   512
      | _ => error "begin_end_prog Apply_Method': uncovered case get_obj"
wneuper@59562
   513
      val itms = if itms <> [] then itms else Chead.complete_metitms oris probl [] ppc
wneuper@59562
   514
      val thy' = get_obj g_domID pt p;
wneuper@59562
   515
      val thy = Celem.assoc_thy thy';
walther@59677
   516
      val srls = Lucin.get_simplifier (pt, pos)
wneuper@59555
   517
(*//----------- hack for funpack: generalise handling of meths which extend problem items -----\\*)
wneuper@59555
   518
val itms =
wneuper@59555
   519
  if mI = ["Biegelinien", "ausBelastung"]
wneuper@59555
   520
  then itms @
wneuper@59597
   521
    [(4, [1], true, "#Given", Model.Cor ((Const ("Biegelinie.Biegelinie", Type ("fun", [Type ("fun", [Type ("Real.real", []), Type ("Real.real", [])]), Type ("Input_Descript.una", [])])),
wneuper@59555
   522
        [Free ("y", Type ("fun", [Type ("Real.real", []), Type ("Real.real", [])]))]),
wneuper@59555
   523
      (Free ("id_fun", Type ("fun", [Type ("Real.real", []), Type ("Real.real", [])])),
wneuper@59555
   524
        [Free ("y", Type ("fun", [Type ("Real.real", []), Type ("Real.real", [])]))] ))),
wneuper@59597
   525
    (5, [1], true, "#Given", Model.Cor ((Const ("Biegelinie.AbleitungBiegelinie", Type ("fun", [Type ("fun", [Type ("Real.real", []), Type ("Real.real", [])]), Type ("Input_Descript.una", [])])),
wneuper@59555
   526
        [Free ("dy", Type ("fun", [Type ("Real.real", []), Type ("Real.real", [])]))]),
wneuper@59555
   527
      (Free ("id_abl", Type ("fun", [Type ("Real.real", []), Type ("Real.real", [])])),
wneuper@59555
   528
        [Free ("dy", Type ("fun", [Type ("Real.real", []), Type ("Real.real", [])]))] )))]
wneuper@59555
   529
  else itms
wneuper@59555
   530
(*\\----------- hack for funpack: generalise handling of meths which extend problem items -----//*)
walther@59677
   531
     val (is, env, ctxt, scr) = case Lucin.init_pstate srls ctxt itms mI of
walther@59685
   532
         (is as Istate.Pstate {env, ...}, ctxt, scr) => (is, env, ctxt, scr)
wneuper@59583
   533
       | _ => error "begin_end_prog Apply_Method': uncovered case init_pstate"
wneuper@59555
   534
     val ini = Lucin.init_form thy scr env;
wneuper@59555
   535
   in 
wneuper@59555
   536
     case ini of
wneuper@59555
   537
       SOME t =>
wneuper@59555
   538
       let
wneuper@59555
   539
         val pos = ((lev_on o lev_dn) p, Frm)
wneuper@59571
   540
	       val tac_ = Tactic.Apply_Method' (mI, SOME t, is, ctxt);
wneuper@59555
   541
	       val (pos, c, _, pt) = Generate.generate1 thy tac_ (is, ctxt) pos pt (* implicit Take *)
wneuper@59555
   542
       in
wneuper@59571
   543
        ([(Tactic.Apply_Method mI, tac_, (pos, (is, ctxt)))], c, (pt, pos))
wneuper@59555
   544
       end
wneuper@59555
   545
     | NONE =>
wneuper@59555
   546
       let
wneuper@59555
   547
         val pt = update_env pt (fst pos) (SOME (is, ctxt))
wneuper@59562
   548
	       val (tacis, c, ptp) = do_solve_step (pt, pos)
walther@59658
   549
       in (tacis @ [(Tactic.Apply_Method mI,
walther@59658
   550
            Tactic.Apply_Method' (mI, NONE, Istate.e_istate, ctxt), (pos, (is, ctxt)))], c, ptp)
wneuper@59555
   551
       end
wneuper@59555
   552
    end
wneuper@59571
   553
  | begin_end_prog (Tactic.Check_Postcond' (pI, _)) _ (pt, (p, p_))  =
wneuper@59555
   554
    let
wneuper@59555
   555
      val pp = par_pblobj pt p
wneuper@59555
   556
      val asm = (case get_obj g_tac pt p of
wneuper@59571
   557
		    Tactic.Check_elementwise _ => (snd o (get_obj g_result pt)) p (*collects and instantiates asms*)
wneuper@59555
   558
		  | _ => get_assumptions_ pt (p, p_))
wneuper@59555
   559
      val metID = get_obj g_metID pt pp;
wneuper@59555
   560
      val {srls = srls, scr = sc, ...} = Specify.get_met metID;
walther@59676
   561
      val (loc, pst, ctxt) = case get_loc pt (p, p_) of
walther@59676
   562
        loc as (Istate.Pstate pst, ctxt) => (loc, pst, ctxt)
wneuper@59562
   563
      | _ => error "begin_end_prog Check_Postcond': uncovered case get_loc"
wneuper@59555
   564
      val thy' = get_obj g_domID pt pp;
wneuper@59555
   565
      val thy = Celem.assoc_thy thy';
walther@59686
   566
      val (_, _, (scval, _)) = determine_next_tactic (thy', srls) (pt, (p, p_)) sc loc;
wneuper@59555
   567
    in
wneuper@59555
   568
      if pp = []
wneuper@59555
   569
      then 
wneuper@59555
   570
	      let
walther@59698
   571
          val is = Istate.Pstate (pst |> Istate.set_act scval |> Istate.set_appy Istate.Skip_)
walther@59675
   572
	        val tac_ = Tactic.Check_Postcond' (pI, (scval, asm))
walther@59675
   573
	          (*extend Tactic.Check_Postcond' (pI, (scval, asm), scsaf) ^^^^^ *)
wneuper@59555
   574
	        val ((p, p_), ps, _, pt) = Generate.generate1 thy tac_ (is, ctxt) (pp, Res) pt;
wneuper@59571
   575
	      in ([(Tactic.Check_Postcond pI, tac_, ((pp, Res), (is, ctxt)))], ps, (pt, (p, p_))) end
wneuper@59555
   576
      else
walther@59675
   577
        let (*resume script of parent PblObj, transfer value of subpbl-script*)
wneuper@59555
   578
          val ppp = par_pblobj pt (lev_up p);
wneuper@59555
   579
	        val thy' = get_obj g_domID pt ppp;
wneuper@59555
   580
          val thy = Celem.assoc_thy thy';
walther@59676
   581
walther@59676
   582
          val (pst', ctxt') = case get_loc pt (pp, Frm) of
walther@59676
   583
            (Istate.Pstate pst', ctxt') => (pst', ctxt')
wneuper@59562
   584
          | _ => error "begin_end_prog Check_Postcond' script of parpbl: uncovered case get_loc"
wneuper@59577
   585
	        val ctxt'' = ContextC.from_subpbl_to_caller ctxt scval ctxt'
wneuper@59571
   586
          val tac_ = Tactic.Check_Postcond' (pI, (scval, asm))
walther@59698
   587
	        val is = Istate.Pstate (pst' |> Istate.set_act scval |> Istate.set_appy Istate.Skip_)
wneuper@59555
   588
          val ((p, p_), ps, _, pt) = Generate.generate1 thy tac_ (is, ctxt'') (pp, Res) pt;
wneuper@59571
   589
        in ([(Tactic.Check_Postcond pI, tac_, ((pp, Res), (is, ctxt'')))], ps, (pt, (p, p_))) end
wneuper@59555
   590
    end
wneuper@59571
   591
  | begin_end_prog (Tactic.End_Proof'') _ ptp = ([], [], ptp)
wneuper@59562
   592
  | begin_end_prog tac_ is (pt, pos) =
wneuper@59555
   593
    let
wneuper@59555
   594
      val pos = case pos of
wneuper@59555
   595
 		   (p, Met) => ((lev_on o lev_dn) p, Frm) (* begin script        *)
wneuper@59555
   596
 		  | (p, Res) => (lev_on p, Res)            (* somewhere in script *)
wneuper@59555
   597
 		  | _ => pos
wneuper@59592
   598
 	    val (pos', c, _, pt) = Generate.generate1 (Celem.assoc_thy "Isac_Knowledge") tac_ is pos pt;
wneuper@59555
   599
    in
walther@59704
   600
      ([(Tactic.input_from_T tac_, tac_, (pos, is))], c, (pt, pos'))
wneuper@59555
   601
    end
walther@59701
   602
(*find the next tac from the script, begin_end_prog will update the ctree*)
wneuper@59569
   603
and do_solve_step (ptp as (pt, pos as (p, p_))) = (* WN1907: ?only for Begin_/End_Detail' DEL!!!*)
wneuper@59555
   604
    if Celem.e_metID = get_obj g_metID pt (par_pblobj pt p)
wneuper@59555
   605
    then ([], [], (pt, (p, p_)))
wneuper@59555
   606
    else 
wneuper@59555
   607
      let 
wneuper@59555
   608
        val thy' = get_obj g_domID pt (par_pblobj pt p);
wneuper@59555
   609
	      val (srls, is, sc) = Lucin.from_pblobj_or_detail' thy' (p,p_) pt;
wneuper@59559
   610
	      val (tac_, is, (t, _)) = determine_next_tactic (thy', srls) (pt, pos) sc is;
wneuper@59569
   611
	      (* TODO here  ^^^  return finished/helpless/ok ?*)
wneuper@59555
   612
	    in case tac_ of
walther@59686
   613
		    Tactic.End_Detail' _ =>
walther@59686
   614
          ([(Tactic.End_Detail, Tactic.End_Detail' (t, [(*FIXME.04*)]), (pos, is))], [], (pt, pos))
walther@59686
   615
	    | _ => begin_end_prog tac_        is                          ptp
wneuper@59555
   616
      end;
walther@59686
   617
walther@59701
   618
(*
walther@59701
   619
  compare inform with ctree.form at current pos by nrls;
walther@59701
   620
  if found, embed the derivation generated during comparison
walther@59701
   621
  if not, let the mat-engine compute the next ctree.form.
walther@59701
   622
walther@59701
   623
  Code's structure is copied from complete_solve
walther@59701
   624
  CAUTION: tacis in returned calcstate' do NOT construct resulting ptp --
walther@59701
   625
           all_modspec etc. has to be inserted at Subproblem'
walther@59701
   626
*)
wneuper@59555
   627
fun compare_step ((tacis, c, ptp as (pt, pos as (p,p_))): Chead.calcstate') ifo =
wneuper@59555
   628
  let
wneuper@59555
   629
    val fo =
wneuper@59555
   630
      case p_ of
walther@59694
   631
        Pos.Frm => Ctree.get_obj Ctree.g_form pt p
walther@59694
   632
			| Pos.Res => (fst o (Ctree.get_obj Ctree.g_result pt)) p
wneuper@59555
   633
			| _ => Rule.e_term (*on PblObj is fo <> ifo*);
wneuper@59555
   634
	  val {nrls, ...} = Specify.get_met (Ctree.get_obj Ctree.g_metID pt (Ctree.par_pblobj pt p))
wneuper@59555
   635
	  val {rew_ord, erls, rules, ...} = Rule.rep_rls nrls
wneuper@59555
   636
	  val (found, der) = Inform.concat_deriv rew_ord erls rules fo ifo; (*<---------------*)
wneuper@59555
   637
  in
wneuper@59555
   638
    if found
wneuper@59555
   639
    then
wneuper@59555
   640
       let
wneuper@59555
   641
         val tacis' = map (Inform.mk_tacis rew_ord erls) der;
wneuper@59555
   642
		     val (c', ptp) = Generate.embed_deriv tacis' ptp;
wneuper@59555
   643
	     in ("ok", (tacis (*@ tacis'?WN050408*), c @ c', ptp)) end
wneuper@59555
   644
     else 
walther@59694
   645
	     if pos = ([], Pos.Res) (*TODO: we should stop earlier with trying subproblems *)
wneuper@59555
   646
	     then ("no derivation found", (tacis, c, ptp): Chead.calcstate') 
wneuper@59555
   647
	     else
wneuper@59555
   648
         let
wneuper@59562
   649
           val cs' as (tacis, c', ptp) = do_solve_step ptp; (*<---------------------*)
wneuper@59555
   650
		       val (tacis, c'', ptp) = case tacis of
wneuper@59571
   651
			       ((Tactic.Subproblem _, _, _)::_) => 
wneuper@59555
   652
			         let
wneuper@59555
   653
                 val ptp as (pt, (p,_)) = Chead.all_modspec ptp (*<--------------------*)
wneuper@59555
   654
				         val mI = Ctree.get_obj Ctree.g_metID pt p
wneuper@59555
   655
			         in
wneuper@59581
   656
			           begin_end_prog (Tactic.Apply_Method' (mI, NONE, Istate.e_istate, ContextC.e_ctxt)) (Istate.e_istate, ContextC.e_ctxt) ptp
wneuper@59555
   657
               end
wneuper@59555
   658
			     | _ => cs';
wneuper@59555
   659
		     in compare_step (tacis, c @ c' @ c'', ptp) ifo end
wneuper@59555
   660
  end
wneuper@59555
   661
walther@59701
   662
(*
walther@59701
   663
  handle a user-input formula, which may be a CAS-command, too.
walther@59701
   664
 * CAS-command: create a calchead, and do 1 step
walther@59701
   665
 * formula, which is no CAS-command:
wneuper@59556
   666
   compare iform with calc-tree.form at pos by equ_nrls and all subsequent pos;
walther@59701
   667
   collect all the Prog_Tac applied by the way; ...???TODO?
walther@59701
   668
   If "no derivation found" then check_error_patterns.
walther@59701
   669
   ALTERNATIVE: check_error_patterns _within_ compare_step: seems too expensive
walther@59701
   670
*)
wneuper@59562
   671
(*                       vvvv           vvvvvv vvvv    NEW args for compare_step *)
walther@59658
   672
fun locate_input_formula (Rule.Prog _)  ((pt, pos) : Ctree.state) (_ : Istate.T) (_ : Proof.context) tm =
walther@59658
   673
     let                                                          
walther@59658
   674
   		val pos_pred = Ctree.lev_back' pos (*f_pred ---"step pos cs"---> f_succ in appendFormula*)
walther@59658
   675
   		val _(*f_pred*) = Ctree.get_curr_formula (pt, pos_pred)
walther@59658
   676
     in
walther@59658
   677
       (*TODO: use prog, istate, ctxt in compare_step ...*)
walther@59658
   678
   		case compare_step ([], [], (pt, pos_pred)) tm of
walther@59658
   679
   		  ("no derivation found", calcstate') => Not_Derivable calcstate'
walther@59658
   680
       | ("ok", (_, _, cstate as (pt', pos'))) => 
walther@59658
   681
           Step (cstate, get_istate pt' pos', get_ctxt pt' pos')
walther@59658
   682
       | _ => raise ERROR "compare_step: uncovered case"
walther@59658
   683
     end
walther@59658
   684
  | locate_input_formula _ _ _ _ _ = error ""
wneuper@59556
   685
walther@59684
   686
(**)end(**)