@@ -94,44 +94,65 @@ exists_maximal_of_chains_bounded
94
94
(h c hc))
95
95
(λ a b c, trans)
96
96
97
- section partial_order
98
- variables [partial_order α]
97
+ section preorder
98
+ variables [preorder α]
99
99
100
- lemma zorn_partial_order (h : ∀ c : set α, is_chain (≤) c → ∃ ub, ∀ a ∈ c, a ≤ ub) :
101
- ∃ m : α, ∀ a, m ≤ a → a = m :=
102
- let ⟨m, hm⟩ := @exists_maximal_of_chains_bounded α (≤) h (λ a b c, le_trans) in
103
- ⟨m, λ a ha, le_antisymm (hm a ha) ha⟩
100
+ theorem zorn_preorder (h : ∀ c : set α, is_chain (≤) c → bdd_above c) :
101
+ ∃ m : α, ∀ a, m ≤ a → a ≤ m :=
102
+ exists_maximal_of_chains_bounded h (λ a b c, le_trans)
104
103
105
- lemma zorn_nonempty_partial_order [nonempty α]
106
- (h : ∀ c : set α, is_chain (≤) c → c.nonempty → ∃ ub, ∀ a ∈ c, a ≤ ub) :
107
- ∃ (m : α), ∀ a, m ≤ a → a = m :=
108
- let ⟨m, hm⟩ := @exists_maximal_of_nonempty_chains_bounded α (≤) _ h (λ a b c, le_trans) in
109
- ⟨m, λ a ha, le_antisymm (hm a ha) ha⟩
104
+ theorem zorn_nonempty_preorder [nonempty α]
105
+ (h : ∀ (c : set α), is_chain (≤) c → c.nonempty → bdd_above c) :
106
+ ∃ (m : α), ∀ a, m ≤ a → a ≤ m :=
107
+ exists_maximal_of_nonempty_chains_bounded h (λ a b c, le_trans)
110
108
111
- lemma zorn_partial_order ₀ (s : set α)
109
+ theorem zorn_preorder ₀ (s : set α)
112
110
(ih : ∀ c ⊆ s, is_chain (≤) c → ∃ ub ∈ s, ∀ z ∈ c, z ≤ ub) :
113
- ∃ m ∈ s, ∀ z ∈ s, m ≤ z → z = m :=
114
- let ⟨⟨m, hms⟩, h⟩ := @zorn_partial_order {m // m ∈ s} _
111
+ ∃ m ∈ s, ∀ z ∈ s, m ≤ z → z ≤ m :=
112
+ let ⟨⟨m, hms⟩, h⟩ := @zorn_preorder s _
115
113
(λ c hc,
116
114
let ⟨ub, hubs, hub⟩ := ih (subtype.val '' c) (λ _ ⟨⟨x, hx⟩, _, h⟩, h ▸ hx)
117
115
(by { rintro _ ⟨p, hpc, rfl⟩ _ ⟨q, hqc, rfl⟩ hpq;
118
116
refine hc hpc hqc (λ t, hpq (subtype.ext_iff.1 t)) })
119
117
in ⟨⟨ub, hubs⟩, λ ⟨y, hy⟩ hc, hub _ ⟨_, hc, rfl⟩⟩)
120
- in ⟨m, hms, λ z hzs hmz, congr_arg subtype.val (h ⟨z, hzs⟩ hmz)⟩
118
+ in ⟨m, hms, λ z hzs hmz, h ⟨z, hzs⟩ hmz⟩
119
+
120
+ theorem zorn_nonempty_preorder₀ (s : set α)
121
+ (ih : ∀ c ⊆ s, is_chain (≤) c → ∀ y ∈ c, ∃ ub ∈ s, ∀ z ∈ c, z ≤ ub) (x : α) (hxs : x ∈ s) :
122
+ ∃ m ∈ s, x ≤ m ∧ ∀ z ∈ s, m ≤ z → z ≤ m :=
123
+ begin
124
+ rcases zorn_preorder₀ {y ∈ s | x ≤ y} (λ c hcs hc, _) with ⟨m, ⟨hms, hxm⟩, hm⟩,
125
+ { exact ⟨m, hms, hxm, λ z hzs hmz, hm _ ⟨hzs, (hxm.trans hmz)⟩ hmz⟩ },
126
+ { rcases c.eq_empty_or_nonempty with rfl|⟨y, hy⟩,
127
+ { exact ⟨x, ⟨hxs, le_rfl⟩, λ z, false.elim⟩ },
128
+ { rcases ih c (λ z hz, (hcs hz).1 ) hc y hy with ⟨z, hzs, hz⟩,
129
+ exact ⟨z, ⟨hzs, (hcs hy).2 .trans $ hz _ hy⟩, hz⟩ } }
130
+ end
131
+
132
+ end preorder
133
+
134
+ section partial_order
135
+ variables [partial_order α]
136
+
137
+ lemma zorn_partial_order (h : ∀ c : set α, is_chain (≤) c → bdd_above c) :
138
+ ∃ m : α, ∀ a, m ≤ a → a = m :=
139
+ let ⟨m, hm⟩ := zorn_preorder h in ⟨m, λ a ha, le_antisymm (hm a ha) ha⟩
140
+
141
+ theorem zorn_nonempty_partial_order [nonempty α]
142
+ (h : ∀ (c : set α), is_chain (≤) c → c.nonempty → bdd_above c) :
143
+ ∃ (m : α), ∀ a, m ≤ a → a = m :=
144
+ let ⟨m, hm⟩ := zorn_nonempty_preorder h in ⟨m, λ a ha, le_antisymm (hm a ha) ha⟩
145
+
146
+ theorem zorn_partial_order₀ (s : set α)
147
+ (ih : ∀ c ⊆ s, is_chain (≤) c → ∃ ub ∈ s, ∀ z ∈ c, z ≤ ub) :
148
+ ∃ m ∈ s, ∀ z ∈ s, m ≤ z → z = m :=
149
+ let ⟨m, hms, hm⟩ := zorn_preorder₀ s ih in ⟨m, hms, λ z hzs hmz, (hm z hzs hmz).antisymm hmz⟩
121
150
122
151
lemma zorn_nonempty_partial_order₀ (s : set α)
123
152
(ih : ∀ c ⊆ s, is_chain (≤) c → ∀ y ∈ c, ∃ ub ∈ s, ∀ z ∈ c, z ≤ ub) (x : α) (hxs : x ∈ s) :
124
153
∃ m ∈ s, x ≤ m ∧ ∀ z ∈ s, m ≤ z → z = m :=
125
- let ⟨⟨m, hms, hxm⟩, h⟩ := @zorn_partial_order {m // m ∈ s ∧ x ≤ m} _
126
- (λ c hc, c.eq_empty_or_nonempty.elim
127
- (λ hce, hce.symm ▸ ⟨⟨x, hxs, le_rfl⟩, λ _, false.elim⟩)
128
- (λ ⟨m, hmc⟩,
129
- let ⟨ub, hubs, hub⟩ := ih (subtype.val '' c) (image_subset_iff.2 $ λ z hzc, z.2 .1 )
130
- (by rintro _ ⟨p, hpc, rfl⟩ _ ⟨q, hqc, rfl⟩ hpq;
131
- exact hc hpc hqc (ne_of_apply_ne _ hpq)) m.1 (mem_image_of_mem _ hmc) in
132
- ⟨⟨ub, hubs, le_trans m.2 .2 $ hub m.1 $ mem_image_of_mem _ hmc⟩,
133
- λ a hac, hub a.1 ⟨a, hac, rfl⟩⟩)) in
134
- ⟨m, hms, hxm, λ z hzs hmz, congr_arg subtype.val $ h ⟨z, hzs, le_trans hxm hmz⟩ hmz⟩
154
+ let ⟨m, hms, hxm, hm⟩ := zorn_nonempty_preorder₀ s ih x hxs
155
+ in ⟨m, hms, hxm, λ z hzs hmz, (hm z hzs hmz).antisymm hmz⟩
135
156
136
157
end partial_order
137
158
0 commit comments