src/Pure/Proof/proof_rewrite_rules.ML
author kleing
Mon, 21 Jun 2004 10:25:57 +0200
changeset 14981 e73f8140af78
parent 13917 a67c9e6570ac
child 15531 08c8dad8e399
permissions -rw-r--r--
Merged in license change from Isabelle2004
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     1
(*  Title:      Pure/Proof/proof_rewrite_rules.ML
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
11539
0f17da240450 tuned headers;
wenzelm
parents: 11522
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     4
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
     5
Simplification functions for proof terms involving meta level rules.
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     6
*)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     7
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     8
signature PROOF_REWRITE_RULES =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     9
sig
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    10
  val rew : bool -> typ list -> Proofterm.proof -> Proofterm.proof option
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    11
  val rprocs : bool -> (string * (typ list -> Proofterm.proof -> Proofterm.proof option)) list
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
    12
  val rewrite_terms : (term -> term) -> Proofterm.proof -> Proofterm.proof
13341
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
    13
  val elim_defs : Sign.sg -> bool -> thm list -> Proofterm.proof -> Proofterm.proof
13608
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
    14
  val elim_vars : (typ -> term) -> Proofterm.proof -> Proofterm.proof
12237
39aeccee9e1c Added setup.
berghofe
parents: 12002
diff changeset
    15
  val setup : (theory -> theory) list
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    16
end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    17
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    18
structure ProofRewriteRules : PROOF_REWRITE_RULES =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    19
struct
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    20
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    21
open Proofterm;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    22
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    23
fun rew b =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    24
  let
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    25
    fun ? x = if b then Some x else None;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    26
    fun ax (prf as PAxm (s, prop, _)) Ts =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    27
      if b then PAxm (s, prop, Some Ts) else prf;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    28
    fun ty T = if b then
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    29
        let val Type (_, [Type (_, [U, _]), _]) = T
13257
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
    30
        in Some U end
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    31
      else None;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    32
    val equal_intr_axm = ax equal_intr_axm [];
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    33
    val equal_elim_axm = ax equal_elim_axm [];
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    34
    val symmetric_axm = ax symmetric_axm [propT];
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    35
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    36
    fun rew' _ (PThm (("ProtoPure.rev_triv_goal", _), _, _, _) % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    37
        (PThm (("ProtoPure.triv_goal", _), _, _, _) % _ %% prf)) = Some prf
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    38
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    39
        (PAxm ("ProtoPure.equal_intr", _, _) % _ % _ %% prf %% _)) = Some prf
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    40
      | rew' _ (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    41
        (PAxm ("ProtoPure.equal_intr", _, _) % A % B %% prf1 %% prf2)) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    42
            Some (equal_intr_axm % B % A %% prf2 %% prf1)
12002
bc9b5bad0e7b Additional rules for simplifying inside "Goal"
berghofe
parents: 11612
diff changeset
    43
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    44
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A) % Some (_ $ B) %%
12002
bc9b5bad0e7b Additional rules for simplifying inside "Goal"
berghofe
parents: 11612
diff changeset
    45
        (PAxm ("ProtoPure.combination", _, _) % Some (Const ("Goal", _)) %
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    46
          _ % _ % _ %% (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    47
        ((tg as PThm (("ProtoPure.triv_goal", _), _, _, _)) % _ %% prf2)) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    48
        Some (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2))
12002
bc9b5bad0e7b Additional rules for simplifying inside "Goal"
berghofe
parents: 11612
diff changeset
    49
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    50
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A) % Some (_ $ B) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    51
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    52
          (PAxm ("ProtoPure.combination", _, _) % Some (Const ("Goal", _)) %
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    53
             _ % _ % _ %% (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1)) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    54
        ((tg as PThm (("ProtoPure.triv_goal", _), _, _, _)) % _ %% prf2)) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    55
        Some (tg %> B %% (equal_elim_axm %> A %> B %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    56
          (symmetric_axm % ? B % ? A %% prf1) %% prf2))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    57
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    58
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
11612
ae8450657bf0 Exchanged % and %%.
berghofe
parents: 11539
diff changeset
    59
        (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
ae8450657bf0 Exchanged % and %%.
berghofe
parents: 11539
diff changeset
    60
          (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==>", _)) % _ % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    61
             (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    62
        let
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    63
          val _ $ A $ C = Envir.beta_norm X;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    64
          val _ $ B $ D = Envir.beta_norm Y
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    65
        in Some (AbsP ("H1", ? X, AbsP ("H2", ? B,
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    66
          equal_elim_axm %> C %> D %% incr_pboundvars 2 0 prf2 %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    67
            (PBound 1 %% (equal_elim_axm %> B %> A %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    68
              (symmetric_axm % ? A % ? B %% incr_pboundvars 2 0 prf1) %% PBound 0)))))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    69
        end
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    70
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    71
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    72
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    73
          (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    74
            (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==>", _)) % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    75
               (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2))) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    76
        let
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    77
          val _ $ A $ C = Envir.beta_norm Y;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    78
          val _ $ B $ D = Envir.beta_norm X
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    79
        in Some (AbsP ("H1", ? X, AbsP ("H2", ? A,
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    80
          equal_elim_axm %> D %> C %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    81
            (symmetric_axm % ? C % ? D %% incr_pboundvars 2 0 prf2)
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    82
              %% (PBound 1 %% (equal_elim_axm %> A %> B %% incr_pboundvars 2 0 prf1 %% PBound 0)))))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    83
        end
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    84
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    85
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
11612
ae8450657bf0 Exchanged % and %%.
berghofe
parents: 11539
diff changeset
    86
        (PAxm ("ProtoPure.combination", _, _) % Some (Const ("all", _)) % _ % _ % _ %%
ae8450657bf0 Exchanged % and %%.
berghofe
parents: 11539
diff changeset
    87
          (PAxm ("ProtoPure.reflexive", _, _) % _) %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    88
            (PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf))) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    89
        let
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    90
          val Const (_, T) $ P = Envir.beta_norm X;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    91
          val _ $ Q = Envir.beta_norm Y;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    92
        in Some (AbsP ("H", ? X, Abst ("x", ty T,
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    93
            equal_elim_axm %> incr_boundvars 1 P $ Bound 0 %> incr_boundvars 1 Q $ Bound 0 %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    94
              (incr_pboundvars 1 1 prf %> Bound 0) %% (PBound 0 %> Bound 0))))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    95
        end
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    96
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    97
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    98
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%        
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    99
          (PAxm ("ProtoPure.combination", _, _) % Some (Const ("all", _)) % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   100
            (PAxm ("ProtoPure.reflexive", _, _) % _) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   101
              (PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf)))) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   102
        let
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   103
          val Const (_, T) $ P = Envir.beta_norm X;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   104
          val _ $ Q = Envir.beta_norm Y;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   105
          val t = incr_boundvars 1 P $ Bound 0;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   106
          val u = incr_boundvars 1 Q $ Bound 0
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   107
        in Some (AbsP ("H", ? X, Abst ("x", ty T,
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   108
          equal_elim_axm %> t %> u %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   109
            (symmetric_axm % ? u % ? t %% (incr_pboundvars 1 1 prf %> Bound 0))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   110
              %% (PBound 0 %> Bound 0))))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   111
        end
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   112
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   113
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some A % Some C %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   114
        (PAxm ("ProtoPure.transitive", _, _) % _ % Some B % _ %% prf1 %% prf2) %% prf3) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   115
           Some (equal_elim_axm %> B %> C %% prf2 %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   116
             (equal_elim_axm %> A %> B %% prf1 %% prf3))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   117
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some A % Some C %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   118
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   119
          (PAxm ("ProtoPure.transitive", _, _) % _ % Some B % _ %% prf1 %% prf2)) %% prf3) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   120
           Some (equal_elim_axm %> B %> C %% (symmetric_axm % ? C % ? B %% prf1) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   121
             (equal_elim_axm %> A %> B %% (symmetric_axm % ? B % ? A %% prf2) %% prf3))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   122
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   123
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   124
        (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf) = Some prf
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   125
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   126
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   127
          (PAxm ("ProtoPure.reflexive", _, _) % _)) %% prf) = Some prf
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   128
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   129
      | rew' _ (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   130
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %% prf)) = Some prf
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   131
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   132
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   133
        (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A $ C) % Some (_ $ B $ D) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   134
          (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   135
            (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   136
              (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3) %% prf4) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   137
          Some (equal_elim_axm %> C %> D %% prf2 %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   138
            (equal_elim_axm %> A %> C %% prf3 %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   139
              (equal_elim_axm %> B %> A %% (symmetric_axm % ? A % ? B %% prf1) %% prf4)))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   140
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   141
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   142
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   143
          (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A $ C) % Some (_ $ B $ D) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   144
            (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   145
              (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   146
                (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3)) %% prf4) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   147
          Some (equal_elim_axm %> A %> B %% prf1 %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   148
            (equal_elim_axm %> C %> A %% (symmetric_axm % ? A % ? C %% prf3) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   149
              (equal_elim_axm %> D %> C %% (symmetric_axm % ? C % ? D %% prf2) %% prf4)))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   150
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   151
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   152
        (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ B $ D) % Some (_ $ A $ C) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   153
          (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   154
            (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   155
              (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   156
                (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3) %% prf4) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   157
          Some (equal_elim_axm %> D %> C %% (symmetric_axm % ? C % ? D %% prf2) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   158
            (equal_elim_axm %> B %> D %% prf3 %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   159
              (equal_elim_axm %> A %> B %% prf1 %% prf4)))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   160
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   161
      | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   162
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   163
          (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ B $ D) % Some (_ $ A $ C) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   164
            (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   165
              (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   166
                (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   167
                  (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3)) %% prf4) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   168
          Some (equal_elim_axm %> B %> A %% (symmetric_axm % ? A % ? B %% prf1) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   169
            (equal_elim_axm %> D %> B %% (symmetric_axm % ? B % ? D %% prf3) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   170
              (equal_elim_axm %> C %> D %% prf2 %% prf4)))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   171
13257
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   172
      | rew' _ ((prf as PAxm ("ProtoPure.combination", _, _) %
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   173
        Some ((eq as Const ("==", T)) $ t) % _ % _ % _) %%
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   174
          (PAxm ("ProtoPure.reflexive", _, _) % _)) =
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   175
        let val (U, V) = (case T of
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   176
          Type (_, [U, V]) => (U, V) | _ => (dummyT, dummyT))
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   177
        in Some (prf %% (ax combination_axm [V, U] %> eq % ? eq % ? t % ? t %%
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   178
          (ax reflexive_axm [T] % ? eq) %% (ax reflexive_axm [U] % ? t)))
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   179
        end
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   180
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   181
      | rew' _ _ = None;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   182
  in rew' end;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   183
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   184
fun rprocs b = [("Pure/meta_equality", rew b)];
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   185
val setup = [Proofterm.add_prf_rprocs (rprocs false)];
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   186
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   187
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   188
(**** apply rewriting function to all terms in proof ****)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   189
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   190
fun rewrite_terms r =
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   191
  let
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   192
    fun rew_term Ts t =
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   193
      let
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   194
        val frees = map Free (variantlist
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   195
          (replicate (length Ts) "x", add_term_names (t, [])) ~~ Ts);
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   196
        val t' = r (subst_bounds (frees, t));
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   197
        fun strip [] t = t
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   198
          | strip (_ :: xs) (Abs (_, _, t)) = strip xs t;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   199
      in
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   200
        strip Ts (foldl (uncurry lambda o Library.swap) (t', frees))
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   201
      end;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   202
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   203
    fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   204
      | rew Ts (prf % Some t) = rew Ts prf % Some (rew_term Ts t)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   205
      | rew Ts (Abst (s, Some T, prf)) = Abst (s, Some T, rew (T :: Ts) prf)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   206
      | rew Ts (AbsP (s, Some t, prf)) = AbsP (s, Some (rew_term Ts t), rew Ts prf)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   207
      | rew _ prf = prf
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   208
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   209
  in rew [] end;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   210
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   211
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   212
(**** eliminate definitions in proof ****)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   213
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   214
fun vars_of t = rev (foldl_aterms
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   215
  (fn (vs, v as Var _) => v ins vs | (vs, _) => vs) ([], t));
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   216
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   217
fun insert_refl defs Ts (prf1 %% prf2) =
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   218
      insert_refl defs Ts prf1 %% insert_refl defs Ts prf2
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   219
  | insert_refl defs Ts (Abst (s, Some T, prf)) =
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   220
      Abst (s, Some T, insert_refl defs (T :: Ts) prf)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   221
  | insert_refl defs Ts (AbsP (s, t, prf)) =
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   222
      AbsP (s, t, insert_refl defs Ts prf)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   223
  | insert_refl defs Ts prf = (case strip_combt prf of
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   224
        (PThm ((s, _), _, prop, Some Ts), ts) =>
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   225
          if s mem defs then
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   226
            let
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   227
              val vs = vars_of prop;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   228
              val tvars = term_tvars prop;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   229
              val (_, rhs) = Logic.dest_equals prop;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   230
              val rhs' = foldl betapply (subst_TVars (map fst tvars ~~ Ts)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   231
                (foldr (fn p => Abs ("", dummyT, abstract_over p)) (vs, rhs)),
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   232
                map the ts);
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   233
            in
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   234
              change_type (Some [fastype_of1 (Ts, rhs')]) reflexive_axm %> rhs'
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   235
            end
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   236
          else prf
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   237
      | (_, []) => prf
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   238
      | (prf', ts) => proof_combt' (insert_refl defs Ts prf', ts));
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   239
13341
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   240
fun elim_defs sign r defs prf =
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   241
  let
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   242
    val tsig = Sign.tsig_of sign;
13341
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   243
    val defs' = map (Logic.dest_equals o prop_of o Drule.abs_def) defs
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   244
    val defnames = map Thm.name_of_thm defs;
13341
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   245
    val f = if not r then I else
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   246
      let
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   247
        val cnames = map (fst o dest_Const o fst) defs';
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   248
        val thms = flat (map (fn (s, ps) =>
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   249
            if s mem defnames then []
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   250
            else map (pair s o Some o fst) (filter_out (fn (p, _) =>
13646
46ed3d042ba5 Ported find_intro/elim to Isar.
nipkow
parents: 13608
diff changeset
   251
              null (term_consts p inter cnames)) ps))
13341
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   252
          (Symtab.dest (thms_of_proof Symtab.empty prf)))
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   253
      in Reconstruct.expand_proof sign thms end
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   254
  in
13341
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   255
    rewrite_terms (Pattern.rewrite_term tsig defs' [])
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   256
      (insert_refl defnames [] (f prf))
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   257
  end;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   258
13608
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   259
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   260
(**** eliminate all variables that don't occur in the proposition ****)
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   261
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   262
fun elim_vars mk_default prf =
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   263
  let
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   264
    val prop = Reconstruct.prop_of prf;
13917
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   265
    val tv = term_vars prop;
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   266
    val tf = term_frees prop;
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   267
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   268
    fun mk_default' T = list_abs
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   269
      (apfst (map (pair "x")) (apsnd mk_default (strip_type T)));
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   270
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   271
    fun elim_varst (t $ u) = elim_varst t $ elim_varst u
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   272
      | elim_varst (Abs (s, T, t)) = Abs (s, T, elim_varst t)
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   273
      | elim_varst (f as Free (_, T)) = if f mem tf then f else mk_default' T
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   274
      | elim_varst (v as Var (_, T)) = if v mem tv then v else mk_default' T
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   275
      | elim_varst t = t
13608
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   276
  in
13917
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   277
    map_proof_terms (fn t => if not (null (term_vars t \\ tv)) orelse
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   278
        not (null (term_frees t \\ tf)) then Envir.beta_norm (elim_varst t)
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   279
      else t) I prf
13608
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   280
  end;
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   281
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   282
end;