| author | wenzelm | 
| Wed, 14 Sep 2016 12:51:40 +0200 | |
| changeset 63866 | 630eaf8fe9f3 | 
| parent 62922 | 96691631c1eb | 
| child 70412 | 64ead6de6212 | 
| permissions | -rw-r--r-- | 
| 11522 | 1  | 
(* Title: Pure/Proof/proof_rewrite_rules.ML  | 
| 11539 | 2  | 
Author: Stefan Berghofer, TU Muenchen  | 
| 11522 | 3  | 
|
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
4  | 
Simplification functions for proof terms involving meta level rules.  | 
| 11522 | 5  | 
*)  | 
6  | 
||
7  | 
signature PROOF_REWRITE_RULES =  | 
|
8  | 
sig  | 
|
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
9  | 
val rew : bool -> typ list -> term option list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option  | 
| 
33722
 
e588744f14da
generalized procs for rewrite_proof: allow skeleton;
 
wenzelm 
parents: 
29271 
diff
changeset
 | 
10  | 
val rprocs : bool ->  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
11  | 
(typ list -> term option list -> Proofterm.proof -> (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  | 
| 62922 | 13  | 
val elim_defs : Proof.context -> bool -> thm list -> Proofterm.proof -> Proofterm.proof  | 
| 13608 | 14  | 
val elim_vars : (typ -> term) -> Proofterm.proof -> Proofterm.proof  | 
| 22280 | 15  | 
val hhf_proof : term -> term -> Proofterm.proof -> Proofterm.proof  | 
16  | 
val un_hhf_proof : term -> term -> Proofterm.proof -> Proofterm.proof  | 
|
| 62922 | 17  | 
val mk_of_sort_proof : Proof.context -> term option list -> typ * sort -> Proofterm.proof list  | 
18  | 
val expand_of_class : Proof.context -> typ list -> term option list -> Proofterm.proof ->  | 
|
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
19  | 
(Proofterm.proof * Proofterm.proof) option  | 
| 11522 | 20  | 
end;  | 
21  | 
||
22  | 
structure ProofRewriteRules : PROOF_REWRITE_RULES =  | 
|
23  | 
struct  | 
|
24  | 
||
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
25  | 
fun rew b _ _ =  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
26  | 
let  | 
| 17137 | 27  | 
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
 | 
28  | 
fun ax (prf as PAxm (s, prop, _)) Ts =  | 
| 15531 | 29  | 
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
 | 
30  | 
fun ty T = if b then  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
31  | 
let val Type (_, [Type (_, [U, _]), _]) = T  | 
| 15531 | 32  | 
in SOME U end  | 
33  | 
else NONE;  | 
|
| 37310 | 34  | 
val equal_intr_axm = ax Proofterm.equal_intr_axm [];  | 
35  | 
val equal_elim_axm = ax Proofterm.equal_elim_axm [];  | 
|
36  | 
val symmetric_axm = ax Proofterm.symmetric_axm [propT];  | 
|
| 11522 | 37  | 
|
| 
28806
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
38  | 
    fun rew' (PThm (_, (("Pure.protectD", _, _), _)) % _ %%
 | 
| 
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
39  | 
        (PThm (_, (("Pure.protectI", _, _), _)) % _ %% prf)) = SOME prf
 | 
| 
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
40  | 
      | rew' (PThm (_, (("Pure.conjunctionD1", _, _), _)) % _ % _ %%
 | 
| 
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
41  | 
        (PThm (_, (("Pure.conjunctionI", _, _), _)) % _ % _ %% prf %% _)) = SOME prf
 | 
| 
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
42  | 
      | rew' (PThm (_, (("Pure.conjunctionD2", _, _), _)) % _ % _ %%
 | 
| 
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
43  | 
        (PThm (_, (("Pure.conjunctionI", _, _), _)) % _ % _ %% _ %% prf)) = SOME prf
 | 
| 26424 | 44  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %%
 | 
45  | 
        (PAxm ("Pure.equal_intr", _, _) % _ % _ %% prf %% _)) = SOME prf
 | 
|
46  | 
      | rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %%
 | 
|
47  | 
        (PAxm ("Pure.equal_intr", _, _) % A % B %% prf1 %% prf2)) =
 | 
|
| 15531 | 48  | 
SOME (equal_intr_axm % B % A %% prf2 %% prf1)  | 
| 12002 | 49  | 
|
| 26424 | 50  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %%
 | 
| 56244 | 51  | 
        (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) %
 | 
| 26424 | 52  | 
          _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %%
 | 
| 
28806
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
53  | 
        ((tg as PThm (_, (("Pure.protectI", _, _), _))) % _ %% prf2)) =
 | 
| 15531 | 54  | 
SOME (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2))  | 
| 12002 | 55  | 
|
| 26424 | 56  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %%
 | 
57  | 
        (PAxm ("Pure.symmetric", _, _) % _ % _ %%
 | 
|
| 56244 | 58  | 
          (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) %
 | 
| 26424 | 59  | 
             _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1)) %%
 | 
| 
28806
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
60  | 
        ((tg as PThm (_, (("Pure.protectI", _, _), _))) % _ %% prf2)) =
 | 
| 15531 | 61  | 
SOME (tg %> B %% (equal_elim_axm %> A %> B %%  | 
| 17137 | 62  | 
(symmetric_axm % ?? B % ?? A %% prf1) %% prf2))  | 
| 11522 | 63  | 
|
| 26424 | 64  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %%
 | 
65  | 
        (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %%
 | 
|
| 56245 | 66  | 
          (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.imp", _)) % _ % _ % _ %%
 | 
| 26424 | 67  | 
             (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) =
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
68  | 
let  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
69  | 
val _ $ A $ C = Envir.beta_norm X;  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
70  | 
val _ $ B $ D = Envir.beta_norm Y  | 
| 17137 | 71  | 
        in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? B,
 | 
| 37310 | 72  | 
Proofterm.equal_elim_axm %> C %> D %% Proofterm.incr_pboundvars 2 0 prf2 %%  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
73  | 
(PBound 1 %% (equal_elim_axm %> B %> A %%  | 
| 37310 | 74  | 
(Proofterm.symmetric_axm % ?? A % ?? B %% Proofterm.incr_pboundvars 2 0 prf1) %%  | 
75  | 
PBound 0)))))  | 
|
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
76  | 
end  | 
| 11522 | 77  | 
|
| 26424 | 78  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %%
 | 
79  | 
        (PAxm ("Pure.symmetric", _, _) % _ % _ %%
 | 
|
80  | 
          (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %%
 | 
|
| 56245 | 81  | 
            (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.imp", _)) % _ % _ % _ %%
 | 
| 26424 | 82  | 
               (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2))) =
 | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
83  | 
let  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
84  | 
val _ $ A $ C = Envir.beta_norm Y;  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
85  | 
val _ $ B $ D = Envir.beta_norm X  | 
| 17137 | 86  | 
        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
 | 
87  | 
equal_elim_axm %> D %> C %%  | 
| 37310 | 88  | 
(symmetric_axm % ?? C % ?? D %% Proofterm.incr_pboundvars 2 0 prf2) %%  | 
89  | 
(PBound 1 %%  | 
|
90  | 
(equal_elim_axm %> A %> B %% Proofterm.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
 | 
91  | 
end  | 
| 11522 | 92  | 
|
| 26424 | 93  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %%
 | 
| 56245 | 94  | 
        (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.all", _)) % _ % _ % _ %%
 | 
| 26424 | 95  | 
          (PAxm ("Pure.reflexive", _, _) % _) %%
 | 
96  | 
            (PAxm ("Pure.abstract_rule", _, _) % _ % _ %% prf))) =
 | 
|
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
97  | 
let  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
98  | 
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
 | 
99  | 
val _ $ Q = Envir.beta_norm Y;  | 
| 17137 | 100  | 
        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
 | 
101  | 
equal_elim_axm %> incr_boundvars 1 P $ Bound 0 %> incr_boundvars 1 Q $ Bound 0 %%  | 
| 37310 | 102  | 
(Proofterm.incr_pboundvars 1 1 prf %> Bound 0) %% (PBound 0 %> Bound 0))))  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
103  | 
end  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
104  | 
|
| 26424 | 105  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %%
 | 
106  | 
        (PAxm ("Pure.symmetric", _, _) % _ % _ %%        
 | 
|
| 56245 | 107  | 
          (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.all", _)) % _ % _ % _ %%
 | 
| 26424 | 108  | 
            (PAxm ("Pure.reflexive", _, _) % _) %%
 | 
109  | 
              (PAxm ("Pure.abstract_rule", _, _) % _ % _ %% prf)))) =
 | 
|
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
110  | 
let  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
111  | 
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
 | 
112  | 
val _ $ Q = Envir.beta_norm Y;  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
113  | 
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
 | 
114  | 
val u = incr_boundvars 1 Q $ Bound 0  | 
| 17137 | 115  | 
        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
 | 
116  | 
equal_elim_axm %> t %> u %%  | 
| 37310 | 117  | 
(symmetric_axm % ?? u % ?? t %% (Proofterm.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
 | 
118  | 
%% (PBound 0 %> Bound 0))))  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
119  | 
end  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
120  | 
|
| 26424 | 121  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % SOME A % SOME C %%
 | 
122  | 
        (PAxm ("Pure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2) %% prf3) =
 | 
|
| 15531 | 123  | 
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
 | 
124  | 
(equal_elim_axm %> A %> B %% prf1 %% prf3))  | 
| 26424 | 125  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % SOME A % SOME C %%
 | 
126  | 
        (PAxm ("Pure.symmetric", _, _) % _ % _ %%
 | 
|
127  | 
          (PAxm ("Pure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2)) %% prf3) =
 | 
|
| 17137 | 128  | 
SOME (equal_elim_axm %> B %> C %% (symmetric_axm % ?? C % ?? B %% prf1) %%  | 
129  | 
(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
 | 
130  | 
|
| 26424 | 131  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %%
 | 
132  | 
        (PAxm ("Pure.reflexive", _, _) % _) %% prf) = SOME prf
 | 
|
133  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %%
 | 
|
134  | 
        (PAxm ("Pure.symmetric", _, _) % _ % _ %%
 | 
|
135  | 
          (PAxm ("Pure.reflexive", _, _) % _)) %% prf) = SOME prf
 | 
|
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
136  | 
|
| 26424 | 137  | 
      | rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %%
 | 
138  | 
        (PAxm ("Pure.symmetric", _, _) % _ % _ %% prf)) = SOME prf
 | 
|
| 11522 | 139  | 
|
| 26424 | 140  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %%
 | 
141  | 
        (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %%
 | 
|
142  | 
          (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %%
 | 
|
| 56245 | 143  | 
            (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %%
 | 
| 26424 | 144  | 
              (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3) %% prf4) =
 | 
| 15531 | 145  | 
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
 | 
146  | 
(equal_elim_axm %> A %> C %% prf3 %%  | 
| 17137 | 147  | 
(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
 | 
148  | 
|
| 26424 | 149  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %%
 | 
150  | 
        (PAxm ("Pure.symmetric", _, _) % _ % _ %%
 | 
|
151  | 
          (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %%
 | 
|
152  | 
            (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %%
 | 
|
| 56245 | 153  | 
              (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %%
 | 
| 26424 | 154  | 
                (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3)) %% prf4) =
 | 
| 15531 | 155  | 
SOME (equal_elim_axm %> A %> B %% prf1 %%  | 
| 17137 | 156  | 
(equal_elim_axm %> C %> A %% (symmetric_axm % ?? A % ?? C %% prf3) %%  | 
157  | 
(equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %% prf4)))  | 
|
| 11522 | 158  | 
|
| 26424 | 159  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %%
 | 
160  | 
        (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %%
 | 
|
161  | 
          (PAxm ("Pure.symmetric", _, _) % _ % _ %%
 | 
|
162  | 
            (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %%
 | 
|
| 56245 | 163  | 
              (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %%
 | 
| 26424 | 164  | 
                (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3) %% prf4) =
 | 
| 17137 | 165  | 
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
 | 
166  | 
(equal_elim_axm %> B %> D %% prf3 %%  | 
| 
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
167  | 
(equal_elim_axm %> A %> B %% prf1 %% prf4)))  | 
| 11522 | 168  | 
|
| 26424 | 169  | 
      | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %%
 | 
170  | 
        (PAxm ("Pure.symmetric", _, _) % _ % _ %%
 | 
|
171  | 
          (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %%
 | 
|
172  | 
            (PAxm ("Pure.symmetric", _, _) % _ % _ %%
 | 
|
173  | 
              (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %%
 | 
|
| 56245 | 174  | 
                (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %%
 | 
| 26424 | 175  | 
                  (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3)) %% prf4) =
 | 
| 17137 | 176  | 
SOME (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %%  | 
177  | 
(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
 | 
178  | 
(equal_elim_axm %> C %> D %% prf2 %% prf4)))  | 
| 11522 | 179  | 
|
| 26424 | 180  | 
      | rew' ((prf as PAxm ("Pure.combination", _, _) %
 | 
| 56245 | 181  | 
        SOME ((eq as Const ("Pure.eq", T)) $ t) % _ % _ % _) %%
 | 
| 26424 | 182  | 
          (PAxm ("Pure.reflexive", _, _) % _)) =
 | 
| 13257 | 183  | 
let val (U, V) = (case T of  | 
184  | 
Type (_, [U, V]) => (U, V) | _ => (dummyT, dummyT))  | 
|
| 37310 | 185  | 
in SOME (prf %% (ax Proofterm.combination_axm [U, V] %> eq % ?? eq % ?? t % ?? t %%  | 
186  | 
(ax Proofterm.reflexive_axm [T] % ?? eq) %% (ax Proofterm.reflexive_axm [U] % ?? t)))  | 
|
| 13257 | 187  | 
end  | 
188  | 
||
| 19309 | 189  | 
| rew' _ = NONE;  | 
| 37310 | 190  | 
in rew' #> Option.map (rpair Proofterm.no_skel) end;  | 
| 
12866
 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 
berghofe 
parents: 
12237 
diff
changeset
 | 
191  | 
|
| 
28806
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
192  | 
fun rprocs b = [rew b];  | 
| 53171 | 193  | 
val _ = Theory.setup (fold Proofterm.add_prf_rproc (rprocs false));  | 
| 11522 | 194  | 
|
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
195  | 
|
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
196  | 
(**** apply rewriting function to all terms in proof ****)  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
197  | 
|
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
198  | 
fun rewrite_terms r =  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
199  | 
let  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
200  | 
fun rew_term Ts t =  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
201  | 
let  | 
| 
29271
 
1d685baea08e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
28806 
diff
changeset
 | 
202  | 
val frees =  | 
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
43322 
diff
changeset
 | 
203  | 
map Free (Name.invent (Term.declare_term_frees t Name.context) "xa" (length Ts) ~~ Ts);  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
204  | 
val t' = r (subst_bounds (frees, t));  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
205  | 
fun strip [] t = t  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
206  | 
| strip (_ :: xs) (Abs (_, _, t)) = strip xs t;  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
207  | 
in  | 
| 19473 | 208  | 
strip Ts (fold lambda frees t')  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
209  | 
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  | 
fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2  | 
| 15531 | 212  | 
| rew Ts (prf % SOME t) = rew Ts prf % SOME (rew_term Ts t)  | 
213  | 
| rew Ts (Abst (s, SOME T, prf)) = Abst (s, SOME T, rew (T :: Ts) prf)  | 
|
214  | 
| 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
 | 
215  | 
| rew _ prf = prf  | 
| 
 
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  | 
in rew [] end;  | 
| 
 
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  | 
|
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
220  | 
(**** eliminate definitions in proof ****)  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
221  | 
|
| 16861 | 222  | 
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
 | 
223  | 
|
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
224  | 
fun insert_refl defs Ts (prf1 %% prf2) =  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
225  | 
let val (prf1', b) = insert_refl defs Ts prf1  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
226  | 
in  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
227  | 
if b then (prf1', true)  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
228  | 
else (prf1' %% fst (insert_refl defs Ts prf2), false)  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
229  | 
end  | 
| 15531 | 230  | 
| insert_refl defs Ts (Abst (s, SOME T, prf)) =  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
231  | 
(Abst (s, SOME T, fst (insert_refl defs (T :: Ts) prf)), false)  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
232  | 
| insert_refl defs Ts (AbsP (s, t, prf)) =  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
233  | 
(AbsP (s, t, fst (insert_refl defs Ts prf)), false)  | 
| 37310 | 234  | 
| insert_refl defs Ts prf =  | 
235  | 
(case Proofterm.strip_combt prf of  | 
|
| 
28806
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
236  | 
(PThm (_, ((s, prop, SOME Ts), _)), ts) =>  | 
| 20664 | 237  | 
if member (op =) defs s then  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
238  | 
let  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
239  | 
val vs = vars_of prop;  | 
| 
36042
 
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
 
krauss 
parents: 
33722 
diff
changeset
 | 
240  | 
val tvars = Term.add_tvars prop [] |> rev;  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
241  | 
val (_, rhs) = Logic.dest_equals (Logic.strip_imp_concl prop);  | 
| 18185 | 242  | 
val rhs' = Term.betapplys (subst_TVars (map fst tvars ~~ Ts)  | 
| 23178 | 243  | 
                (fold_rev (fn x => fn b => Abs ("", dummyT, abstract_over (x, b))) vs rhs),
 | 
| 19466 | 244  | 
map the ts);  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
245  | 
in  | 
| 37310 | 246  | 
(Proofterm.change_type (SOME [fastype_of1 (Ts, rhs')])  | 
247  | 
Proofterm.reflexive_axm %> rhs', true)  | 
|
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
248  | 
end  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
249  | 
else (prf, false)  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
250  | 
| (_, []) => (prf, false)  | 
| 37310 | 251  | 
| (prf', ts) => (Proofterm.proof_combt' (fst (insert_refl defs Ts prf'), ts), false));  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
252  | 
|
| 62922 | 253  | 
fun elim_defs ctxt r defs prf =  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
254  | 
let  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
255  | 
val defs' = map (Logic.dest_equals o  | 
| 59582 | 256  | 
map_types Type.strip_sorts o Thm.prop_of o Drule.abs_def) defs;  | 
| 
36744
 
6e1f3d609a68
renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
 
wenzelm 
parents: 
36042 
diff
changeset
 | 
257  | 
val defnames = map Thm.derivation_name defs;  | 
| 13341 | 258  | 
val f = if not r then I else  | 
259  | 
let  | 
|
260  | 
val cnames = map (fst o dest_Const o fst) defs';  | 
|
| 37310 | 261  | 
val thms = Proofterm.fold_proof_atoms true  | 
| 
28806
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
262  | 
(fn PThm (_, ((name, prop, _), _)) =>  | 
| 
29271
 
1d685baea08e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
28806 
diff
changeset
 | 
263  | 
if member (op =) defnames name orelse  | 
| 
 
1d685baea08e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
28806 
diff
changeset
 | 
264  | 
not (exists_Const (member (op =) cnames o #1) prop)  | 
| 
 
1d685baea08e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
28806 
diff
changeset
 | 
265  | 
then I  | 
| 
28806
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
266  | 
else cons (name, SOME prop)  | 
| 
 
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
267  | 
| _ => I) [prf] [];  | 
| 62922 | 268  | 
in Reconstruct.expand_proof ctxt thms end;  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
269  | 
in  | 
| 62922 | 270  | 
rewrite_terms (Pattern.rewrite_term (Proof_Context.theory_of ctxt) defs' [])  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
271  | 
(fst (insert_refl defnames [] (f prf)))  | 
| 
12906
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
272  | 
end;  | 
| 
 
165f4e1937f4
New function for eliminating definitions in proof term.
 
berghofe 
parents: 
12866 
diff
changeset
 | 
273  | 
|
| 13608 | 274  | 
|
275  | 
(**** eliminate all variables that don't occur in the proposition ****)  | 
|
276  | 
||
277  | 
fun elim_vars mk_default prf =  | 
|
278  | 
let  | 
|
279  | 
val prop = Reconstruct.prop_of prf;  | 
|
| 19309 | 280  | 
val tv = Term.add_vars prop [];  | 
281  | 
val tf = Term.add_frees prop [];  | 
|
282  | 
||
283  | 
fun hidden_variable (Var v) = not (member (op =) tv v)  | 
|
284  | 
| hidden_variable (Free f) = not (member (op =) tf f)  | 
|
285  | 
| hidden_variable _ = false;  | 
|
| 13917 | 286  | 
|
| 
46219
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
287  | 
fun mk_default' T =  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
288  | 
fold_rev (Term.abs o pair "x") (binder_types T) (mk_default (body_type T));  | 
| 13917 | 289  | 
|
290  | 
fun elim_varst (t $ u) = elim_varst t $ elim_varst u  | 
|
291  | 
| elim_varst (Abs (s, T, t)) = Abs (s, T, elim_varst t)  | 
|
| 19309 | 292  | 
| elim_varst (t as Free (x, T)) = if member (op =) tf (x, T) then t else mk_default' T  | 
293  | 
| elim_varst (t as Var (xi, T)) = if member (op =) tv (xi, T) then t else mk_default' T  | 
|
294  | 
| elim_varst t = t;  | 
|
| 13608 | 295  | 
in  | 
| 37310 | 296  | 
Proofterm.map_proof_terms (fn t =>  | 
| 19309 | 297  | 
if Term.exists_subterm hidden_variable t then Envir.beta_norm (elim_varst t) else t) I prf  | 
| 13608 | 298  | 
end;  | 
299  | 
||
| 22280 | 300  | 
|
301  | 
(**** convert between hhf and non-hhf form ****)  | 
|
302  | 
||
303  | 
fun hhf_proof P Q prf =  | 
|
304  | 
let  | 
|
305  | 
val params = Logic.strip_params Q;  | 
|
306  | 
val Hs = Logic.strip_assums_hyp P;  | 
|
307  | 
val Hs' = Logic.strip_assums_hyp Q;  | 
|
308  | 
val k = length Hs;  | 
|
309  | 
val l = length params;  | 
|
| 56245 | 310  | 
    fun mk_prf i j Hs Hs' (Const ("Pure.all", _) $ Abs (_, _, P)) prf =
 | 
| 22280 | 311  | 
mk_prf i (j - 1) Hs Hs' P (prf %> Bound j)  | 
| 56245 | 312  | 
      | mk_prf i j (H :: Hs) (H' :: Hs') (Const ("Pure.imp", _) $ _ $ P) prf =
 | 
| 22280 | 313  | 
mk_prf (i - 1) j Hs Hs' P (prf %% un_hhf_proof H' H (PBound i))  | 
314  | 
| mk_prf _ _ _ _ _ prf = prf  | 
|
315  | 
in  | 
|
316  | 
prf |> Proofterm.incr_pboundvars k l |> mk_prf (k - 1) (l - 1) Hs Hs' P |>  | 
|
317  | 
    fold_rev (fn P => fn prf => AbsP ("H", SOME P, prf)) Hs' |>
 | 
|
318  | 
fold_rev (fn (s, T) => fn prf => Abst (s, SOME T, prf)) params  | 
|
319  | 
end  | 
|
320  | 
and un_hhf_proof P Q prf =  | 
|
321  | 
let  | 
|
322  | 
val params = Logic.strip_params Q;  | 
|
323  | 
val Hs = Logic.strip_assums_hyp P;  | 
|
324  | 
val Hs' = Logic.strip_assums_hyp Q;  | 
|
325  | 
val k = length Hs;  | 
|
326  | 
val l = length params;  | 
|
| 56245 | 327  | 
    fun mk_prf (Const ("Pure.all", _) $ Abs (s, T, P)) prf =
 | 
| 22280 | 328  | 
Abst (s, SOME T, mk_prf P prf)  | 
| 56245 | 329  | 
      | mk_prf (Const ("Pure.imp", _) $ P $ Q) prf =
 | 
| 22280 | 330  | 
          AbsP ("H", SOME P, mk_prf Q prf)
 | 
331  | 
| mk_prf _ prf = prf  | 
|
332  | 
in  | 
|
333  | 
prf |> Proofterm.incr_pboundvars k l |>  | 
|
334  | 
fold (fn i => fn prf => prf %> Bound i) (l - 1 downto 0) |>  | 
|
335  | 
fold (fn ((H, H'), i) => fn prf => prf %% hhf_proof H' H (PBound i))  | 
|
336  | 
(Hs ~~ Hs' ~~ (k - 1 downto 0)) |>  | 
|
337  | 
mk_prf Q  | 
|
338  | 
end;  | 
|
339  | 
||
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
340  | 
|
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
341  | 
(**** expand OfClass proofs ****)  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
342  | 
|
| 62922 | 343  | 
fun mk_of_sort_proof ctxt hs (T, S) =  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
344  | 
let  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
345  | 
val hs' = map  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
346  | 
(fn SOME t => (SOME (Logic.dest_of_class t) handle TERM _ => NONE)  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
347  | 
| NONE => NONE) hs;  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
348  | 
val sorts = AList.coalesce (op =) (rev (map_filter I hs'));  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
349  | 
fun get_sort T = the_default [] (AList.lookup (op =) sorts T);  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
350  | 
val subst = map_atyps  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
351  | 
(fn T as TVar (ixn, _) => TVar (ixn, get_sort T)  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
352  | 
| T as TFree (s, _) => TFree (s, get_sort T));  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
353  | 
fun hyp T_c = case find_index (equal (SOME T_c)) hs' of  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
354  | 
~1 => error "expand_of_class: missing class hypothesis"  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
355  | 
| i => PBound i;  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
356  | 
fun reconstruct prf prop = prf |>  | 
| 62922 | 357  | 
Reconstruct.reconstruct_proof ctxt prop |>  | 
358  | 
      Reconstruct.expand_proof ctxt [("", NONE)] |>
 | 
|
| 37310 | 359  | 
Same.commit (Proofterm.map_proof_same Same.same Same.same hyp)  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
360  | 
in  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
361  | 
map2 reconstruct  | 
| 62922 | 362  | 
(Proofterm.of_sort_proof (Proof_Context.theory_of ctxt)  | 
363  | 
(OfClass o apfst Type.strip_sorts) (subst T, S))  | 
|
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
364  | 
(Logic.mk_of_sort (T, S))  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
365  | 
end;  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
366  | 
|
| 62922 | 367  | 
fun expand_of_class ctxt Ts hs (OfClass (T, c)) =  | 
368  | 
mk_of_sort_proof ctxt hs (T, [c]) |>  | 
|
| 37310 | 369  | 
hd |> rpair Proofterm.no_skel |> SOME  | 
| 62922 | 370  | 
| expand_of_class ctxt Ts hs _ = NONE;  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36744 
diff
changeset
 | 
371  | 
|
| 11522 | 372  | 
end;  |