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