src/Tools/isac/Interpret/solve-step.sml
author wneuper <Walther.Neuper@jku.at>
Wed, 19 Oct 2022 10:43:04 +0200
changeset 60567 bb3140a02f3d
parent 60559 aba19e46dd84
child 60568 dd387c9fa89a
permissions -rw-r--r--
eliminate term2str in src, Prog_Tac.*_adapt_to_type
walther@59920
     1
(* Title:  Specify/solve-step.sml
walther@59920
     2
   Author: Walther Neuper
walther@59920
     3
   (c) due to copyright terms
walther@59920
     4
walther@59920
     5
Code for the solve-phase in analogy to structure Specify_Step for the specify-phase.
walther@59920
     6
*)
walther@59920
     7
walther@59920
     8
signature SOLVE_STEP =
walther@59920
     9
sig
walther@59921
    10
  val check: Tactic.input -> Calc.T -> Applicable.T
walther@59959
    11
  val add: Tactic.T -> Istate_Def.T * Proof.context -> Calc.T -> Test_Out.T
walther@59935
    12
walther@59959
    13
  val add_general: Tactic.T -> Istate_Def.T * Proof.context -> Calc.T -> Test_Out.T
walther@59932
    14
  val s_add_general: State_Steps.T ->
walther@59932
    15
    Ctree.ctree * Pos.pos' list * Pos.pos' -> Ctree.ctree * Pos.pos' list * Pos.pos'
walther@59933
    16
  val add_hard:
walther@59959
    17
    theory -> Tactic.T -> Pos.pos' -> Ctree.ctree -> Test_Out.T
walther@59932
    18
Walther@60544
    19
  val get_ruleset: 'a -> Pos.pos' -> Ctree.ctree ->
Walther@60509
    20
    string * ThyC.id * Rewrite_Ord.id * Rule_Def.rule_set * bool
Walther@60544
    21
  val get_eval: string -> Pos.pos' -> Ctree.ctree -> string * ThyC.id * Eval.ml
Walther@60538
    22
wenzelm@60223
    23
\<^isac_test>\<open>
Walther@60567
    24
  val get_ctxt: Ctree.ctree -> Pos.pos' -> Proof.context
Walther@60539
    25
  val rew_info: Rule_Def.rule_set -> string * Rule_Def.rule_set * Eval.ml_from_prog list
wenzelm@60223
    26
\<close>
walther@59920
    27
end
walther@59920
    28
walther@59920
    29
(**)
walther@59996
    30
structure Solve_Step(**): SOLVE_STEP(**) =
walther@59920
    31
struct
walther@59920
    32
(**)
walther@59920
    33
walther@59935
    34
(** get data from Calc.T **)
walther@59935
    35
walther@59935
    36
(* the source is the parent node, either a problem or a Rule_Set (with inter_steps) *)
walther@59935
    37
fun rew_info (Rule_Def.Repeat {erls, rew_ord = (rew_ord', _), calc = ca, ...}) =
walther@59935
    38
    (rew_ord', erls, ca)
walther@59935
    39
  | rew_info (Rule_Set.Sequence {erls, rew_ord = (rew_ord', _), calc = ca, ...}) =
walther@59935
    40
    (rew_ord', erls, ca)
walther@59935
    41
  | rew_info (Rule_Set.Rrls {erls, rew_ord = (rew_ord', _), calc = ca, ...}) =
walther@59935
    42
    (rew_ord', erls, ca)
walther@59962
    43
  | rew_info rls = raise ERROR ("rew_info called with '" ^ Rule_Set.id rls ^ "'");
walther@59935
    44
Walther@60544
    45
fun get_ruleset _ (pos as (p, _)) pt = 
walther@59935
    46
  let 
Walther@60559
    47
    val (pbl, p', rls', ctxt) = LItool.parent_node pt pos
walther@59935
    48
  in                                                      
walther@59935
    49
    if pbl
walther@59935
    50
    then 
walther@59935
    51
      let 
walther@59935
    52
        val thy' = Ctree.get_obj Ctree.g_domID pt p'
Walther@60559
    53
        val {rew_ord', erls, ...} = MethodC.from_store ctxt (Ctree.get_obj Ctree.g_metID pt p')              
walther@59935
    54
	    in ("OK", thy', rew_ord', erls, false) end
walther@59935
    55
     else 
walther@59935
    56
      let
walther@59935
    57
        val thy' = Ctree.get_obj Ctree.g_domID pt (Ctree.par_pblobj pt p)
walther@59935
    58
		    val (rew_ord', erls, _) = rew_info rls'
walther@59935
    59
		  in ("OK", thy', rew_ord', erls, false) end
walther@59935
    60
  end;
walther@59935
    61
Walther@60544
    62
fun get_eval scrop (pos as (p, _)) pt = 
walther@59935
    63
  let
Walther@60559
    64
    val (pbl, p', rls', ctxt) =  LItool.parent_node pt pos
walther@59935
    65
  in
walther@59935
    66
    if pbl
walther@59935
    67
    then
walther@59935
    68
      let
walther@59935
    69
        val thy' = Ctree.get_obj Ctree.g_domID pt p'
Walther@60559
    70
        val {calc = scr_isa_fns, ...} = MethodC.from_store ctxt (Ctree.get_obj Ctree.g_metID pt p')
walther@59935
    71
        val opt = assoc (scr_isa_fns, scrop)
walther@59935
    72
	    in
walther@59935
    73
	      case opt of
walther@59935
    74
	        SOME isa_fn => ("OK", thy', isa_fn)
Walther@60539
    75
	      | NONE => ("applicable_in Calculate: unknown '" ^ scrop ^ "'", "", ("", Eval.ml_fun_empty))
walther@59935
    76
	    end
walther@59935
    77
    else 
walther@59935
    78
		  let
walther@59935
    79
		    val thy' = Ctree.get_obj Ctree.g_domID pt (Ctree.par_pblobj pt p);
walther@59935
    80
		    val (_, _,(*_,*)scr_isa_fns) = rew_info rls'(*rls*)
walther@59935
    81
		  in
walther@59935
    82
		    case assoc (scr_isa_fns, scrop) of
walther@59935
    83
		      SOME isa_fn => ("OK",thy',isa_fn)
Walther@60539
    84
		    | NONE => ("applicable_in Calculate: unknown '" ^ scrop ^ "'", "", ("", Eval.ml_fun_empty))
walther@59935
    85
		  end
walther@59935
    86
  end;
walther@59935
    87
Walther@60567
    88
(** get context reliably at switch_specify_solve **)
Walther@60567
    89
Walther@60567
    90
fun at_begin_program (is, Pos.Res) = last_elem is = 0
Walther@60567
    91
  | at_begin_program _ = false;
Walther@60567
    92
Walther@60567
    93
(* strange special case at Apply_Method *)
Walther@60567
    94
fun get_ctxt_from_PblObj pt (p_, Pos.Res) =
Walther@60567
    95
    let
Walther@60567
    96
      val pp = Ctree.par_pblobj pt p_ (*drops the "0"*)
Walther@60567
    97
      val {ctxt, ...} = Ctree.get_obj I pt pp |> Ctree.rep_specify_data
Walther@60567
    98
    in ctxt end
Walther@60567
    99
  | get_ctxt_from_PblObj _ _ = raise ERROR "get_ctxt_from_PblObj called by PrfObj or EmptyPtree";
Walther@60567
   100
Walther@60567
   101
fun get_ctxt pt (p_, Pos.Pbl) =
Walther@60567
   102
    let
Walther@60567
   103
      val pp = Ctree.par_pblobj pt p_ (*drops the "0"*)
Walther@60567
   104
      val {ctxt, ...} = Ctree.get_obj I pt pp |> Ctree.rep_specify_data
Walther@60567
   105
    in ctxt end
Walther@60567
   106
  | get_ctxt pt pos =
Walther@60567
   107
    if at_begin_program pos
Walther@60567
   108
    then get_ctxt_from_PblObj pt pos
Walther@60567
   109
    else Ctree.get_ctxt pt pos
Walther@60567
   110
Walther@60567
   111
Walther@60567
   112
walther@59935
   113
(** Solve_Step.check **)
walther@59935
   114
walther@59922
   115
(*
walther@59922
   116
  check tactics (input by the user, mostly) for applicability
walther@59922
   117
  and determine as much of the result of the tactic as possible initially.
walther@59922
   118
*)
walther@59932
   119
fun check (Tactic.Apply_Method mI) (pt, (p, _)) =
walther@59932
   120
      let
walther@59932
   121
        val (dI, pI, probl, ctxt) = case Ctree.get_obj I pt p of
walther@59932
   122
          Ctree.PblObj {origin = (_, (dI, pI, _), _), probl, ctxt, ...} => (dI, pI, probl, ctxt)
Walther@60567
   123
        | _ => raise ERROR "Solve_Step.check Apply_Method: uncovered case Ctree.get_obj"
Walther@60559
   124
        val {where_, ...} = Problem.from_store ctxt pI
Walther@60477
   125
        val pres = map (I_Model.environment probl |> subst_atomic) where_
walther@59932
   126
        val ctxt = if ContextC.is_empty ctxt (*vvvvvvvvvvvvvv DO THAT EARLIER?!?*)
walther@59932
   127
          then ThyC.get_theory dI |> Proof_Context.init_global |> ContextC.insert_assumptions pres
walther@59932
   128
          else ctxt
walther@59932
   129
      in
walther@59932
   130
        Applicable.Yes (Tactic.Apply_Method' (mI, NONE, Istate_Def.empty (*filled later*), ctxt))
walther@59932
   131
      end
Walther@60544
   132
  | check (Tactic.Calculate op_) (cs as (pt, pos as (p, _))) =
walther@59923
   133
      let 
Walther@60544
   134
        val (msg, thy', isa_fn) = get_eval op_ pos pt;
walther@59928
   135
        val f = Calc.current_formula cs;
walther@59923
   136
      in
walther@59923
   137
        if msg = "OK"
walther@59923
   138
        then
Walther@60500
   139
    	    case Rewrite.calculate_ (ThyC.id_to_ctxt thy') isa_fn f of
walther@59923
   140
    	      SOME (f', (id, thm))
walther@59923
   141
    	        => Applicable.Yes (Tactic.Calculate' (thy', op_, f, (f', (id, thm))))
walther@59929
   142
    	    | NONE => Applicable.No ("'calculate " ^ op_ ^ "' not applicable") 
walther@59923
   143
        else Applicable.No msg                                              
walther@59923
   144
      end
walther@59928
   145
  | check (Tactic.Check_Postcond pI) (_, _) = (*TODO: only applicable, if evaluating to True*)
walther@59928
   146
      Applicable.Yes (Tactic.Check_Postcond' (pI, TermC.empty))
walther@59928
   147
  | check (Tactic.Check_elementwise pred) cs =
walther@59923
   148
      let 
walther@59928
   149
        val f = Calc.current_formula cs;
walther@59923
   150
      in
walther@59928
   151
        Applicable.Yes (Tactic.Check_elementwise' (f, pred, (f, [])))
walther@59923
   152
      end
walther@59923
   153
  | check Tactic.Empty_Tac _ = Applicable.No "Empty_Tac is not applicable"
walther@59929
   154
  | check (Tactic.Free_Solve) _ = Applicable.Yes (Tactic.Free_Solve')
walther@59929
   155
  | check Tactic.Or_to_List cs =
walther@59928
   156
       let 
walther@59929
   157
        val f = Calc.current_formula cs;
walther@59929
   158
        val ls = Prog_Expr.or2list f;
walther@59929
   159
      in
walther@59929
   160
        Applicable.Yes (Tactic.Or_to_List' (f, ls))
walther@59923
   161
      end
Walther@60544
   162
  | check (Tactic.Rewrite thm) (cs as (pt, pos as (p, _))) = 
walther@59923
   163
      let
Walther@60544
   164
        val (msg, thy', ro, rls', _) = get_ruleset thm pos pt;
Walther@60506
   165
        val thy = ThyC.get_theory thy';
Walther@60506
   166
        val ctxt = Proof_Context.init_global thy;
walther@59928
   167
        val f = Calc.current_formula cs;
walther@59923
   168
      in
walther@59923
   169
        if msg = "OK" 
walther@59923
   170
        then
Walther@60548
   171
          case Rewrite.rewrite_ ctxt (get_rew_ord ctxt ro) rls' false (snd thm) f of
walther@59929
   172
            SOME (f',asm) => Applicable.Yes (Tactic.Rewrite' (thy', ro, rls', false, thm, f, (f', asm)))
walther@59929
   173
          | NONE => Applicable.No ((thm |> fst |> quote) ^ " not applicable") 
walther@59923
   174
        else Applicable.No msg
walther@59923
   175
      end
Walther@60527
   176
  | check (Tactic.Rewrite_Inst (subs, thm)) (cs as (pt, pos as (p, _))) = 
walther@59921
   177
      let 
walther@59921
   178
        val pp = Ctree.par_pblobj pt p;
Walther@60534
   179
        val ctxt = Ctree.get_loc pt pos |> snd
Walther@60534
   180
        val thy = Proof_Context.theory_of ctxt
Walther@60559
   181
        val {rew_ord' = ro', erls = erls, ...} = MethodC.from_store ctxt (Ctree.get_obj Ctree.g_metID pt pp);
walther@59928
   182
        val f = Calc.current_formula cs;
Walther@60500
   183
        val subst = Subst.T_from_input ctxt subs; (*TODO: input requires parse _: _ -> _ option*)
walther@59921
   184
      in 
Walther@60548
   185
        case Rewrite.rewrite_inst_ ctxt (get_rew_ord ctxt ro') erls false subst (snd thm) f of
walther@59929
   186
          SOME (f', asm) =>
Walther@60534
   187
            Applicable.Yes (Tactic.Rewrite_Inst' 
Walther@60534
   188
              (Context.theory_name thy, ro', erls, false, subst, thm, f, (f', asm)))
walther@59929
   189
        | NONE => Applicable.No (fst thm ^ " not applicable")
walther@59921
   190
      end
Walther@60534
   191
  | check (Tactic.Rewrite_Set rls) (cs as (pt, pos)) =
walther@59921
   192
      let 
Walther@60534
   193
        val ctxt = Ctree.get_loc pt pos |> snd
Walther@60534
   194
        val thy' = ctxt |> Proof_Context.theory_of |> Context.theory_name
walther@59928
   195
        val f = Calc.current_formula cs;
walther@59923
   196
      in
Walther@60543
   197
        case Rewrite.rewrite_set_ ctxt false (get_rls ctxt rls) f of
walther@59921
   198
          SOME (f', asm)
Walther@60543
   199
            => Applicable.Yes (Tactic.Rewrite_Set' (thy', false, get_rls ctxt rls, f, (f', asm)))
walther@59923
   200
          | NONE => Applicable.No (rls ^ " not applicable")
walther@59921
   201
      end
Walther@60534
   202
  | check (Tactic.Rewrite_Set_Inst (subs, rls)) (cs as (pt, pos)) =
walther@59921
   203
      let 
Walther@60534
   204
        val ctxt = Ctree.get_loc pt pos |> snd
Walther@60534
   205
        val thy' = ctxt |> Proof_Context.theory_of |> Context.theory_name
walther@59928
   206
        val f = Calc.current_formula cs;
Walther@60500
   207
    	  val subst = Subst.T_from_input ctxt subs; (*TODO: input requires parse _: _ -> _ option*)
walther@59921
   208
      in 
Walther@60543
   209
        case Rewrite.rewrite_set_inst_ ctxt false subst (get_rls ctxt rls) f of
walther@59928
   210
          SOME (f', asm)
Walther@60543
   211
            => Applicable.Yes
Walther@60543
   212
                 (Tactic.Rewrite_Set_Inst' (thy', false, subst, get_rls ctxt rls, f, (f', asm)))
walther@59921
   213
        | NONE => Applicable.No (rls ^ " not applicable")
walther@59921
   214
      end
walther@59928
   215
  | check (Tactic.Subproblem (domID, pblID)) (_, _) = 
walther@60154
   216
      Applicable.Yes (Tactic.Subproblem' ((domID, pblID, MethodC.id_empty), [], 
walther@59928
   217
			  TermC.empty, [], ContextC.empty, Auto_Prog.subpbl domID pblID))
Walther@60527
   218
  | check (Tactic.Substitute sube) (cs as (pt, pos as (p, _))) =
walther@59928
   219
      let
walther@59928
   220
        val pp = Ctree.par_pblobj pt p
Walther@60534
   221
        val ctxt = Ctree.get_loc pt pos |> snd
Walther@60534
   222
        val thy = Proof_Context.theory_of ctxt
walther@59928
   223
        val f = Calc.current_formula cs;
Walther@60559
   224
		    val {rew_ord', erls, ...} = MethodC.from_store ctxt (Ctree.get_obj Ctree.g_metID pt pp)
Walther@60567
   225
		    val subte = Subst.input_to_terms ctxt sube (*?TODO: input requires parse _: _ -> _ option?*)
Walther@60548
   226
		    val ro = get_rew_ord ctxt rew_ord'
walther@59928
   227
		  in
walther@59928
   228
		    if foldl and_ (true, map TermC.contains_Var subte)
walther@59928
   229
		    then (*1*)
Walther@60534
   230
		      let val f' = subst_atomic (Subst.T_from_string_eqs thy sube) f
walther@59928
   231
		      in if f = f'
walther@59928
   232
		        then Applicable.No (Subst.string_eqs_to_string sube ^ " not applicable")
walther@59928
   233
		        else Applicable.Yes (Tactic.Substitute' (ro, erls, subte, f, f'))
walther@59928
   234
		      end
walther@59928
   235
		    else (*2*)
Walther@60500
   236
		      case Rewrite.rewrite_terms_ ctxt ro erls subte f of
walther@59928
   237
		        SOME (f', _) =>  Applicable.Yes (Tactic.Substitute' (ro, erls, subte, f, f'))
walther@59928
   238
		      | NONE => Applicable.No (Subst.string_eqs_to_string sube ^ " not applicable")
walther@59928
   239
		  end
Walther@60534
   240
  | check (Tactic.Tac id) (cs as (pt, pos)) =
walther@59929
   241
      let 
Walther@60534
   242
        val thy = (Ctree.get_loc pt pos |> snd) |> Proof_Context.theory_of
walther@59929
   243
        val f = Calc.current_formula cs;
walther@59936
   244
      in
walther@59936
   245
        Applicable.Yes (Tactic.Tac_ (thy, UnparseC.term f, id, UnparseC.term f))
walther@59921
   246
      end
Walther@60567
   247
(*/----------------- updated----------------------------------* )
Walther@60567
   248
  | check (Tactic.Take str) (pt, pos) =
Walther@60567
   249
      Applicable.Yes (Tactic.Take'
Walther@60567
   250
        (TermC.parse ((*Solve_Step.*)get_ctxt pt pos) str)) (*always applicable*)
Walther@60567
   251
( *------------------- updated----------------------------------*)
Walther@60567
   252
  | check (Tactic.Take str) (pt, pos) =
Walther@60567
   253
    let
Walther@60567
   254
      val ctxt = (*Solve_Step.*)get_ctxt pt pos
Walther@60567
   255
      val t = Prog_Tac.Take_adapt_to_type ctxt str
Walther@60567
   256
    in Applicable.Yes (Tactic.Take' t) end
Walther@60567
   257
(*\----------------- updated----------------------------------*)
walther@59929
   258
  | check (Tactic.Begin_Trans) cs =
walther@59929
   259
      Applicable.Yes (Tactic.Begin_Trans' (Calc.current_formula cs))
walther@59923
   260
  | check (Tactic.End_Trans) (pt, (p, p_)) = (*TODO: check parent branches*)
walther@59923
   261
    if p_ = Pos.Res 
walther@59923
   262
	  then Applicable.Yes (Tactic.End_Trans' (Ctree.get_obj Ctree.g_result pt p))
walther@59923
   263
    else Applicable.No "'End_Trans' is not applicable at the beginning of a transitive sequence"
walther@59921
   264
  | check Tactic.End_Proof' _ = Applicable.Yes Tactic.End_Proof''
walther@59921
   265
  | check m _ = raise ERROR ("Solve_Step.check called for " ^ Tactic.input_to_string m);
walther@59920
   266
walther@59935
   267
walther@59935
   268
(** Solve_Step.add **)
walther@59935
   269
walther@59932
   270
fun add (Tactic.Apply_Method' (_, topt, is, _)) (_, ctxt) (pt, pos as (p, _)) = 
walther@59932
   271
    (case topt of 
walther@59932
   272
      SOME t => 
walther@59932
   273
        let val (pt, c) = Ctree.cappend_form pt p (is, ctxt) t
walther@59959
   274
        in (pos, c, Test_Out.EmptyMout, pt) end
walther@59959
   275
    | NONE => (pos, [], Test_Out.EmptyMout, pt))
walther@59932
   276
  | add (Tactic.Take' t) l (pt, (p, _)) = (* val (Take' t) = m; *)
walther@59931
   277
    let
walther@59931
   278
      val p =
walther@59931
   279
        let val (ps, p') = split_last p (* no connex to prev.ppobj *)
walther@59931
   280
	      in if p' = 0 then ps @ [1] else p end
walther@59931
   281
      val (pt, c) = Ctree.cappend_form pt p l t
walther@59931
   282
    in
walther@59959
   283
      ((p, Pos.Frm), c, Test_Out.FormKF (UnparseC.term t), pt)
walther@59931
   284
    end
walther@59931
   285
  | add (Tactic.Begin_Trans' t) l (pt, (p, Pos.Frm)) =
walther@59931
   286
    let
walther@59931
   287
      val (pt, c) = Ctree.cappend_form pt p l t
walther@59931
   288
      val pt = Ctree.update_branch pt p Ctree.TransitiveB (*040312*)
walther@59931
   289
      (* replace the old PrfOjb ~~~~~ *)
walther@59931
   290
      val p = (Pos.lev_on o Pos.lev_dn (* starts with [...,0] *)) p
walther@59931
   291
      val (pt, c') = Ctree.cappend_form pt p l t (*FIXME.0402 same istate ???*)
walther@59931
   292
    in
walther@59959
   293
      ((p, Pos.Frm), c @ c', Test_Out.FormKF (UnparseC.term t), pt)
walther@59931
   294
    end
walther@59931
   295
  | add (Tactic.Begin_Trans' t) l (pt, (p, Pos.Res)) = 
walther@59931
   296
    (*append after existing PrfObj    vvvvvvvvvvvvv*)
walther@59931
   297
    add (Tactic.Begin_Trans' t) l (pt, (Pos.lev_on p, Pos.Frm))
walther@59931
   298
  | add (Tactic.End_Trans' tasm) l (pt, (p, _)) =
walther@59931
   299
    let
walther@59931
   300
      val p' = Pos.lev_up p
walther@59931
   301
      val (pt, c) = Ctree.append_result pt p' l tasm Ctree.Complete
walther@59931
   302
    in
walther@59959
   303
      ((p', Pos.Res), c, Test_Out.FormKF "DUMMY" (*term2str t ..ERROR (t) has not been declared*), pt)
walther@59931
   304
    end
walther@59931
   305
  | add (Tactic.Rewrite_Inst' (_, _, _, _, subs', thm', f, (f', asm))) (is, ctxt) (pt, (p, _)) =
walther@59931
   306
    let
walther@59931
   307
      val (pt, c) = Ctree.cappend_atomic pt p (is, ctxt) f
walther@59931
   308
        (Tactic.Rewrite_Inst (Subst.T_to_input subs', thm')) (f',asm) Ctree.Complete;
walther@59931
   309
      val pt = Ctree.update_branch pt p Ctree.TransitiveB
walther@59931
   310
    in
walther@59959
   311
      ((p, Pos.Res), c, Test_Out.FormKF (UnparseC.term f'), pt)
walther@59931
   312
    end
walther@59931
   313
 | add (Tactic.Rewrite' (_, _, _, _, thm', f, (f', asm))) (is, ctxt) (pt, (p, _)) =
walther@59931
   314
   let
walther@59931
   315
     val (pt, c) = Ctree.cappend_atomic pt p (is, ctxt) f (Tactic.Rewrite thm') (f', asm) Ctree.Complete
walther@59931
   316
     val pt = Ctree.update_branch pt p Ctree.TransitiveB
walther@59931
   317
   in
walther@59959
   318
    ((p, Pos.Res), c, Test_Out.FormKF (UnparseC.term f'), pt)
walther@59931
   319
   end
walther@59931
   320
  | add (Tactic.Rewrite_Set_Inst' (_, _, subs', rls', f, (f', asm))) (is, ctxt) (pt, (p, _)) =
walther@59931
   321
    let
walther@59931
   322
      val (pt, c) = Ctree.cappend_atomic pt p (is, ctxt) f 
walther@59931
   323
        (Tactic.Rewrite_Set_Inst (Subst.T_to_input subs', Rule_Set.id rls')) (f', asm) Ctree.Complete
walther@59931
   324
      val pt = Ctree.update_branch pt p Ctree.TransitiveB
walther@59931
   325
    in
walther@59959
   326
      ((p, Pos.Res), c, Test_Out.FormKF (UnparseC.term f'), pt)
walther@59931
   327
    end
walther@59931
   328
  | add (Tactic.Rewrite_Set' (_, _, rls', f, (f', asm))) (is, ctxt) (pt, (p, _)) =
walther@59931
   329
    let
walther@59931
   330
      val (pt, c) = Ctree.cappend_atomic pt p (is, ctxt) f 
walther@59931
   331
        (Tactic.Rewrite_Set (Rule_Set.id rls')) (f', asm) Ctree.Complete
walther@59931
   332
      val pt = Ctree.update_branch pt p Ctree.TransitiveB
walther@59931
   333
    in
walther@59959
   334
      ((p, Pos.Res), c, Test_Out.FormKF (UnparseC.term f'), pt)
walther@59931
   335
    end
walther@59931
   336
  | add (Tactic.Check_Postcond' (_, scval)) l (pt, (p, _)) =
walther@59931
   337
      let
walther@59931
   338
        val (pt, c) = Ctree.append_result pt p l (scval, []) Ctree.Complete
walther@59931
   339
      in
walther@59959
   340
        ((p, Pos.Res), c, Test_Out.FormKF (UnparseC.term scval), pt)
walther@59931
   341
      end
walther@59931
   342
  | add (Tactic.Calculate' (_, op_, f, (f', _))) l (pt, (p, _)) =
walther@59931
   343
      let
walther@59931
   344
        val (pt,c) = Ctree.cappend_atomic pt p l f (Tactic.Calculate op_) (f', []) Ctree.Complete
walther@59931
   345
      in
walther@59959
   346
        ((p, Pos.Res), c, Test_Out.FormKF (UnparseC.term f'), pt)
walther@59931
   347
      end
walther@59931
   348
  | add (Tactic.Check_elementwise' (consts, pred, (f', asm))) l (pt, (p, _)) =
walther@59931
   349
      let
walther@59931
   350
        val (pt,c) = Ctree.cappend_atomic pt p l consts (Tactic.Check_elementwise pred) (f', asm) Ctree.Complete
walther@59931
   351
      in
walther@59959
   352
        ((p, Pos.Res), c, Test_Out.FormKF (UnparseC.term f'), pt)
walther@59931
   353
      end
walther@59931
   354
  | add (Tactic.Or_to_List' (ors, list)) l (pt, (p, _)) =
walther@59931
   355
      let
walther@59931
   356
        val (pt,c) = Ctree.cappend_atomic pt p l ors Tactic.Or_to_List (list, []) Ctree.Complete
walther@59931
   357
      in
walther@59959
   358
        ((p, Pos.Res), c, Test_Out.FormKF (UnparseC.term list), pt)
walther@59931
   359
      end
walther@59931
   360
  | add (Tactic.Substitute' (_, _, subte, t, t')) l (pt, (p, _)) =
walther@59931
   361
      let
walther@59931
   362
        val (pt,c) =
walther@59931
   363
          Ctree.cappend_atomic pt p l t (Tactic.Substitute (Subst.eqs_to_input subte)) (t',[]) Ctree.Complete
walther@59959
   364
        in ((p, Pos.Res), c, Test_Out.FormKF (UnparseC.term t'), pt) 
walther@59931
   365
        end
Walther@60567
   366
  | add (Tactic.Tac_ (_, f, id, f')) l (pt, pos as (p, _)) =
walther@59931
   367
      let
Walther@60567
   368
        val ctxt = Ctree.get_ctxt pt pos
Walther@60567
   369
        val (pt, c) = Ctree.cappend_atomic pt p l
Walther@60567
   370
          (TermC.parse ctxt f) (Tactic.Tac id) (TermC.parse ctxt f', []) Ctree.Complete
walther@59931
   371
      in
walther@59959
   372
        ((p,Pos.Res), c, Test_Out.FormKF f', pt)
walther@59931
   373
      end
walther@59931
   374
  | add (Tactic.Subproblem' ((domID, pblID, metID), oris, hdl, fmz_, ctxt_specify, f))
walther@59931
   375
      (l as (_, ctxt)) (pt, (p, _)) =
walther@59932
   376
      let
walther@59932
   377
  	    val (pt, c) = Ctree.cappend_problem pt p l (fmz_, (domID, pblID, metID))
walther@59932
   378
  	      (oris, (domID, pblID, metID), hdl, ctxt_specify)
walther@60360
   379
  	    val f = Syntax.string_of_term ctxt f
walther@59932
   380
      in
walther@59959
   381
        ((p, Pos.Pbl), c, Test_Out.FormKF f, pt)
walther@59932
   382
      end
walther@59932
   383
  | add m' _ (_, pos) =
walther@59932
   384
      raise ERROR ("Solve_Step.add: not impl.for " ^ Tactic.string_of m' ^ " at " ^ Pos.pos'2str pos)
walther@59932
   385
walther@59932
   386
(* LI switches between solve-phase and specify-phase *)
walther@59932
   387
fun add_general tac ic cs =
walther@59932
   388
  if Tactic.for_specify' tac
walther@59933
   389
  then Specify_Step.add tac ic cs
walther@59932
   390
  else add tac ic cs
walther@59932
   391
walther@59933
   392
(* the order of State_Steps is reversed: insert last element first  *)
walther@59932
   393
fun s_add_general [] ptp = ptp
walther@59932
   394
  | s_add_general tacis (pt, c, _) = 
walther@59931
   395
    let
walther@59932
   396
      val (tacis', (_, tac_, (p, is))) = split_last tacis
walther@59933
   397
	    val (p', c', _, pt') = add_general tac_ is (pt, p)
walther@59931
   398
    in
walther@59932
   399
      s_add_general tacis' (pt', c@c', p')
walther@59931
   400
    end
walther@59932
   401
walther@59933
   402
(* a still undeveloped concept: do a calculation without LI *)
walther@59933
   403
fun add_hard _(*thy*) m' (p, p_) pt =
walther@59933
   404
  let  
walther@59933
   405
    val p = case p_ of
walther@59933
   406
      Pos.Frm => p | Pos.Res => Pos.lev_on p
walther@59962
   407
    | _ => raise ERROR ("generate_hard: call by " ^ Pos.pos'2str (p,p_))
walther@59933
   408
  in
walther@59933
   409
    add_general m' (Istate_Def.empty, ContextC.empty) (pt, (p, p_))
walther@59933
   410
  end
walther@59931
   411
walther@59920
   412
(**)end(**);