Skip to content

Commit bfcef7b

Browse files
affeldt-aistt6sCohenCyril
authored andcommitted
Zorn lemma for inclusion (#978)
* Zorn lemma for inclusion Co-authored-by: Takafumi Saikawa <[email protected]> Co-authored-by: Cyril Cohen <[email protected]>
1 parent d26fb19 commit bfcef7b

File tree

2 files changed

+35
-5
lines changed

2 files changed

+35
-5
lines changed

CHANGELOG_UNRELEASED.md

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -61,6 +61,9 @@
6161
- in `classical_sets.v`:
6262
+ lemmas `properW`, `properxx`
6363

64+
- in `classical_sets.v`:
65+
+ lemma `Zorn_bigcup`
66+
6467
### Changed
6568

6669
- moved from `lebesgue_measure.v` to `real_interval.v`:

classical/classical_sets.v

Lines changed: 32 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -2578,12 +2578,12 @@ Lemma Zorn T (R : T -> T -> Prop) :
25782578
exists t, forall s, R t s -> s = t.
25792579
Proof.
25802580
move=> Rrefl Rtrans Rantisym Rtot_max.
2581-
set totR := ({A : set T | total_on A R}).
2581+
pose totR := {A : set T | total_on A R}.
25822582
set R' := fun A B : totR => sval A `<=` sval B.
25832583
have R'refl A : R' A A by [].
25842584
have R'trans A B C : R' A B -> R' B C -> R' A C by apply: subset_trans.
25852585
have R'antisym A B : R' A B -> R' B A -> A = B.
2586-
rewrite /R'; case: A; case: B => /= B totB A totA sAB sBA.
2586+
rewrite /R'; move: A B => [/= A totA] [/= B totB] sAB sBA.
25872587
by apply: eq_exist; rewrite predeqE=> ?; split=> [/sAB|/sBA].
25882588
have R'tot_lub A : total_on A R' -> exists t, (forall s, A s -> R' s t) /\
25892589
forall r, (forall s, A s -> R' s r) -> R' t r.
@@ -2594,7 +2594,7 @@ have R'tot_lub A : total_on A R' -> exists t, (forall s, A s -> R' s t) /\
25942594
by have /(_ _ _ Cs Ct) := svalP C.
25952595
by have /(_ _ _ Bs Bt) := svalP B.
25962596
exists (exist _ (\bigcup_(B in A) sval B) AUtot); split.
2597-
by move=> B ???; exists B.
2597+
by move=> B ? ? ?; exists B.
25982598
by move=> B Bub ? /= [? /Bub]; apply.
25992599
apply: contrapT => nomax.
26002600
have {}nomax t : exists s, R t s /\ s <> t.
@@ -2609,9 +2609,9 @@ have Astot : total_on (sval A `|` [set s]) R.
26092609
by move=> [/tub Rvt|->]; right=> //; apply: Rtrans Rts.
26102610
move=> [Av|->]; [apply: (svalP A)|left] => //.
26112611
by apply: Rtrans Rts; apply: tub.
2612-
exists (exist _ (sval A `|` [set s]) Astot); split; first by move=> ??; left.
2612+
exists (exist _ (sval A `|` [set s]) Astot); split; first by move=> ? ?; left.
26132613
split=> [AeAs|[B Btot] sAB sBAs].
2614-
have [/tub Rst|] := (pselect (sval A s)); first exact/snet/Rantisym.
2614+
have [/tub Rst|] := pselect (sval A s); first exact/snet/Rantisym.
26152615
by rewrite AeAs /=; apply; right.
26162616
have [Bs|nBs] := pselect (B s).
26172617
by right; apply: eq_exist; rewrite predeqE => r; split=> [/sBAs|[/sAB|->]].
@@ -2620,6 +2620,33 @@ apply: eq_exist; rewrite predeqE => r; split=> [Br|/sAB] //.
26202620
by have /sBAs [|ser] // := Br; rewrite ser in Br.
26212621
Qed.
26222622

2623+
Section Zorn_subset.
2624+
Variables (T : Type) (P : set (set T)).
2625+
2626+
Lemma Zorn_bigcup :
2627+
(forall F : set (set T), F `<=` P -> total_on F subset ->
2628+
P (\bigcup_(X in F) X)) ->
2629+
exists A, P A /\ forall B, A `<` B -> ~ P B.
2630+
Proof.
2631+
move=> totP; pose R (sA sB : P) := sval sA `<=` sval sB.
2632+
have {}totR F (FR : total_on F R) : exists sB, forall sA, F sA -> R sA sB.
2633+
have FP : [set val x | x in F] `<=` P.
2634+
by move=> _ [X FX <-]; apply: set_mem; apply: valP.
2635+
have totF : total_on [set val x | x in F] subset.
2636+
by move=> _ _ [X FX <-] [Y FY <-]; apply: FR.
2637+
exists (SigSub (mem_set (totP _ FP totF))) => A FA; rewrite /R/=.
2638+
exact: (bigcup_sup (imageP val _)).
2639+
have [| | |sA sAmax] := Zorn _ _ _ totR.
2640+
- by move=> ?; exact: subset_refl.
2641+
- by move=> ? ? ?; exact: subset_trans.
2642+
- by move=> [A PA] [B PB]; rewrite /R /= => AB BA; exact/eq_exist/seteqP.
2643+
- exists (val sA); case: sA => A PA /= in sAmax *; split; first exact: set_mem.
2644+
move=> B AB PB; have [BA] := sAmax (SigSub (mem_set PB)) (properW AB).
2645+
by move: AB; rewrite BA; exact: properxx.
2646+
Qed.
2647+
2648+
End Zorn_subset.
2649+
26232650
Definition premaximal T (R : T -> T -> Prop) (t : T) :=
26242651
forall s, R t s -> R s t.
26252652

0 commit comments

Comments
 (0)