@@ -731,23 +731,20 @@ def typein.principal_seg {α : Type u} (r : α → α → Prop) [is_well_order
731
731
732
732
/-! ### Cardinality of ordinals -/
733
733
734
- /-- The cardinal of an ordinal is the cardinal of any
735
- set with that order type. -/
736
- def card (o : ordinal) : cardinal :=
737
- quot.lift_on o (λ a, #a.α) $ λ ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨e⟩, quotient.sound ⟨e.to_equiv⟩
734
+ /-- The cardinal of an ordinal is the cardinality of any type on which a relation with that order
735
+ type is defined. -/
736
+ def card : ordinal → cardinal := quotient.map Well_order.α $ λ ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨e⟩, ⟨e.to_equiv⟩
738
737
739
- @[simp] theorem card_type (r : α → α → Prop ) [is_well_order α r] :
740
- card (type r) = #α := rfl
738
+ @[simp] theorem card_type (r : α → α → Prop ) [is_well_order α r] : card (type r) = #α := rfl
741
739
742
740
@[simp] lemma card_typein {r : α → α → Prop } [wo : is_well_order α r] (x : α) :
743
- #{y // r y x} = (typein r x).card := rfl
741
+ #{y // r y x} = (typein r x).card :=
742
+ rfl
744
743
745
744
theorem card_le_card {o₁ o₂ : ordinal} : o₁ ≤ o₂ → card o₁ ≤ card o₂ :=
746
745
induction_on o₁ $ λ α r _, induction_on o₂ $ λ β s _ ⟨⟨⟨f, _⟩, _⟩⟩, ⟨f⟩
747
746
748
- instance : has_zero ordinal :=
749
- ⟨@type pempty empty_relation _⟩
750
-
747
+ instance : has_zero ordinal := ⟨type $ @empty_relation pempty⟩
751
748
instance : inhabited ordinal := ⟨0 ⟩
752
749
753
750
@[simp] theorem type_eq_zero_of_empty (r) [is_well_order α r] [is_empty α] : type r = 0 :=
@@ -756,8 +753,7 @@ instance : inhabited ordinal := ⟨0⟩
756
753
@[simp] theorem type_eq_zero_iff_is_empty [is_well_order α r] : type r = 0 ↔ is_empty α :=
757
754
⟨λ h, let ⟨s⟩ := type_eq.1 h in s.to_equiv.is_empty, @type_eq_zero_of_empty α r _⟩
758
755
759
- theorem type_ne_zero_iff_nonempty [is_well_order α r] : type r ≠ 0 ↔ nonempty α :=
760
- by simp
756
+ theorem type_ne_zero_iff_nonempty [is_well_order α r] : type r ≠ 0 ↔ nonempty α := by simp
761
757
762
758
theorem type_ne_zero_of_nonempty (r) [is_well_order α r] [h : nonempty α] : type r ≠ 0 :=
763
759
type_ne_zero_iff_nonempty.2 h
@@ -767,39 +763,29 @@ type_ne_zero_iff_nonempty.2 h
767
763
protected theorem zero_le (o : ordinal) : 0 ≤ o :=
768
764
induction_on o $ λ α r _, ⟨⟨⟨embedding.of_is_empty, is_empty_elim⟩, is_empty_elim⟩⟩
769
765
770
- instance : order_bot ordinal :=
771
- ⟨0 , ordinal.zero_le⟩
766
+ instance : order_bot ordinal := ⟨0 , ordinal.zero_le⟩
772
767
773
768
@[simp] protected theorem le_zero {o : ordinal} : o ≤ 0 ↔ o = 0 :=
774
769
by simp only [le_antisymm_iff, ordinal.zero_le, and_true]
775
770
776
771
protected theorem pos_iff_ne_zero {o : ordinal} : 0 < o ↔ o ≠ 0 :=
777
772
by simp only [lt_iff_le_and_ne, ordinal.zero_le, true_and, ne.def, eq_comm]
778
773
774
+ @[simp] theorem out_empty_iff_eq_zero {o : ordinal} : is_empty o.out.α ↔ o = 0 :=
775
+ by rw [←@type_eq_zero_iff_is_empty o.out.α (<), type_lt]
776
+
779
777
lemma eq_zero_of_out_empty (o : ordinal) [h : is_empty o.out.α] : o = 0 :=
780
- begin
781
- by_contra ho,
782
- replace ho := ordinal.pos_iff_ne_zero.2 ho,
783
- rw ←type_lt o at ho,
784
- have α := enum (<) 0 ho,
785
- exact h.elim α
786
- end
778
+ out_empty_iff_eq_zero.1 h
787
779
788
- @[simp] theorem out_empty_iff_eq_zero {o : ordinal} : is_empty o.out.α ↔ o = 0 :=
789
- begin
790
- refine ⟨@eq_zero_of_out_empty o, λ h, ⟨λ i, _⟩⟩,
791
- subst o,
792
- exact (ordinal.zero_le _).not_lt (typein_lt_self i)
793
- end
780
+ instance is_empty_out_zero : is_empty (0 : ordinal).out.α := out_empty_iff_eq_zero.2 rfl
794
781
795
782
@[simp] theorem out_nonempty_iff_ne_zero {o : ordinal} : nonempty o.out.α ↔ o ≠ 0 :=
796
- by rw [←not_iff_not, ←not_is_empty_iff, not_not, not_not, out_empty_iff_eq_zero ]
783
+ by rw [←@type_ne_zero_iff_nonempty o.out.α (<), type_lt ]
797
784
798
- instance is_empty_out_zero : is_empty ( 0 : ordinal).out.α :=
799
- out_empty_iff_eq_zero. 2 rfl
785
+ lemma ne_zero_of_out_nonempty (o : ordinal) [h : nonempty o .out.α] : o ≠ 0 :=
786
+ out_nonempty_iff_ne_zero. 1 h
800
787
801
- instance : has_one ordinal :=
802
- ⟨@type punit empty_relation _⟩
788
+ instance : has_one ordinal := ⟨type $ @empty_relation punit⟩
803
789
804
790
theorem one_eq_type_unit : 1 = @type unit empty_relation _ :=
805
791
quotient.sound ⟨⟨punit_equiv_punit, λ _ _, iff.rfl⟩⟩
0 commit comments