author | wenzelm |
Fri, 02 Sep 2022 23:19:02 +0200 | |
changeset 76041 | 4a1330addb4e |
parent 74266 | 612b7e0d6721 |
child 79175 | 04dfecb9343a |
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 |
70449 | 13 |
val elim_defs : theory -> 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 |
|
70836 | 17 |
val expand_of_sort_proof : theory -> term option list -> typ * sort -> Proofterm.proof list |
70837 | 18 |
val expand_of_class_proof : theory -> term option list -> typ * class -> proof |
70449 | 19 |
val expand_of_class : theory -> 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
|
20 |
(Proofterm.proof * Proofterm.proof) option |
70879 | 21 |
val standard_preproc: (proof * proof) list -> theory -> proof -> proof |
11522 | 22 |
end; |
23 |
||
70840 | 24 |
structure Proof_Rewrite_Rules : PROOF_REWRITE_RULES = |
11522 | 25 |
struct |
26 |
||
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
27 |
fun rew b _ _ = |
12866
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
28 |
let |
17137 | 29 |
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
|
30 |
fun ax (prf as PAxm (s, prop, _)) Ts = |
15531 | 31 |
if b then PAxm (s, prop, SOME Ts) else prf; |
71087 | 32 |
fun ty T = |
33 |
if b then |
|
34 |
(case T of |
|
35 |
Type (_, [Type (_, [U, _]), _]) => SOME U |
|
36 |
| _ => NONE) |
|
15531 | 37 |
else NONE; |
37310 | 38 |
val equal_intr_axm = ax Proofterm.equal_intr_axm []; |
39 |
val equal_elim_axm = ax Proofterm.equal_elim_axm []; |
|
40 |
val symmetric_axm = ax Proofterm.symmetric_axm [propT]; |
|
11522 | 41 |
|
70493 | 42 |
fun rew' (PThm ({name = "Pure.protectD", ...}, _) % _ %% |
43 |
(PThm ({name = "Pure.protectI", ...}, _) % _ %% prf)) = SOME prf |
|
44 |
| rew' (PThm ({name = "Pure.conjunctionD1", ...}, _) % _ % _ %% |
|
45 |
(PThm ({name = "Pure.conjunctionI", ...}, _) % _ % _ %% prf %% _)) = SOME prf |
|
46 |
| rew' (PThm ({name = "Pure.conjunctionD2", ...}, _) % _ % _ %% |
|
47 |
(PThm ({name = "Pure.conjunctionI", ...}, _) % _ % _ %% _ %% prf)) = SOME prf |
|
26424 | 48 |
| rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% |
49 |
(PAxm ("Pure.equal_intr", _, _) % _ % _ %% prf %% _)) = SOME prf |
|
50 |
| rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %% |
|
51 |
(PAxm ("Pure.equal_intr", _, _) % A % B %% prf1 %% prf2)) = |
|
15531 | 52 |
SOME (equal_intr_axm % B % A %% prf2 %% prf1) |
12002 | 53 |
|
26424 | 54 |
| rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %% |
56244 | 55 |
(PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) % |
26424 | 56 |
_ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% |
70493 | 57 |
((tg as PThm ({name = "Pure.protectI", ...}, _)) % _ %% prf2)) = |
15531 | 58 |
SOME (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2)) |
12002 | 59 |
|
26424 | 60 |
| rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %% |
61 |
(PAxm ("Pure.symmetric", _, _) % _ % _ %% |
|
56244 | 62 |
(PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) % |
26424 | 63 |
_ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1)) %% |
70493 | 64 |
((tg as PThm ({name = "Pure.protectI", ...}, _)) % _ %% prf2)) = |
15531 | 65 |
SOME (tg %> B %% (equal_elim_axm %> A %> B %% |
17137 | 66 |
(symmetric_axm % ?? B % ?? A %% prf1) %% prf2)) |
11522 | 67 |
|
26424 | 68 |
| rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% |
69 |
(PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% |
|
56245 | 70 |
(PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.imp", _)) % _ % _ % _ %% |
26424 | 71 |
(PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) = |
12866
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
72 |
let |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
73 |
val _ $ A $ C = Envir.beta_norm X; |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
74 |
val _ $ B $ D = Envir.beta_norm Y |
17137 | 75 |
in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? B, |
37310 | 76 |
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
|
77 |
(PBound 1 %% (equal_elim_axm %> B %> A %% |
37310 | 78 |
(Proofterm.symmetric_axm % ?? A % ?? B %% Proofterm.incr_pboundvars 2 0 prf1) %% |
79 |
PBound 0))))) |
|
12866
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
80 |
end |
11522 | 81 |
|
26424 | 82 |
| rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% |
83 |
(PAxm ("Pure.symmetric", _, _) % _ % _ %% |
|
84 |
(PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% |
|
56245 | 85 |
(PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.imp", _)) % _ % _ % _ %% |
26424 | 86 |
(PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2))) = |
12866
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
87 |
let |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
88 |
val _ $ A $ C = Envir.beta_norm Y; |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
89 |
val _ $ B $ D = Envir.beta_norm X |
17137 | 90 |
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
|
91 |
equal_elim_axm %> D %> C %% |
37310 | 92 |
(symmetric_axm % ?? C % ?? D %% Proofterm.incr_pboundvars 2 0 prf2) %% |
93 |
(PBound 1 %% |
|
94 |
(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
|
95 |
end |
11522 | 96 |
|
71843
07c85c68ff03
added missing preprocessing step for extraction (due to Stefan Berghofer)
Manuel Eberl <eberlm@in.tum.de>
parents:
71777
diff
changeset
|
97 |
| rew' (PAxm ("Pure.combination", _, _) % |
07c85c68ff03
added missing preprocessing step for extraction (due to Stefan Berghofer)
Manuel Eberl <eberlm@in.tum.de>
parents:
71777
diff
changeset
|
98 |
SOME (imp as (imp' as Const ("Pure.imp", T)) $ X) % _ % Y % Z %% |
07c85c68ff03
added missing preprocessing step for extraction (due to Stefan Berghofer)
Manuel Eberl <eberlm@in.tum.de>
parents:
71777
diff
changeset
|
99 |
(PAxm ("Pure.reflexive", _, _) % _) %% prf) = |
07c85c68ff03
added missing preprocessing step for extraction (due to Stefan Berghofer)
Manuel Eberl <eberlm@in.tum.de>
parents:
71777
diff
changeset
|
100 |
SOME (ax Proofterm.combination_axm [propT, propT] %> imp % ?? imp % Y % Z %% |
07c85c68ff03
added missing preprocessing step for extraction (due to Stefan Berghofer)
Manuel Eberl <eberlm@in.tum.de>
parents:
71777
diff
changeset
|
101 |
(ax Proofterm.combination_axm [propT, propT --> propT] %> imp' % ?? imp' % ?? X % ?? X %% |
07c85c68ff03
added missing preprocessing step for extraction (due to Stefan Berghofer)
Manuel Eberl <eberlm@in.tum.de>
parents:
71777
diff
changeset
|
102 |
(ax Proofterm.reflexive_axm [T] % ?? imp') %% |
07c85c68ff03
added missing preprocessing step for extraction (due to Stefan Berghofer)
Manuel Eberl <eberlm@in.tum.de>
parents:
71777
diff
changeset
|
103 |
(ax Proofterm.reflexive_axm [propT] % ?? X)) %% prf) |
07c85c68ff03
added missing preprocessing step for extraction (due to Stefan Berghofer)
Manuel Eberl <eberlm@in.tum.de>
parents:
71777
diff
changeset
|
104 |
|
26424 | 105 |
| rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% |
56245 | 106 |
(PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.all", _)) % _ % _ % _ %% |
26424 | 107 |
(PAxm ("Pure.reflexive", _, _) % _) %% |
108 |
(PAxm ("Pure.abstract_rule", _, _) % _ % _ %% prf))) = |
|
12866
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
109 |
let |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
110 |
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
|
111 |
val _ $ Q = Envir.beta_norm Y; |
17137 | 112 |
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
|
113 |
equal_elim_axm %> incr_boundvars 1 P $ Bound 0 %> incr_boundvars 1 Q $ Bound 0 %% |
37310 | 114 |
(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
|
115 |
end |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
116 |
|
26424 | 117 |
| rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% |
118 |
(PAxm ("Pure.symmetric", _, _) % _ % _ %% |
|
56245 | 119 |
(PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.all", _)) % _ % _ % _ %% |
26424 | 120 |
(PAxm ("Pure.reflexive", _, _) % _) %% |
121 |
(PAxm ("Pure.abstract_rule", _, _) % _ % _ %% prf)))) = |
|
12866
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
122 |
let |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
123 |
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
|
124 |
val _ $ Q = Envir.beta_norm Y; |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
125 |
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
|
126 |
val u = incr_boundvars 1 Q $ Bound 0 |
17137 | 127 |
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
|
128 |
equal_elim_axm %> t %> u %% |
37310 | 129 |
(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
|
130 |
%% (PBound 0 %> Bound 0)))) |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
131 |
end |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
132 |
|
26424 | 133 |
| rew' (PAxm ("Pure.equal_elim", _, _) % SOME A % SOME C %% |
134 |
(PAxm ("Pure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2) %% prf3) = |
|
15531 | 135 |
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
|
136 |
(equal_elim_axm %> A %> B %% prf1 %% prf3)) |
26424 | 137 |
| rew' (PAxm ("Pure.equal_elim", _, _) % SOME A % SOME C %% |
138 |
(PAxm ("Pure.symmetric", _, _) % _ % _ %% |
|
139 |
(PAxm ("Pure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2)) %% prf3) = |
|
17137 | 140 |
SOME (equal_elim_axm %> B %> C %% (symmetric_axm % ?? C % ?? B %% prf1) %% |
141 |
(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
|
142 |
|
26424 | 143 |
| rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% |
144 |
(PAxm ("Pure.reflexive", _, _) % _) %% prf) = SOME prf |
|
145 |
| rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% |
|
146 |
(PAxm ("Pure.symmetric", _, _) % _ % _ %% |
|
147 |
(PAxm ("Pure.reflexive", _, _) % _)) %% prf) = SOME prf |
|
12866
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
148 |
|
26424 | 149 |
| rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %% |
150 |
(PAxm ("Pure.symmetric", _, _) % _ % _ %% prf)) = SOME prf |
|
11522 | 151 |
|
26424 | 152 |
| rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% |
153 |
(PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %% |
|
154 |
(PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% |
|
56245 | 155 |
(PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% |
26424 | 156 |
(PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3) %% prf4) = |
15531 | 157 |
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
|
158 |
(equal_elim_axm %> A %> C %% prf3 %% |
17137 | 159 |
(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
|
160 |
|
26424 | 161 |
| rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% |
162 |
(PAxm ("Pure.symmetric", _, _) % _ % _ %% |
|
163 |
(PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %% |
|
164 |
(PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% |
|
56245 | 165 |
(PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% |
26424 | 166 |
(PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3)) %% prf4) = |
15531 | 167 |
SOME (equal_elim_axm %> A %> B %% prf1 %% |
17137 | 168 |
(equal_elim_axm %> C %> A %% (symmetric_axm % ?? A % ?? C %% prf3) %% |
169 |
(equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %% prf4))) |
|
11522 | 170 |
|
26424 | 171 |
| rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% |
172 |
(PAxm ("Pure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %% |
|
173 |
(PAxm ("Pure.symmetric", _, _) % _ % _ %% |
|
174 |
(PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% |
|
56245 | 175 |
(PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% |
26424 | 176 |
(PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3) %% prf4) = |
17137 | 177 |
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
|
178 |
(equal_elim_axm %> B %> D %% prf3 %% |
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
berghofe
parents:
12237
diff
changeset
|
179 |
(equal_elim_axm %> A %> B %% prf1 %% prf4))) |
11522 | 180 |
|
26424 | 181 |
| rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% |
182 |
(PAxm ("Pure.symmetric", _, _) % _ % _ %% |
|
183 |
(PAxm ("Pure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %% |
|
184 |
(PAxm ("Pure.symmetric", _, _) % _ % _ %% |
|
185 |
(PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% |
|
56245 | 186 |
(PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% |
26424 | 187 |
(PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3)) %% prf4) = |
17137 | 188 |
SOME (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %% |
189 |
(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
|
190 |
(equal_elim_axm %> C %> D %% prf2 %% prf4))) |
11522 | 191 |
|
26424 | 192 |
| rew' ((prf as PAxm ("Pure.combination", _, _) % |
56245 | 193 |
SOME ((eq as Const ("Pure.eq", T)) $ t) % _ % _ % _) %% |
26424 | 194 |
(PAxm ("Pure.reflexive", _, _) % _)) = |
13257 | 195 |
let val (U, V) = (case T of |
196 |
Type (_, [U, V]) => (U, V) | _ => (dummyT, dummyT)) |
|
37310 | 197 |
in SOME (prf %% (ax Proofterm.combination_axm [U, V] %> eq % ?? eq % ?? t % ?? t %% |
198 |
(ax Proofterm.reflexive_axm [T] % ?? eq) %% (ax Proofterm.reflexive_axm [U] % ?? t))) |
|
13257 | 199 |
end |
200 |
||
19309 | 201 |
| rew' _ = NONE; |
37310 | 202 |
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
|
203 |
|
28806
ba0ffe4cfc2b
rewrite_proof: simplified simprocs (no name required);
wenzelm
parents:
26463
diff
changeset
|
204 |
fun rprocs b = [rew b]; |
53171 | 205 |
val _ = Theory.setup (fold Proofterm.add_prf_rproc (rprocs false)); |
11522 | 206 |
|
12906
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
207 |
|
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
208 |
(**** apply rewriting function to all terms in proof ****) |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
209 |
|
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
210 |
fun rewrite_terms r = |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
211 |
let |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
212 |
fun rew_term Ts t = |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
213 |
let |
29271
1d685baea08e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
28806
diff
changeset
|
214 |
val frees = |
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43322
diff
changeset
|
215 |
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
|
216 |
val t' = r (subst_bounds (frees, t)); |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
217 |
fun strip [] t = t |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
218 |
| strip (_ :: xs) (Abs (_, _, t)) = strip xs t; |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
219 |
in |
19473 | 220 |
strip Ts (fold lambda frees t') |
12906
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
221 |
end; |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
222 |
|
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
223 |
fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2 |
15531 | 224 |
| rew Ts (prf % SOME t) = rew Ts prf % SOME (rew_term Ts t) |
225 |
| rew Ts (Abst (s, SOME T, prf)) = Abst (s, SOME T, rew (T :: Ts) prf) |
|
226 |
| 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
|
227 |
| rew _ prf = prf |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
228 |
|
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
229 |
in rew [] end; |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
230 |
|
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
231 |
|
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
232 |
(**** eliminate definitions in proof ****) |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
233 |
|
16861 | 234 |
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
|
235 |
|
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
236 |
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
|
237 |
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
|
238 |
in |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
239 |
if b then (prf1', true) |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
240 |
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
|
241 |
end |
15531 | 242 |
| 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
|
243 |
(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
|
244 |
| 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
|
245 |
(AbsP (s, t, fst (insert_refl defs Ts prf)), false) |
37310 | 246 |
| insert_refl defs Ts prf = |
247 |
(case Proofterm.strip_combt prf of |
|
70493 | 248 |
(PThm ({name = s, prop, types = SOME Ts, ...}, _), ts) => |
20664 | 249 |
if member (op =) defs s then |
12906
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
250 |
let |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
251 |
val vs = vars_of prop; |
74235 | 252 |
val tvars = build_rev (Term.add_tvars prop); |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
253 |
val (_, rhs) = Logic.dest_equals (Logic.strip_imp_concl prop); |
18185 | 254 |
val rhs' = Term.betapplys (subst_TVars (map fst tvars ~~ Ts) |
23178 | 255 |
(fold_rev (fn x => fn b => Abs ("", dummyT, abstract_over (x, b))) vs rhs), |
19466 | 256 |
map the ts); |
12906
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
257 |
in |
70417 | 258 |
(Proofterm.change_types (SOME [fastype_of1 (Ts, rhs')]) |
37310 | 259 |
Proofterm.reflexive_axm %> rhs', true) |
12906
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
260 |
end |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
261 |
else (prf, false) |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
262 |
| (_, []) => (prf, false) |
37310 | 263 |
| (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
|
264 |
|
70449 | 265 |
fun elim_defs thy r defs prf = |
12906
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
266 |
let |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
267 |
val defs' = map (Logic.dest_equals o |
59582 | 268 |
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
|
269 |
val defnames = map Thm.derivation_name defs; |
70826 | 270 |
val prf' = |
271 |
if r then |
|
272 |
let |
|
273 |
val cnames = map (fst o dest_Const o fst) defs'; |
|
70915
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents:
70879
diff
changeset
|
274 |
val expand = Proofterm.fold_proof_atoms true |
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents:
70879
diff
changeset
|
275 |
(fn PThm ({serial, name, prop, ...}, _) => |
70826 | 276 |
if member (op =) defnames name orelse |
277 |
not (exists_Const (member (op =) cnames o #1) prop) |
|
278 |
then I |
|
70915
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents:
70879
diff
changeset
|
279 |
else Inttab.update (serial, "") |
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents:
70879
diff
changeset
|
280 |
| _ => I) [prf] Inttab.empty; |
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents:
70879
diff
changeset
|
281 |
in Proofterm.expand_proof thy (Inttab.lookup expand o #serial) prf end |
70826 | 282 |
else prf; |
12906
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
283 |
in |
70449 | 284 |
rewrite_terms (Pattern.rewrite_term thy defs' []) |
70826 | 285 |
(fst (insert_refl defnames [] prf')) |
12906
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
286 |
end; |
165f4e1937f4
New function for eliminating definitions in proof term.
berghofe
parents:
12866
diff
changeset
|
287 |
|
13608 | 288 |
|
289 |
(**** eliminate all variables that don't occur in the proposition ****) |
|
290 |
||
291 |
fun elim_vars mk_default prf = |
|
292 |
let |
|
70447
755d58b48cec
clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents:
70417
diff
changeset
|
293 |
val prop = Proofterm.prop_of prf; |
74266 | 294 |
val tv = Vars.build (Vars.add_vars prop); |
295 |
val tf = Frees.build (Frees.add_frees prop); |
|
19309 | 296 |
|
74266 | 297 |
fun hidden_variable (Var v) = not (Vars.defined tv v) |
298 |
| hidden_variable (Free f) = not (Frees.defined tf f) |
|
19309 | 299 |
| hidden_variable _ = false; |
13917 | 300 |
|
46219
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents:
43329
diff
changeset
|
301 |
fun mk_default' T = |
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents:
43329
diff
changeset
|
302 |
fold_rev (Term.abs o pair "x") (binder_types T) (mk_default (body_type T)); |
13917 | 303 |
|
304 |
fun elim_varst (t $ u) = elim_varst t $ elim_varst u |
|
305 |
| elim_varst (Abs (s, T, t)) = Abs (s, T, elim_varst t) |
|
74266 | 306 |
| elim_varst (t as Free (x, T)) = if Frees.defined tf (x, T) then t else mk_default' T |
307 |
| elim_varst (t as Var (xi, T)) = if Vars.defined tv (xi, T) then t else mk_default' T |
|
19309 | 308 |
| elim_varst t = t; |
13608 | 309 |
in |
37310 | 310 |
Proofterm.map_proof_terms (fn t => |
19309 | 311 |
if Term.exists_subterm hidden_variable t then Envir.beta_norm (elim_varst t) else t) I prf |
13608 | 312 |
end; |
313 |
||
22280 | 314 |
|
315 |
(**** convert between hhf and non-hhf form ****) |
|
316 |
||
317 |
fun hhf_proof P Q prf = |
|
318 |
let |
|
319 |
val params = Logic.strip_params Q; |
|
320 |
val Hs = Logic.strip_assums_hyp P; |
|
321 |
val Hs' = Logic.strip_assums_hyp Q; |
|
322 |
val k = length Hs; |
|
323 |
val l = length params; |
|
56245 | 324 |
fun mk_prf i j Hs Hs' (Const ("Pure.all", _) $ Abs (_, _, P)) prf = |
22280 | 325 |
mk_prf i (j - 1) Hs Hs' P (prf %> Bound j) |
56245 | 326 |
| mk_prf i j (H :: Hs) (H' :: Hs') (Const ("Pure.imp", _) $ _ $ P) prf = |
22280 | 327 |
mk_prf (i - 1) j Hs Hs' P (prf %% un_hhf_proof H' H (PBound i)) |
328 |
| mk_prf _ _ _ _ _ prf = prf |
|
329 |
in |
|
330 |
prf |> Proofterm.incr_pboundvars k l |> mk_prf (k - 1) (l - 1) Hs Hs' P |> |
|
331 |
fold_rev (fn P => fn prf => AbsP ("H", SOME P, prf)) Hs' |> |
|
332 |
fold_rev (fn (s, T) => fn prf => Abst (s, SOME T, prf)) params |
|
333 |
end |
|
334 |
and un_hhf_proof P Q prf = |
|
335 |
let |
|
336 |
val params = Logic.strip_params Q; |
|
337 |
val Hs = Logic.strip_assums_hyp P; |
|
338 |
val Hs' = Logic.strip_assums_hyp Q; |
|
339 |
val k = length Hs; |
|
340 |
val l = length params; |
|
56245 | 341 |
fun mk_prf (Const ("Pure.all", _) $ Abs (s, T, P)) prf = |
22280 | 342 |
Abst (s, SOME T, mk_prf P prf) |
56245 | 343 |
| mk_prf (Const ("Pure.imp", _) $ P $ Q) prf = |
22280 | 344 |
AbsP ("H", SOME P, mk_prf Q prf) |
345 |
| mk_prf _ prf = prf |
|
346 |
in |
|
347 |
prf |> Proofterm.incr_pboundvars k l |> |
|
348 |
fold (fn i => fn prf => prf %> Bound i) (l - 1 downto 0) |> |
|
349 |
fold (fn ((H, H'), i) => fn prf => prf %% hhf_proof H' H (PBound i)) |
|
350 |
(Hs ~~ Hs' ~~ (k - 1 downto 0)) |> |
|
351 |
mk_prf Q |
|
352 |
end; |
|
353 |
||
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
354 |
|
71777
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
wenzelm
parents:
71087
diff
changeset
|
355 |
(**** expand PClass proofs ****) |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
356 |
|
70836 | 357 |
fun expand_of_sort_proof thy hyps (T, S) = |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
358 |
let |
70836 | 359 |
val of_class_hyps = map (fn SOME t => try Logic.dest_of_class t | NONE => NONE) hyps; |
360 |
fun of_class_index p = |
|
361 |
(case find_index (fn SOME h => h = p | NONE => false) of_class_hyps of |
|
362 |
~1 => raise Fail "expand_of_class_proof: missing class hypothesis" |
|
363 |
| i => PBound i); |
|
364 |
||
365 |
val sorts = AList.coalesce (op =) (rev (map_filter I of_class_hyps)); |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
366 |
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
|
367 |
val subst = map_atyps |
70836 | 368 |
(fn T as TVar (ai, _) => TVar (ai, get_sort T) |
369 |
| T as TFree (a, _) => TFree (a, get_sort T)); |
|
370 |
||
371 |
fun reconstruct prop = |
|
372 |
Proofterm.reconstruct_proof thy prop #> |
|
70915
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents:
70879
diff
changeset
|
373 |
Proofterm.expand_proof thy Proofterm.expand_name_empty #> |
70836 | 374 |
Same.commit (Proofterm.map_proof_same Same.same Same.same of_class_index); |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
375 |
in |
70836 | 376 |
map2 reconstruct (Logic.mk_of_sort (T, S)) |
70458 | 377 |
(Proofterm.of_sort_proof (Sign.classes_of thy) (Thm.classrel_proof thy) (Thm.arity_proof thy) |
71777
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
wenzelm
parents:
71087
diff
changeset
|
378 |
(PClass 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
|
379 |
end; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
380 |
|
70837 | 381 |
fun expand_of_class_proof thy hyps (T, c) = |
382 |
hd (expand_of_sort_proof thy hyps (T, [c])); |
|
383 |
||
71777
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
wenzelm
parents:
71087
diff
changeset
|
384 |
fun expand_of_class thy (_: typ list) hyps (PClass (T, c)) = |
70837 | 385 |
SOME (expand_of_class_proof thy hyps (T, c), Proofterm.no_skel) |
70836 | 386 |
| expand_of_class _ _ _ _ = NONE; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36744
diff
changeset
|
387 |
|
70879 | 388 |
|
389 |
(* standard preprocessor *) |
|
390 |
||
391 |
fun standard_preproc rews thy = |
|
392 |
Proofterm.rewrite_proof thy (rews, rprocs true @ [expand_of_class thy]); |
|
393 |
||
11522 | 394 |
end; |