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