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