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