src/Tools/isac/Interpret/solve-step.sml
author Walther Neuper <walther.neuper@jku.at>
Mon, 04 May 2020 10:19:16 +0200
changeset 59933 92214be419b2
parent 59932 87336f3b021f
child 59935 16927a749dd7
permissions -rw-r--r--
spearate Specify_Step.add
     1 (* Title:  Specify/solve-step.sml
     2    Author: Walther Neuper
     3    (c) due to copyright terms
     4 
     5 Code for the solve-phase in analogy to structure Specify_Step for the specify-phase.
     6 *)
     7 
     8 signature SOLVE_STEP =
     9 sig
    10   val check: Tactic.input -> Calc.T -> Applicable.T
    11   val add: Tactic.T -> Istate_Def.T * Proof.context -> Calc.T -> Generate.test_out
    12   val add_general: Tactic.T -> Istate_Def.T * Proof.context -> Calc.T -> Generate.test_out
    13   val s_add_general: State_Steps.T ->
    14     Ctree.ctree * Pos.pos' list * Pos.pos' -> Ctree.ctree * Pos.pos' list * Pos.pos'
    15   val add_hard:
    16     theory -> Tactic.T -> Pos.pos' -> Ctree.ctree -> Generate.test_out
    17 
    18 (* ---- for tests only: shifted from below to remove the Warning "unused" at fun.def. --------- *)
    19   (*NONE*)                                                     
    20 (*/-------------------------------------------------------- ! aktivate for Test_Isac BEGIN ---\* )
    21   (*NONE*)                                                     
    22 ( *\--- ! aktivate for Test_Isac END ----------------------------------------------------------/*)
    23 end
    24 
    25 (**)
    26 structure Solve_Step(** ): SOLVE_STEP( **) =
    27 struct
    28 (**)
    29 
    30 (*
    31   check tactics (input by the user, mostly) for applicability
    32   and determine as much of the result of the tactic as possible initially.
    33 *)
    34 fun check (Tactic.Apply_Method mI) (pt, (p, _)) =
    35       let
    36         val (dI, pI, probl, ctxt) = case Ctree.get_obj I pt p of
    37           Ctree.PblObj {origin = (_, (dI, pI, _), _), probl, ctxt, ...} => (dI, pI, probl, ctxt)
    38         | _ => raise ERROR "Specify_Step.check Apply_Method: uncovered case Ctree.get_obj"
    39         val {where_, ...} = Specify.get_pbt pI
    40         val pres = map (Model.mk_env probl |> subst_atomic) where_
    41         val ctxt = if ContextC.is_empty ctxt (*vvvvvvvvvvvvvv DO THAT EARLIER?!?*)
    42           then ThyC.get_theory dI |> Proof_Context.init_global |> ContextC.insert_assumptions pres
    43           else ctxt
    44       in
    45         Applicable.Yes (Tactic.Apply_Method' (mI, NONE, Istate_Def.empty (*filled later*), ctxt))
    46       end
    47   | check (Tactic.Calculate op_) (cs as (pt, (p, _))) =
    48       let 
    49         val (msg, thy', isa_fn) = ApplicableOLD.from_pblobj_or_detail_calc op_ p pt;
    50         val f = Calc.current_formula cs;
    51       in
    52         if msg = "OK"
    53         then
    54     	    case Rewrite.calculate_ (ThyC.get_theory thy') isa_fn f of
    55     	      SOME (f', (id, thm))
    56     	        => Applicable.Yes (Tactic.Calculate' (thy', op_, f, (f', (id, thm))))
    57     	    | NONE => Applicable.No ("'calculate " ^ op_ ^ "' not applicable") 
    58         else Applicable.No msg                                              
    59       end
    60   | check (Tactic.Check_Postcond pI) (_, _) = (*TODO: only applicable, if evaluating to True*)
    61       Applicable.Yes (Tactic.Check_Postcond' (pI, TermC.empty))
    62   | check (Tactic.Check_elementwise pred) cs =
    63       let 
    64         val f = Calc.current_formula cs;
    65       in
    66         Applicable.Yes (Tactic.Check_elementwise' (f, pred, (f, [])))
    67       end
    68   | check Tactic.Empty_Tac _ = Applicable.No "Empty_Tac is not applicable"
    69   | check (Tactic.Free_Solve) _ = Applicable.Yes (Tactic.Free_Solve')
    70   | check Tactic.Or_to_List cs =
    71        let 
    72         val f = Calc.current_formula cs;
    73         val ls = Prog_Expr.or2list f;
    74       in
    75         Applicable.Yes (Tactic.Or_to_List' (f, ls))
    76       end
    77   | check (Tactic.Rewrite thm) (cs as (pt, (p, _))) = 
    78       let
    79         val (msg, thy', ro, rls', _) = ApplicableOLD.from_pblobj_or_detail_thm thm p pt;
    80         val thy = ThyC.get_theory thy';
    81         val f = Calc.current_formula cs;
    82       in
    83         if msg = "OK" 
    84         then
    85           case Rewrite.rewrite_ thy (Rewrite_Ord.assoc_rew_ord ro) rls' false (snd thm) f of
    86             SOME (f',asm) => Applicable.Yes (Tactic.Rewrite' (thy', ro, rls', false, thm, f, (f', asm)))
    87           | NONE => Applicable.No ((thm |> fst |> quote) ^ " not applicable") 
    88         else Applicable.No msg
    89       end
    90   | check (Tactic.Rewrite_Inst (subs, thm)) (cs as (pt, (p, _))) = 
    91       let 
    92         val pp = Ctree.par_pblobj pt p;
    93         val thy' = Ctree.get_obj Ctree.g_domID pt pp;
    94         val thy = ThyC.get_theory thy';
    95         val {rew_ord' = ro', erls = erls, ...} = Specify.get_met (Ctree.get_obj Ctree.g_metID pt pp);
    96         val f = Calc.current_formula cs;
    97         val subst = Subst.T_from_input thy subs; (*TODO: input requires parse _: _ -> _ option*)
    98       in 
    99         case Rewrite.rewrite_inst_ thy (Rewrite_Ord.assoc_rew_ord ro') erls false subst (snd thm) f of
   100           SOME (f', asm) =>
   101             Applicable.Yes (Tactic.Rewrite_Inst' (thy', ro', erls, false, subst, thm, f, (f', asm)))
   102         | NONE => Applicable.No (fst thm ^ " not applicable")
   103       end
   104   | check (Tactic.Rewrite_Set rls) (cs as (pt, (p, _))) =
   105       let 
   106         val pp = Ctree.par_pblobj pt p; 
   107         val thy' = Ctree.get_obj Ctree.g_domID pt pp;
   108         val f = Calc.current_formula cs;
   109       in
   110         case Rewrite.rewrite_set_ (ThyC.get_theory thy') false (assoc_rls rls) f of
   111           SOME (f', asm)
   112             => Applicable.Yes (Tactic.Rewrite_Set' (thy', false, assoc_rls rls, f, (f', asm)))
   113           | NONE => Applicable.No (rls ^ " not applicable")
   114       end
   115   | check (Tactic.Rewrite_Set_Inst (subs, rls)) (cs as (pt, (p, _))) =
   116       let 
   117         val pp = Ctree.par_pblobj pt p;
   118         val thy' = Ctree.get_obj Ctree.g_domID pt pp;
   119         val thy = ThyC.get_theory thy';
   120         val f = Calc.current_formula cs;
   121     	  val subst = Subst.T_from_input thy subs; (*TODO: input requires parse _: _ -> _ option*)
   122       in 
   123         case Rewrite.rewrite_set_inst_ thy false subst (assoc_rls rls) f of
   124           SOME (f', asm)
   125             => Applicable.Yes (Tactic.Rewrite_Set_Inst' (thy', false, subst, assoc_rls rls, f, (f', asm)))
   126         | NONE => Applicable.No (rls ^ " not applicable")
   127       end
   128   | check (Tactic.Subproblem (domID, pblID)) (_, _) = 
   129       Applicable.Yes (Tactic.Subproblem' ((domID, pblID, Method.id_empty), [], 
   130 			  TermC.empty, [], ContextC.empty, Auto_Prog.subpbl domID pblID))
   131    | check (Tactic.Substitute sube) (cs as (pt, (p, _))) =
   132       let
   133         val pp = Ctree.par_pblobj pt p
   134         val thy = ThyC.get_theory (Ctree.get_obj Ctree.g_domID pt pp)
   135         val f = Calc.current_formula cs;
   136 		    val {rew_ord', erls, ...} = Specify.get_met (Ctree.get_obj Ctree.g_metID pt pp)
   137 		    val subte = Subst.input_to_terms sube (*TODO: input requires parse _: _ -> _ option*)
   138 		    val subst = Subst.T_from_string_eqs thy sube
   139 		    val ro = Rewrite_Ord.assoc_rew_ord rew_ord'
   140 		  in
   141 		    if foldl and_ (true, map TermC.contains_Var subte)
   142 		    then (*1*)
   143 		      let val f' = subst_atomic subst f
   144 		      in if f = f'
   145 		        then Applicable.No (Subst.string_eqs_to_string sube ^ " not applicable")
   146 		        else Applicable.Yes (Tactic.Substitute' (ro, erls, subte, f, f'))
   147 		      end
   148 		    else (*2*)
   149 		      case Rewrite.rewrite_terms_ thy ro erls subte f of
   150 		        SOME (f', _) =>  Applicable.Yes (Tactic.Substitute' (ro, erls, subte, f, f'))
   151 		      | NONE => Applicable.No (Subst.string_eqs_to_string sube ^ " not applicable")
   152 		  end
   153   | check (Tactic.Tac id) (cs as (pt, (p, _))) =
   154       let 
   155         val pp = Ctree.par_pblobj pt p; 
   156         val thy' = Ctree.get_obj Ctree.g_domID pt pp;
   157         val thy = ThyC.get_theory thy';
   158         val f = Calc.current_formula cs;
   159       in case id of
   160         "subproblem_equation_dummy" =>
   161     	  if TermC.is_expliceq f
   162     	  then Applicable.Yes (Tactic.Tac_ (thy, UnparseC.term f, id, "subproblem_equation_dummy (" ^ UnparseC.term f ^ ")"))
   163     	  else Applicable.No "applicable only to equations made explicit"
   164       | "solve_equation_dummy" =>
   165     	  let val (id', f') = ApplicableOLD.split_dummy (UnparseC.term f);
   166     	  in
   167     	    if id' <> "subproblem_equation_dummy"
   168     	    then Applicable.No "no subproblem"
   169     	    else if (ThyC.to_ctxt thy, f') |-> TermC.parseNEW |> the |> TermC.is_expliceq
   170     		    then Applicable.Yes (Tactic.Tac_ (thy, UnparseC.term f, id, "[" ^ f' ^ "]"))
   171     		    else error ("Solve_Step.check: f= " ^ f')
   172         end
   173       | _ => Applicable.Yes (Tactic.Tac_ (thy, UnparseC.term f, id, UnparseC.term f))
   174       end
   175   | check (Tactic.Take str) _ = Applicable.Yes (Tactic.Take' (TermC.str2term str)) (* always applicable ?*)
   176   | check (Tactic.Begin_Trans) cs =
   177       Applicable.Yes (Tactic.Begin_Trans' (Calc.current_formula cs))
   178   | check (Tactic.End_Trans) (pt, (p, p_)) = (*TODO: check parent branches*)
   179     if p_ = Pos.Res 
   180 	  then Applicable.Yes (Tactic.End_Trans' (Ctree.get_obj Ctree.g_result pt p))
   181     else Applicable.No "'End_Trans' is not applicable at the beginning of a transitive sequence"
   182   | check Tactic.End_Proof' _ = Applicable.Yes Tactic.End_Proof''
   183   | check m _ = raise ERROR ("Solve_Step.check called for " ^ Tactic.input_to_string m);
   184 
   185 fun add (Tactic.Apply_Method' (_, topt, is, _)) (_, ctxt) (pt, pos as (p, _)) = 
   186     (case topt of 
   187       SOME t => 
   188         let val (pt, c) = Ctree.cappend_form pt p (is, ctxt) t
   189         in (pos, c, Generate.EmptyMout, pt) end
   190     | NONE => (pos, [], Generate.EmptyMout, pt))
   191   | add (Tactic.Take' t) l (pt, (p, _)) = (* val (Take' t) = m; *)
   192     let
   193       val p =
   194         let val (ps, p') = split_last p (* no connex to prev.ppobj *)
   195 	      in if p' = 0 then ps @ [1] else p end
   196       val (pt, c) = Ctree.cappend_form pt p l t
   197     in
   198       ((p, Pos.Frm), c, Generate.FormKF (UnparseC.term t), pt)
   199     end
   200   | add (Tactic.Begin_Trans' t) l (pt, (p, Pos.Frm)) =
   201     let
   202       val (pt, c) = Ctree.cappend_form pt p l t
   203       val pt = Ctree.update_branch pt p Ctree.TransitiveB (*040312*)
   204       (* replace the old PrfOjb ~~~~~ *)
   205       val p = (Pos.lev_on o Pos.lev_dn (* starts with [...,0] *)) p
   206       val (pt, c') = Ctree.cappend_form pt p l t (*FIXME.0402 same istate ???*)
   207     in
   208       ((p, Pos.Frm), c @ c', Generate.FormKF (UnparseC.term t), pt)
   209     end
   210   | add (Tactic.Begin_Trans' t) l (pt, (p, Pos.Res)) = 
   211     (*append after existing PrfObj    vvvvvvvvvvvvv*)
   212     add (Tactic.Begin_Trans' t) l (pt, (Pos.lev_on p, Pos.Frm))
   213   | add (Tactic.End_Trans' tasm) l (pt, (p, _)) =
   214     let
   215       val p' = Pos.lev_up p
   216       val (pt, c) = Ctree.append_result pt p' l tasm Ctree.Complete
   217     in
   218       ((p', Pos.Res), c, Generate.FormKF "DUMMY" (*term2str t ..ERROR (t) has not been declared*), pt)
   219     end
   220   | add (Tactic.Rewrite_Inst' (_, _, _, _, subs', thm', f, (f', asm))) (is, ctxt) (pt, (p, _)) =
   221     let
   222       val (pt, c) = Ctree.cappend_atomic pt p (is, ctxt) f
   223         (Tactic.Rewrite_Inst (Subst.T_to_input subs', thm')) (f',asm) Ctree.Complete;
   224       val pt = Ctree.update_branch pt p Ctree.TransitiveB
   225     in
   226       ((p, Pos.Res), c, Generate.FormKF (UnparseC.term f'), pt)
   227     end
   228  | add (Tactic.Rewrite' (_, _, _, _, thm', f, (f', asm))) (is, ctxt) (pt, (p, _)) =
   229    let
   230      val (pt, c) = Ctree.cappend_atomic pt p (is, ctxt) f (Tactic.Rewrite thm') (f', asm) Ctree.Complete
   231      val pt = Ctree.update_branch pt p Ctree.TransitiveB
   232    in
   233     ((p, Pos.Res), c, Generate.FormKF (UnparseC.term f'), pt)
   234    end
   235   | add (Tactic.Rewrite_Set_Inst' (_, _, subs', rls', f, (f', asm))) (is, ctxt) (pt, (p, _)) =
   236     let
   237       val (pt, c) = Ctree.cappend_atomic pt p (is, ctxt) f 
   238         (Tactic.Rewrite_Set_Inst (Subst.T_to_input subs', Rule_Set.id rls')) (f', asm) Ctree.Complete
   239       val pt = Ctree.update_branch pt p Ctree.TransitiveB
   240     in
   241       ((p, Pos.Res), c, Generate.FormKF (UnparseC.term f'), pt)
   242     end
   243   | add (Tactic.Rewrite_Set' (_, _, rls', f, (f', asm))) (is, ctxt) (pt, (p, _)) =
   244     let
   245       val (pt, c) = Ctree.cappend_atomic pt p (is, ctxt) f 
   246         (Tactic.Rewrite_Set (Rule_Set.id rls')) (f', asm) Ctree.Complete
   247       val pt = Ctree.update_branch pt p Ctree.TransitiveB
   248     in
   249       ((p, Pos.Res), c, Generate.FormKF (UnparseC.term f'), pt)
   250     end
   251   | add (Tactic.Check_Postcond' (_, scval)) l (pt, (p, _)) =
   252       let
   253         val (pt, c) = Ctree.append_result pt p l (scval, []) Ctree.Complete
   254       in
   255         ((p, Pos.Res), c, Generate.FormKF (UnparseC.term scval), pt)
   256       end
   257   | add (Tactic.Calculate' (_, op_, f, (f', _))) l (pt, (p, _)) =
   258       let
   259         val (pt,c) = Ctree.cappend_atomic pt p l f (Tactic.Calculate op_) (f', []) Ctree.Complete
   260       in
   261         ((p, Pos.Res), c, Generate.FormKF (UnparseC.term f'), pt)
   262       end
   263   | add (Tactic.Check_elementwise' (consts, pred, (f', asm))) l (pt, (p, _)) =
   264       let
   265         val (pt,c) = Ctree.cappend_atomic pt p l consts (Tactic.Check_elementwise pred) (f', asm) Ctree.Complete
   266       in
   267         ((p, Pos.Res), c, Generate.FormKF (UnparseC.term f'), pt)
   268       end
   269   | add (Tactic.Or_to_List' (ors, list)) l (pt, (p, _)) =
   270       let
   271         val (pt,c) = Ctree.cappend_atomic pt p l ors Tactic.Or_to_List (list, []) Ctree.Complete
   272       in
   273         ((p, Pos.Res), c, Generate.FormKF (UnparseC.term list), pt)
   274       end
   275   | add (Tactic.Substitute' (_, _, subte, t, t')) l (pt, (p, _)) =
   276       let
   277         val (pt,c) =
   278           Ctree.cappend_atomic pt p l t (Tactic.Substitute (Subst.eqs_to_input subte)) (t',[]) Ctree.Complete
   279         in ((p, Pos.Res), c, Generate.FormKF (UnparseC.term t'), pt) 
   280         end
   281   | add (Tactic.Tac_ (_, f, id, f')) l (pt, (p, _)) =
   282       let
   283         val (pt, c) = Ctree.cappend_atomic pt p l (TermC.str2term f) (Tactic.Tac id) (TermC.str2term f', []) Ctree.Complete
   284       in
   285         ((p,Pos.Res), c, Generate.FormKF f', pt)
   286       end
   287   | add (Tactic.Subproblem' ((domID, pblID, metID), oris, hdl, fmz_, ctxt_specify, f))
   288       (l as (_, ctxt)) (pt, (p, _)) =
   289       let
   290   	    val (pt, c) = Ctree.cappend_problem pt p l (fmz_, (domID, pblID, metID))
   291   	      (oris, (domID, pblID, metID), hdl, ctxt_specify)
   292   	    val f = Syntax.string_of_term (ThyC.to_ctxt (Proof_Context.theory_of ctxt)) f
   293       in
   294         ((p, Pos.Pbl), c, Generate.FormKF f, pt)
   295       end
   296   | add m' _ (_, pos) =
   297       raise ERROR ("Solve_Step.add: not impl.for " ^ Tactic.string_of m' ^ " at " ^ Pos.pos'2str pos)
   298 
   299 (* LI switches between solve-phase and specify-phase *)
   300 fun add_general tac ic cs =
   301   if Tactic.for_specify' tac
   302   then Specify_Step.add tac ic cs
   303   else add tac ic cs
   304 
   305 (* the order of State_Steps is reversed: insert last element first  *)
   306 fun s_add_general [] ptp = ptp
   307   | s_add_general tacis (pt, c, _) = 
   308     let
   309       val (tacis', (_, tac_, (p, is))) = split_last tacis
   310 	    val (p', c', _, pt') = add_general tac_ is (pt, p)
   311     in
   312       s_add_general tacis' (pt', c@c', p')
   313     end
   314 
   315 (* a still undeveloped concept: do a calculation without LI *)
   316 fun add_hard _(*thy*) m' (p, p_) pt =
   317   let  
   318     val p = case p_ of
   319       Pos.Frm => p | Pos.Res => Pos.lev_on p
   320     | _ => error ("generate_hard: call by " ^ Pos.pos'2str (p,p_))
   321   in
   322     add_general m' (Istate_Def.empty, ContextC.empty) (pt, (p, p_))
   323   end
   324 
   325 (**)end(**);