author | wenzelm |
Tue, 08 Jan 2002 21:02:15 +0100 | |
changeset 12678 | 4d36d8df29fa |
parent 12667 | 7e6eaaa125f2 |
child 12763 | 6cecd9dfd53f |
permissions | -rw-r--r-- |
2469 | 1 |
(* Title: ZF/AC/OrdQuant.thy |
2 |
ID: $Id$ |
|
3 |
Authors: Krzysztof Grabczewski and L C Paulson |
|
4 |
||
5 |
Quantifiers and union operator for ordinals. |
|
6 |
*) |
|
7 |
||
12620 | 8 |
theory OrdQuant = Ordinal: |
2469 | 9 |
|
12620 | 10 |
constdefs |
2469 | 11 |
|
12 |
(* Ordinal Quantifiers *) |
|
12620 | 13 |
oall :: "[i, i => o] => o" |
14 |
"oall(A, P) == ALL x. x<A --> P(x)" |
|
15 |
||
16 |
oex :: "[i, i => o] => o" |
|
17 |
"oex(A, P) == EX x. x<A & P(x)" |
|
2469 | 18 |
|
19 |
(* Ordinal Union *) |
|
12620 | 20 |
OUnion :: "[i, i => i] => i" |
21 |
"OUnion(i,B) == {z: UN x:i. B(x). Ord(i)}" |
|
2469 | 22 |
|
23 |
syntax |
|
12620 | 24 |
"@oall" :: "[idt, i, o] => o" ("(3ALL _<_./ _)" 10) |
25 |
"@oex" :: "[idt, i, o] => o" ("(3EX _<_./ _)" 10) |
|
26 |
"@OUNION" :: "[idt, i, i] => i" ("(3UN _<_./ _)" 10) |
|
2469 | 27 |
|
28 |
translations |
|
29 |
"ALL x<a. P" == "oall(a, %x. P)" |
|
30 |
"EX x<a. P" == "oex(a, %x. P)" |
|
31 |
"UN x<a. B" == "OUnion(a, %x. B)" |
|
32 |
||
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
6093
diff
changeset
|
33 |
syntax (xsymbols) |
12620 | 34 |
"@oall" :: "[idt, i, o] => o" ("(3\<forall>_<_./ _)" 10) |
35 |
"@oex" :: "[idt, i, o] => o" ("(3\<exists>_<_./ _)" 10) |
|
36 |
"@OUNION" :: "[idt, i, i] => i" ("(3\<Union>_<_./ _)" 10) |
|
37 |
||
38 |
||
12667 | 39 |
declare Ord_Un [intro,simp,TC] |
40 |
declare Ord_UN [intro,simp,TC] |
|
41 |
declare Ord_Union [intro,simp,TC] |
|
12620 | 42 |
|
43 |
(** These mostly belong to theory Ordinal **) |
|
44 |
||
45 |
lemma Union_upper_le: |
|
46 |
"\<lbrakk>j: J; i\<le>j; Ord(\<Union>(J))\<rbrakk> \<Longrightarrow> i \<le> \<Union>J" |
|
47 |
apply (subst Union_eq_UN) |
|
48 |
apply (rule UN_upper_le) |
|
49 |
apply auto |
|
50 |
done |
|
51 |
||
12667 | 52 |
lemma zero_not_Limit [iff]: "~ Limit(0)" |
53 |
by (simp add: Limit_def) |
|
54 |
||
55 |
lemma Limit_has_1: "Limit(i) \<Longrightarrow> 1 < i" |
|
56 |
by (blast intro: Limit_has_0 Limit_has_succ) |
|
57 |
||
58 |
lemma Limit_Union [rule_format]: "\<lbrakk>I \<noteq> 0; \<forall>i\<in>I. Limit(i)\<rbrakk> \<Longrightarrow> Limit(\<Union>I)" |
|
59 |
apply (simp add: Limit_def lt_def) |
|
60 |
apply (blast intro!: equalityI) |
|
61 |
done |
|
62 |
||
12620 | 63 |
lemma increasing_LimitI: "\<lbrakk>0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y\<rbrakk> \<Longrightarrow> Limit(l)" |
64 |
apply (simp add: Limit_def lt_Ord2) |
|
65 |
apply clarify |
|
66 |
apply (drule_tac i=y in ltD) |
|
67 |
apply (blast intro: lt_trans1 succ_leI ltI lt_Ord2) |
|
68 |
done |
|
69 |
||
70 |
lemma UN_upper_lt: |
|
71 |
"\<lbrakk>a\<in> A; i < b(a); Ord(\<Union>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x\<in>A. b(x))" |
|
72 |
by (unfold lt_def, blast) |
|
73 |
||
74 |
lemma lt_imp_0_lt: "j<i \<Longrightarrow> 0<i" |
|
75 |
by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord]) |
|
76 |
||
77 |
lemma Ord_set_cases: |
|
78 |
"\<forall>i\<in>I. Ord(i) \<Longrightarrow> I=0 \<or> \<Union>(I) \<in> I \<or> (\<Union>(I) \<notin> I \<and> Limit(\<Union>(I)))" |
|
79 |
apply (clarify elim!: not_emptyE) |
|
80 |
apply (cases "\<Union>(I)" rule: Ord_cases) |
|
81 |
apply (blast intro: Ord_Union) |
|
82 |
apply (blast intro: subst_elem) |
|
83 |
apply auto |
|
84 |
apply (clarify elim!: equalityE succ_subsetE) |
|
85 |
apply (simp add: Union_subset_iff) |
|
86 |
apply (subgoal_tac "B = succ(j)", blast ) |
|
87 |
apply (rule le_anti_sym) |
|
88 |
apply (simp add: le_subset_iff) |
|
89 |
apply (simp add: ltI) |
|
90 |
done |
|
91 |
||
92 |
lemma Ord_Union_eq_succD: "[|\<forall>x\<in>X. Ord(x); \<Union>X = succ(j)|] ==> succ(j) \<in> X" |
|
93 |
by (drule Ord_set_cases, auto) |
|
94 |
||
95 |
(*See also Transset_iff_Union_succ*) |
|
96 |
lemma Ord_Union_succ_eq: "Ord(i) \<Longrightarrow> \<Union>(succ(i)) = i" |
|
97 |
by (blast intro: Ord_trans) |
|
2540 | 98 |
|
12620 | 99 |
lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) \<Longrightarrow> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)" |
100 |
by (auto simp: lt_def Ord_Union) |
|
101 |
||
102 |
lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i Un j" |
|
103 |
by (simp add: lt_def) |
|
104 |
||
105 |
lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i Un j" |
|
106 |
by (simp add: lt_def) |
|
107 |
||
108 |
lemma Ord_OUN [intro,simp]: |
|
109 |
"\<lbrakk>!!x. x<A \<Longrightarrow> Ord(B(x))\<rbrakk> \<Longrightarrow> Ord(\<Union>x<A. B(x))" |
|
110 |
by (simp add: OUnion_def ltI Ord_UN) |
|
111 |
||
112 |
lemma OUN_upper_lt: |
|
113 |
"\<lbrakk>a<A; i < b(a); Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x<A. b(x))" |
|
114 |
by (unfold OUnion_def lt_def, blast ) |
|
115 |
||
116 |
lemma OUN_upper_le: |
|
117 |
"\<lbrakk>a<A; i\<le>b(a); Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i \<le> (\<Union>x<A. b(x))" |
|
118 |
apply (unfold OUnion_def) |
|
119 |
apply auto |
|
120 |
apply (rule UN_upper_le ) |
|
121 |
apply (auto simp add: lt_def) |
|
122 |
done |
|
2469 | 123 |
|
12620 | 124 |
lemma Limit_OUN_eq: "Limit(i) ==> (UN x<i. x) = i" |
125 |
by (simp add: OUnion_def Limit_Union_eq Limit_is_Ord) |
|
126 |
||
127 |
(* No < version; consider (UN i:nat.i)=nat *) |
|
128 |
lemma OUN_least: |
|
129 |
"(!!x. x<A ==> B(x) \<subseteq> C) ==> (UN x<A. B(x)) \<subseteq> C" |
|
130 |
by (simp add: OUnion_def UN_least ltI) |
|
131 |
||
132 |
(* No < version; consider (UN i:nat.i)=nat *) |
|
133 |
lemma OUN_least_le: |
|
134 |
"[| Ord(i); !!x. x<A ==> b(x) \<le> i |] ==> (UN x<A. b(x)) \<le> i" |
|
135 |
by (simp add: OUnion_def UN_least_le ltI Ord_0_le) |
|
136 |
||
137 |
lemma le_implies_OUN_le_OUN: |
|
138 |
"[| !!x. x<A ==> c(x) \<le> d(x) |] ==> (UN x<A. c(x)) \<le> (UN x<A. d(x))" |
|
139 |
by (blast intro: OUN_least_le OUN_upper_le le_Ord2 Ord_OUN) |
|
140 |
||
141 |
lemma OUN_UN_eq: |
|
142 |
"(!!x. x:A ==> Ord(B(x))) |
|
143 |
==> (UN z < (UN x:A. B(x)). C(z)) = (UN x:A. UN z < B(x). C(z))" |
|
144 |
by (simp add: OUnion_def) |
|
145 |
||
146 |
lemma OUN_Union_eq: |
|
147 |
"(!!x. x:X ==> Ord(x)) |
|
148 |
==> (UN z < Union(X). C(z)) = (UN x:X. UN z < x. C(z))" |
|
149 |
by (simp add: OUnion_def) |
|
150 |
||
2469 | 151 |
end |