author | huffman |
Tue, 06 Sep 2005 23:11:09 +0200 | |
changeset 17295 | fadf6e1faa16 |
parent 16417 | 9bc16273c2d4 |
child 18372 | 2bffdf62fe7f |
permissions | -rw-r--r-- |
12229 | 1 |
(* Title: ZF/Induct/Brouwer.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1994 University of Cambridge |
|
5 |
*) |
|
6 |
||
7 |
header {* Infinite branching datatype definitions *} |
|
8 |
||
16417 | 9 |
theory Brouwer imports Main_ZFC begin |
12229 | 10 |
|
11 |
subsection {* The Brouwer ordinals *} |
|
12 |
||
13 |
consts |
|
14 |
brouwer :: i |
|
15 |
||
12610 | 16 |
datatype \<subseteq> "Vfrom(0, csucc(nat))" |
17 |
"brouwer" = Zero | Suc ("b \<in> brouwer") | Lim ("h \<in> nat -> brouwer") |
|
12229 | 18 |
monos Pi_mono |
13137 | 19 |
type_intros inf_datatype_intros |
12229 | 20 |
|
21 |
lemma brouwer_unfold: "brouwer = {0} + brouwer + (nat -> brouwer)" |
|
22 |
by (fast intro!: brouwer.intros [unfolded brouwer.con_defs] |
|
23 |
elim: brouwer.cases [unfolded brouwer.con_defs]) |
|
24 |
||
25 |
lemma brouwer_induct2: |
|
12610 | 26 |
"[| b \<in> brouwer; |
12229 | 27 |
P(Zero); |
12610 | 28 |
!!b. [| b \<in> brouwer; P(b) |] ==> P(Suc(b)); |
29 |
!!h. [| h \<in> nat -> brouwer; \<forall>i \<in> nat. P(h`i) |
|
12229 | 30 |
|] ==> P(Lim(h)) |
31 |
|] ==> P(b)" |
|
32 |
-- {* A nicer induction rule than the standard one. *} |
|
33 |
proof - |
|
34 |
case rule_context |
|
12610 | 35 |
assume "b \<in> brouwer" |
12229 | 36 |
thus ?thesis |
37 |
apply induct |
|
38 |
apply (assumption | rule rule_context)+ |
|
39 |
apply (fast elim: fun_weaken_type) |
|
40 |
apply (fast dest: apply_type) |
|
41 |
done |
|
42 |
qed |
|
43 |
||
44 |
||
45 |
subsection {* The Martin-Löf wellordering type *} |
|
46 |
||
47 |
consts |
|
48 |
Well :: "[i, i => i] => i" |
|
49 |
||
12610 | 50 |
datatype \<subseteq> "Vfrom(A \<union> (\<Union>x \<in> A. B(x)), csucc(nat \<union> |\<Union>x \<in> A. B(x)|))" |
12229 | 51 |
-- {* The union with @{text nat} ensures that the cardinal is infinite. *} |
12610 | 52 |
"Well(A, B)" = Sup ("a \<in> A", "f \<in> B(a) -> Well(A, B)") |
12229 | 53 |
monos Pi_mono |
13137 | 54 |
type_intros le_trans [OF UN_upper_cardinal le_nat_Un_cardinal] inf_datatype_intros |
12229 | 55 |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
13137
diff
changeset
|
56 |
lemma Well_unfold: "Well(A, B) = (\<Sigma> x \<in> A. B(x) -> Well(A, B))" |
12229 | 57 |
by (fast intro!: Well.intros [unfolded Well.con_defs] |
58 |
elim: Well.cases [unfolded Well.con_defs]) |
|
59 |
||
60 |
||
61 |
lemma Well_induct2: |
|
12610 | 62 |
"[| w \<in> Well(A, B); |
63 |
!!a f. [| a \<in> A; f \<in> B(a) -> Well(A,B); \<forall>y \<in> B(a). P(f`y) |
|
12229 | 64 |
|] ==> P(Sup(a,f)) |
65 |
|] ==> P(w)" |
|
66 |
-- {* A nicer induction rule than the standard one. *} |
|
67 |
proof - |
|
68 |
case rule_context |
|
12610 | 69 |
assume "w \<in> Well(A, B)" |
12229 | 70 |
thus ?thesis |
71 |
apply induct |
|
72 |
apply (assumption | rule rule_context)+ |
|
73 |
apply (fast elim: fun_weaken_type) |
|
74 |
apply (fast dest: apply_type) |
|
75 |
done |
|
76 |
qed |
|
77 |
||
12610 | 78 |
lemma Well_bool_unfold: "Well(bool, \<lambda>x. x) = 1 + (1 -> Well(bool, \<lambda>x. x))" |
12229 | 79 |
-- {* In fact it's isomorphic to @{text nat}, but we need a recursion operator *} |
80 |
-- {* for @{text Well} to prove this. *} |
|
81 |
apply (rule Well_unfold [THEN trans]) |
|
82 |
apply (simp add: Sigma_bool Pi_empty1 succ_def) |
|
83 |
done |
|
84 |
||
85 |
end |