src/Tools/isac/Interpret/derive.sml
author wneuper <Walther.Neuper@jku.at>
Fri, 03 Feb 2023 12:05:30 +0100
changeset 60668 b5e3823f1cbd
parent 60655 f73460617c3d
child 60675 d841c720d288
permissions -rw-r--r--
eliminate use of Thy_Info 16: eliminate UnparseC.terms in src/*, too
walther@59906
     1
(* Title:  Interpret/derive.sml
walther@59934
     2
   Author: Walther Neuper
walther@59906
     3
   (c) due to copyright terms
walther@59906
     4
walther@59934
     5
Try to make (term * rule * result) steps (= derivation) by use of a Rule_Set.T.
walther@59934
     6
Two purposes:
walther@59934
     7
(1) derive steps from a given term towards another give term
walther@59934
     8
(2) term transformations, which cannot be done by rewriting, e.g cancellation of polynomials.
walther@59906
     9
*)
walther@59906
    10
walther@59907
    11
signature DERIVE =
walther@59906
    12
sig
walther@59934
    13
  type rule_result
walther@59934
    14
  type step
walther@59907
    15
  type derivation
walther@59934
    16
Walther@60509
    17
  val do_one : Proof.context -> Rule_Set.T -> Rule.rule list -> Rewrite_Ord.function ->
walther@59907
    18
    term option -> term -> derivation
Walther@60509
    19
  val steps_reverse : Proof.context -> Rule_Set.T -> Rule.rule list -> Rewrite_Ord.function ->
walther@59934
    20
    term option -> term -> rule_result list
Walther@60525
    21
  val steps : Proof.context -> Rewrite_Ord.T -> Rule_Set.T -> Rule.rule list -> term -> term ->
walther@59934
    22
    bool * derivation
walther@59932
    23
  val embed: State_Steps.T -> Calc.T -> Pos.pos' list * Calc.T
wenzelm@60223
    24
\<^isac_test>\<open>
Walther@60611
    25
  val trtas2str : Proof.context -> (term * Rule.rule * (term * term list)) list -> string
Walther@60611
    26
  val deriv2str : Proof.context -> (term * Rule.rule * (term * term list)) list -> string
Walther@60630
    27
  val rev_deriv' : Proof.context -> 'a * Rule.rule * ('b * 'c) -> 'b * Rule.rule * ('a * 'c)
wenzelm@60223
    28
\<close>
walther@59906
    29
end
walther@59906
    30
walther@59906
    31
(**)
walther@59907
    32
structure Derive(**): DERIVE(**) =
walther@59906
    33
struct
walther@59906
    34
(**)
walther@59907
    35
walther@59907
    36
(** the triple for a step **)
walther@59907
    37
walther@59934
    38
type rule_result = Rule.rule * Calc.result;
walther@59934
    39
type step = term * Rule.rule * Calc.result;
walther@59934
    40
type derivation = step list;
walther@59906
    41
Walther@60611
    42
fun trta2str ctxt (t, r, (t', a)) =
Walther@60611
    43
  "\n(" ^ UnparseC.term_in_ctxt ctxt t ^ ", " ^ Rule.to_string_short r ^ 
Walther@60668
    44
  ", (" ^ UnparseC.term_in_ctxt ctxt t' ^ ", " ^ UnparseC.terms_in_ctxt ctxt a ^ "))"
Walther@60611
    45
fun trtas2str ctxt trtas = (strs2str o (map (trta2str ctxt))) trtas
walther@59906
    46
val deriv2str = trtas2str
walther@59906
    47
walther@59934
    48
walther@59907
    49
(** make one triple towards the goal term **)
walther@59906
    50
Walther@60501
    51
fun msg_1 ctxt rts =
Walther@60501
    52
  (tracing ("do_one exceeds " ^ int2str (Config.get ctxt rewrite_limit) ^ "with derivation =\n");
Walther@60611
    53
   tracing (deriv2str ctxt rts));
Walther@60500
    54
fun msg_2 ctxt thmid =
Walther@60500
    55
  if not (Config.get ctxt rewrite_trace) then () else tracing ("### trying thm \"" ^ thmid ^ "\"");
Walther@60500
    56
fun msg_3 ctxt t' =
Walther@60592
    57
  if Config.get ctxt rewrite_trace then tracing ("=== rewrites to: " ^ UnparseC.term_in_ctxt ctxt t') else ();
Walther@60500
    58
fun msg_4 ctxt op_ =
Walther@60500
    59
  if not (Config.get ctxt rewrite_trace) then () else tracing ("### trying calc. \"" ^ op_^"\"");
Walther@60500
    60
fun msg_5 ctxt t' =
Walther@60592
    61
  if not (Config.get ctxt rewrite_trace) then () else tracing("=== calc. to: " ^ UnparseC.term_in_ctxt ctxt t')
walther@59906
    62
walther@59934
    63
Walther@60586
    64
fun do_one ctxt asm_rls rs ro goal tt = 
Walther@60629
    65
  let
walther@59907
    66
    datatype switch = Appl | Noap (* TODO: unify with version in Rewrite *)
walther@59906
    67
    fun rew_once _ rts t Noap [] = 
walther@59907
    68
        (case goal of NONE => rts | SOME _ =>
Walther@60592
    69
          raise ERROR ("Derive.do_one: no derivation for " ^ UnparseC.term_in_ctxt ctxt t))
walther@59906
    70
      | rew_once lim rts t Appl [] = rew_once lim rts t Noap rs
walther@59906
    71
      | rew_once lim rts t apno rs' =
walther@59906
    72
        (case goal of 
walther@59906
    73
          NONE => rew_or_calc lim rts t apno rs'
walther@59906
    74
        | SOME g => if g = t then rts else rew_or_calc lim rts t apno rs')
walther@59906
    75
    and rew_or_calc lim rts t apno (rrs' as (r :: rs')) =
walther@59906
    76
      if lim < 0 
Walther@60501
    77
      then (msg_1 ctxt rts; rts)
walther@59906
    78
      else
walther@59906
    79
        (case r of
walther@59906
    80
          Rule.Thm (thmid, tm) => 
Walther@60500
    81
            (msg_2 ctxt thmid;
Walther@60586
    82
            case Rewrite.rewrite_ ctxt ro asm_rls true tm t of
walther@59906
    83
              NONE => rew_once lim rts t apno rs'
walther@59906
    84
            | SOME (t', a') =>
Walther@60500
    85
              (msg_3 ctxt t'; rew_once (lim - 1) (rts @ [(t, r, (t', a'))]) t' Appl rrs'))
walther@59906
    86
        | Rule.Eval (c as (op_, _)) => 
Walther@60500
    87
            (msg_4 ctxt op_;
Walther@60519
    88
            case Eval.adhoc_thm ctxt c t of
walther@59906
    89
              NONE => rew_once lim rts t apno rs'
walther@59906
    90
            | SOME (thmid, tm) => 
walther@59906
    91
              (let
Walther@60586
    92
                val (t', a') = case Rewrite.rewrite_ ctxt ro asm_rls true tm t of
walther@59906
    93
                  SOME ta => ta
walther@59907
    94
                | NONE => raise ERROR "adhoc_thm: NONE"
Walther@60500
    95
                val _ = msg_5 ctxt t'
walther@59906
    96
                val r' = Rule.Thm (thmid, tm)
walther@59906
    97
              in rew_once (lim - 1) (rts @ [(t, r', (t', a'))]) t' Appl rrs' end) 
walther@60269
    98
                handle Rewrite.NO_REWRITE => raise ERROR "derive_norm, Eval: no rewrite")
walther@59907
    99
        | Rule.Rls_ rls =>
Walther@60500
   100
          (case Rewrite.rewrite_set_ ctxt true rls t of
walther@59906
   101
            NONE => rew_once lim rts t apno rs'
walther@59906
   102
          | SOME (t', a') => rew_once (lim - 1) (rts @ [(t, r, (t', a'))]) t' Appl rrs')
Walther@60647
   103
        | rule => raise ERROR ("rew_once: uncovered case " ^ Rule.to_string ctxt rule))
walther@59907
   104
    | rew_or_calc _ _ _ _ [] = raise ERROR "rew_or_calc: called with []"
Walther@60501
   105
  in rew_once (Config.get ctxt rewrite_limit) [] tt Noap rs end
walther@59906
   106
walther@59907
   107
walther@59907
   108
(** concatenate several steps in revers order **)
walther@59907
   109
Walther@60644
   110
fun rev_deriv ctxt (t, r, (_, a)) = (ThmC.make_sym_rule ctxt r, (t, a));
Walther@60630
   111
fun steps_reverse ctxt asm_rls rs ro goal t =
Walther@60644
   112
    (rev o (map (rev_deriv ctxt))) (do_one ctxt asm_rls rs ro goal t)
walther@59906
   113
walther@59907
   114
walther@59907
   115
(** concatenate several steps **)
walther@59907
   116
Walther@60644
   117
fun rev_deriv' ctxt (t, r, (t', a)) = (t', ThmC.make_sym_rule ctxt r, (t, a));
walther@59906
   118
Walther@60528
   119
(* case fo = ifo excluded already in inform *)
Walther@60586
   120
fun steps ctxt rew_ord asm_rls rules fo ifo =
Walther@60630
   121
  let
walther@59906
   122
    fun derivat ([]:(term * Rule.rule * (term * term list)) list) = TermC.empty
walther@59906
   123
      | derivat dt = (#1 o #3 o last_elem) dt
walther@59907
   124
    fun equal (_, _, (t1, _)) (_, _, (t2, _)) = t1 = t2
Walther@60586
   125
    val  fod = do_one ctxt asm_rls rules (snd rew_ord) NONE  fo
Walther@60586
   126
    val ifod = do_one ctxt asm_rls rules (snd rew_ord) NONE ifo
walther@59906
   127
  in 
walther@59906
   128
    case (fod, ifod) of
walther@59906
   129
      ([], []) => if fo = ifo then (true, []) else (false, [])
walther@59906
   130
    | (fod, []) => if derivat fod = ifo then (true, fod) (*ifo is normal form*) else (false, [])
Walther@60630
   131
    | ([], ifod) => if fo = derivat ifod then (true, ((map (rev_deriv' ctxt)) o rev) ifod) else (false, [])
walther@59906
   132
    | (fod, ifod) =>
walther@59906
   133
      if derivat fod = derivat ifod (*common normal form found*) then
walther@59906
   134
        let 
walther@59906
   135
          val (fod', rifod') = dropwhile' equal (rev fod) (rev ifod)
Walther@60630
   136
        in (true, fod' @ (map (rev_deriv' ctxt) rifod')) end
walther@59906
   137
      else (false, [])
walther@59906
   138
  end
walther@59906
   139
walther@59934
   140
(** embed a derivation into the Ctree **)
walther@59934
   141
walther@59932
   142
fun embed tacis (pt, pos as (p, Pos.Frm)) =
Walther@60527
   143
  (*inform at Frm: replace the whole PrfObj by a Transitive-ProfObj
walther@59932
   144
    and transfer the istate (from _after_ compare_deriv) from Frm to Res*)
walther@59932
   145
    let
walther@59932
   146
    	val (ist, ctxt) = case Ctree.get_obj Ctree.g_loc pt p of
walther@59932
   147
    	  (SOME (ist, ctxt), _) => (ist, ctxt)
walther@59934
   148
      | (NONE, _) => raise ERROR "Derive.embed Frm: uncovered case get_obj"
Walther@60655
   149
      val (res, asm) = ((State_Steps.result ctxt) o last_elem) tacis
walther@59932
   150
    	val form =  Ctree.get_obj  Ctree.g_form pt p
walther@59932
   151
      (*val p = lev_on p; ---------------only difference to (..,Res) below*)
walther@59932
   152
    	val tacis = (Tactic.Begin_Trans, Tactic.Begin_Trans' form, (pos, (Istate_Def.Uistate, ctxt))) ::
walther@59932
   153
    		(State_Steps.insert_pos ((Pos.lev_on o Pos.lev_dn) p) tacis) @ [(Tactic.End_Trans, Tactic.End_Trans' (res, asm),
walther@59932
   154
    			(Pos.pos_plus (length tacis) (Pos.lev_dn p, Pos.Res), (Ctree.new_val res ist, ctxt)))]
Walther@60586
   155
    	val {rew_rls, ...} = MethodC.from_store ctxt (Ctree.get_obj Ctree.g_metID pt (Ctree.par_pblobj pt p))
walther@59932
   156
    	val (pt, c, pos as (p, _)) = Solve_Step.s_add_general (rev tacis) (pt, [], (p, Pos.Res))
Walther@60586
   157
    	val pt = Ctree.update_tac pt p (Tactic.Derive (Rule_Set.id rew_rls))
walther@59932
   158
    	val pt = Ctree.update_branch pt p Ctree.TransitiveB
walther@59932
   159
    in (c, (pt, pos)) end
walther@59932
   160
  | embed tacis (pt, (p, Pos.Res)) =
Walther@60527
   161
    (*inform at Res: append a Transitive-PrfObj
walther@59932
   162
      and transfer the istate (from _after_ compare_deriv) from Res to new Res*)
walther@59932
   163
    let
walther@59932
   164
    	val (ist, ctxt) = case Ctree.get_obj Ctree.g_loc pt p of
walther@59932
   165
    	  (_, SOME (ist, ctxt)) => (ist, ctxt)
walther@59934
   166
      | (_, NONE) => raise ERROR "Derive.embed Frm: uncovered case get_obj"
Walther@60655
   167
      val (res, asm) = ((State_Steps.result ctxt) o last_elem) tacis
walther@59932
   168
    	val (f, _) = Ctree.get_obj Ctree.g_result pt p
walther@59932
   169
    	val p = Pos.lev_on p(*---------------only difference to (..,Frm) above*);
walther@59932
   170
    	val tacis = (Tactic.Begin_Trans, Tactic.Begin_Trans' f, ((p, Pos.Frm), (Istate_Def.Uistate, ctxt))) ::
walther@59932
   171
    		(State_Steps.insert_pos ((Pos.lev_on o Pos.lev_dn) p) tacis) @ [(Tactic.End_Trans, Tactic.End_Trans' (res, asm), 
walther@59932
   172
    			(Pos.pos_plus (length tacis) (Pos.lev_dn p, Pos.Res), (Ctree.new_val res ist, ctxt)))];
Walther@60586
   173
    	val {rew_rls, ...} = MethodC.from_store ctxt (Ctree.get_obj Ctree.g_metID pt (Ctree.par_pblobj pt p))
walther@59932
   174
    	val (pt, c, pos as (p, _)) = Solve_Step.s_add_general (rev tacis) (pt, [], (p, Pos.Res))
Walther@60586
   175
    	val pt = Ctree.update_tac pt p (Tactic.Derive (Rule_Set.id rew_rls))
walther@59932
   176
    	val pt = Ctree.update_branch pt p Ctree.TransitiveB
walther@59932
   177
    in (c, (pt, pos)) end
walther@59934
   178
  | embed _ _ = raise ERROR "Derive.embed: uncovered definition"
walther@59932
   179
walther@59906
   180
(**)end(**)