Disjoint union of types #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file proves basic results about the sum type α ⊕ β.
α ⊕ β is the type made of a copy of α and a copy of β. It is also called disjoint union.
Main declarations #
sum.get_left: Retrieves the left content ofx : α ⊕ βor returnsnoneif it's coming from the right.sum.get_right: Retrieves the right content ofx : α ⊕ βor returnsnoneif it's coming from the left.sum.is_left: Returns whetherx : α ⊕ βcomes from the left component or not.sum.is_right: Returns whetherx : α ⊕ βcomes from the right component or not.sum.map: Mapsα ⊕ βtoγ ⊕ δcomponent-wise.sum.elim: Nondependent eliminator/induction principle forα ⊕ β.sum.swap: Mapsα ⊕ βtoβ ⊕ αby swapping components.sum.lex: Lexicographic order onα ⊕ βinduced by a relation onαand a relation onβ.
Notes #
The definition of sum takes values in Type*. This effectively forbids Prop- valued sum types.
To this effect, we have psum, which takes value in Sort* and carries a more complicated
universe signature in consequence. The Prop version is or.
Check if a sum is inl and if so, retrieve its contents.
Equations
- (sum.inr _x).get_left = option.none
- (sum.inl a).get_left = option.some a
Check if a sum is inr and if so, retrieve its contents.
Equations
- (sum.inr b).get_right = option.some b
- (sum.inl _x).get_right = option.none
- inl : ∀ {α : Type u} {β : Type v} {γ : Type u_1} {δ : Type u_2} {r : α → γ → Prop} {s : β → δ → Prop} {a : α} {c : γ}, r a c → sum.lift_rel r s (sum.inl a) (sum.inl c)
- inr : ∀ {α : Type u} {β : Type v} {γ : Type u_1} {δ : Type u_2} {r : α → γ → Prop} {s : β → δ → Prop} {b : β} {d : δ}, s b d → sum.lift_rel r s (sum.inr b) (sum.inr d)
Lifts pointwise two relations between α and γ and between β and δ to a relation between
α ⊕ β and γ ⊕ δ.
Instances for sum.lift_rel
Equations
- sum.lift_rel.decidable (sum.inr b) (sum.inr d) = decidable_of_iff' (s b d) sum.lift_rel_inr_inr
- sum.lift_rel.decidable (sum.inr b) (sum.inl c) = decidable.is_false sum.not_lift_rel_inr_inl
- sum.lift_rel.decidable (sum.inl a) (sum.inr d) = decidable.is_false sum.not_lift_rel_inl_inr
- sum.lift_rel.decidable (sum.inl a) (sum.inl c) = decidable_of_iff' (r a c) sum.lift_rel_inl_inl
- inl : ∀ {α : Type u} {β : Type v} {r : α → α → Prop} {s : β → β → Prop} {a₁ a₂ : α}, r a₁ a₂ → sum.lex r s (sum.inl a₁) (sum.inl a₂)
- inr : ∀ {α : Type u} {β : Type v} {r : α → α → Prop} {s : β → β → Prop} {b₁ b₂ : β}, s b₁ b₂ → sum.lex r s (sum.inr b₁) (sum.inr b₂)
- sep : ∀ {α : Type u} {β : Type v} {r : α → α → Prop} {s : β → β → Prop} (a : α) (b : β), sum.lex r s (sum.inl a) (sum.inr b)
Lexicographic order for sum. Sort all the inl a before the inr b, otherwise use the
respective order on α or β.
Equations
- sum.lex.decidable_rel (sum.inr b) (sum.inr d) = decidable_of_iff' (s b d) sum.lex_inr_inr
- sum.lex.decidable_rel (sum.inr b) (sum.inl c) = decidable.is_false sum.lex_inr_inl
- sum.lex.decidable_rel (sum.inl a) (sum.inr d) = decidable.is_true _
- sum.lex.decidable_rel (sum.inl a) (sum.inl c) = decidable_of_iff' (r a c) sum.lex_inl_inl
Ternary sum #
Abbreviations for the maps from the summands to α ⊕ β ⊕ γ. This is useful for pattern-matching.