| author | wenzelm | 
| Tue, 21 Mar 2006 12:18:20 +0100 | |
| changeset 19309 | 8ea43e9ad83a | 
| parent 19126 | a3cf88213ea5 | 
| child 19466 | 29bc35832a77 | 
| permissions | -rw-r--r-- | 
| 11522 | 1  | 
(* Title: Pure/Proof/proof_rewrite_rules.ML  | 
2  | 
ID: $Id$  | 
|
| 11539 | 3  | 
Author: Stefan Berghofer, TU Muenchen  | 
| 11522 | 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 | 6  | 
*)  | 
7  | 
||
8  | 
signature PROOF_REWRITE_RULES =  | 
|
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  | 
| 17203 | 13  | 
val elim_defs : theory -> bool -> thm list -> Proofterm.proof -> Proofterm.proof  | 
| 13608 | 14  | 
val elim_vars : (typ -> term) -> Proofterm.proof -> Proofterm.proof  | 
| 11522 | 15  | 
end;  | 
16  | 
||
17  | 
structure ProofRewriteRules : PROOF_REWRITE_RULES =  | 
|
18  | 
struct  | 
|
19  | 
||
20  | 
open Proofterm;  | 
|
21  | 
||
| 19309 | 22  | 
fun rew b _ =  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
23  | 
let  | 
| 17137 | 24  | 
fun ?? x = if b then SOME x else NONE;  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
25  | 
fun ax (prf as PAxm (s, prop, _)) Ts =  | 
| 15531 | 26  | 
if b then PAxm (s, prop, SOME Ts) else prf;  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
27  | 
fun ty T = if b then  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
28  | 
let val Type (_, [Type (_, [U, _]), _]) = T  | 
| 15531 | 29  | 
in SOME U end  | 
30  | 
else NONE;  | 
|
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
31  | 
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
 | 
32  | 
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
 | 
33  | 
val symmetric_axm = ax symmetric_axm [propT];  | 
| 11522 | 34  | 
|
| 19309 | 35  | 
    fun rew' (PThm (("ProtoPure.protectD", _), _, _, _) % _ %%
 | 
| 18024 | 36  | 
        (PThm (("ProtoPure.protectI", _), _, _, _) % _ %% prf)) = SOME prf
 | 
| 19309 | 37  | 
      | rew' (PThm (("ProtoPure.conjunctionD1", _), _, _, _) % _ % _ %%
 | 
| 19126 | 38  | 
        (PThm (("ProtoPure.conjunctionI", _), _, _, _) % _ % _ %% prf %% _)) = SOME prf
 | 
| 19309 | 39  | 
      | rew' (PThm (("ProtoPure.conjunctionD2", _), _, _, _) % _ % _ %%
 | 
| 19126 | 40  | 
        (PThm (("ProtoPure.conjunctionI", _), _, _, _) % _ % _ %% _ %% prf)) = SOME prf
 | 
| 19309 | 41  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 15531 | 42  | 
        (PAxm ("ProtoPure.equal_intr", _, _) % _ % _ %% prf %% _)) = SOME prf
 | 
| 19309 | 43  | 
      | rew' (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
44  | 
        (PAxm ("ProtoPure.equal_intr", _, _) % A % B %% prf1 %% prf2)) =
 | 
| 15531 | 45  | 
SOME (equal_intr_axm % B % A %% prf2 %% prf1)  | 
| 12002 | 46  | 
|
| 19309 | 47  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %%
 | 
| 18024 | 48  | 
        (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("prop", _)) %
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
49  | 
          _ % _ % _ %% (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %%
 | 
| 18024 | 50  | 
        ((tg as PThm (("ProtoPure.protectI", _), _, _, _)) % _ %% prf2)) =
 | 
| 15531 | 51  | 
SOME (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2))  | 
| 12002 | 52  | 
|
| 19309 | 53  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
54  | 
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 18024 | 55  | 
          (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("prop", _)) %
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
56  | 
             _ % _ % _ %% (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1)) %%
 | 
| 18024 | 57  | 
        ((tg as PThm (("ProtoPure.protectI", _), _, _, _)) % _ %% prf2)) =
 | 
| 15531 | 58  | 
SOME (tg %> B %% (equal_elim_axm %> A %> B %%  | 
| 17137 | 59  | 
(symmetric_axm % ?? B % ?? A %% prf1) %% prf2))  | 
| 11522 | 60  | 
|
| 19309 | 61  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
 | 
| 11612 | 62  | 
        (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 15531 | 63  | 
          (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==>", _)) % _ % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
64  | 
             (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) =
 | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
65  | 
let  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
66  | 
val _ $ A $ C = Envir.beta_norm X;  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
67  | 
val _ $ B $ D = Envir.beta_norm Y  | 
| 17137 | 68  | 
        in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? B,
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
69  | 
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
 | 
70  | 
(PBound 1 %% (equal_elim_axm %> B %> A %%  | 
| 17137 | 71  | 
(symmetric_axm % ?? A % ?? B %% incr_pboundvars 2 0 prf1) %% PBound 0)))))  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
72  | 
end  | 
| 11522 | 73  | 
|
| 19309 | 74  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
75  | 
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
76  | 
          (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 15531 | 77  | 
            (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==>", _)) % _ % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
78  | 
               (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2))) =
 | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
79  | 
let  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
80  | 
val _ $ A $ C = Envir.beta_norm Y;  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
81  | 
val _ $ B $ D = Envir.beta_norm X  | 
| 17137 | 82  | 
        in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? A,
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
83  | 
equal_elim_axm %> D %> C %%  | 
| 17137 | 84  | 
(symmetric_axm % ?? C % ?? D %% incr_pboundvars 2 0 prf2)  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
85  | 
%% (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
 | 
86  | 
end  | 
| 11522 | 87  | 
|
| 19309 | 88  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
 | 
| 15531 | 89  | 
        (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("all", _)) % _ % _ % _ %%
 | 
| 11612 | 90  | 
          (PAxm ("ProtoPure.reflexive", _, _) % _) %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
91  | 
            (PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf))) =
 | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
92  | 
let  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
93  | 
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
 | 
94  | 
val _ $ Q = Envir.beta_norm Y;  | 
| 17137 | 95  | 
        in SOME (AbsP ("H", ?? X, Abst ("x", ty T,
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
96  | 
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
 | 
97  | 
(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
 | 
98  | 
end  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
99  | 
|
| 19309 | 100  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
101  | 
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%        
 | 
| 15531 | 102  | 
          (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("all", _)) % _ % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
103  | 
            (PAxm ("ProtoPure.reflexive", _, _) % _) %%
 | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
104  | 
              (PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf)))) =
 | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
105  | 
let  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
106  | 
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
 | 
107  | 
val _ $ Q = Envir.beta_norm Y;  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
108  | 
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
 | 
109  | 
val u = incr_boundvars 1 Q $ Bound 0  | 
| 17137 | 110  | 
        in SOME (AbsP ("H", ?? X, Abst ("x", ty T,
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
111  | 
equal_elim_axm %> t %> u %%  | 
| 17137 | 112  | 
(symmetric_axm % ?? u % ?? t %% (incr_pboundvars 1 1 prf %> Bound 0))  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
113  | 
%% (PBound 0 %> Bound 0))))  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
114  | 
end  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
115  | 
|
| 19309 | 116  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME A % SOME C %%
 | 
| 15531 | 117  | 
        (PAxm ("ProtoPure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2) %% prf3) =
 | 
118  | 
SOME (equal_elim_axm %> B %> C %% prf2 %%  | 
|
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
119  | 
(equal_elim_axm %> A %> B %% prf1 %% prf3))  | 
| 19309 | 120  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME A % SOME C %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
121  | 
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 15531 | 122  | 
          (PAxm ("ProtoPure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2)) %% prf3) =
 | 
| 17137 | 123  | 
SOME (equal_elim_axm %> B %> C %% (symmetric_axm % ?? C % ?? B %% prf1) %%  | 
124  | 
(equal_elim_axm %> A %> B %% (symmetric_axm % ?? B % ?? A %% prf2) %% prf3))  | 
|
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
125  | 
|
| 19309 | 126  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 15531 | 127  | 
        (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf) = SOME prf
 | 
| 19309 | 128  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
129  | 
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 15531 | 130  | 
          (PAxm ("ProtoPure.reflexive", _, _) % _)) %% prf) = SOME prf
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
131  | 
|
| 19309 | 132  | 
      | rew' (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 15531 | 133  | 
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %% prf)) = SOME prf
 | 
| 11522 | 134  | 
|
| 19309 | 135  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 15531 | 136  | 
        (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
137  | 
          (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 15531 | 138  | 
            (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
139  | 
              (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3) %% prf4) =
 | 
| 15531 | 140  | 
SOME (equal_elim_axm %> C %> D %% prf2 %%  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
141  | 
(equal_elim_axm %> A %> C %% prf3 %%  | 
| 17137 | 142  | 
(equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %% prf4)))  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
143  | 
|
| 19309 | 144  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
145  | 
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 15531 | 146  | 
          (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
147  | 
            (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 15531 | 148  | 
              (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
149  | 
                (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3)) %% prf4) =
 | 
| 15531 | 150  | 
SOME (equal_elim_axm %> A %> B %% prf1 %%  | 
| 17137 | 151  | 
(equal_elim_axm %> C %> A %% (symmetric_axm % ?? A % ?? C %% prf3) %%  | 
152  | 
(equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %% prf4)))  | 
|
| 11522 | 153  | 
|
| 19309 | 154  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 15531 | 155  | 
        (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
156  | 
          (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
157  | 
            (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 15531 | 158  | 
              (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
159  | 
                (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3) %% prf4) =
 | 
| 17137 | 160  | 
SOME (equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %%  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
161  | 
(equal_elim_axm %> B %> D %% prf3 %%  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
162  | 
(equal_elim_axm %> A %> B %% prf1 %% prf4)))  | 
| 11522 | 163  | 
|
| 19309 | 164  | 
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
165  | 
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 15531 | 166  | 
          (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
167  | 
            (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
168  | 
              (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 15531 | 169  | 
                (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
170  | 
                  (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3)) %% prf4) =
 | 
| 17137 | 171  | 
SOME (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %%  | 
172  | 
(equal_elim_axm %> D %> B %% (symmetric_axm % ?? B % ?? D %% prf3) %%  | 
|
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
173  | 
(equal_elim_axm %> C %> D %% prf2 %% prf4)))  | 
| 11522 | 174  | 
|
| 19309 | 175  | 
      | rew' ((prf as PAxm ("ProtoPure.combination", _, _) %
 | 
| 15531 | 176  | 
        SOME ((eq as Const ("==", T)) $ t) % _ % _ % _) %%
 | 
| 13257 | 177  | 
          (PAxm ("ProtoPure.reflexive", _, _) % _)) =
 | 
178  | 
let val (U, V) = (case T of  | 
|
179  | 
Type (_, [U, V]) => (U, V) | _ => (dummyT, dummyT))  | 
|
| 17137 | 180  | 
in SOME (prf %% (ax combination_axm [V, U] %> eq % ?? eq % ?? t % ?? t %%  | 
181  | 
(ax reflexive_axm [T] % ?? eq) %% (ax reflexive_axm [U] % ?? t)))  | 
|
| 13257 | 182  | 
end  | 
183  | 
||
| 19309 | 184  | 
| rew' _ = NONE;  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
185  | 
in rew' end;  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
186  | 
|
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
187  | 
fun rprocs b = [("Pure/meta_equality", rew b)];
 | 
| 18708 | 188  | 
val _ = Context.add_setup (Proofterm.add_prf_rprocs (rprocs false));  | 
| 11522 | 189  | 
|
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
190  | 
|
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
191  | 
(**** apply rewriting function to all terms in proof ****)  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
192  | 
|
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
193  | 
fun rewrite_terms r =  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
194  | 
let  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
195  | 
fun rew_term Ts t =  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
196  | 
let  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
197  | 
val frees = map Free (variantlist  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
198  | 
(replicate (length Ts) "x", add_term_names (t, [])) ~~ Ts);  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
199  | 
val t' = r (subst_bounds (frees, t));  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
200  | 
fun strip [] t = t  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
201  | 
| strip (_ :: xs) (Abs (_, _, t)) = strip xs t;  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
202  | 
in  | 
| 15570 | 203  | 
strip Ts (Library.foldl (uncurry lambda o Library.swap) (t', frees))  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
204  | 
end;  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
205  | 
|
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
206  | 
fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2  | 
| 15531 | 207  | 
| rew Ts (prf % SOME t) = rew Ts prf % SOME (rew_term Ts t)  | 
208  | 
| rew Ts (Abst (s, SOME T, prf)) = Abst (s, SOME T, rew (T :: Ts) prf)  | 
|
209  | 
| rew Ts (AbsP (s, SOME t, prf)) = AbsP (s, SOME (rew_term Ts t), rew Ts prf)  | 
|
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
210  | 
| rew _ prf = prf  | 
| 
 
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  | 
in rew [] end;  | 
| 
 
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  | 
|
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
215  | 
(**** eliminate definitions in proof ****)  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
216  | 
|
| 16861 | 217  | 
fun vars_of t = rev (fold_aterms (fn v as Var _ => insert (op =) v | _ => I) t []);  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
218  | 
|
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
219  | 
fun insert_refl defs Ts (prf1 %% prf2) =  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
220  | 
insert_refl defs Ts prf1 %% insert_refl defs Ts prf2  | 
| 15531 | 221  | 
| insert_refl defs Ts (Abst (s, SOME T, prf)) =  | 
222  | 
Abst (s, SOME T, insert_refl defs (T :: Ts) prf)  | 
|
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
223  | 
| insert_refl defs Ts (AbsP (s, t, prf)) =  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
224  | 
AbsP (s, t, insert_refl defs Ts prf)  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
225  | 
| insert_refl defs Ts prf = (case strip_combt prf of  | 
| 15531 | 226  | 
(PThm ((s, _), _, prop, SOME Ts), ts) =>  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
227  | 
if s mem defs then  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
228  | 
let  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
229  | 
val vs = vars_of prop;  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
230  | 
val tvars = term_tvars prop;  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
231  | 
val (_, rhs) = Logic.dest_equals prop;  | 
| 18185 | 232  | 
val rhs' = Term.betapplys (subst_TVars (map fst tvars ~~ Ts)  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
233  | 
                (foldr (fn p => Abs ("", dummyT, abstract_over p)) rhs vs),
 | 
| 15570 | 234  | 
map valOf ts);  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
235  | 
in  | 
| 15531 | 236  | 
change_type (SOME [fastype_of1 (Ts, rhs')]) reflexive_axm %> rhs'  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
237  | 
end  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
238  | 
else prf  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
239  | 
| (_, []) => prf  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
240  | 
| (prf', ts) => proof_combt' (insert_refl defs Ts prf', ts));  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
241  | 
|
| 17203 | 242  | 
fun elim_defs thy r defs prf =  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
243  | 
let  | 
| 13341 | 244  | 
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
 | 
245  | 
val defnames = map Thm.name_of_thm defs;  | 
| 13341 | 246  | 
val f = if not r then I else  | 
247  | 
let  | 
|
248  | 
val cnames = map (fst o dest_Const o fst) defs';  | 
|
| 15570 | 249  | 
val thms = List.concat (map (fn (s, ps) =>  | 
| 13341 | 250  | 
if s mem defnames then []  | 
| 15531 | 251  | 
else map (pair s o SOME o fst) (filter_out (fn (p, _) =>  | 
| 13646 | 252  | 
null (term_consts p inter cnames)) ps))  | 
| 17018 | 253  | 
(Symtab.dest (thms_of_proof prf Symtab.empty)))  | 
| 17203 | 254  | 
in Reconstruct.expand_proof thy thms end  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
255  | 
in  | 
| 17203 | 256  | 
rewrite_terms (Pattern.rewrite_term thy defs' [])  | 
| 13341 | 257  | 
(insert_refl defnames [] (f prf))  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
258  | 
end;  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
259  | 
|
| 13608 | 260  | 
|
261  | 
(**** eliminate all variables that don't occur in the proposition ****)  | 
|
262  | 
||
263  | 
fun elim_vars mk_default prf =  | 
|
264  | 
let  | 
|
265  | 
val prop = Reconstruct.prop_of prf;  | 
|
| 19309 | 266  | 
val tv = Term.add_vars prop [];  | 
267  | 
val tf = Term.add_frees prop [];  | 
|
268  | 
||
269  | 
fun hidden_variable (Var v) = not (member (op =) tv v)  | 
|
270  | 
| hidden_variable (Free f) = not (member (op =) tf f)  | 
|
271  | 
| hidden_variable _ = false;  | 
|
| 13917 | 272  | 
|
273  | 
fun mk_default' T = list_abs  | 
|
274  | 
(apfst (map (pair "x")) (apsnd mk_default (strip_type T)));  | 
|
275  | 
||
276  | 
fun elim_varst (t $ u) = elim_varst t $ elim_varst u  | 
|
277  | 
| elim_varst (Abs (s, T, t)) = Abs (s, T, elim_varst t)  | 
|
| 19309 | 278  | 
| elim_varst (t as Free (x, T)) = if member (op =) tf (x, T) then t else mk_default' T  | 
279  | 
| elim_varst (t as Var (xi, T)) = if member (op =) tv (xi, T) then t else mk_default' T  | 
|
280  | 
| elim_varst t = t;  | 
|
| 13608 | 281  | 
in  | 
| 19309 | 282  | 
map_proof_terms (fn t =>  | 
283  | 
if Term.exists_subterm hidden_variable t then Envir.beta_norm (elim_varst t) else t) I prf  | 
|
| 13608 | 284  | 
end;  | 
285  | 
||
| 11522 | 286  | 
end;  |