src/Tools/isac/Interpret/generate.sml
author Walther Neuper <wneuper@ist.tugraz.at>
Wed, 14 Dec 2016 10:45:41 +0100
changeset 59267 aab874fdd910
parent 59266 56762e8a672e
child 59268 c988bdecd7be
permissions -rw-r--r--
redesigned inout, mout (since 2003 for interface Kernel - Java)

cf 56762e8a672e note (1)
     1 (* use"ME/generate.sml";
     2    use"generate.sml";
     3    *)
     4 signature CALC_TREE =
     5 sig (*vvv request into signature is incremental *)
     6   (* for calchead.sml *)
     7   type taci
     8   val e_taci : taci
     9   datatype edit = EdUndef | Protect | Write
    10   eqtype indent
    11   datatype nest = Closed | Nundef | Open
    12   datatype pblmet = Method of metID | Problem of pblID | Upblmet
    13   datatype inout =
    14      Error_ of string
    15   | FormKF of cellID * edit * indent * nest * cterm'
    16   | PpcKF of cellID * edit * indent * nest * (pblmet * item ppc) 
    17   | RefineKF of match list
    18   | RefinedKF of pblID * (itm list * (bool * term) list)
    19   datatype mout = EmptyMout | Error' of inout | Form' of inout | Problems of inout
    20   val generate1 : theory -> tac_ -> istate * Proof.context ->
    21     posel list * pos_ -> ptree -> pos' * pos' list * mout * ptree
    22   val init_pbl : (string * (term * 'a)) list -> itm list
    23   val init_pbl' : (string * (term * term)) list -> itm list
    24   (* for appl.sml *)
    25   (* for script.sml *)
    26   (* for solve.sml *)
    27   val generate_hard : theory -> tac_ -> pos * pos_ -> ptree -> pos' * pos' list * mout * ptree
    28   val init_istate : tac -> term -> istate
    29   (* for inform.sml *)
    30   val embed_deriv : taci list -> ptree * pos' -> pos' list * (ptree * pos')
    31   (* for mathengine.sml *)
    32   val generate : (tac * tac_ * ((posel list * pos_) * (istate * Proof.context))) list ->
    33      ptree * pos' list * pos' -> ptree * pos' list * pos'
    34   (* for interface.sml *)
    35   val generate_inconsistent_rew :
    36      subs option * thm'' ->
    37      term -> istate * Proof.context -> pos * pos_ -> ptree -> ptree * (posel list * pos_)
    38 (*/-------------------------------------------------------- ! aktivate for Test_Isac BEGIN ---\* )
    39 ( *\--- ! aktivate for Test_Isac END ----------------------------------------------------------/*)
    40 end
    41 
    42 (**)
    43 structure Ctree(*: CALC_TREE*) =
    44 (**)
    45 struct
    46 (* initialize istate for Detail_Set *)
    47 fun init_istate (Rewrite_Set rls) t =
    48     (case assoc_rls rls of
    49       Rrls {scr = Rfuns {init_state = ii, ...}, ...} => RrlsState (ii t)
    50     | Rls {scr = EmptyScr, ...} => 
    51       error ("interSteps>..>init_istate: \"" ^ rls ^ "\" has EmptyScr." ^
    52         "use prep_rls' for storing rule-sets !")
    53     | Rls {scr = Prog s, ...} => (ScrState ([(one_scr_arg s, t)], [], NONE, e_term, Sundef, true))
    54     | Seq {scr=EmptyScr,...} => 
    55       error ("interSteps>..>init_istate: \"" ^ rls ^ "\" has EmptyScr." ^
    56 		    "use prep_rls' for storing rule-sets !")
    57     | Seq {scr = Prog s,...} => (ScrState ([(one_scr_arg s, t)], [], NONE, e_term, Sundef, true))
    58     | _ => error "init_istate Rewrite_Set: uncovered case assoc_rls")
    59   | init_istate (Rewrite_Set_Inst (subs, rls)) t =
    60     let
    61       val v = case subs2subst (assoc_thy "Isac") subs of
    62         (_, v) :: _ => v
    63       | _ => error "init_istate: uncovered case "
    64     (*...we suppose the substitution of only _one_ bound variable*)
    65     in case assoc_rls rls of
    66       Rls {scr = EmptyScr, ...} => 
    67         error ("interSteps>..>init_istate: \""^rls^"\" has EmptyScr." ^
    68           "use prep_rls' for storing rule-sets !")
    69 	  | Rls {scr = Prog s, ...} =>
    70 	    let val (form, bdv) = two_scr_arg s
    71 	    in (ScrState ([(form, t), (bdv, v)], [], NONE, e_term, Sundef,true))
    72 	    end
    73 	  | Seq {scr = EmptyScr, ...} => 
    74 	    error ("interSteps>..>init_istate: \""^rls^"\" has EmptyScr." ^
    75 	      "use prep_rls' for storing rule-sets !")
    76 	  | Seq {scr = Prog s,...} =>
    77 	    let val (form, bdv) = two_scr_arg s
    78 	    in (ScrState ([(form, t), (bdv, v)],[], NONE, e_term, Sundef,true))
    79 	    end
    80     | _ => error "init_istate Rewrite_Set_Inst: uncovered case assoc_rls"
    81     end
    82   | init_istate tac _ = error ("init_istate: uncovered definition for " ^ tac2str tac)
    83 
    84 (* a taci holds alle information required to build a node in the calc-tree;
    85    a taci is assumed to be used efficiently such that the calc-tree
    86    resulting from applying a taci need not be stored separately;
    87    see "type calcstate" *)
    88 (*TODO.WN0504 redesign ??? or redesign generate ?? see "fun generate"
    89   TODO.WN0512 ? redesign this _list_:
    90   # only used for [Apply_Method + (Take or Subproblem)], i.e. for initacs
    91   # the latter problem may be resolved automatically if "fun autocalc" is 
    92     not any more used for the specify-phase and for changing the phases*)
    93 type taci = 
    94   (tac *                        (* for comparison with input tac             *)      
    95    tac_ *                       (* for ptree generation                      *)
    96    (pos' *                      (* after applying tac_, for ptree generation *)
    97     (istate * Proof.context)))  (* after applying tac_, for ptree generation *)
    98 val e_taci = (Empty_Tac, Empty_Tac_, (e_pos', (e_istate, e_ctxt))): taci
    99 fun taci2str ((tac, tac_, (pos', (istate, _))):taci) =
   100   "( " ^ tac2str tac ^ ", " ^ tac_2str tac_ ^ ", ( " ^ pos'2str pos' ^ ", " ^
   101   istate2str istate ^ " ))"
   102 fun tacis2str tacis = (strs2str o (map (linefeed o taci2str))) tacis
   103 
   104 datatype pblmet =     (*%^%*)
   105   Upblmet             (*undefined*)
   106 | Problem of pblID    (*%^%*)
   107 | Method of metID;    (*%^%*)
   108 fun pblmet2str (Problem pblID) = "Problem " ^ strs2str pblID (*%^%*)
   109   | pblmet2str (Method metID) = "Method " ^ metID2str metID (*%^%*)
   110   | pblmet2str x = error ("pblmet2str: uncovered definition " ^ pblmet2str x)
   111 
   112 (*3.5.00: TODO: foppFK eliminated in interface FE-KE !!!*)
   113 datatype foppFK =                  (* in DG cases div 2 *)
   114   EmptyFoppFK         (*DG internal*)
   115 | FormFK of cterm'
   116 | PpcFK of cterm' ppc
   117 fun foppFK2str (FormFK ct') ="FormFK " ^ ct'
   118   | foppFK2str (PpcFK  ppc) ="PpcFK " ^ ppc2str ppc
   119   | foppFK2str EmptyFoppFK  ="EmptyFoppFK"
   120 
   121 datatype nest = Open | Closed | Nundef;
   122 fun nest2str Open = "Open"
   123   | nest2str Closed = "Closed"
   124   | nest2str Nundef = "Nundef"
   125 
   126 type indent = int;
   127 datatype edit = EdUndef | Write | Protect;
   128                                    (* bridge --> kernel *)
   129                                    (* bridge <-> kernel *)
   130 (* needed in dialog.sml *)         (* bridge <-- kernel *)
   131 fun edit2str EdUndef = "EdUndef"
   132   | edit2str Write = "Write"
   133   | edit2str Protect = "Protect";
   134 
   135 datatype inout = (*FIXME.WN1105 drop this: was required for proto0 with dialog in sml*)
   136   Error_ of string                                             (*<--*)
   137 | FormKF of cellID * edit * indent * nest * cterm'             (*<--*)
   138 | PpcKF of cellID * edit * indent * nest * (pblmet * item ppc) (*<--*)
   139 | RefineKF of match list                                       (*<--*)
   140 | RefinedKF of (pblID * ((itm list) * ((bool * term) list)))   (*<--*)
   141 
   142 fun inout2str End_Proof = "End_Proof"
   143   | inout2str (Error_  s) = "Error_ "^s
   144   | inout2str (FormKF (cellID, edit, indent, nest, ct')) =  
   145 	       "FormKF ("^(string_of_int cellID)^","
   146 	       ^(edit2str edit)^","^(string_of_int indent)^","
   147 	       ^(nest2str nest)^",("
   148 	       ^ct' ^")"
   149   | inout2str (PpcKF (cellID, edit, indent, nest, (pm,itemppc))) =
   150 	       "PpcKF ("^(string_of_int cellID)^","
   151 	       ^(edit2str edit)^","^(string_of_int indent)^","
   152 	       ^(nest2str nest)^",("
   153 	       ^(pblmet2str pm)^","^(itemppc2str itemppc)^"))"
   154   | inout2str (RefineKF ms)  = "RefineKF "^(matchs2str ms)
   155   | inout2str _ = error "inout2str: uncovered definition"
   156 fun inouts2str ios = (strs2str' o (map inout2str)) ios
   157 
   158 (*
   159   datatype mout = EmptyMout | Error' of inout | Form' of inout | Problems of inout
   160 *)
   161 datatype mout =
   162   FormKF of cterm'
   163 | PpcKF of (pblmet * item ppc) 
   164 | RefinedKF of pblID * (itm list * (bool * term) list)
   165 | Error' of string
   166 | EmptyMout
   167 
   168 fun mout2str (FormKF cterm') = "FormKF " ^ cterm'
   169   | mout2str (PpcKF  (pm, itemppc)) = "PpcKF (" ^ pblmet2str pm ^ "," ^ itemppc2str itemppc ^ ")"
   170   | mout2str (RefinedKF  (pblID, ls)) = "mout2str: RefinedKF not impl."
   171   | mout2str (Error'  str) = "Error' " ^ str
   172   | mout2str (EmptyMout    ) = "EmptyMout"
   173 
   174 (* init pbl with ...,dsc,empty | [] *)
   175 fun init_pbl pbt = 
   176   let
   177     fun pbt2itm (f, (d, _)) = ((0, [], false, f, Inc ((d, []), (e_term, []))) : itm)
   178   in map pbt2itm pbt end
   179 
   180 (* take formal parameters from pbt, for transfer from pbl/met-hierarchy *)
   181 fun init_pbl' pbt = 
   182   let 
   183     fun pbt2itm (f, (d, t)) = ((0, [], false, f, Inc((d, [t]), (e_term, []))) : itm)
   184   in map pbt2itm pbt end
   185 
   186 (*generate 1 ppobj in ptree*)
   187 (*TODO.WN0501: take calcstate as an argument (see embed_derive etc.)?specify?*)
   188 fun generate1 thy (Add_Given' (_, itmlist)) _ (pos as (p, p_)) pt = 
   189     (pos: pos', [], PpcKF (Upblmet, itms2itemppc thy [][]),
   190        case p_ of
   191          Pbl => update_pbl pt p itmlist
   192 	     | Met => update_met pt p itmlist
   193        | _ => error ("uncovered case " ^ pos_2str p_))
   194     (* WN110515 probably declare_constraints with input item (without dsc) --
   195       -- important when specifying without formalisation *)
   196   | generate1 thy (Add_Find' (_, itmlist)) _ (pos as (p, p_)) pt = 
   197     (pos, [], (PpcKF (Upblmet, itms2itemppc thy [] [])),
   198        case p_ of
   199          Pbl => update_pbl pt p itmlist
   200 	     | Met => update_met pt p itmlist
   201        | _ => error ("uncovered case " ^ pos_2str p_))
   202     (*WN110515 probably declare_constraints with input item (without dsc)*)
   203   | generate1 thy (Add_Relation' (_, itmlist)) _ (pos as (p, p_)) pt = 
   204     (pos, [],  PpcKF (Upblmet, itms2itemppc thy [] []),
   205        case p_ of
   206          Pbl => update_pbl pt p itmlist
   207 	     | Met => update_met pt p itmlist
   208        | _ => error ("uncovered case " ^ pos_2str p_))
   209   | generate1 thy (Specify_Theory' domID) _ (pos as (p,_)) pt = 
   210     (pos, [] , PpcKF  (Upblmet, itms2itemppc thy [] []),
   211       update_domID pt p domID)
   212   | generate1 thy (Specify_Problem' (pI, (_, (itms, _)))) _ (p, _) pt = 
   213     let
   214       val pt = update_pbl pt p itms
   215       val pt = update_pblID pt p pI
   216     in
   217       ((p, Pbl), [], PpcKF (Upblmet, itms2itemppc thy [] []), pt)
   218     end
   219   | generate1 thy (Specify_Method' (mID, oris, itms)) _ (p, _) pt = 
   220     let
   221       val pt = update_oris pt p oris
   222       val pt = update_met pt p itms
   223       val pt = update_metID pt p mID
   224     in
   225       ((p, Met), [], PpcKF (Upblmet, itms2itemppc thy [] []), pt)
   226     end
   227   | generate1 thy (Model_Problem' (_, itms, met)) _ (pos as (p,_)) pt =
   228     let 
   229       val pt = update_pbl pt p itms
   230 	    val pt = update_met pt p met
   231     in
   232       (pos, [], PpcKF (Upblmet, itms2itemppc thy [] []), pt)
   233     end
   234   | generate1 thy (Refine_Tacitly' (_, pIre, domID, metID, pbl)) _ (pos as (p, _)) pt = 
   235     let
   236       val pt = update_pbl pt p pbl
   237       val pt = update_orispec pt p (domID, pIre, metID)
   238     in
   239       (pos, [], PpcKF (Upblmet, itms2itemppc thy [] []), pt)
   240     end
   241   | generate1 thy (Refine_Problem' (pI, _)) _ (pos as (p, _)) pt =
   242     let
   243       val (dI, _, mI) = get_obj g_spec pt p
   244       val pt = update_spec pt p (dI, pI, mI)
   245     in
   246       (pos, [], PpcKF (Upblmet, itms2itemppc thy [] []), pt)
   247     end
   248   | generate1 _ (Apply_Method' (_, topt, is, ctxt)) _ (pos as (p, _)) pt = 
   249     (case topt of 
   250       SOME t => 
   251         let val (pt, c) = cappend_form pt p (is, ctxt) t
   252         in (pos, c, EmptyMout, pt) end
   253     | NONE => (pos, [], EmptyMout, update_env pt p (SOME (is, ctxt))))
   254   | generate1 _ (Take' t) l (p, _) pt = (* val (Take' t) = m; *)
   255     let
   256       val p =
   257         let val (ps, p') = split_last p (* no connex to prev.ppobj *)
   258 	      in if p' = 0 then ps @ [1] else p end
   259       val (pt, c) = cappend_form pt p l t
   260     in
   261       ((p, Frm): pos', c, FormKF (term2str t), pt)
   262     end
   263   | generate1 _ (Begin_Trans' t) l (p, Frm) pt =
   264     let
   265       val (pt, c) = cappend_form pt p l t
   266       val pt = update_branch pt p TransitiveB (*040312*)
   267       (* replace the old PrfOjb ~~~~~ *)
   268       val p = (lev_on o lev_dn (* starts with [...,0] *)) p
   269       val (pt, c') = cappend_form pt p l t (*FIXME.0402 same istate ???*)
   270     in
   271       ((p, Frm), c @ c', FormKF (term2str t), pt)
   272     end
   273   | generate1 thy (Begin_Trans' t) l (p, Res) pt = 
   274     (*append after existing PrfObj    vvvvvvvvvvvvv*)
   275     generate1 thy (Begin_Trans' t) l (lev_on p, Frm) pt
   276   | generate1 _ (End_Trans' tasm) l (p, _) pt =
   277     let
   278       val p' = lev_up p
   279       val (pt, c) = append_result pt p' l tasm Complete
   280     in
   281       ((p', Res), c, FormKF (term2str t), pt)
   282     end
   283   | generate1 _ (Rewrite_Inst' (_, _, _, _, subs', thm', f, (f', asm))) (is, ctxt) (p, _) pt =
   284     let
   285       val (pt, c) = cappend_atomic pt p (is, insert_assumptions asm ctxt) f
   286         (Rewrite_Inst (subst2subs subs', thm')) (f',asm) Complete;
   287       val pt = update_branch pt p TransitiveB
   288     in
   289       ((p, Res), c, FormKF (term2str f'), pt)
   290     end
   291  | generate1 _ (Rewrite' (_, _, _, _, thm', f, (f', asm))) (is, ctxt) (p, _) pt =
   292    let
   293      val (pt, c) = cappend_atomic pt p (is, insert_assumptions asm ctxt) f
   294        (Rewrite thm') (f', asm) Complete
   295      val pt = update_branch pt p TransitiveB
   296    in
   297     ((p, Res), c, FormKF (term2str f'), pt)
   298    end
   299   | generate1 thy (Rewrite_Asm' all) l p pt = generate1 thy (Rewrite' all) l p pt
   300   | generate1 _ (Rewrite_Set_Inst' (_, _, subs', rls', f, (f', asm))) (is, ctxt) (p, _) pt =
   301     let
   302       val (pt, c) = cappend_atomic pt p (is, insert_assumptions asm ctxt) f 
   303         (Rewrite_Set_Inst (subst2subs subs', id_rls rls')) (f', asm) Complete
   304       val pt = update_branch pt p TransitiveB
   305     in
   306       ((p, Res), c, FormKF (term2str f'), pt)
   307     end
   308   | generate1 thy (Detail_Set_Inst' (_, _, subs, rls, f, (_, asm))) (is, ctxt) (p, _) pt =
   309     let
   310       val ctxt' = insert_assumptions asm ctxt
   311       val (pt, _) = cappend_form pt p (is, ctxt') f 
   312       val pt = update_branch pt p TransitiveB
   313       val is = init_istate (Rewrite_Set_Inst (subst2subs subs, id_rls rls)) f 
   314       val tac_ = Apply_Method' (e_metID, SOME t, is, ctxt')
   315       val pos' = ((lev_on o lev_dn) p, Frm)
   316     in
   317       generate1 thy tac_ (is, ctxt') pos' pt (*implicit Take*)
   318     end
   319   | generate1 _ (Rewrite_Set' (_, _, rls', f, (f', asm))) (is, ctxt) (p, _) pt =
   320     let
   321       val (pt, c) = cappend_atomic pt p (is, insert_assumptions asm ctxt) f 
   322         (Rewrite_Set (id_rls rls')) (f',asm) Complete
   323       val pt = update_branch pt p TransitiveB
   324     in
   325       ((p, Res), c, FormKF (term2str f'), pt)
   326     end
   327   | generate1 thy (Detail_Set' (_, _, rls, f, (_, asm))) (is, ctxt) (p, _) pt =
   328     let
   329       val ctxt' = insert_assumptions asm ctxt
   330       val (pt, _) = cappend_form pt p (is, ctxt') f 
   331       val pt = update_branch pt p TransitiveB
   332       val is = init_istate (Rewrite_Set (id_rls rls)) f
   333       val tac_ = Apply_Method' (e_metID, SOME t, is, ctxt')
   334       val pos' = ((lev_on o lev_dn) p, Frm)
   335     in
   336       generate1 thy tac_ (is, ctxt') pos' pt (*implicit Take*)
   337     end
   338   | generate1 _ (Check_Postcond' (_, (scval, asm))) l (p, _) pt =
   339       let
   340         val (pt, c) = append_result pt p l (scval, asm) Complete
   341       in
   342         ((p, Res), c, FormKF (term2str scval), pt)
   343       end
   344   | generate1 _ (Calculate' (_, op_, f, (f', _))) l (p, _) pt =
   345       let
   346         val (pt,c) = cappend_atomic pt p l f (Calculate op_) (f', []) Complete
   347       in
   348         ((p, Res), c, FormKF (term2str f'), pt)
   349       end
   350   | generate1 _ (Check_elementwise' (consts, pred, (f', asm))) l (p, _) pt =
   351       let
   352         val (pt,c) = cappend_atomic pt p l consts (Check_elementwise pred) (f',asm) Complete
   353       in
   354         ((p, Res), c, FormKF (term2str f'), pt)
   355       end
   356   | generate1 _ (Or_to_List' (ors, list)) l (p, _) pt =
   357       let
   358         val (pt,c) = cappend_atomic pt p l ors Or_to_List (list, []) Complete
   359       in
   360         ((p, Res), c, FormKF (term2str list), pt)
   361       end
   362   | generate1 _ (Substitute' (_, _, subte, t, t')) l (p, _) pt =
   363       let
   364         val (pt,c) =
   365           cappend_atomic pt p l t (Substitute (subte2sube subte)) (t',[]) Complete
   366         in ((p, Res), c, FormKF (term2str t'), pt) 
   367         end
   368   | generate1 _ (Tac_ (_, f, id, f')) l (p, _) pt =
   369       let
   370         val (pt, c) = cappend_atomic pt p l (str2term f) (Tac id) (str2term f', []) Complete
   371       in
   372         ((p,Res), c, FormKF f', pt)
   373       end
   374   | generate1 thy (Subproblem' ((domID, pblID, metID), oris, hdl, fmz_, ctxt, f)) l (p, _) pt =
   375     let
   376 	    val (pt, c) = cappend_problem pt p l (fmz_, (domID, pblID, metID)) (oris, (domID, pblID, metID), hdl)
   377       val pt = update_ctxt pt p ctxt
   378 	    val f = Syntax.string_of_term (thy2ctxt thy) f
   379     in
   380       ((p, Pbl), c, FormKF f, pt)
   381     end
   382   | generate1 _ m' _ _ _ = error ("generate1: not impl.for " ^ tac_2str m')
   383 
   384 fun generate_inconsistent_rew (subs_opt, thm') f' (is, ctxt) (pos as (p,_)) pt =
   385   let
   386     val f = get_curr_formula (pt, pos)
   387     val pos' as (p', _) = (lev_on p, Res)
   388     val (pt, _) = case subs_opt of
   389       NONE => cappend_atomic pt p' (is, insert_assumptions [] ctxt) f
   390         (Rewrite thm') (f', []) Inconsistent
   391     | SOME subs => cappend_atomic pt p' (is, insert_assumptions [] ctxt) f
   392         (Rewrite_Inst (subs, thm')) (f', []) Inconsistent
   393     val pt = update_branch pt p' TransitiveB
   394   in (pt, pos') end
   395 
   396 fun generate_hard thy m' (p,p_) pt =
   397   let  
   398     val p = case p_ of
   399       Frm => p | Res => lev_on p
   400     | _ => error ("generate_hard: call by " ^ pos'2str (p,p_))
   401   in generate1 thy m' (e_istate, e_ctxt) (p,p_) pt end
   402 
   403 (* tacis are in reverse order from nxt_solve_/specify_: last = fst to insert *)
   404 fun generate ([]: taci list) ptp = ptp
   405   | generate tacis (pt, c, _: pos'(*!dropped!WN0504redesign generate/tacis?*))= 
   406     let
   407       val (tacis', (_, tac_, (p, is))) = split_last tacis
   408 	    val (p',c',_,pt') = generate1 (assoc_thy "Isac") tac_ is p pt
   409     in
   410       generate tacis' (pt', c@c', p')
   411     end
   412 
   413 (* update pos in tacis for embedding by generate *)
   414 fun insert_pos (_: pos) [] = []
   415   | insert_pos (p: pos) (((tac, tac_, (_, ist)): taci) :: tacis) = 
   416     ((tac, tac_, ((p, Res), ist)): taci) :: (insert_pos (lev_on p) tacis)
   417 
   418 fun res_from_taci (_, Rewrite' (_, _, _, _, _, _,(res, asm)), _) = (res, asm)
   419   | res_from_taci (_, Rewrite_Set' (_, _, _, _, (res, asm)), _) = (res, asm)
   420   | res_from_taci (_, tac_, _) = error ("res_from_taci: called with" ^ tac_2str tac_)
   421 
   422 (* embed the tacis created by a '_deriv'ation; sys.form <> input.form
   423   tacis are in order, thus are reverted for generate *)
   424 fun embed_deriv (tacis: taci list) (pt, pos as (p, Frm): pos') =
   425   (*inform at Frm: replace the whole PrfObj by a Transitive-ProfObj FIXME?0402
   426     and transfer the istate (from _after_ compare_deriv) from Frm to Res*)
   427     let
   428       val (res, asm) = (res_from_taci o last_elem) tacis
   429     	val (ist, ctxt) = case get_obj g_loc pt p of
   430     	  (SOME (ist, ctxt), _) => (ist, ctxt)
   431       | (NONE, _) => error "embed_deriv Frm: uncovered case get_obj"
   432     	val form = get_obj g_form pt p
   433           (*val p = lev_on p; ---------------only difference to (..,Res) below*)
   434     	val tacis = (Begin_Trans, Begin_Trans' form, (pos, (Uistate, ctxt))) ::
   435     		(insert_pos ((lev_on o lev_dn) p) tacis) @ [(End_Trans, End_Trans' (res, asm),
   436     			(pos_plus (length tacis) (lev_dn p, Res), (new_val res ist, ctxt)))]
   437     	val {nrls, ...} = get_met (get_obj g_metID pt (par_pblobj pt p))
   438     	val (pt, c, pos as (p, _)) = generate (rev tacis) (pt, [], (p, Res))
   439     	val pt = update_tac pt p (Derive (id_rls nrls))
   440     	val pt = update_branch pt p TransitiveB
   441     in (c, (pt, pos: pos')) end
   442   | embed_deriv tacis (pt, (p, Res)) =
   443     (*inform at Res: append a Transitive-PrfObj FIXME?0402 other branch-types ?
   444       and transfer the istate (from _after_ compare_deriv) from Res to new Res*)
   445     let val (res, asm) = (res_from_taci o last_elem) tacis
   446     	val (ist, ctxt) = case get_obj g_loc pt p of
   447     	  (_, SOME (ist, ctxt)) => (ist, ctxt)
   448       | (_, NONE) => error "embed_deriv Frm: uncovered case get_obj"
   449     	val (f, _) = get_obj g_result pt p
   450     	val p = lev_on p(*---------------only difference to (..,Frm) above*);
   451     	val tacis = (Begin_Trans, Begin_Trans' f, ((p, Frm), (Uistate, ctxt))) ::
   452     		(insert_pos ((lev_on o lev_dn) p) tacis) @ [(End_Trans, End_Trans' (res, asm), 
   453     			(pos_plus (length tacis) (lev_dn p, Res), (new_val res ist, ctxt)))];
   454     	val {nrls, ...} = get_met (get_obj g_metID pt (par_pblobj pt p))
   455     	val (pt, c, pos as (p, _)) = generate (rev tacis) (pt, [], (p, Res))
   456     	val pt = update_tac pt p (Derive (id_rls nrls))
   457     	val pt = update_branch pt p TransitiveB
   458     in (c, (pt, pos)) end
   459   | embed_deriv _ _ = error "embed_deriv: uncovered definition"
   460 end