author | haftmann |
Tue, 16 Oct 2007 23:12:45 +0200 | |
changeset 25062 | af5ef0d4d655 |
parent 23467 | d1b97708d5eb |
child 26391 | 6e8aa5a4eb82 |
permissions | -rw-r--r-- |
17441 | 1 |
(* Title: CTT/CTT.thy |
0 | 2 |
ID: $Id$ |
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
*) |
|
6 |
||
17441 | 7 |
header {* Constructive Type Theory *} |
0 | 8 |
|
17441 | 9 |
theory CTT |
10 |
imports Pure |
|
19761 | 11 |
uses "~~/src/Provers/typedsimp.ML" ("rew.ML") |
17441 | 12 |
begin |
13 |
||
14 |
typedecl i |
|
15 |
typedecl t |
|
16 |
typedecl o |
|
0 | 17 |
|
18 |
consts |
|
19 |
(*Types*) |
|
17441 | 20 |
F :: "t" |
21 |
T :: "t" (*F is empty, T contains one element*) |
|
0 | 22 |
contr :: "i=>i" |
23 |
tt :: "i" |
|
24 |
(*Natural numbers*) |
|
25 |
N :: "t" |
|
26 |
succ :: "i=>i" |
|
27 |
rec :: "[i, i, [i,i]=>i] => i" |
|
28 |
(*Unions*) |
|
17441 | 29 |
inl :: "i=>i" |
30 |
inr :: "i=>i" |
|
0 | 31 |
when :: "[i, i=>i, i=>i]=>i" |
32 |
(*General Sum and Binary Product*) |
|
33 |
Sum :: "[t, i=>t]=>t" |
|
17441 | 34 |
fst :: "i=>i" |
35 |
snd :: "i=>i" |
|
0 | 36 |
split :: "[i, [i,i]=>i] =>i" |
37 |
(*General Product and Function Space*) |
|
38 |
Prod :: "[t, i=>t]=>t" |
|
14765 | 39 |
(*Types*) |
22808 | 40 |
Plus :: "[t,t]=>t" (infixr "+" 40) |
0 | 41 |
(*Equality type*) |
42 |
Eq :: "[t,i,i]=>t" |
|
43 |
eq :: "i" |
|
44 |
(*Judgements*) |
|
45 |
Type :: "t => prop" ("(_ type)" [10] 5) |
|
10467
e6e7205e9e91
x-symbol support for Pi, Sigma, -->, : (membership)
paulson
parents:
3837
diff
changeset
|
46 |
Eqtype :: "[t,t]=>prop" ("(_ =/ _)" [10,10] 5) |
0 | 47 |
Elem :: "[i, t]=>prop" ("(_ /: _)" [10,10] 5) |
10467
e6e7205e9e91
x-symbol support for Pi, Sigma, -->, : (membership)
paulson
parents:
3837
diff
changeset
|
48 |
Eqelem :: "[i,i,t]=>prop" ("(_ =/ _ :/ _)" [10,10,10] 5) |
0 | 49 |
Reduce :: "[i,i]=>prop" ("Reduce[_,_]") |
50 |
(*Types*) |
|
14765 | 51 |
|
0 | 52 |
(*Functions*) |
53 |
lambda :: "(i => i) => i" (binder "lam " 10) |
|
22808 | 54 |
app :: "[i,i]=>i" (infixl "`" 60) |
0 | 55 |
(*Natural numbers*) |
56 |
"0" :: "i" ("0") |
|
57 |
(*Pairing*) |
|
58 |
pair :: "[i,i]=>i" ("(1<_,/_>)") |
|
59 |
||
14765 | 60 |
syntax |
19761 | 61 |
"_PROD" :: "[idt,t,t]=>t" ("(3PROD _:_./ _)" 10) |
62 |
"_SUM" :: "[idt,t,t]=>t" ("(3SUM _:_./ _)" 10) |
|
0 | 63 |
translations |
19761 | 64 |
"PROD x:A. B" == "Prod(A, %x. B)" |
65 |
"SUM x:A. B" == "Sum(A, %x. B)" |
|
66 |
||
67 |
abbreviation |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
68 |
Arrow :: "[t,t]=>t" (infixr "-->" 30) where |
19761 | 69 |
"A --> B == PROD _:A. B" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
70 |
abbreviation |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
71 |
Times :: "[t,t]=>t" (infixr "*" 50) where |
19761 | 72 |
"A * B == SUM _:A. B" |
0 | 73 |
|
21210 | 74 |
notation (xsymbols) |
21524 | 75 |
lambda (binder "\<lambda>\<lambda>" 10) and |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
76 |
Elem ("(_ /\<in> _)" [10,10] 5) and |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
77 |
Eqelem ("(2_ =/ _ \<in>/ _)" [10,10,10] 5) and |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
78 |
Arrow (infixr "\<longrightarrow>" 30) and |
19761 | 79 |
Times (infixr "\<times>" 50) |
17441 | 80 |
|
21210 | 81 |
notation (HTML output) |
21524 | 82 |
lambda (binder "\<lambda>\<lambda>" 10) and |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
83 |
Elem ("(_ /\<in> _)" [10,10] 5) and |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
84 |
Eqelem ("(2_ =/ _ \<in>/ _)" [10,10,10] 5) and |
19761 | 85 |
Times (infixr "\<times>" 50) |
17441 | 86 |
|
10467
e6e7205e9e91
x-symbol support for Pi, Sigma, -->, : (membership)
paulson
parents:
3837
diff
changeset
|
87 |
syntax (xsymbols) |
21524 | 88 |
"_PROD" :: "[idt,t,t] => t" ("(3\<Pi> _\<in>_./ _)" 10) |
89 |
"_SUM" :: "[idt,t,t] => t" ("(3\<Sigma> _\<in>_./ _)" 10) |
|
10467
e6e7205e9e91
x-symbol support for Pi, Sigma, -->, : (membership)
paulson
parents:
3837
diff
changeset
|
90 |
|
14565 | 91 |
syntax (HTML output) |
21524 | 92 |
"_PROD" :: "[idt,t,t] => t" ("(3\<Pi> _\<in>_./ _)" 10) |
93 |
"_SUM" :: "[idt,t,t] => t" ("(3\<Sigma> _\<in>_./ _)" 10) |
|
14565 | 94 |
|
17441 | 95 |
axioms |
0 | 96 |
|
97 |
(*Reduction: a weaker notion than equality; a hack for simplification. |
|
98 |
Reduce[a,b] means either that a=b:A for some A or else that "a" and "b" |
|
99 |
are textually identical.*) |
|
100 |
||
101 |
(*does not verify a:A! Sound because only trans_red uses a Reduce premise |
|
102 |
No new theorems can be proved about the standard judgements.*) |
|
17441 | 103 |
refl_red: "Reduce[a,a]" |
104 |
red_if_equal: "a = b : A ==> Reduce[a,b]" |
|
105 |
trans_red: "[| a = b : A; Reduce[b,c] |] ==> a = c : A" |
|
0 | 106 |
|
107 |
(*Reflexivity*) |
|
108 |
||
17441 | 109 |
refl_type: "A type ==> A = A" |
110 |
refl_elem: "a : A ==> a = a : A" |
|
0 | 111 |
|
112 |
(*Symmetry*) |
|
113 |
||
17441 | 114 |
sym_type: "A = B ==> B = A" |
115 |
sym_elem: "a = b : A ==> b = a : A" |
|
0 | 116 |
|
117 |
(*Transitivity*) |
|
118 |
||
17441 | 119 |
trans_type: "[| A = B; B = C |] ==> A = C" |
120 |
trans_elem: "[| a = b : A; b = c : A |] ==> a = c : A" |
|
0 | 121 |
|
17441 | 122 |
equal_types: "[| a : A; A = B |] ==> a : B" |
123 |
equal_typesL: "[| a = b : A; A = B |] ==> a = b : B" |
|
0 | 124 |
|
125 |
(*Substitution*) |
|
126 |
||
17441 | 127 |
subst_type: "[| a : A; !!z. z:A ==> B(z) type |] ==> B(a) type" |
128 |
subst_typeL: "[| a = c : A; !!z. z:A ==> B(z) = D(z) |] ==> B(a) = D(c)" |
|
0 | 129 |
|
17441 | 130 |
subst_elem: "[| a : A; !!z. z:A ==> b(z):B(z) |] ==> b(a):B(a)" |
131 |
subst_elemL: |
|
0 | 132 |
"[| a=c : A; !!z. z:A ==> b(z)=d(z) : B(z) |] ==> b(a)=d(c) : B(a)" |
133 |
||
134 |
||
135 |
(*The type N -- natural numbers*) |
|
136 |
||
17441 | 137 |
NF: "N type" |
138 |
NI0: "0 : N" |
|
139 |
NI_succ: "a : N ==> succ(a) : N" |
|
140 |
NI_succL: "a = b : N ==> succ(a) = succ(b) : N" |
|
0 | 141 |
|
17441 | 142 |
NE: |
143 |
"[| p: N; a: C(0); !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] |
|
3837 | 144 |
==> rec(p, a, %u v. b(u,v)) : C(p)" |
0 | 145 |
|
17441 | 146 |
NEL: |
147 |
"[| p = q : N; a = c : C(0); |
|
148 |
!!u v. [| u: N; v: C(u) |] ==> b(u,v) = d(u,v): C(succ(u)) |] |
|
3837 | 149 |
==> rec(p, a, %u v. b(u,v)) = rec(q,c,d) : C(p)" |
0 | 150 |
|
17441 | 151 |
NC0: |
152 |
"[| a: C(0); !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] |
|
3837 | 153 |
==> rec(0, a, %u v. b(u,v)) = a : C(0)" |
0 | 154 |
|
17441 | 155 |
NC_succ: |
156 |
"[| p: N; a: C(0); |
|
157 |
!!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] ==> |
|
3837 | 158 |
rec(succ(p), a, %u v. b(u,v)) = b(p, rec(p, a, %u v. b(u,v))) : C(succ(p))" |
0 | 159 |
|
160 |
(*The fourth Peano axiom. See page 91 of Martin-Lof's book*) |
|
17441 | 161 |
zero_ne_succ: |
0 | 162 |
"[| a: N; 0 = succ(a) : N |] ==> 0: F" |
163 |
||
164 |
||
165 |
(*The Product of a family of types*) |
|
166 |
||
17441 | 167 |
ProdF: "[| A type; !!x. x:A ==> B(x) type |] ==> PROD x:A. B(x) type" |
0 | 168 |
|
17441 | 169 |
ProdFL: |
170 |
"[| A = C; !!x. x:A ==> B(x) = D(x) |] ==> |
|
3837 | 171 |
PROD x:A. B(x) = PROD x:C. D(x)" |
0 | 172 |
|
17441 | 173 |
ProdI: |
3837 | 174 |
"[| A type; !!x. x:A ==> b(x):B(x)|] ==> lam x. b(x) : PROD x:A. B(x)" |
0 | 175 |
|
17441 | 176 |
ProdIL: |
177 |
"[| A type; !!x. x:A ==> b(x) = c(x) : B(x)|] ==> |
|
3837 | 178 |
lam x. b(x) = lam x. c(x) : PROD x:A. B(x)" |
0 | 179 |
|
17441 | 180 |
ProdE: "[| p : PROD x:A. B(x); a : A |] ==> p`a : B(a)" |
181 |
ProdEL: "[| p=q: PROD x:A. B(x); a=b : A |] ==> p`a = q`b : B(a)" |
|
0 | 182 |
|
17441 | 183 |
ProdC: |
184 |
"[| a : A; !!x. x:A ==> b(x) : B(x)|] ==> |
|
3837 | 185 |
(lam x. b(x)) ` a = b(a) : B(a)" |
0 | 186 |
|
17441 | 187 |
ProdC2: |
3837 | 188 |
"p : PROD x:A. B(x) ==> (lam x. p`x) = p : PROD x:A. B(x)" |
0 | 189 |
|
190 |
||
191 |
(*The Sum of a family of types*) |
|
192 |
||
17441 | 193 |
SumF: "[| A type; !!x. x:A ==> B(x) type |] ==> SUM x:A. B(x) type" |
194 |
SumFL: |
|
3837 | 195 |
"[| A = C; !!x. x:A ==> B(x) = D(x) |] ==> SUM x:A. B(x) = SUM x:C. D(x)" |
0 | 196 |
|
17441 | 197 |
SumI: "[| a : A; b : B(a) |] ==> <a,b> : SUM x:A. B(x)" |
198 |
SumIL: "[| a=c:A; b=d:B(a) |] ==> <a,b> = <c,d> : SUM x:A. B(x)" |
|
0 | 199 |
|
17441 | 200 |
SumE: |
201 |
"[| p: SUM x:A. B(x); !!x y. [| x:A; y:B(x) |] ==> c(x,y): C(<x,y>) |] |
|
3837 | 202 |
==> split(p, %x y. c(x,y)) : C(p)" |
0 | 203 |
|
17441 | 204 |
SumEL: |
205 |
"[| p=q : SUM x:A. B(x); |
|
206 |
!!x y. [| x:A; y:B(x) |] ==> c(x,y)=d(x,y): C(<x,y>)|] |
|
3837 | 207 |
==> split(p, %x y. c(x,y)) = split(q, % x y. d(x,y)) : C(p)" |
0 | 208 |
|
17441 | 209 |
SumC: |
210 |
"[| a: A; b: B(a); !!x y. [| x:A; y:B(x) |] ==> c(x,y): C(<x,y>) |] |
|
3837 | 211 |
==> split(<a,b>, %x y. c(x,y)) = c(a,b) : C(<a,b>)" |
0 | 212 |
|
17441 | 213 |
fst_def: "fst(a) == split(a, %x y. x)" |
214 |
snd_def: "snd(a) == split(a, %x y. y)" |
|
0 | 215 |
|
216 |
||
217 |
(*The sum of two types*) |
|
218 |
||
17441 | 219 |
PlusF: "[| A type; B type |] ==> A+B type" |
220 |
PlusFL: "[| A = C; B = D |] ==> A+B = C+D" |
|
0 | 221 |
|
17441 | 222 |
PlusI_inl: "[| a : A; B type |] ==> inl(a) : A+B" |
223 |
PlusI_inlL: "[| a = c : A; B type |] ==> inl(a) = inl(c) : A+B" |
|
0 | 224 |
|
17441 | 225 |
PlusI_inr: "[| A type; b : B |] ==> inr(b) : A+B" |
226 |
PlusI_inrL: "[| A type; b = d : B |] ==> inr(b) = inr(d) : A+B" |
|
0 | 227 |
|
17441 | 228 |
PlusE: |
229 |
"[| p: A+B; !!x. x:A ==> c(x): C(inl(x)); |
|
230 |
!!y. y:B ==> d(y): C(inr(y)) |] |
|
3837 | 231 |
==> when(p, %x. c(x), %y. d(y)) : C(p)" |
0 | 232 |
|
17441 | 233 |
PlusEL: |
234 |
"[| p = q : A+B; !!x. x: A ==> c(x) = e(x) : C(inl(x)); |
|
235 |
!!y. y: B ==> d(y) = f(y) : C(inr(y)) |] |
|
3837 | 236 |
==> when(p, %x. c(x), %y. d(y)) = when(q, %x. e(x), %y. f(y)) : C(p)" |
0 | 237 |
|
17441 | 238 |
PlusC_inl: |
239 |
"[| a: A; !!x. x:A ==> c(x): C(inl(x)); |
|
240 |
!!y. y:B ==> d(y): C(inr(y)) |] |
|
3837 | 241 |
==> when(inl(a), %x. c(x), %y. d(y)) = c(a) : C(inl(a))" |
0 | 242 |
|
17441 | 243 |
PlusC_inr: |
244 |
"[| b: B; !!x. x:A ==> c(x): C(inl(x)); |
|
245 |
!!y. y:B ==> d(y): C(inr(y)) |] |
|
3837 | 246 |
==> when(inr(b), %x. c(x), %y. d(y)) = d(b) : C(inr(b))" |
0 | 247 |
|
248 |
||
249 |
(*The type Eq*) |
|
250 |
||
17441 | 251 |
EqF: "[| A type; a : A; b : A |] ==> Eq(A,a,b) type" |
252 |
EqFL: "[| A=B; a=c: A; b=d : A |] ==> Eq(A,a,b) = Eq(B,c,d)" |
|
253 |
EqI: "a = b : A ==> eq : Eq(A,a,b)" |
|
254 |
EqE: "p : Eq(A,a,b) ==> a = b : A" |
|
0 | 255 |
|
256 |
(*By equality of types, can prove C(p) from C(eq), an elimination rule*) |
|
17441 | 257 |
EqC: "p : Eq(A,a,b) ==> p = eq : Eq(A,a,b)" |
0 | 258 |
|
259 |
(*The type F*) |
|
260 |
||
17441 | 261 |
FF: "F type" |
262 |
FE: "[| p: F; C type |] ==> contr(p) : C" |
|
263 |
FEL: "[| p = q : F; C type |] ==> contr(p) = contr(q) : C" |
|
0 | 264 |
|
265 |
(*The type T |
|
266 |
Martin-Lof's book (page 68) discusses elimination and computation. |
|
267 |
Elimination can be derived by computation and equality of types, |
|
268 |
but with an extra premise C(x) type x:T. |
|
269 |
Also computation can be derived from elimination. *) |
|
270 |
||
17441 | 271 |
TF: "T type" |
272 |
TI: "tt : T" |
|
273 |
TE: "[| p : T; c : C(tt) |] ==> c : C(p)" |
|
274 |
TEL: "[| p = q : T; c = d : C(tt) |] ==> c = d : C(p)" |
|
275 |
TC: "p : T ==> p = tt : T" |
|
0 | 276 |
|
19761 | 277 |
|
278 |
subsection "Tactics and derived rules for Constructive Type Theory" |
|
279 |
||
280 |
(*Formation rules*) |
|
281 |
lemmas form_rls = NF ProdF SumF PlusF EqF FF TF |
|
282 |
and formL_rls = ProdFL SumFL PlusFL EqFL |
|
283 |
||
284 |
(*Introduction rules |
|
285 |
OMITTED: EqI, because its premise is an eqelem, not an elem*) |
|
286 |
lemmas intr_rls = NI0 NI_succ ProdI SumI PlusI_inl PlusI_inr TI |
|
287 |
and intrL_rls = NI_succL ProdIL SumIL PlusI_inlL PlusI_inrL |
|
288 |
||
289 |
(*Elimination rules |
|
290 |
OMITTED: EqE, because its conclusion is an eqelem, not an elem |
|
291 |
TE, because it does not involve a constructor *) |
|
292 |
lemmas elim_rls = NE ProdE SumE PlusE FE |
|
293 |
and elimL_rls = NEL ProdEL SumEL PlusEL FEL |
|
294 |
||
295 |
(*OMITTED: eqC are TC because they make rewriting loop: p = un = un = ... *) |
|
296 |
lemmas comp_rls = NC0 NC_succ ProdC SumC PlusC_inl PlusC_inr |
|
297 |
||
298 |
(*rules with conclusion a:A, an elem judgement*) |
|
299 |
lemmas element_rls = intr_rls elim_rls |
|
300 |
||
301 |
(*Definitions are (meta)equality axioms*) |
|
302 |
lemmas basic_defs = fst_def snd_def |
|
303 |
||
304 |
(*Compare with standard version: B is applied to UNSIMPLIFIED expression! *) |
|
305 |
lemma SumIL2: "[| c=a : A; d=b : B(a) |] ==> <c,d> = <a,b> : Sum(A,B)" |
|
306 |
apply (rule sym_elem) |
|
307 |
apply (rule SumIL) |
|
308 |
apply (rule_tac [!] sym_elem) |
|
309 |
apply assumption+ |
|
310 |
done |
|
311 |
||
312 |
lemmas intrL2_rls = NI_succL ProdIL SumIL2 PlusI_inlL PlusI_inrL |
|
313 |
||
314 |
(*Exploit p:Prod(A,B) to create the assumption z:B(a). |
|
315 |
A more natural form of product elimination. *) |
|
316 |
lemma subst_prodE: |
|
317 |
assumes "p: Prod(A,B)" |
|
318 |
and "a: A" |
|
319 |
and "!!z. z: B(a) ==> c(z): C(z)" |
|
320 |
shows "c(p`a): C(p`a)" |
|
321 |
apply (rule prems ProdE)+ |
|
322 |
done |
|
323 |
||
324 |
||
325 |
subsection {* Tactics for type checking *} |
|
326 |
||
327 |
ML {* |
|
328 |
||
329 |
local |
|
330 |
||
331 |
fun is_rigid_elem (Const("CTT.Elem",_) $ a $ _) = not(is_Var (head_of a)) |
|
332 |
| is_rigid_elem (Const("CTT.Eqelem",_) $ a $ _ $ _) = not(is_Var (head_of a)) |
|
333 |
| is_rigid_elem (Const("CTT.Type",_) $ a) = not(is_Var (head_of a)) |
|
334 |
| is_rigid_elem _ = false |
|
335 |
||
336 |
in |
|
337 |
||
338 |
(*Try solving a:A or a=b:A by assumption provided a is rigid!*) |
|
339 |
val test_assume_tac = SUBGOAL(fn (prem,i) => |
|
340 |
if is_rigid_elem (Logic.strip_assums_concl prem) |
|
341 |
then assume_tac i else no_tac) |
|
342 |
||
343 |
fun ASSUME tf i = test_assume_tac i ORELSE tf i |
|
344 |
||
345 |
end; |
|
346 |
||
347 |
*} |
|
348 |
||
349 |
(*For simplification: type formation and checking, |
|
350 |
but no equalities between terms*) |
|
351 |
lemmas routine_rls = form_rls formL_rls refl_type element_rls |
|
352 |
||
353 |
ML {* |
|
354 |
local |
|
355 |
val routine_rls = thms "routine_rls"; |
|
356 |
val form_rls = thms "form_rls"; |
|
357 |
val intr_rls = thms "intr_rls"; |
|
358 |
val element_rls = thms "element_rls"; |
|
359 |
val equal_rls = form_rls @ element_rls @ thms "intrL_rls" @ |
|
360 |
thms "elimL_rls" @ thms "refl_elem" |
|
361 |
in |
|
362 |
||
363 |
fun routine_tac rls prems = ASSUME (filt_resolve_tac (prems @ rls) 4); |
|
364 |
||
365 |
(*Solve all subgoals "A type" using formation rules. *) |
|
366 |
val form_tac = REPEAT_FIRST (ASSUME (filt_resolve_tac(form_rls) 1)); |
|
367 |
||
368 |
(*Type checking: solve a:A (a rigid, A flexible) by intro and elim rules. *) |
|
369 |
fun typechk_tac thms = |
|
370 |
let val tac = filt_resolve_tac (thms @ form_rls @ element_rls) 3 |
|
371 |
in REPEAT_FIRST (ASSUME tac) end |
|
372 |
||
373 |
(*Solve a:A (a flexible, A rigid) by introduction rules. |
|
374 |
Cannot use stringtrees (filt_resolve_tac) since |
|
375 |
goals like ?a:SUM(A,B) have a trivial head-string *) |
|
376 |
fun intr_tac thms = |
|
377 |
let val tac = filt_resolve_tac(thms@form_rls@intr_rls) 1 |
|
378 |
in REPEAT_FIRST (ASSUME tac) end |
|
379 |
||
380 |
(*Equality proving: solve a=b:A (where a is rigid) by long rules. *) |
|
381 |
fun equal_tac thms = |
|
382 |
REPEAT_FIRST (ASSUME (filt_resolve_tac (thms @ equal_rls) 3)) |
|
0 | 383 |
|
17441 | 384 |
end |
19761 | 385 |
|
386 |
*} |
|
387 |
||
388 |
||
389 |
subsection {* Simplification *} |
|
390 |
||
391 |
(*To simplify the type in a goal*) |
|
392 |
lemma replace_type: "[| B = A; a : A |] ==> a : B" |
|
393 |
apply (rule equal_types) |
|
394 |
apply (rule_tac [2] sym_type) |
|
395 |
apply assumption+ |
|
396 |
done |
|
397 |
||
398 |
(*Simplify the parameter of a unary type operator.*) |
|
399 |
lemma subst_eqtyparg: |
|
23467 | 400 |
assumes 1: "a=c : A" |
401 |
and 2: "!!z. z:A ==> B(z) type" |
|
19761 | 402 |
shows "B(a)=B(c)" |
403 |
apply (rule subst_typeL) |
|
404 |
apply (rule_tac [2] refl_type) |
|
23467 | 405 |
apply (rule 1) |
406 |
apply (erule 2) |
|
19761 | 407 |
done |
408 |
||
409 |
(*Simplification rules for Constructive Type Theory*) |
|
410 |
lemmas reduction_rls = comp_rls [THEN trans_elem] |
|
411 |
||
412 |
ML {* |
|
413 |
local |
|
414 |
val EqI = thm "EqI"; |
|
415 |
val EqE = thm "EqE"; |
|
416 |
val NE = thm "NE"; |
|
417 |
val FE = thm "FE"; |
|
418 |
val ProdI = thm "ProdI"; |
|
419 |
val SumI = thm "SumI"; |
|
420 |
val SumE = thm "SumE"; |
|
421 |
val PlusE = thm "PlusE"; |
|
422 |
val PlusI_inl = thm "PlusI_inl"; |
|
423 |
val PlusI_inr = thm "PlusI_inr"; |
|
424 |
val subst_prodE = thm "subst_prodE"; |
|
425 |
val intr_rls = thms "intr_rls"; |
|
426 |
in |
|
427 |
||
428 |
(*Converts each goal "e : Eq(A,a,b)" into "a=b:A" for simplification. |
|
429 |
Uses other intro rules to avoid changing flexible goals.*) |
|
430 |
val eqintr_tac = REPEAT_FIRST (ASSUME (filt_resolve_tac(EqI::intr_rls) 1)) |
|
431 |
||
432 |
(** Tactics that instantiate CTT-rules. |
|
433 |
Vars in the given terms will be incremented! |
|
434 |
The (rtac EqE i) lets them apply to equality judgements. **) |
|
435 |
||
436 |
fun NE_tac (sp: string) i = |
|
437 |
TRY (rtac EqE i) THEN res_inst_tac [ ("p",sp) ] NE i |
|
438 |
||
439 |
fun SumE_tac (sp: string) i = |
|
440 |
TRY (rtac EqE i) THEN res_inst_tac [ ("p",sp) ] SumE i |
|
441 |
||
442 |
fun PlusE_tac (sp: string) i = |
|
443 |
TRY (rtac EqE i) THEN res_inst_tac [ ("p",sp) ] PlusE i |
|
444 |
||
445 |
(** Predicate logic reasoning, WITH THINNING!! Procedures adapted from NJ. **) |
|
446 |
||
447 |
(*Finds f:Prod(A,B) and a:A in the assumptions, concludes there is z:B(a) *) |
|
448 |
fun add_mp_tac i = |
|
449 |
rtac subst_prodE i THEN assume_tac i THEN assume_tac i |
|
450 |
||
451 |
(*Finds P-->Q and P in the assumptions, replaces implication by Q *) |
|
452 |
fun mp_tac i = etac subst_prodE i THEN assume_tac i |
|
453 |
||
454 |
(*"safe" when regarded as predicate calculus rules*) |
|
455 |
val safe_brls = sort (make_ord lessb) |
|
456 |
[ (true,FE), (true,asm_rl), |
|
457 |
(false,ProdI), (true,SumE), (true,PlusE) ] |
|
458 |
||
459 |
val unsafe_brls = |
|
460 |
[ (false,PlusI_inl), (false,PlusI_inr), (false,SumI), |
|
461 |
(true,subst_prodE) ] |
|
462 |
||
463 |
(*0 subgoals vs 1 or more*) |
|
464 |
val (safe0_brls, safep_brls) = |
|
465 |
List.partition (curry (op =) 0 o subgoals_of_brl) safe_brls |
|
466 |
||
467 |
fun safestep_tac thms i = |
|
468 |
form_tac ORELSE |
|
469 |
resolve_tac thms i ORELSE |
|
470 |
biresolve_tac safe0_brls i ORELSE mp_tac i ORELSE |
|
471 |
DETERM (biresolve_tac safep_brls i) |
|
472 |
||
473 |
fun safe_tac thms i = DEPTH_SOLVE_1 (safestep_tac thms i) |
|
474 |
||
475 |
fun step_tac thms = safestep_tac thms ORELSE' biresolve_tac unsafe_brls |
|
476 |
||
477 |
(*Fails unless it solves the goal!*) |
|
478 |
fun pc_tac thms = DEPTH_SOLVE_1 o (step_tac thms) |
|
479 |
||
480 |
end |
|
481 |
*} |
|
482 |
||
483 |
use "rew.ML" |
|
484 |
||
485 |
||
486 |
subsection {* The elimination rules for fst/snd *} |
|
487 |
||
488 |
lemma SumE_fst: "p : Sum(A,B) ==> fst(p) : A" |
|
489 |
apply (unfold basic_defs) |
|
490 |
apply (erule SumE) |
|
491 |
apply assumption |
|
492 |
done |
|
493 |
||
494 |
(*The first premise must be p:Sum(A,B) !!*) |
|
495 |
lemma SumE_snd: |
|
496 |
assumes major: "p: Sum(A,B)" |
|
497 |
and "A type" |
|
498 |
and "!!x. x:A ==> B(x) type" |
|
499 |
shows "snd(p) : B(fst(p))" |
|
500 |
apply (unfold basic_defs) |
|
501 |
apply (rule major [THEN SumE]) |
|
502 |
apply (rule SumC [THEN subst_eqtyparg, THEN replace_type]) |
|
503 |
apply (tactic {* typechk_tac (thms "prems") *}) |
|
504 |
done |
|
505 |
||
506 |
end |