author | skalberg |
Thu, 03 Mar 2005 12:43:01 +0100 | |
changeset 15570 | 8d8c70b41bab |
parent 15531 | 08c8dad8e399 |
child 19798 | 94f12468bbba |
permissions | -rw-r--r-- |
13404 | 1 |
(* Title: HOL/Tools/rewrite_hol_proof.ML |
2 |
ID: $Id$ |
|
3 |
Author: Stefan Berghofer, TU Muenchen |
|
4 |
||
5 |
Rewrite rules for HOL proofs |
|
6 |
*) |
|
7 |
||
8 |
signature REWRITE_HOL_PROOF = |
|
9 |
sig |
|
10 |
val rews: (Proofterm.proof * Proofterm.proof) list |
|
11 |
val elim_cong: typ list -> Proofterm.proof -> Proofterm.proof option |
|
12 |
end; |
|
13 |
||
14 |
structure RewriteHOLProof : REWRITE_HOL_PROOF = |
|
15 |
struct |
|
16 |
||
17 |
open Proofterm; |
|
18 |
||
19 |
val rews = map (pairself (ProofSyntax.proof_of_term (the_context ()) Symtab.empty true) o |
|
20 |
Logic.dest_equals o Logic.varify o ProofSyntax.read_term (the_context ()) propT) |
|
21 |
||
22 |
(** eliminate meta-equality rules **) |
|
23 |
||
24 |
["(equal_elim % x1 % x2 %% \ |
|
25 |
\ (combination % TYPE('T1) % TYPE('T2) % Trueprop % x3 % A % B %% \ |
|
26 |
\ (axm.reflexive % TYPE('T3) % x4) %% prf1) %% prf2) == \ |
|
27 |
\ (iffD1 % A % B %% \ |
|
28 |
\ (meta_eq_to_obj_eq % TYPE(bool) % A % B %% prf1) %% prf2)", |
|
29 |
||
30 |
"(equal_elim % x1 % x2 %% (axm.symmetric % TYPE('T1) % x3 % x4 %% \ |
|
31 |
\ (combination % TYPE('T2) % TYPE('T3) % Trueprop % x5 % A % B %% \ |
|
32 |
\ (axm.reflexive % TYPE('T4) % x6) %% prf1)) %% prf2) == \ |
|
33 |
\ (iffD2 % A % B %% \ |
|
34 |
\ (meta_eq_to_obj_eq % TYPE(bool) % A % B %% prf1) %% prf2)", |
|
35 |
||
36 |
"(meta_eq_to_obj_eq % TYPE('U) % x1 % x2 %% \ |
|
37 |
\ (combination % TYPE('U) % TYPE('T) % f % g % x % y %% prf1 %% prf2)) == \ |
|
38 |
\ (cong % TYPE('U) % TYPE('T) % f % g % x % y %% \ |
|
39 |
\ (meta_eq_to_obj_eq % TYPE('T => 'U) % f % g %% prf1) %% \ |
|
40 |
\ (meta_eq_to_obj_eq % TYPE('T) % x % y %% prf2))", |
|
41 |
||
42 |
"(meta_eq_to_obj_eq % TYPE('T) % x1 % x2 %% \ |
|
43 |
\ (axm.transitive % TYPE('T) % x % y % z %% prf1 %% prf2)) == \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
44 |
\ (HOL.trans % TYPE('T) % x % y % z %% \ |
13404 | 45 |
\ (meta_eq_to_obj_eq % TYPE('T) % x % y %% prf1) %% \ |
46 |
\ (meta_eq_to_obj_eq % TYPE('T) % y % z %% prf2))", |
|
47 |
||
48 |
"(meta_eq_to_obj_eq % TYPE('T) % x % x %% (axm.reflexive % TYPE('T) % x)) == \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
49 |
\ (HOL.refl % TYPE('T) % x)", |
13404 | 50 |
|
51 |
"(meta_eq_to_obj_eq % TYPE('T) % x % y %% \ |
|
52 |
\ (axm.symmetric % TYPE('T) % x % y %% prf)) == \ |
|
53 |
\ (sym % TYPE('T) % x % y %% (meta_eq_to_obj_eq % TYPE('T) % x % y %% prf))", |
|
54 |
||
55 |
"(meta_eq_to_obj_eq % TYPE('T => 'U) % x1 % x2 %% \ |
|
56 |
\ (abstract_rule % TYPE('U) % TYPE('T) % f % g %% prf)) == \ |
|
57 |
\ (ext % TYPE('U) % TYPE('T) % f % g %% \ |
|
58 |
\ (Lam (x::'T). meta_eq_to_obj_eq % TYPE('U) % f x % g x %% (prf % x)))", |
|
59 |
||
60 |
"(meta_eq_to_obj_eq % TYPE('T) % x % y %% \ |
|
61 |
\ (eq_reflection % TYPE('T) % x % y %% prf)) == prf", |
|
62 |
||
63 |
"(meta_eq_to_obj_eq % TYPE('T1) % x1 % x2 %% (equal_elim % x3 % x4 %% \ |
|
64 |
\ (combination % TYPE(prop) % TYPE('T) % x7 % x8 % C % D %% \ |
|
65 |
\ (combination % TYPE('T3) % TYPE('T) % op == % op == % A % B %% \ |
|
66 |
\ (axm.reflexive % TYPE('T4) % op ==) %% prf1) %% prf2) %% prf3)) == \ |
|
67 |
\ (iffD1 % A = C % B = D %% \ |
|
68 |
\ (cong % TYPE(bool) % TYPE('T::type) % op = A % op = B % C % D %% \ |
|
69 |
\ (cong % TYPE('T=>bool) % TYPE('T) % \ |
|
70 |
\ (op = :: 'T=>'T=>bool) % (op = :: 'T=>'T=>bool) % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
71 |
\ (HOL.refl % TYPE('T=>'T=>bool) % (op = :: 'T=>'T=>bool)) %% \ |
13404 | 72 |
\ (meta_eq_to_obj_eq % TYPE('T) % A % B %% prf1)) %% \ |
73 |
\ (meta_eq_to_obj_eq % TYPE('T) % C % D %% prf2)) %% \ |
|
74 |
\ (meta_eq_to_obj_eq % TYPE('T) % C % D %% prf3))", |
|
75 |
||
76 |
"(meta_eq_to_obj_eq % TYPE('T1) % x1 % x2 %% (equal_elim % x3 % x4 %% \ |
|
77 |
\ (axm.symmetric % TYPE('T2) % x5 % x6 %% \ |
|
78 |
\ (combination % TYPE(prop) % TYPE('T) % x7 % x8 % C % D %% \ |
|
79 |
\ (combination % TYPE('T3) % TYPE('T) % op == % op == % A % B %% \ |
|
80 |
\ (axm.reflexive % TYPE('T4) % op ==) %% prf1) %% prf2)) %% prf3)) == \ |
|
81 |
\ (iffD2 % A = C % B = D %% \ |
|
82 |
\ (cong % TYPE(bool) % TYPE('T::type) % op = A % op = B % C % D %% \ |
|
83 |
\ (cong % TYPE('T=>bool) % TYPE('T) % \ |
|
84 |
\ (op = :: 'T=>'T=>bool) % (op = :: 'T=>'T=>bool) % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
85 |
\ (HOL.refl % TYPE('T=>'T=>bool) % (op = :: 'T=>'T=>bool)) %% \ |
13404 | 86 |
\ (meta_eq_to_obj_eq % TYPE('T) % A % B %% prf1)) %% \ |
87 |
\ (meta_eq_to_obj_eq % TYPE('T) % C % D %% prf2)) %% \ |
|
88 |
\ (meta_eq_to_obj_eq % TYPE('T) % C % D %% prf3))", |
|
89 |
||
90 |
(** rewriting on bool: insert proper congruence rules for logical connectives **) |
|
91 |
||
92 |
(* All *) |
|
93 |
||
94 |
"(iffD1 % All P % All Q %% (cong % TYPE('T1) % TYPE('T2) % All % All % P % Q %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
95 |
\ (HOL.refl % TYPE('T3) % x1) %% (ext % TYPE(bool) % TYPE('a) % x2 % x3 %% prf)) %% prf') == \ |
13404 | 96 |
\ (allI % TYPE('a) % Q %% \ |
97 |
\ (Lam x. \ |
|
98 |
\ iffD1 % P x % Q x %% (prf % x) %% \ |
|
99 |
\ (spec % TYPE('a) % P % x %% prf')))", |
|
100 |
||
101 |
"(iffD2 % All P % All Q %% (cong % TYPE('T1) % TYPE('T2) % All % All % P % Q %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
102 |
\ (HOL.refl % TYPE('T3) % x1) %% (ext % TYPE(bool) % TYPE('a) % x2 % x3 %% prf)) %% prf') == \ |
13404 | 103 |
\ (allI % TYPE('a) % P %% \ |
104 |
\ (Lam x. \ |
|
105 |
\ iffD2 % P x % Q x %% (prf % x) %% \ |
|
106 |
\ (spec % TYPE('a) % ?Q % x %% prf')))", |
|
107 |
||
108 |
(* Ex *) |
|
109 |
||
110 |
"(iffD1 % Ex P % Ex Q %% (cong % TYPE('T1) % TYPE('T2) % Ex % Ex % P % Q %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
111 |
\ (HOL.refl % TYPE('T3) % x1) %% (ext % TYPE(bool) % TYPE('a) % x2 % x3 %% prf)) %% prf') == \ |
13404 | 112 |
\ (exE % TYPE('a) % P % EX x. Q x %% prf' %% \ |
113 |
\ (Lam x H : P x. \ |
|
114 |
\ exI % TYPE('a) % Q % x %% \ |
|
115 |
\ (iffD1 % P x % Q x %% (prf % x) %% H)))", |
|
116 |
||
117 |
"(iffD2 % Ex P % Ex Q %% (cong % TYPE('T1) % TYPE('T2) % Ex % Ex % P % Q %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
118 |
\ (HOL.refl % TYPE('T3) % x1) %% (ext % TYPE(bool) % TYPE('a) % x2 % x3 %% prf)) %% prf') == \ |
13404 | 119 |
\ (exE % TYPE('a) % Q % EX x. P x %% prf' %% \ |
120 |
\ (Lam x H : Q x. \ |
|
121 |
\ exI % TYPE('a) % P % x %% \ |
|
122 |
\ (iffD2 % P x % Q x %% (prf % x) %% H)))", |
|
123 |
||
124 |
(* & *) |
|
125 |
||
126 |
"(iffD1 % A & C % B & D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% \ |
|
127 |
\ (cong % TYPE('T3) % TYPE('T4) % op & % op & % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
128 |
\ (HOL.refl % TYPE('T5) % op &) %% prf1) %% prf2) %% prf3) == \ |
13404 | 129 |
\ (conjI % B % D %% \ |
130 |
\ (iffD1 % A % B %% prf1 %% (conjunct1 % A % C %% prf3)) %% \ |
|
131 |
\ (iffD1 % C % D %% prf2 %% (conjunct2 % A % C %% prf3)))", |
|
132 |
||
133 |
"(iffD2 % A & C % B & D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% \ |
|
134 |
\ (cong % TYPE('T3) % TYPE('T4) % op & % op & % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
135 |
\ (HOL.refl % TYPE('T5) % op &) %% prf1) %% prf2) %% prf3) == \ |
13404 | 136 |
\ (conjI % A % C %% \ |
137 |
\ (iffD2 % A % B %% prf1 %% (conjunct1 % B % D %% prf3)) %% \ |
|
138 |
\ (iffD2 % C % D %% prf2 %% (conjunct2 % B % D %% prf3)))", |
|
139 |
||
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
140 |
"(cong % TYPE(bool) % TYPE(bool) % op & A % op & A % B % C %% \ |
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
141 |
\ (HOL.refl % TYPE(bool=>bool) % op & A)) == \ |
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
142 |
\ (cong % TYPE(bool) % TYPE(bool) % op & A % op & A % B % C %% \ |
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
143 |
\ (cong % TYPE(bool=>bool) % TYPE(bool) % \ |
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
144 |
\ (op & :: bool=>bool=>bool) % (op & :: bool=>bool=>bool) % A % A %% \ |
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
145 |
\ (HOL.refl % TYPE(bool=>bool=>bool) % (op & :: bool=>bool=>bool)) %% \ |
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
146 |
\ (HOL.refl % TYPE(bool) % A)))", |
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
147 |
|
13404 | 148 |
(* | *) |
149 |
||
150 |
"(iffD1 % A | C % B | D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% \ |
|
151 |
\ (cong % TYPE('T3) % TYPE('T4) % op | % op | % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
152 |
\ (HOL.refl % TYPE('T5) % op | ) %% prf1) %% prf2) %% prf3) == \ |
13404 | 153 |
\ (disjE % A % C % B | D %% prf3 %% \ |
154 |
\ (Lam H : A. disjI1 % B % D %% (iffD1 % A % B %% prf1 %% H)) %% \ |
|
155 |
\ (Lam H : C. disjI2 % D % B %% (iffD1 % C % D %% prf2 %% H)))", |
|
156 |
||
157 |
"(iffD2 % A | C % B | D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% \ |
|
158 |
\ (cong % TYPE('T3) % TYPE('T4) % op | % op | % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
159 |
\ (HOL.refl % TYPE('T5) % op | ) %% prf1) %% prf2) %% prf3) == \ |
13404 | 160 |
\ (disjE % B % D % A | C %% prf3 %% \ |
161 |
\ (Lam H : B. disjI1 % A % C %% (iffD2 % A % B %% prf1 %% H)) %% \ |
|
162 |
\ (Lam H : D. disjI2 % C % A %% (iffD2 % C % D %% prf2 %% H)))", |
|
163 |
||
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
164 |
"(cong % TYPE(bool) % TYPE(bool) % op | A % op | A % B % C %% \ |
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
165 |
\ (HOL.refl % TYPE(bool=>bool) % op | A)) == \ |
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
166 |
\ (cong % TYPE(bool) % TYPE(bool) % op | A % op | A % B % C %% \ |
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
167 |
\ (cong % TYPE(bool=>bool) % TYPE(bool) % \ |
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
168 |
\ (op | :: bool=>bool=>bool) % (op | :: bool=>bool=>bool) % A % A %% \ |
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
169 |
\ (HOL.refl % TYPE(bool=>bool=>bool) % (op | :: bool=>bool=>bool)) %% \ |
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
170 |
\ (HOL.refl % TYPE(bool) % A)))", |
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
171 |
|
13404 | 172 |
(* --> *) |
173 |
||
174 |
"(iffD1 % A --> C % B --> D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% \ |
|
175 |
\ (cong % TYPE('T3) % TYPE('T4) % op --> % op --> % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
176 |
\ (HOL.refl % TYPE('T5) % op --> ) %% prf1) %% prf2) %% prf3) == \ |
13404 | 177 |
\ (impI % B % D %% (Lam H: B. iffD1 % C % D %% prf2 %% \ |
178 |
\ (mp % A % C %% prf3 %% (iffD2 % A % B %% prf1 %% H))))", |
|
179 |
||
180 |
"(iffD2 % A --> C % B --> D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% \ |
|
181 |
\ (cong % TYPE('T3) % TYPE('T4) % op --> % op --> % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
182 |
\ (HOL.refl % TYPE('T5) % op --> ) %% prf1) %% prf2) %% prf3) == \ |
13404 | 183 |
\ (impI % A % C %% (Lam H: A. iffD2 % C % D %% prf2 %% \ |
184 |
\ (mp % B % D %% prf3 %% (iffD1 % A % B %% prf1 %% H))))", |
|
185 |
||
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
186 |
"(cong % TYPE(bool) % TYPE(bool) % op --> A % op --> A % B % C %% \ |
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
187 |
\ (HOL.refl % TYPE(bool=>bool) % op --> A)) == \ |
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
188 |
\ (cong % TYPE(bool) % TYPE(bool) % op --> A % op --> A % B % C %% \ |
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
189 |
\ (cong % TYPE(bool=>bool) % TYPE(bool) % \ |
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
190 |
\ (op --> :: bool=>bool=>bool) % (op --> :: bool=>bool=>bool) % A % A %% \ |
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
191 |
\ (HOL.refl % TYPE(bool=>bool=>bool) % (op --> :: bool=>bool=>bool)) %% \ |
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
192 |
\ (HOL.refl % TYPE(bool) % A)))", |
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
193 |
|
13404 | 194 |
(* ~ *) |
195 |
||
196 |
"(iffD1 % ~ P % ~ Q %% (cong % TYPE('T1) % TYPE('T2) % Not % Not % P % Q %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
197 |
\ (HOL.refl % TYPE('T3) % Not) %% prf1) %% prf2) == \ |
13404 | 198 |
\ (notI % Q %% (Lam H: Q. \ |
199 |
\ notE % P % False %% prf2 %% (iffD2 % P % Q %% prf1 %% H)))", |
|
200 |
||
201 |
"(iffD2 % ~ P % ~ Q %% (cong % TYPE('T1) % TYPE('T2) % Not % Not % P % Q %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
202 |
\ (HOL.refl % TYPE('T3) % Not) %% prf1) %% prf2) == \ |
13404 | 203 |
\ (notI % P %% (Lam H: P. \ |
204 |
\ notE % Q % False %% prf2 %% (iffD1 % P % Q %% prf1 %% H)))", |
|
205 |
||
206 |
(* = *) |
|
207 |
||
208 |
"(iffD1 % B % D %% \ |
|
209 |
\ (iffD1 % A = C % B = D %% (cong % TYPE('T1) % TYPE(bool) % x1 % x2 % C % D %% \ |
|
210 |
\ (cong % TYPE('T2) % TYPE(bool) % op = % op = % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
211 |
\ (HOL.refl % TYPE('T3) % op =) %% prf1) %% prf2) %% prf3) %% prf4) == \ |
13404 | 212 |
\ (iffD1 % C % D %% prf2 %% \ |
213 |
\ (iffD1 % A % C %% prf3 %% (iffD2 % A % B %% prf1 %% prf4)))", |
|
214 |
||
215 |
"(iffD2 % B % D %% \ |
|
216 |
\ (iffD1 % A = C % B = D %% (cong % TYPE('T1) % TYPE(bool) % x1 % x2 % C % D %% \ |
|
217 |
\ (cong % TYPE('T2) % TYPE(bool) % op = % op = % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
218 |
\ (HOL.refl % TYPE('T3) % op =) %% prf1) %% prf2) %% prf3) %% prf4) == \ |
13404 | 219 |
\ (iffD1 % A % B %% prf1 %% \ |
220 |
\ (iffD2 % A % C %% prf3 %% (iffD2 % C % D %% prf2 %% prf4)))", |
|
221 |
||
222 |
"(iffD1 % A % C %% \ |
|
223 |
\ (iffD2 % A = C % B = D %% (cong % TYPE('T1) % TYPE(bool) % x1 % x2 % C % D %% \ |
|
224 |
\ (cong % TYPE('T2) % TYPE(bool) % op = % op = % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
225 |
\ (HOL.refl % TYPE('T3) % op =) %% prf1) %% prf2) %% prf3) %% prf4)== \ |
13404 | 226 |
\ (iffD2 % C % D %% prf2 %% \ |
227 |
\ (iffD1 % B % D %% prf3 %% (iffD1 % A % B %% prf1 %% prf4)))", |
|
228 |
||
229 |
"(iffD2 % A % C %% \ |
|
230 |
\ (iffD2 % A = C % B = D %% (cong % TYPE('T1) % TYPE(bool) % x1 % x2 % C % D %% \ |
|
231 |
\ (cong % TYPE('T2) % TYPE(bool) % op = % op = % A % B %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
232 |
\ (HOL.refl % TYPE('T3) % op =) %% prf1) %% prf2) %% prf3) %% prf4) == \ |
13404 | 233 |
\ (iffD2 % A % B %% prf1 %% \ |
234 |
\ (iffD2 % B % D %% prf3 %% (iffD1 % C % D %% prf2 %% prf4)))", |
|
235 |
||
236 |
"(cong % TYPE(bool) % TYPE(bool) % op = A % op = A % B % C %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
237 |
\ (HOL.refl % TYPE(bool=>bool) % op = A)) == \ |
13404 | 238 |
\ (cong % TYPE(bool) % TYPE(bool) % op = A % op = A % B % C %% \ |
239 |
\ (cong % TYPE(bool=>bool) % TYPE(bool) % \ |
|
240 |
\ (op = :: bool=>bool=>bool) % (op = :: bool=>bool=>bool) % A % A %% \ |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
241 |
\ (HOL.refl % TYPE(bool=>bool=>bool) % (op = :: bool=>bool=>bool)) %% \ |
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
242 |
\ (HOL.refl % TYPE(bool) % A)))", |
13404 | 243 |
|
13916
f078a758e5d8
elim_cong now eta-expands proofs on the fly if required.
berghofe
parents:
13602
diff
changeset
|
244 |
(** transitivity, reflexivity, and symmetry **) |
f078a758e5d8
elim_cong now eta-expands proofs on the fly if required.
berghofe
parents:
13602
diff
changeset
|
245 |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
246 |
"(iffD1 % A % C %% (HOL.trans % TYPE(bool) % A % B % C %% prf1 %% prf2) %% prf3) == \ |
13404 | 247 |
\ (iffD1 % B % C %% prf2 %% (iffD1 % A % B %% prf1 %% prf3))", |
248 |
||
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
249 |
"(iffD2 % A % C %% (HOL.trans % TYPE(bool) % A % B % C %% prf1 %% prf2) %% prf3) == \ |
13404 | 250 |
\ (iffD2 % A % B %% prf1 %% (iffD2 % B % C %% prf2 %% prf3))", |
251 |
||
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
252 |
"(iffD1 % A % A %% (HOL.refl % TYPE(bool) % A) %% prf) == prf", |
13404 | 253 |
|
15530
6f43714517ee
Fully qualified refl and trans to avoid confusion with theorems
berghofe
parents:
14981
diff
changeset
|
254 |
"(iffD2 % A % A %% (HOL.refl % TYPE(bool) % A) %% prf) == prf", |
13404 | 255 |
|
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
256 |
"(iffD1 % A % B %% (sym % TYPE(bool) % B % A %% prf)) == (iffD2 % B % A %% prf)", |
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
257 |
|
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
258 |
"(iffD2 % A % B %% (sym % TYPE(bool) % B % A %% prf)) == (iffD1 % B % A %% prf)", |
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
259 |
|
13404 | 260 |
(** normalization of HOL proofs **) |
261 |
||
262 |
"(mp % A % B %% (impI % A % B %% prf)) == prf", |
|
263 |
||
264 |
"(impI % A % B %% (mp % A % B %% prf)) == prf", |
|
265 |
||
266 |
"(spec % TYPE('a) % P % x %% (allI % TYPE('a) % P %% prf)) == prf % x", |
|
267 |
||
268 |
"(allI % TYPE('a) % P %% (Lam x::'a. spec % TYPE('a) % P % x %% prf)) == prf", |
|
269 |
||
13602
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
270 |
"(exE % TYPE('a) % P % Q %% (exI % TYPE('a) % P % x %% prf1) %% prf2) == (prf2 % x %% prf1)", |
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
271 |
|
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
272 |
"(exE % TYPE('a) % P % Q %% prf %% (exI % TYPE('a) % P)) == prf", |
4cecd1e0f4a9
- additional congruence rules for boolean operators
berghofe
parents:
13404
diff
changeset
|
273 |
|
13404 | 274 |
"(disjE % P % Q % R %% (disjI1 % P % Q %% prf1) %% prf2 %% prf3) == (prf2 %% prf1)", |
275 |
||
276 |
"(disjE % P % Q % R %% (disjI2 % Q % P %% prf1) %% prf2 %% prf3) == (prf3 %% prf1)", |
|
277 |
||
278 |
"(conjunct1 % P % Q %% (conjI % P % Q %% prf1 %% prf2)) == prf1", |
|
279 |
||
280 |
"(conjunct2 % P % Q %% (conjI % P % Q %% prf1 %% prf2)) == prf2", |
|
281 |
||
282 |
"(iffD1 % A % B %% (iffI % A % B %% prf1 %% prf2)) == prf1", |
|
283 |
||
284 |
"(iffD2 % A % B %% (iffI % A % B %% prf1 %% prf2)) == prf2"]; |
|
285 |
||
286 |
||
287 |
(** Replace congruence rules by substitution rules **) |
|
288 |
||
15531 | 289 |
fun strip_cong ps (PThm (("HOL.cong", _), _, _, _) % _ % _ % SOME x % SOME y %% |
13404 | 290 |
prf1 %% prf2) = strip_cong (((x, y), prf2) :: ps) prf1 |
15531 | 291 |
| strip_cong ps (PThm (("HOL.refl", _), _, _, _) % SOME f) = SOME (f, ps) |
292 |
| strip_cong _ _ = NONE; |
|
13404 | 293 |
|
294 |
val subst_prf = fst (strip_combt (#2 (#der (rep_thm subst)))); |
|
295 |
val sym_prf = fst (strip_combt (#2 (#der (rep_thm sym)))); |
|
296 |
||
297 |
fun make_subst Ts prf xs (_, []) = prf |
|
298 |
| make_subst Ts prf xs (f, ((x, y), prf') :: ps) = |
|
299 |
let val T = fastype_of1 (Ts, x) |
|
300 |
in if x aconv y then make_subst Ts prf (xs @ [x]) (f, ps) |
|
15531 | 301 |
else change_type (SOME [T]) subst_prf %> x %> y %> |
13404 | 302 |
Abs ("z", T, list_comb (incr_boundvars 1 f, |
303 |
map (incr_boundvars 1) xs @ Bound 0 :: |
|
304 |
map (incr_boundvars 1 o snd o fst) ps)) %% prf' %% |
|
305 |
make_subst Ts prf (xs @ [x]) (f, ps) |
|
306 |
end; |
|
307 |
||
308 |
fun make_sym Ts ((x, y), prf) = |
|
15531 | 309 |
((y, x), change_type (SOME [fastype_of1 (Ts, x)]) sym_prf %> x %> y %% prf); |
13404 | 310 |
|
13916
f078a758e5d8
elim_cong now eta-expands proofs on the fly if required.
berghofe
parents:
13602
diff
changeset
|
311 |
fun mk_AbsP P t = AbsP ("H", P, t); |
f078a758e5d8
elim_cong now eta-expands proofs on the fly if required.
berghofe
parents:
13602
diff
changeset
|
312 |
|
13404 | 313 |
fun elim_cong Ts (PThm (("HOL.iffD1", _), _, _, _) % _ % _ %% prf1 %% prf2) = |
15570 | 314 |
Option.map (make_subst Ts prf2 []) (strip_cong [] prf1) |
13916
f078a758e5d8
elim_cong now eta-expands proofs on the fly if required.
berghofe
parents:
13602
diff
changeset
|
315 |
| elim_cong Ts (PThm (("HOL.iffD1", _), _, _, _) % P % _ %% prf) = |
15570 | 316 |
Option.map (mk_AbsP P o make_subst Ts (PBound 0) []) |
13916
f078a758e5d8
elim_cong now eta-expands proofs on the fly if required.
berghofe
parents:
13602
diff
changeset
|
317 |
(strip_cong [] (incr_pboundvars 1 0 prf)) |
13404 | 318 |
| elim_cong Ts (PThm (("HOL.iffD2", _), _, _, _) % _ % _ %% prf1 %% prf2) = |
15570 | 319 |
Option.map (make_subst Ts prf2 [] o |
13404 | 320 |
apsnd (map (make_sym Ts))) (strip_cong [] prf1) |
13916
f078a758e5d8
elim_cong now eta-expands proofs on the fly if required.
berghofe
parents:
13602
diff
changeset
|
321 |
| elim_cong Ts (PThm (("HOL.iffD2", _), _, _, _) % _ % P %% prf) = |
15570 | 322 |
Option.map (mk_AbsP P o make_subst Ts (PBound 0) [] o |
13916
f078a758e5d8
elim_cong now eta-expands proofs on the fly if required.
berghofe
parents:
13602
diff
changeset
|
323 |
apsnd (map (make_sym Ts))) (strip_cong [] (incr_pboundvars 1 0 prf)) |
15531 | 324 |
| elim_cong _ _ = NONE; |
13404 | 325 |
|
326 |
end; |