From a057441eb14fe0e65a5588084e7e2c1193b82579 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 4 May 2022 11:10:45 +0000 Subject: [PATCH] feat(order/basic): Notation for `order_dual` (#13798) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `αᵒᵈ` as notation for `order_dual α` and replace current uses. [Zulip poll](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Notation.20for.20order_dual/near/280629129) --- src/algebra/big_operators/multiset.lean | 2 +- src/algebra/big_operators/order.lean | 6 +- src/algebra/bounds.lean | 4 +- src/algebra/group_power/order.lean | 11 +- src/algebra/indicator_function.lean | 4 +- src/algebra/order/archimedean.lean | 3 +- src/algebra/order/field.lean | 4 +- src/algebra/order/group.lean | 28 ++-- src/algebra/order/module.lean | 6 +- src/algebra/order/monoid.lean | 80 +++++----- src/algebra/order/ring.lean | 28 ++-- src/algebra/order/smul.lean | 20 +-- src/algebra/order/with_zero.lean | 8 +- src/algebra/support.lean | 2 +- .../prime_spectrum/basic.lean | 6 +- .../projective_spectrum/topology.lean | 6 +- .../box_integral/partition/filter.lean | 4 +- src/analysis/convex/basic.lean | 18 +-- src/analysis/convex/extrema.lean | 5 +- src/analysis/convex/function.lean | 18 +-- src/analysis/convex/jensen.lean | 8 +- src/analysis/convex/quasiconvex.lean | 4 +- src/analysis/convex/strict.lean | 3 +- src/analysis/inner_product_space/basic.lean | 2 +- .../normed_space/lattice_ordered_group.lean | 6 +- src/combinatorics/pigeonhole.lean | 12 +- src/data/fin/basic.lean | 2 +- src/data/fin/tuple/basic.lean | 2 +- src/data/finset/lattice.lean | 147 ++++++++---------- src/data/finset/locally_finite.lean | 4 +- src/data/finset/pi_induction.lean | 2 +- src/data/finset/sigma.lean | 2 +- src/data/fintype/basic.lean | 12 +- src/data/list/big_operators.lean | 2 +- src/data/list/min_max.lean | 35 ++--- src/data/nat/lattice.lean | 4 +- src/data/nat/pairing.lean | 2 +- src/data/ordmap/ordset.lean | 20 +-- src/data/real/ennreal.lean | 8 +- src/data/real/ereal.lean | 2 +- src/data/set/finite.lean | 7 +- src/data/set/function.lean | 4 +- src/data/set/intervals/basic.lean | 7 +- src/data/set/intervals/infinite.lean | 3 +- src/data/set/intervals/pi.lean | 2 +- src/data/set/intervals/with_bot_top.lean | 4 +- src/data/sum/order.lean | 6 +- src/field_theory/galois.lean | 2 +- src/linear_algebra/affine_space/ordered.lean | 20 +-- src/linear_algebra/basic.lean | 2 +- src/linear_algebra/prod.lean | 2 +- .../constructions/borel_space.lean | 16 +- src/measure_theory/function/ess_sup.lean | 12 +- .../function/locally_integrable.lean | 6 +- .../integral/interval_integral.lean | 2 +- src/measure_theory/lattice.lean | 12 +- src/measure_theory/measurable_space_def.lean | 2 +- src/measure_theory/measure/measure_space.lean | 3 +- src/measure_theory/pi_system.lean | 2 +- src/order/antisymmetrization.lean | 2 +- src/order/atoms.lean | 26 ++-- src/order/basic.lean | 39 ++--- src/order/boolean_algebra.lean | 2 +- src/order/bounded.lean | 30 ++-- src/order/bounded_order.lean | 22 +-- src/order/bounds.lean | 59 ++++--- src/order/category/BoolAlg.lean | 2 +- src/order/category/BoundedDistribLattice.lean | 2 +- src/order/category/BoundedLattice.lean | 2 +- src/order/category/BoundedOrder.lean | 2 +- src/order/category/CompleteLattice.lean | 2 +- src/order/category/DistribLattice.lean | 2 +- src/order/category/FinBoolAlg.lean | 2 +- src/order/category/FinPartialOrder.lean | 2 +- src/order/category/Lattice.lean | 3 +- src/order/category/LinearOrder.lean | 2 +- src/order/category/NonemptyFinLinOrd.lean | 4 +- src/order/category/PartialOrder.lean | 2 +- src/order/category/Preorder.lean | 2 +- src/order/category/Semilattice.lean | 4 +- src/order/circular.lean | 22 +-- src/order/compare.lean | 4 +- src/order/complete_boolean_algebra.lean | 30 ++-- src/order/complete_lattice.lean | 119 +++++++------- src/order/concept.lean | 2 +- src/order/conditionally_complete_lattice.lean | 77 +++++---- src/order/cover.lean | 5 +- src/order/directed.lean | 6 +- src/order/filter/at_top_bot.lean | 129 ++++++++------- src/order/filter/basic.lean | 2 +- src/order/filter/cofinite.lean | 4 +- src/order/filter/extr.lean | 7 +- src/order/fixed_points.lean | 6 +- src/order/grade.lean | 16 +- src/order/hom/basic.lean | 36 ++--- src/order/hom/bounded.lean | 16 +- src/order/hom/complete_lattice.lean | 16 +- src/order/hom/lattice.lean | 31 ++-- src/order/iterate.lean | 8 +- src/order/lattice.lean | 61 ++++---- src/order/liminf_limsup.lean | 40 ++--- src/order/locally_finite.lean | 6 +- src/order/max.lean | 22 ++- src/order/min_max.lean | 19 +-- src/order/modular_lattice.lean | 10 +- src/order/monotone.lean | 13 +- src/order/order_dual.lean | 59 +++---- src/order/pfilter.lean | 6 +- src/order/rel_classes.lean | 6 +- src/order/rel_iso.lean | 2 +- src/order/succ_pred/basic.lean | 22 ++- src/order/succ_pred/relation.lean | 14 +- src/order/upper_lower.lean | 8 +- src/order/well_founded.lean | 2 +- src/order/well_founded_set.lean | 3 +- src/order/zorn.lean | 5 +- src/ring_theory/artinian.lean | 8 +- src/ring_theory/nullstellensatz.lean | 2 +- src/ring_theory/valuation/basic.lean | 9 +- .../valuation/valuation_subring.lean | 3 +- src/topology/algebra/order/basic.lean | 111 ++++++------- src/topology/algebra/order/compact.lean | 27 ++-- src/topology/algebra/order/extend_from.lean | 2 +- .../algebra/order/intermediate_value.lean | 6 +- src/topology/algebra/order/left_right.lean | 2 +- src/topology/algebra/order/liminf_limsup.lean | 10 +- .../algebra/order/monotone_continuity.lean | 7 +- .../algebra/order/monotone_convergence.lean | 49 +++--- src/topology/continuous_function/ordered.lean | 4 +- src/topology/local_extr.lean | 2 +- src/topology/order.lean | 24 ++- src/topology/order/lattice.lean | 8 +- src/topology/semicontinuous.lean | 28 ++-- 133 files changed, 909 insertions(+), 1096 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index 19a9cf026dff9..fdeecd970c4e6 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -318,7 +318,7 @@ prod_le_prod_of_rel_le $ rel_map_left.2 $ rel_refl_of_refl_on h @[to_additive] lemma prod_le_sum_prod (f : α → α) (h : ∀ x, x ∈ s → x ≤ f x) : s.prod ≤ (s.map f).prod := -@prod_map_le_prod (order_dual α) _ _ f h +@prod_map_le_prod αᵒᵈ _ _ f h @[to_additive card_nsmul_le_sum] lemma pow_card_le_prod (h : ∀ x ∈ s, a ≤ x) : a ^ s.card ≤ s.prod := diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 08bad25be2983..8c7030a2ecf03 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -161,7 +161,7 @@ end @[to_additive sum_eq_zero_iff_of_nonneg] lemma prod_eq_one_iff_of_le_one' : (∀ i ∈ s, f i ≤ 1) → (∏ i in s, f i = 1 ↔ ∀ i ∈ s, f i = 1) := -@prod_eq_one_iff_of_one_le' _ (order_dual N) _ _ _ +@prod_eq_one_iff_of_one_le' _ Nᵒᵈ _ _ _ @[to_additive single_le_sum] lemma single_le_prod' (hf : ∀ i ∈ s, 1 ≤ f i) {a} (h : a ∈ s) : f a ≤ (∏ x in s, f x) := @@ -181,7 +181,7 @@ end @[to_additive card_nsmul_le_sum] lemma pow_card_le_prod (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : n ^ s.card ≤ s.prod f := -@finset.prod_le_pow_card _ (order_dual N) _ _ _ _ h +@finset.prod_le_pow_card _ Nᵒᵈ _ _ _ _ h lemma card_bUnion_le_card_mul [decidable_eq β] (s : finset ι) (f : ι → finset β) (n : ℕ) (h : ∀ a ∈ s, (f a).card ≤ n) : @@ -204,7 +204,7 @@ calc (∏ y in t, ∏ x in s.filter (λ x, g x = y), f x) ≤ lemma prod_le_prod_fiberwise_of_prod_fiber_le_one' {t : finset ι'} {g : ι → ι'} {f : ι → N} (h : ∀ y ∉ t, (∏ x in s.filter (λ x, g x = y), f x) ≤ 1) : (∏ x in s, f x) ≤ ∏ y in t, ∏ x in s.filter (λ x, g x = y), f x := -@prod_fiberwise_le_prod_of_one_le_prod_fiber' _ (order_dual N) _ _ _ _ _ _ _ h +@prod_fiberwise_le_prod_of_one_le_prod_fiber' _ Nᵒᵈ _ _ _ _ _ _ _ h end ordered_comm_monoid diff --git a/src/algebra/bounds.lean b/src/algebra/bounds.lean index 4ee7d179402ec..50eab7e89589c 100644 --- a/src/algebra/bounds.lean +++ b/src/algebra/bounds.lean @@ -69,11 +69,11 @@ image2_subset_iff.2 $ λ x hx y hy, mul_mem_upper_bounds_mul hx hy @[to_additive] lemma mul_mem_lower_bounds_mul {s t : set M} {a b : M} (ha : a ∈ lower_bounds s) (hb : b ∈ lower_bounds t) : a * b ∈ lower_bounds (s * t) := -@mul_mem_upper_bounds_mul (order_dual M) _ _ _ _ _ _ _ _ ha hb +@mul_mem_upper_bounds_mul Mᵒᵈ _ _ _ _ _ _ _ _ ha hb @[to_additive] lemma subset_lower_bounds_mul (s t : set M) : lower_bounds s * lower_bounds t ⊆ lower_bounds (s * t) := -@subset_upper_bounds_mul (order_dual M) _ _ _ _ _ _ +@subset_upper_bounds_mul Mᵒᵈ _ _ _ _ _ _ @[to_additive] lemma bdd_above.mul {s t : set M} (hs : bdd_above s) (ht : bdd_above t) : bdd_above (s * t) := diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index 04d0ff5c15b5f..aa2ab3534f216 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -34,8 +34,7 @@ theorem one_le_pow_of_one_le' {a : M} (H : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n | (k + 1) := by { rw pow_succ, exact one_le_mul H (one_le_pow_of_one_le' k) } @[to_additive nsmul_nonpos] -theorem pow_le_one' {a : M} (H : a ≤ 1) (n : ℕ) : a ^ n ≤ 1 := -@one_le_pow_of_one_le' (order_dual M) _ _ _ _ H n +lemma pow_le_one' {a : M} (H : a ≤ 1) (n : ℕ) : a ^ n ≤ 1 := @one_le_pow_of_one_le' Mᵒᵈ _ _ _ _ H n @[to_additive nsmul_le_nsmul] theorem pow_le_pow' {a : M} {n m : ℕ} (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := @@ -45,7 +44,7 @@ calc a ^ n ≤ a ^ n * a ^ k : le_mul_of_one_le_right' (one_le_pow_of_one_le' ha @[to_additive nsmul_le_nsmul_of_nonpos] theorem pow_le_pow_of_le_one' {a : M} {n m : ℕ} (ha : a ≤ 1) (h : n ≤ m) : a ^ m ≤ a ^ n := -@pow_le_pow' (order_dual M) _ _ _ _ _ _ ha h +@pow_le_pow' Mᵒᵈ _ _ _ _ _ _ ha h @[to_additive nsmul_pos] theorem one_lt_pow' {a : M} (ha : 1 < a) {k : ℕ} (hk : k ≠ 0) : 1 < a ^ k := @@ -59,8 +58,8 @@ begin end @[to_additive nsmul_neg] -theorem pow_lt_one' {a : M} (ha : a < 1) {k : ℕ} (hk : k ≠ 0) : a ^ k < 1 := -@one_lt_pow' (order_dual M) _ _ _ _ ha k hk +lemma pow_lt_one' {a : M} (ha : a < 1) {k : ℕ} (hk : k ≠ 0) : a ^ k < 1 := +@one_lt_pow' Mᵒᵈ _ _ _ _ ha k hk @[to_additive nsmul_lt_nsmul] theorem pow_lt_pow' [covariant_class M M (*) (<)] {a : M} {n m : ℕ} (ha : 1 < a) (h : n < m) : @@ -88,7 +87,7 @@ lemma one_le_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 ≤ x ^ n ↔ 1 ≤ x @[to_additive] lemma pow_le_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n ≤ 1 ↔ x ≤ 1 := -@one_le_pow_iff (order_dual M) _ _ _ _ _ hn +@one_le_pow_iff Mᵒᵈ _ _ _ _ _ hn @[to_additive nsmul_pos_iff] lemma one_lt_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 < x ^ n ↔ 1 < x := diff --git a/src/algebra/indicator_function.lean b/src/algebra/indicator_function.lean index 2dffdf02a2538..2762a65b1b36b 100644 --- a/src/algebra/indicator_function.lean +++ b/src/algebra/indicator_function.lean @@ -482,7 +482,7 @@ end @[to_additive] lemma le_mul_indicator_apply {y} (hfg : a ∈ s → y ≤ g a) (hf : a ∉ s → y ≤ 1) : y ≤ mul_indicator s g a := -@mul_indicator_apply_le' α (order_dual M) ‹_› _ _ _ _ _ hfg hf +@mul_indicator_apply_le' α Mᵒᵈ ‹_› _ _ _ _ _ hfg hf @[to_additive] lemma le_mul_indicator (hfg : ∀ a ∈ s, f a ≤ g a) (hf : ∀ a ∉ s, f a ≤ 1) : f ≤ mul_indicator s g := @@ -573,7 +573,7 @@ end lemma indicator_nonpos_le_indicator {β} [linear_order β] [has_zero β] (s : set α) (f : α → β) : {x | f x ≤ 0}.indicator f ≤ s.indicator f := -@indicator_le_indicator_nonneg α (order_dual β) _ _ s f +@indicator_le_indicator_nonneg α βᵒᵈ _ _ s f end set diff --git a/src/algebra/order/archimedean.lean b/src/algebra/order/archimedean.lean index 1dfd6928e874a..3a307b1cb2281 100644 --- a/src/algebra/order/archimedean.lean +++ b/src/algebra/order/archimedean.lean @@ -37,8 +37,7 @@ such that `0 < y` there exists a natural number `n` such that `x ≤ n • y`. - class archimedean (α) [ordered_add_comm_monoid α] : Prop := (arch : ∀ (x : α) {y}, 0 < y → ∃ n : ℕ, x ≤ n • y) -instance order_dual.archimedean [ordered_add_comm_group α] [archimedean α] : - archimedean (order_dual α) := +instance order_dual.archimedean [ordered_add_comm_group α] [archimedean α] : archimedean αᵒᵈ := ⟨λ x y hy, let ⟨n, hn⟩ := archimedean.arch (-x : α) (neg_pos.2 hy) in ⟨n, by rwa [neg_nsmul, neg_le_neg_iff] at hn⟩⟩ diff --git a/src/algebra/order/field.lean b/src/algebra/order/field.lean index 28f70e4a0629e..a2c1521c88c89 100644 --- a/src/algebra/order/field.lean +++ b/src/algebra/order/field.lean @@ -706,11 +706,11 @@ eq.symm $ monotone.map_max (λ x y, div_le_div_of_le hc) lemma min_div_div_right_of_nonpos {c : α} (hc : c ≤ 0) (a b : α) : min (a / c) (b / c) = (max a b) / c := -eq.symm $ @monotone.map_max α (order_dual α) _ _ _ _ _ (λ x y, div_le_div_of_nonpos_of_le hc) +eq.symm $ @monotone.map_max α αᵒᵈ _ _ _ _ _ (λ x y, div_le_div_of_nonpos_of_le hc) lemma max_div_div_right_of_nonpos {c : α} (hc : c ≤ 0) (a b : α) : max (a / c) (b / c) = (min a b) / c := -eq.symm $ @monotone.map_min α (order_dual α) _ _ _ _ _ (λ x y, div_le_div_of_nonpos_of_le hc) +eq.symm $ @monotone.map_min α αᵒᵈ _ _ _ _ _ (λ x y, div_le_div_of_nonpos_of_le hc) lemma abs_div (a b : α) : |a / b| = |a| / |b| := (abs_hom : α →*₀ α).map_div a b diff --git a/src/algebra/order/group.lean b/src/algebra/order/group.lean index ef151b9840a5d..af87bcf29f2f8 100644 --- a/src/algebra/order/group.lean +++ b/src/algebra/order/group.lean @@ -91,14 +91,14 @@ instance ordered_comm_group.has_exists_mul_of_le (α : Type u) has_exists_mul_of_le α := ⟨λ a b hab, ⟨b * a⁻¹, (mul_inv_cancel_comm_assoc a b).symm⟩⟩ -@[to_additive] instance [h : has_inv α] : has_inv (order_dual α) := h -@[to_additive] instance [h : has_div α] : has_div (order_dual α) := h -@[to_additive] instance [h : has_involutive_inv α] : has_involutive_inv (order_dual α) := h -@[to_additive] instance [h : div_inv_monoid α] : div_inv_monoid (order_dual α) := h -@[to_additive] instance [h : group α] : group (order_dual α) := h -@[to_additive] instance [h : comm_group α] : comm_group (order_dual α) := h - -@[to_additive] instance [ordered_comm_group α] : ordered_comm_group (order_dual α) := +@[to_additive] instance [h : has_inv α] : has_inv αᵒᵈ := h +@[to_additive] instance [h : has_div α] : has_div αᵒᵈ := h +@[to_additive] instance [h : has_involutive_inv α] : has_involutive_inv αᵒᵈ := h +@[to_additive] instance [h : div_inv_monoid α] : div_inv_monoid αᵒᵈ := h +@[to_additive] instance [h : group α] : group αᵒᵈ := h +@[to_additive] instance [h : comm_group α] : comm_group αᵒᵈ := h + +@[to_additive] instance [ordered_comm_group α] : ordered_comm_group αᵒᵈ := { .. order_dual.ordered_comm_monoid, .. order_dual.group } section group @@ -292,7 +292,7 @@ variable (α) /-- `x ↦ x⁻¹` as an order-reversing equivalence. -/ @[to_additive "`x ↦ -x` as an order-reversing equivalence.", simps] -def order_iso.inv : α ≃o order_dual α := +def order_iso.inv : α ≃o αᵒᵈ := { to_equiv := (equiv.inv α).trans order_dual.to_dual, map_rel_iff' := λ a b, @inv_le_inv_iff α _ _ _ _ _ _ } @@ -853,7 +853,7 @@ calc a ≤ b * (b⁻¹ * c) : h _ (lt_inv_mul_iff_lt.mpr hc) @[to_additive] lemma le_of_forall_lt_one_mul_le (h : ∀ ε < 1, a * ε ≤ b) : a ≤ b := -@le_of_forall_one_lt_le_mul (order_dual α) _ _ _ _ _ _ h +@le_of_forall_one_lt_le_mul αᵒᵈ _ _ _ _ _ _ h @[to_additive] lemma le_of_forall_one_lt_div_le (h : ∀ ε : α, 1 < ε → a / ε ≤ b) : a ≤ b := @@ -866,7 +866,7 @@ lemma le_iff_forall_one_lt_le_mul : a ≤ b ↔ ∀ ε, 1 < ε → a ≤ b * ε @[to_additive] lemma le_iff_forall_lt_one_mul_le : a ≤ b ↔ ∀ ε < 1, a * ε ≤ b := -@le_iff_forall_one_lt_le_mul (order_dual α) _ _ _ _ _ _ +@le_iff_forall_one_lt_le_mul αᵒᵈ _ _ _ _ _ _ end densely_ordered @@ -897,7 +897,7 @@ multiplication is monotone. -/ class linear_ordered_comm_group (α : Type u) extends ordered_comm_group α, linear_order α @[to_additive] instance [linear_ordered_comm_group α] : - linear_ordered_comm_group (order_dual α) := + linear_ordered_comm_group αᵒᵈ := { .. order_dual.ordered_comm_group, .. order_dual.linear_order α } section linear_ordered_comm_group @@ -933,11 +933,11 @@ mul_lt_mul_left' h c @[to_additive min_neg_neg] lemma min_inv_inv' (a b : α) : min (a⁻¹) (b⁻¹) = (max a b)⁻¹ := -eq.symm $ @monotone.map_max α (order_dual α) _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr +eq.symm $ @monotone.map_max α αᵒᵈ _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr @[to_additive max_neg_neg] lemma max_inv_inv' (a b : α) : max (a⁻¹) (b⁻¹) = (min a b)⁻¹ := -eq.symm $ @monotone.map_min α (order_dual α) _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr +eq.symm $ @monotone.map_min α αᵒᵈ _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr @[to_additive min_sub_sub_right] lemma min_div_div_right' (a b c : α) : min (a / c) (b / c) = min a b / c := diff --git a/src/algebra/order/module.lean b/src/algebra/order/module.lean index 45cf0ea3e7b90..6c236a59482b4 100644 --- a/src/algebra/order/module.lean +++ b/src/algebra/order/module.lean @@ -29,7 +29,7 @@ variables {k M N : Type*} namespace order_dual -instance [semiring k] [ordered_add_comm_monoid M] [module k M] : module k (order_dual M) := +instance [semiring k] [ordered_add_comm_monoid M] [module k M] : module k Mᵒᵈ := { add_smul := λ r s x, order_dual.rec (add_smul _ _) x, zero_smul := λ m, order_dual.rec (zero_smul _) m } @@ -110,7 +110,7 @@ calc ... = 0 : smul_zero' M c lemma smul_nonneg_of_nonpos_of_nonpos (hc : c ≤ 0) (ha : a ≤ 0) : 0 ≤ c • a := -@smul_nonpos_of_nonpos_of_nonneg k (order_dual M) _ _ _ _ _ _ hc ha +@smul_nonpos_of_nonpos_of_nonneg k Mᵒᵈ _ _ _ _ _ _ hc ha alias smul_pos_iff_of_neg ↔ _ smul_pos_of_neg_of_neg alias smul_neg_iff_of_pos ↔ _ smul_neg_of_pos_of_neg @@ -149,7 +149,7 @@ end variables (M) /-- Left scalar multiplication as an order isomorphism. -/ -@[simps] def order_iso.smul_left_dual {c : k} (hc : c < 0) : M ≃o order_dual M := +@[simps] def order_iso.smul_left_dual {c : k} (hc : c < 0) : M ≃o Mᵒᵈ := { to_fun := λ b, order_dual.to_dual (c • b), inv_fun := λ b, c⁻¹ • (order_dual.of_dual b), left_inv := inv_smul_smul₀ hc.ne, diff --git a/src/algebra/order/monoid.lean b/src/algebra/order/monoid.lean index 66f35e177ee4b..f03f7d9a1c43b 100644 --- a/src/algebra/order/monoid.lean +++ b/src/algebra/order/monoid.lean @@ -672,90 +672,90 @@ end linear_ordered_cancel_comm_monoid namespace order_dual -@[to_additive] instance [h : has_mul α] : has_mul (order_dual α) := h -@[to_additive] instance [h : has_one α] : has_one (order_dual α) := h -@[to_additive] instance [h : semigroup α] : semigroup (order_dual α) := h -@[to_additive] instance [h : comm_semigroup α] : comm_semigroup (order_dual α) := h -@[to_additive] instance [h : mul_one_class α] : mul_one_class (order_dual α) := h -@[to_additive] instance [h : monoid α] : monoid (order_dual α) := h -@[to_additive] instance [h : comm_monoid α] : comm_monoid (order_dual α) := h -@[to_additive] instance [h : left_cancel_monoid α] : left_cancel_monoid (order_dual α) := h -@[to_additive] instance [h : right_cancel_monoid α] : right_cancel_monoid (order_dual α) := h -@[to_additive] instance [h : cancel_monoid α] : cancel_monoid (order_dual α) := h -@[to_additive] instance [h : cancel_comm_monoid α] : cancel_comm_monoid (order_dual α) := h -instance [h : mul_zero_class α] : mul_zero_class (order_dual α) := h -instance [h : mul_zero_one_class α] : mul_zero_one_class (order_dual α) := h -instance [h : monoid_with_zero α] : monoid_with_zero (order_dual α) := h -instance [h : comm_monoid_with_zero α] : comm_monoid_with_zero (order_dual α) := h -instance [h : cancel_comm_monoid_with_zero α] : cancel_comm_monoid_with_zero (order_dual α) := h +@[to_additive] instance [h : has_mul α] : has_mul αᵒᵈ := h +@[to_additive] instance [h : has_one α] : has_one αᵒᵈ := h +@[to_additive] instance [h : semigroup α] : semigroup αᵒᵈ := h +@[to_additive] instance [h : comm_semigroup α] : comm_semigroup αᵒᵈ := h +@[to_additive] instance [h : mul_one_class α] : mul_one_class αᵒᵈ := h +@[to_additive] instance [h : monoid α] : monoid αᵒᵈ := h +@[to_additive] instance [h : comm_monoid α] : comm_monoid αᵒᵈ := h +@[to_additive] instance [h : left_cancel_monoid α] : left_cancel_monoid αᵒᵈ := h +@[to_additive] instance [h : right_cancel_monoid α] : right_cancel_monoid αᵒᵈ := h +@[to_additive] instance [h : cancel_monoid α] : cancel_monoid αᵒᵈ := h +@[to_additive] instance [h : cancel_comm_monoid α] : cancel_comm_monoid αᵒᵈ := h +instance [h : mul_zero_class α] : mul_zero_class αᵒᵈ := h +instance [h : mul_zero_one_class α] : mul_zero_one_class αᵒᵈ := h +instance [h : monoid_with_zero α] : monoid_with_zero αᵒᵈ := h +instance [h : comm_monoid_with_zero α] : comm_monoid_with_zero αᵒᵈ := h +instance [h : cancel_comm_monoid_with_zero α] : cancel_comm_monoid_with_zero αᵒᵈ := h @[to_additive] instance contravariant_class_mul_le [has_le α] [has_mul α] [c : contravariant_class α α (*) (≤)] : - contravariant_class (order_dual α) (order_dual α) (*) (≤) := + contravariant_class αᵒᵈ αᵒᵈ (*) (≤) := ⟨c.1.flip⟩ @[to_additive] instance covariant_class_mul_le [has_le α] [has_mul α] [c : covariant_class α α (*) (≤)] : - covariant_class (order_dual α) (order_dual α) (*) (≤) := + covariant_class αᵒᵈ αᵒᵈ (*) (≤) := ⟨c.1.flip⟩ @[to_additive] instance contravariant_class_swap_mul_le [has_le α] [has_mul α] [c : contravariant_class α α (swap (*)) (≤)] : - contravariant_class (order_dual α) (order_dual α) (swap (*)) (≤) := + contravariant_class αᵒᵈ αᵒᵈ (swap (*)) (≤) := ⟨c.1.flip⟩ @[to_additive] instance covariant_class_swap_mul_le [has_le α] [has_mul α] [c : covariant_class α α (swap (*)) (≤)] : - covariant_class (order_dual α) (order_dual α) (swap (*)) (≤) := + covariant_class αᵒᵈ αᵒᵈ (swap (*)) (≤) := ⟨c.1.flip⟩ @[to_additive] instance contravariant_class_mul_lt [has_lt α] [has_mul α] [c : contravariant_class α α (*) (<)] : - contravariant_class (order_dual α) (order_dual α) (*) (<) := + contravariant_class αᵒᵈ αᵒᵈ (*) (<) := ⟨c.1.flip⟩ @[to_additive] instance covariant_class_mul_lt [has_lt α] [has_mul α] [c : covariant_class α α (*) (<)] : - covariant_class (order_dual α) (order_dual α) (*) (<) := + covariant_class αᵒᵈ αᵒᵈ (*) (<) := ⟨c.1.flip⟩ @[to_additive] instance contravariant_class_swap_mul_lt [has_lt α] [has_mul α] [c : contravariant_class α α (swap (*)) (<)] : - contravariant_class (order_dual α) (order_dual α) (swap (*)) (<) := + contravariant_class αᵒᵈ αᵒᵈ (swap (*)) (<) := ⟨c.1.flip⟩ @[to_additive] instance covariant_class_swap_mul_lt [has_lt α] [has_mul α] [c : covariant_class α α (swap (*)) (<)] : - covariant_class (order_dual α) (order_dual α) (swap (*)) (<) := + covariant_class αᵒᵈ αᵒᵈ (swap (*)) (<) := ⟨c.1.flip⟩ @[to_additive] -instance [ordered_comm_monoid α] : ordered_comm_monoid (order_dual α) := +instance [ordered_comm_monoid α] : ordered_comm_monoid αᵒᵈ := { mul_le_mul_left := λ a b h c, mul_le_mul_left' h c, .. order_dual.partial_order α, .. order_dual.comm_monoid } @[to_additive ordered_cancel_add_comm_monoid.to_contravariant_class] instance ordered_cancel_comm_monoid.to_contravariant_class [ordered_cancel_comm_monoid α] : - contravariant_class (order_dual α) (order_dual α) has_mul.mul has_le.le := + contravariant_class αᵒᵈ αᵒᵈ has_mul.mul has_le.le := { elim := λ a b c bc, (ordered_cancel_comm_monoid.le_of_mul_le_mul_left a c b (dual_le.mp bc)) } @[to_additive] -instance [ordered_cancel_comm_monoid α] : ordered_cancel_comm_monoid (order_dual α) := +instance [ordered_cancel_comm_monoid α] : ordered_cancel_comm_monoid αᵒᵈ := { le_of_mul_le_mul_left := λ a b c : α, le_of_mul_le_mul_left', .. order_dual.ordered_comm_monoid, .. order_dual.cancel_comm_monoid } @[to_additive] instance [linear_ordered_cancel_comm_monoid α] : - linear_ordered_cancel_comm_monoid (order_dual α) := + linear_ordered_cancel_comm_monoid αᵒᵈ := { .. order_dual.linear_order α, .. order_dual.ordered_cancel_comm_monoid } @[to_additive] instance [linear_ordered_comm_monoid α] : - linear_ordered_comm_monoid (order_dual α) := + linear_ordered_comm_monoid αᵒᵈ := { .. order_dual.linear_order α, .. order_dual.ordered_comm_monoid } @@ -1089,7 +1089,7 @@ lemma coe_bit1 [has_one α] {a : α} : ((bit1 a : α) : with_bot α) = bit1 a := lemma add_ne_bot : a + b ≠ ⊥ ↔ a ≠ ⊥ ∧ b ≠ ⊥ := with_top.add_ne_top lemma bot_lt_add [partial_order α] {a b : with_bot α} : ⊥ < a + b ↔ ⊥ < a ∧ ⊥ < b := -@with_top.add_lt_top (order_dual α) _ _ _ _ +@with_top.add_lt_top αᵒᵈ _ _ _ _ lemma add_eq_coe : a + b = x ↔ ∃ (a' b' : α), ↑a' = a ∧ ↑b' = b ∧ a' + b' = x := with_top.add_eq_coe @@ -1100,35 +1100,35 @@ variables [preorder α] instance covariant_class_add_le [covariant_class α α (+) (≤)] : covariant_class (with_bot α) (with_bot α) (+) (≤) := -@order_dual.covariant_class_add_le (with_top $ order_dual α) _ _ _ +@order_dual.covariant_class_add_le (with_top αᵒᵈ) _ _ _ instance covariant_class_swap_add_le [covariant_class α α (swap (+)) (≤)] : covariant_class (with_bot α) (with_bot α) (swap (+)) (≤) := -@order_dual.covariant_class_swap_add_le (with_top $ order_dual α) _ _ _ +@order_dual.covariant_class_swap_add_le (with_top αᵒᵈ) _ _ _ instance contravariant_class_add_lt [contravariant_class α α (+) (<)] : contravariant_class (with_bot α) (with_bot α) (+) (<) := -@order_dual.contravariant_class_add_lt (with_top $ order_dual α) _ _ _ +@order_dual.contravariant_class_add_lt (with_top αᵒᵈ) _ _ _ instance contravariant_class_swap_add_lt [contravariant_class α α (swap (+)) (<)] : contravariant_class (with_bot α) (with_bot α) (swap (+)) (<) := -@order_dual.contravariant_class_swap_add_lt (with_top $ order_dual α) _ _ _ +@order_dual.contravariant_class_swap_add_lt (with_top αᵒᵈ) _ _ _ protected lemma le_of_add_le_add_left [contravariant_class α α (+) (≤)] (ha : a ≠ ⊥) (h : a + b ≤ a + c) : b ≤ c := -@with_top.le_of_add_le_add_left (order_dual α) _ _ _ _ _ _ ha h +@with_top.le_of_add_le_add_left αᵒᵈ _ _ _ _ _ _ ha h protected lemma le_of_add_le_add_right [contravariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊥) (h : b + a ≤ c + a) : b ≤ c := -@with_top.le_of_add_le_add_right (order_dual α) _ _ _ _ _ _ ha h +@with_top.le_of_add_le_add_right αᵒᵈ _ _ _ _ _ _ ha h protected lemma add_lt_add_left [covariant_class α α (+) (<)] (ha : a ≠ ⊥) (h : b < c) : a + b < a + c := -@with_top.add_lt_add_left (order_dual α) _ _ _ _ _ _ ha h +@with_top.add_lt_add_left αᵒᵈ _ _ _ _ _ _ ha h protected lemma add_lt_add_right [covariant_class α α (swap (+)) (<)] (ha : a ≠ ⊥) (h : b < c) : b + a < c + a := -@with_top.add_lt_add_right (order_dual α) _ _ _ _ _ _ ha h +@with_top.add_lt_add_right αᵒᵈ _ _ _ _ _ _ ha h protected lemma add_le_add_iff_left [covariant_class α α (+) (≤)] [contravariant_class α α (+) (≤)] (ha : a ≠ ⊥) : a + b ≤ a + c ↔ b ≤ c := @@ -1148,11 +1148,11 @@ protected lemma add_lt_add_iff_right [covariant_class α α (swap (+)) (<)] protected lemma add_lt_add_of_le_of_lt [covariant_class α α (+) (<)] [covariant_class α α (swap (+)) (≤)] (hb : b ≠ ⊥) (hab : a ≤ b) (hcd : c < d) : a + c < b + d := -@with_top.add_lt_add_of_le_of_lt (order_dual α) _ _ _ _ _ _ _ _ hb hab hcd +@with_top.add_lt_add_of_le_of_lt αᵒᵈ _ _ _ _ _ _ _ _ hb hab hcd protected lemma add_lt_add_of_lt_of_le [covariant_class α α (+) (≤)] [covariant_class α α (swap (+)) (<)] (hd : d ≠ ⊥) (hab : a < b) (hcd : c ≤ d) : a + c < b + d := -@with_top.add_lt_add_of_lt_of_le (order_dual α) _ _ _ _ _ _ _ _ hd hab hcd +@with_top.add_lt_add_of_lt_of_le αᵒᵈ _ _ _ _ _ _ _ _ hd hab hcd end has_add end with_bot diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index fef4937123adf..5e964988d52e9 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -95,20 +95,20 @@ namespace order_dual /-! Note that `order_dual` does not satisfy any of the ordered ring typeclasses due to the `zero_le_one` field. -/ -instance [h : distrib α] : distrib (order_dual α) := h -instance [has_mul α] [h : has_distrib_neg α] : has_distrib_neg (order_dual α) := h -instance [h : non_unital_non_assoc_semiring α] : non_unital_non_assoc_semiring (order_dual α) := h -instance [h : non_unital_semiring α] : non_unital_semiring (order_dual α) := h -instance [h : non_assoc_semiring α] : non_assoc_semiring (order_dual α) := h -instance [h : semiring α] : semiring (order_dual α) := h -instance [h : non_unital_comm_semiring α] : non_unital_comm_semiring (order_dual α) := h -instance [h : comm_semiring α] : comm_semiring (order_dual α) := h -instance [h : non_unital_non_assoc_ring α] : non_unital_non_assoc_ring (order_dual α) := h -instance [h : non_unital_ring α] : non_unital_ring (order_dual α) := h -instance [h : non_assoc_ring α] : non_assoc_ring (order_dual α) := h -instance [h : ring α] : ring (order_dual α) := h -instance [h : non_unital_comm_ring α] : non_unital_comm_ring (order_dual α) := h -instance [h : comm_ring α] : comm_ring (order_dual α) := h +instance [h : distrib α] : distrib αᵒᵈ := h +instance [has_mul α] [h : has_distrib_neg α] : has_distrib_neg αᵒᵈ := h +instance [h : non_unital_non_assoc_semiring α] : non_unital_non_assoc_semiring αᵒᵈ := h +instance [h : non_unital_semiring α] : non_unital_semiring αᵒᵈ := h +instance [h : non_assoc_semiring α] : non_assoc_semiring αᵒᵈ := h +instance [h : semiring α] : semiring αᵒᵈ := h +instance [h : non_unital_comm_semiring α] : non_unital_comm_semiring αᵒᵈ := h +instance [h : comm_semiring α] : comm_semiring αᵒᵈ := h +instance [h : non_unital_non_assoc_ring α] : non_unital_non_assoc_ring αᵒᵈ := h +instance [h : non_unital_ring α] : non_unital_ring αᵒᵈ := h +instance [h : non_assoc_ring α] : non_assoc_ring αᵒᵈ := h +instance [h : ring α] : ring αᵒᵈ := h +instance [h : non_unital_comm_ring α] : non_unital_comm_ring αᵒᵈ := h +instance [h : comm_ring α] : comm_ring αᵒᵈ := h end order_dual diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index 85bc3a0a26666..f7f874104681f 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -49,40 +49,36 @@ namespace order_dual variables {R M : Type*} -instance [has_scalar R M] : has_scalar R (order_dual M) := -{ smul := λ k x, order_dual.rec (λ x', (k • x' : M)) x } +instance [has_scalar R M] : has_scalar R Mᵒᵈ := ⟨λ k x, order_dual.rec (λ x', (k • x' : M)) x⟩ -instance [has_zero R] [add_zero_class M] [h : smul_with_zero R M] : - smul_with_zero R (order_dual M) := +instance [has_zero R] [add_zero_class M] [h : smul_with_zero R M] : smul_with_zero R Mᵒᵈ := { zero_smul := λ m, order_dual.rec (zero_smul _) m, smul_zero := λ r, order_dual.rec (smul_zero' _) r, ..order_dual.has_scalar } -instance [monoid R] [mul_action R M] : mul_action R (order_dual M) := +instance [monoid R] [mul_action R M] : mul_action R Mᵒᵈ := { one_smul := λ m, order_dual.rec (one_smul _) m, mul_smul := λ r, order_dual.rec mul_smul r, ..order_dual.has_scalar } instance [monoid_with_zero R] [add_monoid M] [mul_action_with_zero R M] : - mul_action_with_zero R (order_dual M) := + mul_action_with_zero R Mᵒᵈ := { ..order_dual.mul_action, ..order_dual.smul_with_zero } instance [monoid_with_zero R] [add_monoid M] [distrib_mul_action R M] : - distrib_mul_action R (order_dual M) := + distrib_mul_action R Mᵒᵈ := { smul_add := λ k a, order_dual.rec (λ a' b, order_dual.rec (smul_add _ _) b) a, smul_zero := λ r, order_dual.rec smul_zero r } instance [ordered_semiring R] [ordered_add_comm_monoid M] [smul_with_zero R M] [ordered_smul R M] : - ordered_smul R (order_dual M) := + ordered_smul R Mᵒᵈ := { smul_lt_smul_of_pos := λ a b, @ordered_smul.smul_lt_smul_of_pos R M _ _ _ _ b a, lt_of_smul_lt_smul_of_pos := λ a b, @ordered_smul.lt_of_smul_lt_smul_of_pos R M _ _ _ _ b a } @[simp] lemma to_dual_smul [has_scalar R M] {c : R} {a : M} : to_dual (c • a) = c • to_dual a := rfl - -@[simp] lemma of_dual_smul [has_scalar R M] {c : R} {a : order_dual M} : - of_dual (c • a) = c • of_dual a := +@[simp] lemma of_dual_smul [has_scalar R M] {c : R} {a : Mᵒᵈ} : of_dual (c • a) = c • of_dual a := rfl end order_dual @@ -109,7 +105,7 @@ calc (0 : M) = c • (0 : M) : (smul_zero' M c).symm ... ≤ c • a : smul_le_smul_of_nonneg ha hc lemma smul_nonpos_of_nonneg_of_nonpos (hc : 0 ≤ c) (ha : a ≤ 0) : c • a ≤ 0 := -@smul_nonneg R (order_dual M) _ _ _ _ _ _ hc ha +@smul_nonneg R Mᵒᵈ _ _ _ _ _ _ hc ha lemma eq_of_smul_eq_smul_of_pos_of_le (h₁ : c • a = c • b) (hc : 0 < c) (hle : a ≤ b) : a = b := diff --git a/src/algebra/order/with_zero.lean b/src/algebra/order/with_zero.lean index 8a199a22621e6..33eba665df98d 100644 --- a/src/algebra/order/with_zero.lean +++ b/src/algebra/order/with_zero.lean @@ -34,7 +34,7 @@ variables {α : Type*} variables {a b c d x y z : α} instance [linear_ordered_add_comm_monoid_with_top α] : - linear_ordered_comm_monoid_with_zero (multiplicative (order_dual α)) := + linear_ordered_comm_monoid_with_zero (multiplicative αᵒᵈ) := { zero := multiplicative.of_add (⊤ : α), zero_mul := top_add, mul_zero := add_top, @@ -43,7 +43,7 @@ instance [linear_ordered_add_comm_monoid_with_top α] : ..multiplicative.linear_order } instance [linear_ordered_add_comm_group_with_top α] : - linear_ordered_comm_group_with_zero (multiplicative (order_dual α)) := + linear_ordered_comm_group_with_zero (multiplicative αᵒᵈ) := { inv_zero := linear_ordered_add_comm_group_with_top.neg_top, mul_inv_cancel := linear_ordered_add_comm_group_with_top.add_neg_cancel, ..multiplicative.div_inv_monoid, @@ -196,7 +196,7 @@ lemma ne_zero_of_lt (h : b < a) : a ≠ 0 := lemma pow_pos_iff [no_zero_divisors α] {n : ℕ} (hn : 0 < n) : 0 < a ^ n ↔ 0 < a := by simp_rw [zero_lt_iff, pow_ne_zero_iff hn] -instance : linear_ordered_add_comm_monoid_with_top (additive (order_dual α)) := +instance : linear_ordered_add_comm_monoid_with_top (additive αᵒᵈ) := { top := (0 : α), top_add' := λ a, (zero_mul a : (0 : α) * a = 0), le_top := λ _, zero_le', @@ -294,7 +294,7 @@ by rw [div_eq_mul_inv, le_mul_inv_iff₀ hc] lemma div_le_iff₀ (hc : c ≠ 0) : a/c ≤ b ↔ a ≤ b*c := by rw [div_eq_mul_inv, mul_inv_le_iff₀ hc] -instance : linear_ordered_add_comm_group_with_top (additive (order_dual α)) := +instance : linear_ordered_add_comm_group_with_top (additive αᵒᵈ) := { neg_top := inv_zero, add_neg_cancel := λ a ha, mul_inv_cancel ha, ..additive.sub_neg_monoid, diff --git a/src/algebra/support.lean b/src/algebra/support.lean index d900375d0d67d..9b9b9dc8198b0 100644 --- a/src/algebra/support.lean +++ b/src/algebra/support.lean @@ -127,7 +127,7 @@ end @[to_additive] lemma mul_support_infi [conditionally_complete_lattice M] [nonempty ι] (f : ι → α → M) : mul_support (λ x, ⨅ i, f i x) ⊆ ⋃ i, mul_support (f i) := -@mul_support_supr _ (order_dual M) ι ⟨(1:M)⟩ _ _ f +@mul_support_supr _ Mᵒᵈ ι ⟨(1:M)⟩ _ _ f @[to_additive] lemma mul_support_comp_subset {g : M → N} (hg : g 1 = 1) (f : α → M) : mul_support (g ∘ f) ⊆ mul_support f := diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index 1898859ae7fe1..d45eb115408ec 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -161,14 +161,12 @@ section gc variable (R) /-- `zero_locus` and `vanishing_ideal` form a galois connection. -/ -lemma gc : @galois_connection - (ideal R) (order_dual (set (prime_spectrum R))) _ _ +lemma gc : @galois_connection (ideal R) (set (prime_spectrum R))ᵒᵈ _ _ (λ I, zero_locus I) (λ t, vanishing_ideal t) := λ I t, subset_zero_locus_iff_le_vanishing_ideal t I /-- `zero_locus` and `vanishing_ideal` form a galois connection. -/ -lemma gc_set : @galois_connection - (set R) (order_dual (set (prime_spectrum R))) _ _ +lemma gc_set : @galois_connection (set R) (set (prime_spectrum R))ᵒᵈ _ _ (λ s, zero_locus s) (λ t, vanishing_ideal t) := have ideal_gc : galois_connection (ideal.span) coe := (submodule.gi R R).gc, by simpa [zero_locus_span, function.comp] using ideal_gc.compose (gc R) diff --git a/src/algebraic_geometry/projective_spectrum/topology.lean b/src/algebraic_geometry/projective_spectrum/topology.lean index 31c1aceae6351..68e82cd8de7e5 100644 --- a/src/algebraic_geometry/projective_spectrum/topology.lean +++ b/src/algebraic_geometry/projective_spectrum/topology.lean @@ -130,19 +130,19 @@ lemma subset_zero_locus_iff_le_vanishing_ideal (t : set (projective_spectrum variable (𝒜) /-- `zero_locus` and `vanishing_ideal` form a galois connection. -/ lemma gc_ideal : @galois_connection - (ideal A) (order_dual (set (projective_spectrum 𝒜))) _ _ + (ideal A) (set (projective_spectrum 𝒜))ᵒᵈ _ _ (λ I, zero_locus 𝒜 I) (λ t, (vanishing_ideal t).to_ideal) := λ I t, subset_zero_locus_iff_le_vanishing_ideal t I /-- `zero_locus` and `vanishing_ideal` form a galois connection. -/ lemma gc_set : @galois_connection - (set A) (order_dual (set (projective_spectrum 𝒜))) _ _ + (set A) (set (projective_spectrum 𝒜))ᵒᵈ _ _ (λ s, zero_locus 𝒜 s) (λ t, vanishing_ideal t) := have ideal_gc : galois_connection (ideal.span) coe := (submodule.gi A _).gc, by simpa [zero_locus_span, function.comp] using galois_connection.compose ideal_gc (gc_ideal 𝒜) lemma gc_homogeneous_ideal : @galois_connection - (homogeneous_ideal 𝒜) (order_dual (set (projective_spectrum 𝒜))) _ _ + (homogeneous_ideal 𝒜) (set (projective_spectrum 𝒜))ᵒᵈ _ _ (λ I, zero_locus 𝒜 I) (λ t, (vanishing_ideal t)) := λ I t, by simpa [show I.to_ideal ≤ (vanishing_ideal t).to_ideal ↔ I ≤ (vanishing_ideal t), from iff.rfl] using subset_zero_locus_iff_le_vanishing_ideal t I.to_ideal diff --git a/src/analysis/box_integral/partition/filter.lean b/src/analysis/box_integral/partition/filter.lean index 2c876aa786f47..b73ff899457a2 100644 --- a/src/analysis/box_integral/partition/filter.lean +++ b/src/analysis/box_integral/partition/filter.lean @@ -193,7 +193,7 @@ variables {l l₁ l₂ : integration_params} namespace integration_params /-- Auxiliary equivalence with a product type used to lift an order. -/ -def equiv_prod : integration_params ≃ bool × order_dual bool × order_dual bool := +def equiv_prod : integration_params ≃ bool × boolᵒᵈ × boolᵒᵈ := { to_fun := λ l, ⟨l.1, order_dual.to_dual l.2, order_dual.to_dual l.3⟩, inv_fun := λ l, ⟨l.1, order_dual.of_dual l.2.1, order_dual.of_dual l.2.2⟩, left_inv := λ ⟨a, b, c⟩, rfl, @@ -203,7 +203,7 @@ instance : partial_order integration_params := partial_order.lift equiv_prod equiv_prod.injective /-- Auxiliary `order_iso` with a product type used to lift a `bounded_order` structure. -/ -def iso_prod : integration_params ≃o bool × order_dual bool × order_dual bool := +def iso_prod : integration_params ≃o bool × boolᵒᵈ × boolᵒᵈ := ⟨equiv_prod, λ ⟨x, y, z⟩, iff.rfl⟩ instance : bounded_order integration_params := diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index c3b4b10c672a0..162b16ef7cba2 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -696,8 +696,7 @@ calc : add_le_add (smul_le_smul_of_nonneg hx ha) (smul_le_smul_of_nonneg hy hb) ... = r : convex.combo_self hab _ -lemma convex_Ici (r : β) : convex 𝕜 (Ici r) := -@convex_Iic 𝕜 (order_dual β) _ _ _ _ r +lemma convex_Ici (r : β) : convex 𝕜 (Ici r) := @convex_Iic 𝕜 βᵒᵈ _ _ _ _ r lemma convex_Icc (r s : β) : convex 𝕜 (Icc r s) := Ici_inter_Iic.subst ((convex_Ici r).inter $ convex_Iic s) @@ -736,8 +735,7 @@ begin ... = r : convex.combo_self hab _ end -lemma convex_Ioi (r : β) : convex 𝕜 (Ioi r) := -@convex_Iio 𝕜 (order_dual β) _ _ _ _ r +lemma convex_Ioi (r : β) : convex 𝕜 (Ioi r) := @convex_Iio 𝕜 βᵒᵈ _ _ _ _ r lemma convex_Ioo (r s : β) : convex 𝕜 (Ioo r s) := Ioi_inter_Iio.subst ((convex_Ioi r).inter $ convex_Iio s) @@ -786,27 +784,27 @@ lemma monotone_on.convex_lt (hf : monotone_on f s) (hs : convex 𝕜 s) (r : β) lemma monotone_on.convex_ge (hf : monotone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | r ≤ f x} := -@monotone_on.convex_le 𝕜 (order_dual E) (order_dual β) _ _ _ _ _ _ _ hf.dual hs r +@monotone_on.convex_le 𝕜 Eᵒᵈ βᵒᵈ _ _ _ _ _ _ _ hf.dual hs r lemma monotone_on.convex_gt (hf : monotone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | r < f x} := -@monotone_on.convex_lt 𝕜 (order_dual E) (order_dual β) _ _ _ _ _ _ _ hf.dual hs r +@monotone_on.convex_lt 𝕜 Eᵒᵈ βᵒᵈ _ _ _ _ _ _ _ hf.dual hs r lemma antitone_on.convex_le (hf : antitone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | f x ≤ r} := -@monotone_on.convex_ge 𝕜 E (order_dual β) _ _ _ _ _ _ _ hf hs r +@monotone_on.convex_ge 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r lemma antitone_on.convex_lt (hf : antitone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | f x < r} := -@monotone_on.convex_gt 𝕜 E (order_dual β) _ _ _ _ _ _ _ hf hs r +@monotone_on.convex_gt 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r lemma antitone_on.convex_ge (hf : antitone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | r ≤ f x} := -@monotone_on.convex_le 𝕜 E (order_dual β) _ _ _ _ _ _ _ hf hs r +@monotone_on.convex_le 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r lemma antitone_on.convex_gt (hf : antitone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | r < f x} := -@monotone_on.convex_lt 𝕜 E (order_dual β) _ _ _ _ _ _ _ hf hs r +@monotone_on.convex_lt 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r lemma monotone.convex_le (hf : monotone f) (r : β) : convex 𝕜 {x | f x ≤ r} := diff --git a/src/analysis/convex/extrema.lean b/src/analysis/convex/extrema.lean index 5bc221d88a1c9..340be610d6774 100644 --- a/src/analysis/convex/extrema.lean +++ b/src/analysis/convex/extrema.lean @@ -74,8 +74,7 @@ end lemma is_max_on.of_is_local_max_on_of_concave_on {f : E → β} {a : E} (a_in_s : a ∈ s) (h_localmax: is_local_max_on f s a) (h_conc : concave_on ℝ s f) : is_max_on f s a := -@is_min_on.of_is_local_min_on_of_convex_on - _ (order_dual β) _ _ _ _ _ _ _ _ s f a a_in_s h_localmax h_conc +@is_min_on.of_is_local_min_on_of_convex_on _ βᵒᵈ _ _ _ _ _ _ _ _ s f a a_in_s h_localmax h_conc /-- A local minimum of a convex function is a global minimum. -/ lemma is_min_on.of_is_local_min_of_convex_univ {f : E → β} {a : E} @@ -86,4 +85,4 @@ lemma is_min_on.of_is_local_min_of_convex_univ {f : E → β} {a : E} /-- A local maximum of a concave function is a global maximum. -/ lemma is_max_on.of_is_local_max_of_convex_univ {f : E → β} {a : E} (h_local_max : is_local_max f a) (h_conc : concave_on ℝ univ f) : ∀ x, f x ≤ f a := -@is_min_on.of_is_local_min_of_convex_univ _ (order_dual β) _ _ _ _ _ _ _ _ f a h_local_max h_conc +@is_min_on.of_is_local_min_of_convex_univ _ βᵒᵈ _ _ _ _ _ _ _ _ f a h_local_max h_conc diff --git a/src/analysis/convex/function.lean b/src/analysis/convex/function.lean index bcc1464ef2596..94cd592ec2d38 100644 --- a/src/analysis/convex/function.lean +++ b/src/analysis/convex/function.lean @@ -131,7 +131,7 @@ lemma convex_on_const (c : β) (hs : convex 𝕜 s) : convex_on 𝕜 s (λ x:E, ⟨hs, λ x y _ _ a b _ _ hab, (convex.combo_self hab c).ge⟩ lemma concave_on_const (c : β) (hs : convex 𝕜 s) : concave_on 𝕜 s (λ x:E, c) := -@convex_on_const _ _ (order_dual β) _ _ _ _ _ _ c hs +@convex_on_const _ _ βᵒᵈ _ _ _ _ _ _ c hs lemma convex_on_of_convex_epigraph (h : convex 𝕜 {p : E × β | p.1 ∈ s ∧ f p.1 ≤ p.2}) : convex_on 𝕜 s f := @@ -140,7 +140,7 @@ lemma convex_on_of_convex_epigraph (h : convex 𝕜 {p : E × β | p.1 ∈ s ∧ lemma concave_on_of_convex_hypograph (h : convex 𝕜 {p : E × β | p.1 ∈ s ∧ p.2 ≤ f p.1}) : concave_on 𝕜 s f := -@convex_on_of_convex_epigraph 𝕜 E (order_dual β) _ _ _ _ _ _ _ h +@convex_on_of_convex_epigraph 𝕜 E βᵒᵈ _ _ _ _ _ _ _ h end module @@ -180,7 +180,7 @@ lemma convex_on_iff_convex_epigraph : lemma concave_on_iff_convex_hypograph : concave_on 𝕜 s f ↔ convex 𝕜 {p : E × β | p.1 ∈ s ∧ p.2 ≤ f p.1} := -@convex_on_iff_convex_epigraph 𝕜 E (order_dual β) _ _ _ _ _ _ _ f +@convex_on_iff_convex_epigraph 𝕜 E βᵒᵈ _ _ _ _ _ _ _ f end ordered_smul @@ -234,7 +234,7 @@ lemma concave_on_iff_forall_pos {s : set E} {f : E → β} : concave_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → a + b = 1 → a • f x + b • f y ≤ f (a • x + b • y) := -@convex_on_iff_forall_pos 𝕜 E (order_dual β) _ _ _ _ _ _ _ +@convex_on_iff_forall_pos 𝕜 E βᵒᵈ _ _ _ _ _ _ _ lemma convex_on_iff_pairwise_pos {s : set E} {f : E → β} : convex_on 𝕜 s f ↔ convex 𝕜 s ∧ @@ -253,7 +253,7 @@ lemma concave_on_iff_pairwise_pos {s : set E} {f : E → β} : concave_on 𝕜 s f ↔ convex 𝕜 s ∧ s.pairwise (λ x y, ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → a + b = 1 → a • f x + b • f y ≤ f (a • x + b • y)) := -@convex_on_iff_pairwise_pos 𝕜 E (order_dual β) _ _ _ _ _ _ _ +@convex_on_iff_pairwise_pos 𝕜 E βᵒᵈ _ _ _ _ _ _ _ /-- A linear map is convex. -/ lemma linear_map.convex_on (f : E →ₗ[𝕜] β) {s : set E} (hs : convex 𝕜 s) : convex_on 𝕜 s f := @@ -314,7 +314,7 @@ main use case is `E = ℝ` however one can apply it, e.g., to `ℝ^n` with lexic lemma linear_order.concave_on_of_lt (hs : convex 𝕜 s) (hf : ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → x < y → ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → a + b = 1 → a • f x + b • f y ≤ f (a • x + b • y)) : concave_on 𝕜 s f := -@linear_order.convex_on_of_lt _ _ (order_dual β) _ _ _ _ _ _ s f hs hf +@linear_order.convex_on_of_lt _ _ βᵒᵈ _ _ _ _ _ _ s f hs hf /-- For a function on a convex set in a linearly ordered space (where the order and the algebraic structures aren't necessarily compatible), in order to prove that it is convex, it suffices to @@ -337,7 +337,7 @@ main use case is `E = 𝕜` however one can apply it, e.g., to `𝕜^n` with lex lemma linear_order.strict_concave_on_of_lt (hs : convex 𝕜 s) (hf : ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → x < y → ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → a + b = 1 → a • f x + b • f y < f (a • x + b • y)) : strict_concave_on 𝕜 s f := -@linear_order.strict_convex_on_of_lt _ _ (order_dual β) _ _ _ _ _ _ _ _ hs hf +@linear_order.strict_convex_on_of_lt _ _ βᵒᵈ _ _ _ _ _ _ _ _ hs hf end linear_order end module @@ -824,7 +824,7 @@ end⟩ lemma concave_on_iff_div {f : E → β} : concave_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → ∀ ⦃a b : 𝕜⦄, 0 ≤ a → 0 ≤ b → 0 < a + b → (a/(a+b)) • f x + (b/(a+b)) • f y ≤ f ((a/(a+b)) • x + (b/(a+b)) • y) := -@convex_on_iff_div _ _ (order_dual β) _ _ _ _ _ _ _ +@convex_on_iff_div _ _ βᵒᵈ _ _ _ _ _ _ _ lemma strict_convex_on_iff_div {f : E → β} : strict_convex_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → x ≠ y → ∀ ⦃a b : 𝕜⦄, 0 < a @@ -844,7 +844,7 @@ end⟩ lemma strict_concave_on_iff_div {f : E → β} : strict_concave_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → x ≠ y → ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → (a/(a+b)) • f x + (b/(a+b)) • f y < f ((a/(a+b)) • x + (b/(a+b)) • y) := -@strict_convex_on_iff_div _ _ (order_dual β) _ _ _ _ _ _ _ +@strict_convex_on_iff_div _ _ βᵒᵈ _ _ _ _ _ _ _ end has_scalar end ordered_add_comm_monoid diff --git a/src/analysis/convex/jensen.lean b/src/analysis/convex/jensen.lean index a39c01298cb18..a548eb26f0c6b 100644 --- a/src/analysis/convex/jensen.lean +++ b/src/analysis/convex/jensen.lean @@ -52,7 +52,7 @@ end lemma concave_on.le_map_center_mass (hf : concave_on 𝕜 s f) (h₀ : ∀ i ∈ t, 0 ≤ w i) (h₁ : 0 < ∑ i in t, w i) (hmem : ∀ i ∈ t, p i ∈ s) : t.center_mass w (f ∘ p) ≤ f (t.center_mass w p) := -@convex_on.map_center_mass_le 𝕜 E (order_dual β) _ _ _ _ _ _ _ _ _ _ _ _ hf h₀ h₁ hmem +@convex_on.map_center_mass_le 𝕜 E βᵒᵈ _ _ _ _ _ _ _ _ _ _ _ _ hf h₀ h₁ hmem /-- Convex **Jensen's inequality**, `finset.sum` version. -/ lemma convex_on.map_sum_le (hf : convex_on 𝕜 s f) (h₀ : ∀ i ∈ t, 0 ≤ w i) (h₁ : ∑ i in t, w i = 1) @@ -65,7 +65,7 @@ by simpa only [center_mass, h₁, inv_one, one_smul] lemma concave_on.le_map_sum (hf : concave_on 𝕜 s f) (h₀ : ∀ i ∈ t, 0 ≤ w i) (h₁ : ∑ i in t, w i = 1) (hmem : ∀ i ∈ t, p i ∈ s) : ∑ i in t, w i • f (p i) ≤ f (∑ i in t, w i • p i) := -@convex_on.map_sum_le 𝕜 E (order_dual β) _ _ _ _ _ _ _ _ _ _ _ _ hf h₀ h₁ hmem +@convex_on.map_sum_le 𝕜 E βᵒᵈ _ _ _ _ _ _ _ _ _ _ _ _ hf h₀ h₁ hmem end jensen @@ -100,7 +100,7 @@ end lemma concave_on.exists_le_of_center_mass (h : concave_on 𝕜 s f) (hw₀ : ∀ i ∈ t, 0 ≤ w i) (hw₁ : 0 < ∑ i in t, w i) (hp : ∀ i ∈ t, p i ∈ s) : ∃ i ∈ t, f (p i) ≤ f (t.center_mass w p) := -@convex_on.exists_ge_of_center_mass 𝕜 E (order_dual β) _ _ _ _ _ _ _ _ _ _ _ _ h hw₀ hw₁ hp +@convex_on.exists_ge_of_center_mass 𝕜 E βᵒᵈ _ _ _ _ _ _ _ _ _ _ _ _ h hw₀ hw₁ hp /-- Maximum principle for convex functions. If a function `f` is convex on the convex hull of `s`, then the eventual maximum of `f` on `convex_hull 𝕜 s` lies in `s`. -/ @@ -118,6 +118,6 @@ end then the eventual minimum of `f` on `convex_hull 𝕜 s` lies in `s`. -/ lemma concave_on.exists_le_of_mem_convex_hull (hf : concave_on 𝕜 (convex_hull 𝕜 s) f) {x} (hx : x ∈ convex_hull 𝕜 s) : ∃ y ∈ s, f y ≤ f x := -@convex_on.exists_ge_of_mem_convex_hull 𝕜 E (order_dual β) _ _ _ _ _ _ _ _ hf _ hx +@convex_on.exists_ge_of_mem_convex_hull 𝕜 E βᵒᵈ _ _ _ _ _ _ _ _ hf _ hx end maximum_principle diff --git a/src/analysis/convex/quasiconvex.lean b/src/analysis/convex/quasiconvex.lean index a61ed4d00ffc2..6fd19188f29d9 100644 --- a/src/analysis/convex/quasiconvex.lean +++ b/src/analysis/convex/quasiconvex.lean @@ -80,7 +80,7 @@ lemma convex.quasiconvex_on_of_convex_le (hs : convex 𝕜 s) (h : ∀ r, convex lemma convex.quasiconcave_on_of_convex_ge (hs : convex 𝕜 s) (h : ∀ r, convex 𝕜 {x | r ≤ f x}) : quasiconcave_on 𝕜 s f := -@convex.quasiconvex_on_of_convex_le 𝕜 E (order_dual β) _ _ _ _ _ _ hs h +@convex.quasiconvex_on_of_convex_le 𝕜 E βᵒᵈ _ _ _ _ _ _ hs h lemma quasiconvex_on.convex [is_directed β (≤)] (hf : quasiconvex_on 𝕜 s f) : convex 𝕜 s := λ x y hx hy a b ha hb hab, @@ -122,7 +122,7 @@ lemma quasiconcave_on_iff_min_le : quasiconcave_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → ∀ ⦃a b : 𝕜⦄, 0 ≤ a → 0 ≤ b → a + b = 1 → min (f x) (f y) ≤ f (a • x + b • y) := -@quasiconvex_on_iff_le_max 𝕜 E (order_dual β) _ _ _ _ _ _ +@quasiconvex_on_iff_le_max 𝕜 E βᵒᵈ _ _ _ _ _ _ lemma quasilinear_on_iff_mem_interval : quasilinear_on 𝕜 s f ↔ convex 𝕜 s ∧ diff --git a/src/analysis/convex/strict.lean b/src/analysis/convex/strict.lean index 935c86d37a588..fae867240a565 100644 --- a/src/analysis/convex/strict.lean +++ b/src/analysis/convex/strict.lean @@ -153,8 +153,7 @@ begin { exact add_lt_add (smul_lt_smul_of_pos hx ha) (smul_lt_smul_of_pos hy hb) } end -lemma strict_convex_Ici (r : β) : strict_convex 𝕜 (Ici r) := -@strict_convex_Iic 𝕜 (order_dual β) _ _ _ _ _ _ r +lemma strict_convex_Ici (r : β) : strict_convex 𝕜 (Ici r) := @strict_convex_Iic 𝕜 βᵒᵈ _ _ _ _ _ _ r lemma strict_convex_Icc (r s : β) : strict_convex 𝕜 (Icc r s) := (strict_convex_Ici r).inter $ strict_convex_Iic s diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 3f3ef98cff338..7414b5e7a66f7 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -2203,7 +2203,7 @@ variables (𝕜 E) /-- `submodule.orthogonal` gives a `galois_connection` between `submodule 𝕜 E` and its `order_dual`. -/ lemma submodule.orthogonal_gc : - @galois_connection (submodule 𝕜 E) (order_dual $ submodule 𝕜 E) _ _ + @galois_connection (submodule 𝕜 E) (submodule 𝕜 E)ᵒᵈ _ _ submodule.orthogonal submodule.orthogonal := λ K₁ K₂, ⟨λ h v hv u hu, submodule.inner_left_of_mem_orthogonal hv (h hu), λ h v hv u hu, submodule.inner_left_of_mem_orthogonal hv (h hu)⟩ diff --git a/src/analysis/normed_space/lattice_ordered_group.lean b/src/analysis/normed_space/lattice_ordered_group.lean index b6046e7258ea8..447bccf3bda2c 100644 --- a/src/analysis/normed_space/lattice_ordered_group.lean +++ b/src/analysis/normed_space/lattice_ordered_group.lean @@ -63,7 +63,7 @@ instance normed_lattice_add_comm_group_to_ordered_add_comm_group {α : Type*} Let `α` be a normed group with a partial order. Then the order dual is also a normed group. -/ @[priority 100] -- see Note [lower instance priority] -instance {α : Type*} : Π [normed_group α], normed_group (order_dual α) := id +instance {α : Type*} : Π [normed_group α], normed_group αᵒᵈ := id variables {α : Type*} [normed_lattice_add_comm_group α] open lattice_ordered_comm_group @@ -84,7 +84,7 @@ Let `α` be a normed lattice ordered group, then the order dual is also a normed lattice ordered group. -/ @[priority 100] -- see Note [lower instance priority] -instance : normed_lattice_add_comm_group (order_dual α) := +instance : normed_lattice_add_comm_group αᵒᵈ := { add_le_add_left := begin intros a b h₁ c, rw ← order_dual.dual_le, @@ -150,7 +150,7 @@ end instance normed_lattice_add_comm_group_has_continuous_sup {α : Type*} [normed_lattice_add_comm_group α] : has_continuous_sup α := -order_dual.has_continuous_sup (order_dual α) +order_dual.has_continuous_sup αᵒᵈ /-- Let `α` be a normed lattice ordered group. Then `α` is a topological lattice in the norm topology. diff --git a/src/combinatorics/pigeonhole.lean b/src/combinatorics/pigeonhole.lean index 4b1a755f04ea4..7c3655380f457 100644 --- a/src/combinatorics/pigeonhole.lean +++ b/src/combinatorics/pigeonhole.lean @@ -124,7 +124,7 @@ than `b`. -/ lemma exists_sum_fiber_lt_of_maps_to_of_sum_lt_nsmul (hf : ∀ a ∈ s, f a ∈ t) (hb : (∑ x in s, w x) < t.card • b) : ∃ y ∈ t, (∑ x in s.filter (λ x, f x = y), w x) < b := -@exists_lt_sum_fiber_of_maps_to_of_nsmul_lt_sum α β (order_dual M) _ _ _ _ _ _ _ hf hb +@exists_lt_sum_fiber_of_maps_to_of_nsmul_lt_sum α β Mᵒᵈ _ _ _ _ _ _ _ hf hb /-- The pigeonhole principle for finitely many pigeons counted by weight, strict inequality version: if the total weight of a finite set of pigeons is greater than `n • b`, they are sorted into some @@ -147,7 +147,7 @@ is less than `b`. -/ lemma exists_sum_fiber_lt_of_sum_fiber_nonneg_of_sum_lt_nsmul (ht : ∀ y ∉ t, (0:M) ≤ ∑ x in s.filter (λ x, f x = y), w x) (hb : (∑ x in s, w x) < t.card • b) : ∃ y ∈ t, (∑ x in s.filter (λ x, f x = y), w x) < b := -@exists_lt_sum_fiber_of_sum_fiber_nonpos_of_nsmul_lt_sum α β (order_dual M) _ _ _ _ _ _ _ ht hb +@exists_lt_sum_fiber_of_sum_fiber_nonpos_of_nsmul_lt_sum α β Mᵒᵈ _ _ _ _ _ _ _ ht hb /-! #### Non-strict inequality versions @@ -169,7 +169,7 @@ this pigeonhole is less than or equal to `b`. -/ lemma exists_sum_fiber_le_of_maps_to_of_sum_le_nsmul (hf : ∀ a ∈ s, f a ∈ t) (ht : t.nonempty) (hb : (∑ x in s, w x) ≤ t.card • b) : ∃ y ∈ t, (∑ x in s.filter (λ x, f x = y), w x) ≤ b := -@exists_le_sum_fiber_of_maps_to_of_nsmul_le_sum α β (order_dual M) _ _ _ _ _ _ _ hf ht hb +@exists_le_sum_fiber_of_maps_to_of_nsmul_le_sum α β Mᵒᵈ _ _ _ _ _ _ _ hf ht hb /-- The pigeonhole principle for finitely many pigeons counted by weight, non-strict inequality version: if the total weight of a finite set of pigeons is greater than or equal to `n • b`, they @@ -194,7 +194,7 @@ lemma exists_sum_fiber_le_of_sum_fiber_nonneg_of_sum_le_nsmul (hf : ∀ y ∉ t, (0:M) ≤ ∑ x in s.filter (λ x, f x = y), w x) (ht : t.nonempty) (hb : (∑ x in s, w x) ≤ t.card • b) : ∃ y ∈ t, (∑ x in s.filter (λ x, f x = y), w x) ≤ b := -@exists_le_sum_fiber_of_sum_fiber_nonpos_of_nsmul_le_sum α β (order_dual M) _ _ _ _ _ _ _ hf ht hb +@exists_le_sum_fiber_of_sum_fiber_nonpos_of_nsmul_le_sum α β Mᵒᵈ _ _ _ _ _ _ _ hf ht hb end @@ -356,7 +356,7 @@ version: there is a pigeonhole with the total weight of pigeons in it less than the total number of pigeonholes times `b` is greater than the total weight of all pigeons. -/ lemma exists_sum_fiber_lt_of_sum_lt_nsmul (hb : (∑ x, w x) < card β • b) : ∃ y, (∑ x in univ.filter (λ x, f x = y), w x) < b := -@exists_lt_sum_fiber_of_nsmul_lt_sum α β (order_dual M) _ _ _ _ _ _ _ hb +@exists_lt_sum_fiber_of_nsmul_lt_sum α β Mᵒᵈ _ _ _ _ _ _ _ hb /-- The pigeonhole principle for finitely many pigeons of different weights, non-strict inequality version: there is a pigeonhole with the total weight of pigeons in it less than or equal to `b` @@ -364,7 +364,7 @@ provided that the total number of pigeonholes times `b` is greater than or equal of all pigeons. -/ lemma exists_sum_fiber_le_of_sum_le_nsmul [nonempty β] (hb : (∑ x, w x) ≤ card β • b) : ∃ y, (∑ x in univ.filter (λ x, f x = y), w x) ≤ b := -@exists_le_sum_fiber_of_nsmul_le_sum α β (order_dual M) _ _ _ _ _ _ _ _ hb +@exists_le_sum_fiber_of_nsmul_le_sum α β Mᵒᵈ _ _ _ _ _ _ _ _ hb end diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index ec6e542f08e0b..a5f7a3788fc1f 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -1192,7 +1192,7 @@ begin end /-- By sending `x` to `last n - x`, `fin n` is order-equivalent to its `order_dual`. -/ -def _root_.order_iso.fin_equiv : ∀ {n}, order_dual (fin n) ≃o fin n +def _root_.order_iso.fin_equiv : ∀ {n}, (fin n)ᵒᵈ ≃o fin n | 0 := ⟨⟨elim0, elim0, elim0, elim0⟩, elim0⟩ | (n+1) := order_iso.symm $ { to_fun := λ x, last n - x, diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index c333485f05f3a..97a15f219fdf6 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -167,7 +167,7 @@ forall_fin_succ.trans $ and_congr iff.rfl $ forall_congr $ λ j, by simp [tail] lemma cons_le [Π i, preorder (α i)] {x : α 0} {q : Π i, α i} {p : Π i : fin n, α i.succ} : cons x p ≤ q ↔ x ≤ q 0 ∧ p ≤ tail q := -@le_cons _ (λ i, order_dual (α i)) _ x q p +@le_cons _ (λ i, (α i)ᵒᵈ) _ x q p lemma cons_le_cons [Π i, preorder (α i)] {x₀ y₀ : α 0} {x y : Π i : fin n, α (i.succ)} : cons x₀ x ≤ cons y₀ y ↔ x₀ ≤ y₀ ∧ x ≤ y := diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 5880a4729d6f1..a594502965903 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -257,7 +257,7 @@ lemma inf_def : s.inf f = (s.1.map f).inf := rfl fold_empty @[simp] lemma inf_cons {b : β} (h : b ∉ s) : (cons b s h).inf f = f b ⊓ s.inf f := -@sup_cons (order_dual α) _ _ _ _ _ _ h +@sup_cons αᵒᵈ _ _ _ _ _ _ h @[simp] lemma inf_insert [decidable_eq β] {b : β} : (insert b s : finset β).inf f = f b ⊓ s.inf f := fold_insert_idem @@ -274,25 +274,25 @@ fold_map inf_singleton lemma inf_union [decidable_eq β] : (s₁ ∪ s₂).inf f = s₁.inf f ⊓ s₂.inf f := -@sup_union (order_dual α) _ _ _ _ _ _ _ +@sup_union αᵒᵈ _ _ _ _ _ _ _ lemma inf_inf : s.inf (f ⊓ g) = s.inf f ⊓ s.inf g := -@sup_sup (order_dual α) _ _ _ _ _ _ +@sup_sup αᵒᵈ _ _ _ _ _ _ theorem inf_congr {f g : β → α} (hs : s₁ = s₂) (hfg : ∀a∈s₂, f a = g a) : s₁.inf f = s₂.inf g := by subst hs; exact finset.fold_congr hfg @[simp] lemma inf_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : (s.bUnion t).inf f = s.inf (λ x, (t x).inf f) := -@sup_bUnion (order_dual α) _ _ _ _ _ _ _ _ +@sup_bUnion αᵒᵈ _ _ _ _ _ _ _ _ lemma inf_const {s : finset β} (h : s.nonempty) (c : α) : s.inf (λ _, c) = c := -@sup_const (order_dual α) _ _ _ _ h _ +@sup_const αᵒᵈ _ _ _ _ h _ -@[simp] lemma inf_top (s : finset β) : s.inf (λ _, ⊤) = (⊤ : α) := @sup_bot (order_dual α) _ _ _ _ +@[simp] lemma inf_top (s : finset β) : s.inf (λ _, ⊤) = (⊤ : α) := @sup_bot αᵒᵈ _ _ _ _ lemma le_inf_iff {a : α} : a ≤ s.inf f ↔ ∀ b ∈ s, a ≤ f b := -@sup_le_iff (order_dual α) _ _ _ _ _ _ +@sup_le_iff αᵒᵈ _ _ _ _ _ _ lemma inf_le {b : β} (hb : b ∈ s) : s.inf f ≤ f b := le_inf_iff.1 le_rfl _ hb @@ -307,22 +307,22 @@ lemma inf_mono (h : s₁ ⊆ s₂) : s₂.inf f ≤ s₁.inf f := le_inf $ assume b hb, inf_le (h hb) lemma inf_attach (s : finset β) (f : β → α) : s.attach.inf (λ x, f x) = s.inf f := -@sup_attach (order_dual α) _ _ _ _ _ +@sup_attach αᵒᵈ _ _ _ _ _ lemma inf_comm (s : finset β) (t : finset γ) (f : β → γ → α) : s.inf (λ b, t.inf (f b)) = t.inf (λ c, s.inf (λ b, f b c)) := -@sup_comm (order_dual α) _ _ _ _ _ _ _ +@sup_comm αᵒᵈ _ _ _ _ _ _ _ lemma inf_product_left (s : finset β) (t : finset γ) (f : β × γ → α) : (s.product t).inf f = s.inf (λ i, t.inf $ λ i', f ⟨i, i'⟩) := -@sup_product_left (order_dual α) _ _ _ _ _ _ _ +@sup_product_left αᵒᵈ _ _ _ _ _ _ _ lemma inf_product_right (s : finset β) (t : finset γ) (f : β × γ → α) : (s.product t).inf f = t.inf (λ i', s.inf $ λ i, f ⟨i, i'⟩) := -@sup_product_right (order_dual α) _ _ _ _ _ _ _ +@sup_product_right αᵒᵈ _ _ _ _ _ _ _ @[simp] lemma inf_erase_top [decidable_eq α] (s : finset α) : (s.erase ⊤).inf id = s.inf id := -@sup_erase_bot (order_dual α) _ _ _ _ +@sup_erase_bot αᵒᵈ _ _ _ _ lemma sup_sdiff_left {α β : Type*} [boolean_algebra α] (s : finset β) (f : β → α) (a : α) : s.sup (λ b, a \ f b) = a \ s.inf f := @@ -353,18 +353,18 @@ end lemma comp_inf_eq_inf_comp [semilattice_inf γ] [order_top γ] {s : finset β} {f : β → α} (g : α → γ) (g_inf : ∀ x y, g (x ⊓ y) = g x ⊓ g y) (top : g ⊤ = ⊤) : g (s.inf f) = s.inf (g ∘ f) := -@comp_sup_eq_sup_comp (order_dual α) _ (order_dual γ) _ _ _ _ _ _ _ g_inf top +@comp_sup_eq_sup_comp αᵒᵈ _ γᵒᵈ _ _ _ _ _ _ _ g_inf top /-- Computing `inf` in a subtype (closed under `inf`) is the same as computing it in `α`. -/ lemma inf_coe {P : α → Prop} {Ptop : P ⊤} {Pinf : ∀{{x y}}, P x → P y → P (x ⊓ y)} (t : finset β) (f : β → {x : α // P x}) : (@inf _ _ (subtype.semilattice_inf Pinf) (subtype.order_top Ptop) t f : α) = t.inf (λ x, f x) := -@sup_coe (order_dual α) _ _ _ _ Ptop Pinf t f +@sup_coe αᵒᵈ _ _ _ _ Ptop Pinf t f lemma inf_induction {p : α → Prop} (ht : p ⊤) (hp : ∀ a₁, p a₁ → ∀ a₂, p a₂ → p (a₁ ⊓ a₂)) (hs : ∀ b ∈ s, p (f b)) : p (s.inf f) := -@sup_induction (order_dual α) _ _ _ _ _ _ ht hp hs +@sup_induction αᵒᵈ _ _ _ _ _ _ ht hp hs lemma inf_mem (s : set α) (w₁ : ⊤ ∈ s) (w₂ : ∀ x y ∈ s, x ⊓ y ∈ s) @@ -375,7 +375,7 @@ lemma inf_mem @[simp] lemma inf_eq_top_iff (f : β → α) (S : finset β) : S.inf f = ⊤ ↔ ∀ s ∈ S, f s = ⊤ := -@finset.sup_eq_bot_iff (order_dual α) _ _ _ _ _ +@finset.sup_eq_bot_iff αᵒᵈ _ _ _ _ _ end inf @@ -410,11 +410,11 @@ variables [order_top α] lemma inf_sup_distrib_left (s : finset ι) (f : ι → α) (a : α) : a ⊔ s.inf f = s.inf (λ i, a ⊔ f i) := -@sup_inf_distrib_left (order_dual α) _ _ _ _ _ _ +@sup_inf_distrib_left αᵒᵈ _ _ _ _ _ _ lemma inf_sup_distrib_right (s : finset ι) (f : ι → α) (a : α) : s.inf f ⊔ a = s.inf (λ i, f i ⊔ a) := -@sup_inf_distrib_right (order_dual α) _ _ _ _ _ _ +@sup_inf_distrib_right αᵒᵈ _ _ _ _ _ _ end order_top end distrib_lattice @@ -455,22 +455,21 @@ lemma comp_inf_eq_inf_comp_of_is_total [semilattice_inf β] [order_top β] (g : comp_inf_eq_inf_comp g mono_g.map_inf top @[simp] protected lemma inf_le_iff (ha : a < ⊤) : s.inf f ≤ a ↔ ∃ b ∈ s, f b ≤ a := -@finset.le_sup_iff (order_dual α) _ _ _ _ _ _ ha +@finset.le_sup_iff αᵒᵈ _ _ _ _ _ _ ha -@[simp] protected lemma inf_lt_iff : s.inf f < a ↔ (∃ b ∈ s, f b < a) := -@finset.lt_sup_iff (order_dual α) _ _ _ _ _ _ +@[simp] protected lemma inf_lt_iff : s.inf f < a ↔ ∃ b ∈ s, f b < a := +@finset.lt_sup_iff αᵒᵈ _ _ _ _ _ _ @[simp] protected lemma lt_inf_iff (ha : a < ⊤) : a < s.inf f ↔ ∀ b ∈ s, a < f b := -@finset.sup_lt_iff (order_dual α) _ _ _ _ _ _ ha +@finset.sup_lt_iff αᵒᵈ _ _ _ _ _ _ ha end order_top end linear_order -lemma inf_eq_infi [complete_lattice β] (s : finset α) (f : α → β) : s.inf f = (⨅a∈s, f a) := -@sup_eq_supr _ (order_dual β) _ _ _ +lemma inf_eq_infi [complete_lattice β] (s : finset α) (f : α → β) : s.inf f = ⨅ a ∈ s, f a := +@sup_eq_supr _ βᵒᵈ _ _ _ -lemma inf_id_eq_Inf [complete_lattice α] (s : finset α) : s.inf id = Inf s := -@sup_id_eq_Sup (order_dual α) _ _ +lemma inf_id_eq_Inf [complete_lattice α] (s : finset α) : s.inf id = Inf s := @sup_id_eq_Sup αᵒᵈ _ _ lemma inf_id_set_eq_sInter (s : finset (set α)) : s.inf id = ⋂₀(↑s) := inf_id_eq_Inf _ @@ -479,7 +478,7 @@ inf_id_eq_Inf _ inf_eq_infi _ _ lemma inf_eq_Inf_image [complete_lattice β] (s : finset α) (f : α → β) : s.inf f = Inf (f '' s) := -@sup_eq_Sup_image _ (order_dual β) _ _ _ +@sup_eq_Sup_image _ βᵒᵈ _ _ _ section sup' variables [semilattice_sup α] @@ -579,52 +578,48 @@ variables [semilattice_inf α] lemma inf_of_mem {s : finset β} (f : β → α) {b : β} (h : b ∈ s) : ∃ (a : α), s.inf (coe ∘ f : β → with_top α) = ↑a := -@sup_of_mem (order_dual α) _ _ _ f _ h +@sup_of_mem αᵒᵈ _ _ _ f _ h /-- Given nonempty finset `s` then `s.inf' H f` is the infimum of its image under `f` in (possibly unbounded) meet-semilattice `α`, where `H` is a proof of nonemptiness. If `α` has a top element you may instead use `finset.inf` which does not require `s` nonempty. -/ def inf' (s : finset β) (H : s.nonempty) (f : β → α) : α := -@sup' (order_dual α) _ _ s H f +@sup' αᵒᵈ _ _ s H f -variables {s : finset β} (H : s.nonempty) (f : β → α) +variables {s : finset β} (H : s.nonempty) (f : β → α) {a : α} {b : β} @[simp] lemma coe_inf' : ((s.inf' H f : α) : with_top α) = s.inf (coe ∘ f) := -@coe_sup' (order_dual α) _ _ _ H f +@coe_sup' αᵒᵈ _ _ _ H f @[simp] lemma inf'_cons {b : β} {hb : b ∉ s} {h : (cons b s hb).nonempty} : (cons b s hb).inf' h f = f b ⊓ s.inf' H f := -@sup'_cons (order_dual α) _ _ _ H f _ _ _ +@sup'_cons αᵒᵈ _ _ _ H f _ _ _ @[simp] lemma inf'_insert [decidable_eq β] {b : β} {h : (insert b s).nonempty} : (insert b s).inf' h f = f b ⊓ s.inf' H f := -@sup'_insert (order_dual α) _ _ _ H f _ _ _ +@sup'_insert αᵒᵈ _ _ _ H f _ _ _ @[simp] lemma inf'_singleton {b : β} {h : ({b} : finset β).nonempty} : ({b} : finset β).inf' h f = f b := rfl -lemma le_inf' {a : α} (hs : ∀ b ∈ s, a ≤ f b) : a ≤ s.inf' H f := -@sup'_le (order_dual α) _ _ _ H f _ hs - -lemma inf'_le {b : β} (h : b ∈ s) : s.inf' ⟨b, h⟩ f ≤ f b := -@le_sup' (order_dual α) _ _ _ f _ h +lemma le_inf' (hs : ∀ b ∈ s, a ≤ f b) : a ≤ s.inf' H f := @sup'_le αᵒᵈ _ _ _ H f _ hs +lemma inf'_le (h : b ∈ s) : s.inf' ⟨b, h⟩ f ≤ f b := @le_sup' αᵒᵈ _ _ _ f _ h -@[simp] lemma inf'_const (a : α) : s.inf' H (λ b, a) = a := -@sup'_const (order_dual α) _ _ _ _ _ +@[simp] lemma inf'_const (a : α) : s.inf' H (λ b, a) = a := @sup'_const αᵒᵈ _ _ _ _ _ lemma inf'_bUnion [decidable_eq β] {s : finset γ} (Hs : s.nonempty) {t : γ → finset β} (Ht : ∀ b, (t b).nonempty) : (s.bUnion t).inf' (Hs.bUnion (λ b _, Ht b)) f = s.inf' Hs (λ b, (t b).inf' (Ht b) f) := -@sup'_bUnion (order_dual α) _ _ _ _ _ _ Hs _ Ht +@sup'_bUnion αᵒᵈ _ _ _ _ _ _ Hs _ Ht lemma comp_inf'_eq_inf'_comp [semilattice_inf γ] {s : finset β} (H : s.nonempty) {f : β → α} (g : α → γ) (g_inf : ∀ x y, g (x ⊓ y) = g x ⊓ g y) : g (s.inf' H f) = s.inf' H (g ∘ f) := -@comp_sup'_eq_sup'_comp (order_dual α) _ (order_dual γ) _ _ _ H f g g_inf +@comp_sup'_eq_sup'_comp αᵒᵈ _ γᵒᵈ _ _ _ H f g g_inf lemma inf'_induction {p : α → Prop} (hp : ∀ a₁, p a₁ → ∀ a₂, p a₂ → p (a₁ ⊓ a₂)) (hs : ∀ b ∈ s, p (f b)) : p (s.inf' H f) := -@sup'_induction (order_dual α) _ _ _ H f _ hp hs +@sup'_induction αᵒᵈ _ _ _ H f _ hp hs lemma inf'_mem (s : set α) (w : ∀ x y ∈ s, x ⊓ y ∈ s) {ι : Type*} (t : finset ι) (H : t.nonempty) (p : ι → α) (h : ∀ i ∈ t, p i ∈ s) : @@ -632,7 +627,7 @@ lemma inf'_mem (s : set α) (w : ∀ x y ∈ s, x ⊓ y ∈ s) @[congr] lemma inf'_congr {t : finset β} {f g : β → α} (h₁ : s = t) (h₂ : ∀ x ∈ s, f x = g x) : s.inf' H f = t.inf' (h₁ ▸ H) g := -@sup'_congr (order_dual α) _ _ _ H _ _ _ h₁ h₂ +@sup'_congr αᵒᵈ _ _ _ H _ _ _ h₁ h₂ end inf' @@ -656,15 +651,15 @@ section inf variables [semilattice_inf α] [order_top α] lemma inf'_eq_inf {s : finset β} (H : s.nonempty) (f : β → α) : s.inf' H f = s.inf f := -@sup'_eq_sup (order_dual α) _ _ _ _ H f +@sup'_eq_sup αᵒᵈ _ _ _ _ H f lemma inf_closed_of_inf_closed {s : set α} (t : finset α) (htne : t.nonempty) (h_subset : ↑t ⊆ s) (h : ∀ a b ∈ s, a ⊓ b ∈ s) : t.inf id ∈ s := -@sup_closed_of_sup_closed (order_dual α) _ _ _ t htne h_subset h +@sup_closed_of_sup_closed αᵒᵈ _ _ _ t htne h_subset h lemma coe_inf_of_nonempty {s : finset β} (h : s.nonempty) (f : β → α): (↑(s.inf f) : with_top α) = s.inf (λ i, f i) := -@coe_sup_of_nonempty (order_dual α) _ _ _ _ h f +@coe_sup_of_nonempty αᵒᵈ _ _ _ _ h f end inf @@ -684,7 +679,7 @@ variables {C : β → Type*} [Π (b : β), semilattice_inf (C b)] [Π (b : β), @[simp] protected lemma inf_apply (s : finset α) (f : α → Π (b : β), C b) (b : β) : s.inf f b = s.inf (λ a, f a b) := -@finset.sup_apply _ _ (λ b, order_dual (C b)) _ _ s f b +@finset.sup_apply _ _ (λ b, (C b)ᵒᵈ) _ _ s f b end inf @@ -704,7 +699,7 @@ variables {C : β → Type*} [Π (b : β), semilattice_inf (C b)] @[simp] protected lemma inf'_apply {s : finset α} (H : s.nonempty) (f : α → Π (b : β), C b) (b : β) : s.inf' H f b = s.inf' H (λ a, f a b) := -@finset.sup'_apply _ _ (λ b, order_dual (C b)) _ _ H f b +@finset.sup'_apply _ _ (λ b, (C b)ᵒᵈ) _ _ H f b end inf' @@ -729,14 +724,9 @@ begin exact ball_congr (λ b hb, with_bot.coe_lt_coe), end -@[simp] lemma inf'_le_iff : s.inf' H f ≤ a ↔ ∃ i ∈ s, f i ≤ a := -@le_sup'_iff (order_dual α) _ _ _ H f _ - -@[simp] lemma inf'_lt_iff : s.inf' H f < a ↔ ∃ i ∈ s, f i < a := -@lt_sup'_iff (order_dual α) _ _ _ H f _ - -@[simp] lemma lt_inf'_iff : a < s.inf' H f ↔ ∀ i ∈ s, a < f i := -@sup'_lt_iff (order_dual α) _ _ _ H f _ +@[simp] lemma inf'_le_iff : s.inf' H f ≤ a ↔ ∃ i ∈ s, f i ≤ a := @le_sup'_iff αᵒᵈ _ _ _ H f _ +@[simp] lemma inf'_lt_iff : s.inf' H f < a ↔ ∃ i ∈ s, f i < a := @lt_sup'_iff αᵒᵈ _ _ _ H f _ +@[simp] lemma lt_inf'_iff : a < s.inf' H f ↔ ∀ i ∈ s, a < f i := @sup'_lt_iff αᵒᵈ _ _ _ H f _ lemma exists_mem_eq_sup' (f : ι → α) : ∃ i, i ∈ s ∧ s.sup' H f = f i := begin @@ -750,7 +740,7 @@ begin end lemma exists_mem_eq_inf' (f : ι → α) : ∃ i, i ∈ s ∧ s.inf' H f = f i := -@exists_mem_eq_sup' (order_dual α) _ _ _ H f +@exists_mem_eq_sup' αᵒᵈ _ _ _ H f lemma exists_mem_eq_sup [order_bot α] (s : finset ι) (h : s.nonempty) (f : ι → α) : ∃ i, i ∈ s ∧ s.sup f = f i := @@ -758,7 +748,7 @@ sup'_eq_sup h f ▸ exists_mem_eq_sup' h f lemma exists_mem_eq_inf [order_top α] (s : finset ι) (h : s.nonempty) (f : ι → α) : ∃ i, i ∈ s ∧ s.inf f = f i := -@exists_mem_eq_sup (order_dual α) _ _ _ _ h f +@exists_mem_eq_sup αᵒᵈ _ _ _ _ h f end linear_order @@ -839,8 +829,7 @@ theorem min_eq_none {s : finset α} : s.min = none ↔ s = ∅ := (λ H, let ⟨a, ha⟩ := min_of_nonempty H in by rw h at ha; cases ha), λ h, h.symm ▸ min_empty⟩ -theorem mem_of_min {s : finset α} : ∀ {a : α}, a ∈ s.min → a ∈ s := -@mem_of_max (order_dual α) _ s +theorem mem_of_min {s : finset α} : ∀ {a : α}, a ∈ s.min → a ∈ s := @mem_of_max αᵒᵈ _ s theorem min_le_of_mem {s : finset α} {a b : α} (h₁ : b ∈ s) (h₂ : a ∈ s.min) : a ≤ b := by rcases @inf_le (with_top α) _ _ _ _ _ _ h₁ _ rfl with ⟨b', hb, ab⟩; @@ -862,7 +851,7 @@ def max' (s : finset α) (H : s.nonempty) : α := let ⟨k, hk⟩ := H in let ⟨b, hb⟩ := max_of_mem hk in by simp at hb; simp [hb] -variables (s : finset α) (H : s.nonempty) +variables (s : finset α) (H : s.nonempty) {x : α} theorem min'_mem : s.min' H ∈ s := mem_of_min $ by simp [min'] @@ -894,14 +883,12 @@ is_lub_le_iff (is_greatest_max' s H).is_lub @[simp] lemma max'_lt_iff {x} : s.max' H < x ↔ ∀ y ∈ s, y < x := ⟨λ Hlt y hy, (s.le_max' y hy).trans_lt Hlt, λ H, H _ $ s.max'_mem _⟩ -@[simp] lemma lt_min'_iff {x} : x < s.min' H ↔ ∀ y ∈ s, x < y := -@max'_lt_iff (order_dual α) _ _ H _ +@[simp] lemma lt_min'_iff : x < s.min' H ↔ ∀ y ∈ s, x < y := @max'_lt_iff αᵒᵈ _ _ H _ lemma max'_eq_sup' : s.max' H = s.sup' H id := eq_of_forall_ge_iff $ λ a, (max'_le_iff _ _).trans (sup'_le_iff _ _).symm -lemma min'_eq_inf' : s.min' H = s.inf' H id := -@max'_eq_sup' (order_dual α) _ s H +lemma min'_eq_inf' : s.min' H = s.inf' H id := @max'_eq_sup' αᵒᵈ _ s H /-- `{a}.max' _` is `a`. -/ @[simp] lemma max'_singleton (a : α) : @@ -970,7 +957,7 @@ lt_of_le_of_ne (le_max' _ _ (mem_of_mem_erase ha)) $ ne_of_mem_of_not_mem ha $ n lemma min'_lt_of_mem_erase_min' [decidable_eq α] {a : α} (ha : a ∈ s.erase (s.min' H)) : s.min' H < a := -@lt_max'_of_mem_erase_max' (order_dual α) _ s H _ a ha +@lt_max'_of_mem_erase_max' αᵒᵈ _ s H _ a ha @[simp] lemma max'_image [linear_order β] {f : α → β} (hf : monotone f) (s : finset α) (h : (s.image f).nonempty) : @@ -1019,7 +1006,7 @@ end @[elab_as_eliminator] lemma induction_on_min [decidable_eq α] {p : finset α → Prop} (s : finset α) (h0 : p ∅) (step : ∀ a s, (∀ x ∈ s, a < x) → p s → p (insert a s)) : p s := -@induction_on_max (order_dual α) _ _ _ s h0 step +@induction_on_max αᵒᵈ _ _ _ s h0 step end max_min @@ -1061,7 +1048,7 @@ ordered type : a predicate is true on all `s : finset α` provided that: lemma induction_on_min_value [decidable_eq ι] (f : ι → α) {p : finset ι → Prop} (s : finset ι) (h0 : p ∅) (step : ∀ a s, a ∉ s → (∀ x ∈ s, f a ≤ f x) → p s → p (insert a s)) : p s := -@induction_on_max_value (order_dual α) ι _ _ _ _ s h0 step +@induction_on_max_value αᵒᵈ ι _ _ _ _ s h0 step end max_min_induction_value @@ -1078,7 +1065,7 @@ end lemma exists_min_image (s : finset β) (f : β → α) (h : s.nonempty) : ∃ x ∈ s, ∀ x' ∈ s, f x ≤ f x' := -@exists_max_image (order_dual α) β _ s f h +@exists_max_image αᵒᵈ β _ s f h end exists_max_min end finset @@ -1175,16 +1162,15 @@ by rw [← supr_eq_supr_finset, ← equiv.plift.surjective.supr_comp]; refl /-- Infimum of `s i`, `i : ι`, is equal to the infimum over `t : finset ι` of infima `⨅ i ∈ t, s i`. This version assumes `ι` is a `Type*`. See `infi_eq_infi_finset'` for a version that works for `ι : Sort*`. -/ -lemma infi_eq_infi_finset (s : ι → α) : - (⨅i, s i) = (⨅t:finset ι, ⨅i∈t, s i) := -@supr_eq_supr_finset (order_dual α) _ _ _ +lemma infi_eq_infi_finset (s : ι → α) : (⨅ i, s i) = ⨅ (t : finset ι) (i ∈ t), s i := +@supr_eq_supr_finset αᵒᵈ _ _ _ /-- Infimum of `s i`, `i : ι`, is equal to the infimum over `t : finset ι` of infima `⨅ i ∈ t, s i`. This version works for `ι : Sort*`. See `infi_eq_infi_finset` for a version that assumes `ι : Type*` but has no `plift`s. -/ lemma infi_eq_infi_finset' (s : ι' → α) : (⨅i, s i) = (⨅t:finset (plift ι'), ⨅i∈t, s (plift.down i)) := -@supr_eq_supr_finset' (order_dual α) _ _ _ +@supr_eq_supr_finset' αᵒᵈ _ _ _ end lattice @@ -1273,9 +1259,8 @@ lemma supr_option_to_finset (o : option α) (f : α → β) : (⨆ x ∈ o.to_finset, f x) = ⨆ x ∈ o, f x := by simp -lemma infi_option_to_finset (o : option α) (f : α → β) : - (⨅ x ∈ o.to_finset, f x) = ⨅ x ∈ o, f x := -@supr_option_to_finset _ (order_dual β) _ _ _ +lemma infi_option_to_finset (o : option α) (f : α → β) : (⨅ x ∈ o.to_finset, f x) = ⨅ x ∈ o, f x := +@supr_option_to_finset _ βᵒᵈ _ _ _ variables [decidable_eq α] @@ -1285,7 +1270,7 @@ by simp [supr_or, supr_sup_eq] theorem infi_union {f : α → β} {s t : finset α} : (⨅ x ∈ s ∪ t, f x) = (⨅ x ∈ s, f x) ⊓ (⨅ x ∈ t, f x) := -@supr_union α (order_dual β) _ _ _ _ _ +@supr_union α βᵒᵈ _ _ _ _ _ lemma supr_insert (a : α) (s : finset α) (t : α → β) : (⨆ x ∈ insert a s, t x) = t a ⊔ (⨆ x ∈ s, t x) := @@ -1293,7 +1278,7 @@ by { rw insert_eq, simp only [supr_union, finset.supr_singleton] } lemma infi_insert (a : α) (s : finset α) (t : α → β) : (⨅ x ∈ insert a s, t x) = t a ⊓ (⨅ x ∈ s, t x) := -@supr_insert α (order_dual β) _ _ _ _ _ +@supr_insert α βᵒᵈ _ _ _ _ _ lemma supr_finset_image {f : γ → α} {g : α → β} {s : finset γ} : (⨆ x ∈ s.image f, g x) = (⨆ y ∈ s, g (f y)) := @@ -1320,7 +1305,7 @@ end lemma infi_insert_update {x : α} {t : finset α} (f : α → β) {s : β} (hx : x ∉ t) : (⨅ (i ∈ insert x t), update f x s i) = (s ⊓ ⨅ (i ∈ t), f i) := -@supr_insert_update α (order_dual β) _ _ _ _ f _ hx +@supr_insert_update α βᵒᵈ _ _ _ _ f _ hx lemma supr_bUnion (s : finset γ) (t : γ → finset α) (f : α → β) : (⨆ y ∈ s.bUnion t, f y) = ⨆ (x ∈ s) (y ∈ t x), f y := @@ -1328,7 +1313,7 @@ by simp [@supr_comm _ α, supr_and] lemma infi_bUnion (s : finset γ) (t : γ → finset α) (f : α → β) : (⨅ y ∈ s.bUnion t, f y) = ⨅ (x ∈ s) (y ∈ t x), f y := -@supr_bUnion _ (order_dual β) _ _ _ _ _ _ +@supr_bUnion _ βᵒᵈ _ _ _ _ _ _ end lattice diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index 462782b038ba2..ebe1b669ddf15 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -316,7 +316,7 @@ begin end lemma card_Ioc_eq_card_Icc_sub_one (a b : α) : (Ioc a b).card = (Icc a b).card - 1 := -@card_Ico_eq_card_Icc_sub_one (order_dual α) _ _ _ _ +@card_Ico_eq_card_Icc_sub_one αᵒᵈ _ _ _ _ lemma card_Ioo_eq_card_Ico_sub_one (a b : α) : (Ioo a b).card = (Ico a b).card - 1 := begin @@ -330,7 +330,7 @@ begin end lemma card_Ioo_eq_card_Ioc_sub_one (a b : α) : (Ioo a b).card = (Ioc a b).card - 1 := -@card_Ioo_eq_card_Ico_sub_one (order_dual α) _ _ _ _ +@card_Ioo_eq_card_Ico_sub_one αᵒᵈ _ _ _ _ lemma card_Ioo_eq_card_Icc_sub_two (a b : α) : (Ioo a b).card = (Icc a b).card - 2 := by { rw [card_Ioo_eq_card_Ico_sub_one, card_Ico_eq_card_Icc_sub_one], refl } diff --git a/src/data/finset/pi_induction.lean b/src/data/finset/pi_induction.lean index 8cf3ac1a2f5b1..23fd8ef8cb3bb 100644 --- a/src/data/finset/pi_induction.lean +++ b/src/data/finset/pi_induction.lean @@ -92,6 +92,6 @@ lemma induction_on_pi_min [Π i, linear_order (α i)] {p : (Π i, finset (α i)) (step : ∀ (g : Π i, finset (α i)) (i : ι) (x : α i), (∀ y ∈ g i, x < y) → p g → p (update g i (insert x (g i)))) : p f := -@induction_on_pi_max ι (λ i, order_dual (α i)) _ _ _ _ _ _ h0 step +@induction_on_pi_max ι (λ i, (α i)ᵒᵈ) _ _ _ _ _ _ h0 step end finset diff --git a/src/data/finset/sigma.lean b/src/data/finset/sigma.lean index 9456e4dfb452d..6b24cf5fc1a91 100644 --- a/src/data/finset/sigma.lean +++ b/src/data/finset/sigma.lean @@ -71,7 +71,7 @@ end lemma inf_sigma [semilattice_inf β] [order_top β] : (s.sigma t).inf f = s.inf (λ i, (t i).inf $ λ b, f ⟨i, b⟩) := -@sup_sigma _ _ (order_dual β) _ _ _ _ _ +@sup_sigma _ _ βᵒᵈ _ _ _ _ _ end sigma diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 6a0ad4418444b..f323070c5ba65 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -218,7 +218,7 @@ lemma sup_univ_eq_supr [complete_lattice β] (f : α → β) : finset.univ.sup f /-- A special case of `finset.inf_eq_infi` that omits the useless `x ∈ univ` binder. -/ lemma inf_univ_eq_infi [complete_lattice β] (f : α → β) : finset.univ.inf f = infi f := -sup_univ_eq_supr (by exact f : α → order_dual β) +sup_univ_eq_supr (by exact f : α → βᵒᵈ) @[simp] lemma fold_inf_univ [semilattice_inf α] [order_bot α] (a : α) : finset.univ.fold (⊓) a (λ x, x) = ⊥ := @@ -227,7 +227,7 @@ eq_bot_iff.2 $ ((finset.fold_op_rel_iff_and $ @_root_.le_inf_iff α _).1 le_rfl) @[simp] lemma fold_sup_univ [semilattice_sup α] [order_top α] (a : α) : finset.univ.fold (⊔) a (λ x, x) = ⊤ := -@fold_inf_univ (order_dual α) ‹fintype α› _ _ _ +@fold_inf_univ αᵒᵈ ‹fintype α› _ _ _ end finset @@ -891,10 +891,10 @@ fintype.of_equiv _ equiv.plift.symm fintype.card (plift α) = fintype.card α := fintype.of_equiv_card _ -instance (α : Type*) [fintype α] : fintype (order_dual α) := ‹fintype α› +instance (α : Type*) [fintype α] : fintype αᵒᵈ := ‹fintype α› -@[simp] lemma fintype.card_order_dual (α : Type*) [fintype α] : - fintype.card (order_dual α) = fintype.card α := rfl +@[simp] lemma fintype.card_order_dual (α : Type*) [fintype α] : fintype.card αᵒᵈ = fintype.card α := +rfl instance (α : Type*) [fintype α] : fintype (lex α) := ‹fintype α› @@ -1788,7 +1788,7 @@ end lemma finset.exists_maximal {α : Type*} [preorder α] (s : finset α) (h : s.nonempty) : ∃ m ∈ s, ∀ x ∈ s, ¬ (m < x) := -@finset.exists_minimal (order_dual α) _ s h +@finset.exists_minimal αᵒᵈ _ s h namespace infinite diff --git a/src/data/list/big_operators.lean b/src/data/list/big_operators.lean index f1b6b382975a9..bb70195dae307 100644 --- a/src/data/list/big_operators.lean +++ b/src/data/list/big_operators.lean @@ -235,7 +235,7 @@ lemma pow_card_le_prod [preorder M] [covariant_class M M (function.swap (*)) (≤)] [covariant_class M M (*) (≤)] (l : list M) (n : M) (h : ∀ (x ∈ l), n ≤ x) : n ^ l.length ≤ l.prod := -@prod_le_pow_card (order_dual M) _ _ _ _ l n h +@prod_le_pow_card Mᵒᵈ _ _ _ _ l n h @[to_additive exists_lt_of_sum_lt] lemma exists_lt_of_prod_lt' [linear_order M] [covariant_class M M (function.swap (*)) (≤)] [covariant_class M M (*) (≤)] {l : list ι} diff --git a/src/data/list/min_max.lean b/src/data/list/min_max.lean index f38af432cc5ec..c79f23e5b76c8 100644 --- a/src/data/list/min_max.lean +++ b/src/data/list/min_max.lean @@ -34,8 +34,7 @@ l.foldl (argmax₂ f) none /-- `argmin f l` returns `some a`, where `a` of `l` that minimises `f a`. If there are `a b` such that `f a = f b`, it returns whichever of `a` or `b` comes first in the list. `argmin f []` = none` -/ -def argmin (f : α → β) (l : list α) := -@argmax _ (order_dual β) _ f l +def argmin (f : α → β) (l : list α) := @argmax _ βᵒᵈ _ f l @[simp] lemma argmax_two_self (f : α → β) (a : α) : argmax₂ f (some a) a = a := if_pos le_rfl @@ -98,19 +97,19 @@ theorem argmax_mem {f : α → β} : Π {l : list α} {m : α}, m ∈ argmax f l | (hd::tl) m := by simpa [argmax, argmax₂] using foldl_argmax₂_mem f tl hd m theorem argmin_mem {f : α → β} : Π {l : list α} {m : α}, m ∈ argmin f l → m ∈ l := -@argmax_mem _ (order_dual β) _ _ +@argmax_mem _ βᵒᵈ _ _ @[simp] theorem argmax_eq_none {f : α → β} {l : list α} : l.argmax f = none ↔ l = [] := by simp [argmax] @[simp] theorem argmin_eq_none {f : α → β} {l : list α} : l.argmin f = none ↔ l = [] := -@argmax_eq_none _ (order_dual β) _ _ _ +@argmax_eq_none _ βᵒᵈ _ _ _ theorem le_argmax_of_mem {f : α → β} {a m : α} {l : list α} : a ∈ l → m ∈ argmax f l → f a ≤ f m := le_of_foldl_argmax₂ theorem argmin_le_of_mem {f : α → β} {a m : α} {l : list α} : a ∈ l → m ∈ argmin f l → f m ≤ f a:= -@le_argmax_of_mem _ (order_dual β) _ _ _ _ _ +@le_argmax_of_mem _ βᵒᵈ _ _ _ _ _ theorem argmax_concat (f : α → β) (a : α) (l : list α) : argmax f (l ++ [a]) = option.cases_on (argmax f l) (some a) (λ c, if f a ≤ f c then some c else some a) := @@ -118,7 +117,7 @@ by rw [argmax, argmax]; simp [argmax₂] theorem argmin_concat (f : α → β) (a : α) (l : list α) : argmin f (l ++ [a]) = option.cases_on (argmin f l) (some a) (λ c, if f c ≤ f a then some c else some a) := -@argmax_concat _ (order_dual β) _ _ _ _ +@argmax_concat _ βᵒᵈ _ _ _ _ theorem argmax_cons (f : α → β) (a : α) (l : list α) : argmax f (a :: l) = option.cases_on (argmax f l) (some a) (λ c, if f c ≤ f a then some a else some c) := @@ -144,7 +143,7 @@ list.reverse_rec_on l rfl $ theorem argmin_cons (f : α → β) (a : α) (l : list α) : argmin f (a :: l) = option.cases_on (argmin f l) (some a) (λ c, if f a ≤ f c then some a else some c) := -@argmax_cons _ (order_dual β) _ _ _ _ +@argmax_cons _ βᵒᵈ _ _ _ _ theorem index_of_argmax [decidable_eq α] {f : α → β} : Π {l : list α} {m : α}, m ∈ argmax f l → ∀ {a}, a ∈ l → f m ≤ f a → l.index_of m ≤ l.index_of a @@ -169,7 +168,7 @@ end theorem index_of_argmin [decidable_eq α] {f : α → β} : Π {l : list α} {m : α}, m ∈ argmin f l → ∀ {a}, a ∈ l → f a ≤ f m → l.index_of m ≤ l.index_of a := -@index_of_argmax _ (order_dual β) _ _ _ +@index_of_argmax _ βᵒᵈ _ _ _ theorem mem_argmax_iff [decidable_eq α] {f : α → β} {m : α} {l : list α} : m ∈ argmax f l ↔ m ∈ l ∧ (∀ a ∈ l, f a ≤ f m) ∧ @@ -191,7 +190,7 @@ theorem argmax_eq_some_iff [decidable_eq α] {f : α → β} {m : α} {l : list theorem mem_argmin_iff [decidable_eq α] {f : α → β} {m : α} {l : list α} : m ∈ argmin f l ↔ m ∈ l ∧ (∀ a ∈ l, f m ≤ f a) ∧ (∀ a ∈ l, f a ≤ f m → l.index_of m ≤ l.index_of a) := -@mem_argmax_iff _ (order_dual β) _ _ _ _ _ +@mem_argmax_iff _ βᵒᵈ _ _ _ _ _ theorem argmin_eq_some_iff [decidable_eq α] {f : α → β} {m : α} {l : list α} : argmin f l = some m ↔ m ∈ l ∧ (∀ a ∈ l, f m ≤ f a) ∧ @@ -236,7 +235,7 @@ option.cases_on (maximum l) (λ _ h, absurd ha ((h rfl).symm ▸ not_mem_nil _)) (@maximum_eq_none _ _ l).1 theorem le_minimum_of_mem' {a : α} {l : list α} (ha : a ∈ l) : minimum l ≤ (a : with_top α) := -@le_maximum_of_mem' (order_dual α) _ _ _ ha +@le_maximum_of_mem' αᵒᵈ _ _ _ ha theorem maximum_concat (a : α) (l : list α) : maximum (l ++ [a]) = max (maximum l) a := begin @@ -250,14 +249,14 @@ begin end theorem minimum_concat (a : α) (l : list α) : minimum (l ++ [a]) = min (minimum l) a := -@maximum_concat (order_dual α) _ _ _ +@maximum_concat αᵒᵈ _ _ _ theorem maximum_cons (a : α) (l : list α) : maximum (a :: l) = max a (maximum l) := list.reverse_rec_on l (by simp [@max_eq_left (with_bot α) _ _ _ bot_le]) (λ tl hd ih, by rw [← cons_append, maximum_concat, ih, maximum_concat, max_assoc]) theorem minimum_cons (a : α) (l : list α) : minimum (a :: l) = min a (minimum l) := -@maximum_cons (order_dual α) _ _ _ +@maximum_cons αᵒᵈ _ _ _ theorem maximum_eq_coe_iff {m : α} {l : list α} : maximum l = m ↔ m ∈ l ∧ (∀ a ∈ l, a ≤ m) := @@ -273,7 +272,7 @@ end theorem minimum_eq_coe_iff {m : α} {l : list α} : minimum l = m ↔ m ∈ l ∧ (∀ a ∈ l, m ≤ a) := -@maximum_eq_coe_iff (order_dual α) _ _ _ +@maximum_eq_coe_iff αᵒᵈ _ _ _ section fold @@ -281,7 +280,7 @@ variables {M : Type*} [canonically_linear_ordered_add_monoid M] /-! Note: since there is no typeclass typeclass dual to `canonically_linear_ordered_add_monoid α` we cannot express these lemmas generally for -`minimum`; instead we are limited to doing so on `order_dual α`. -/ +`minimum`; instead we are limited to doing so on `αᵒᵈ`. -/ lemma maximum_eq_coe_foldr_max_of_ne_nil (l : list M) (h : l ≠ []) : l.maximum = (l.foldr max ⊥ : M) := @@ -294,8 +293,8 @@ begin { simp [IH h] } } end -lemma minimum_eq_coe_foldr_min_of_ne_nil (l : list (order_dual M)) (h : l ≠ []) : - l.minimum = (l.foldr min ⊤ : order_dual M) := +lemma minimum_eq_coe_foldr_min_of_ne_nil (l : list Mᵒᵈ) (h : l ≠ []) : + l.minimum = (l.foldr min ⊤ : Mᵒᵈ) := maximum_eq_coe_foldr_max_of_ne_nil l h lemma maximum_nat_eq_coe_foldr_max_of_ne_nil (l : list ℕ) (h : l ≠ []) : @@ -312,9 +311,7 @@ begin simpa [hy] using IH } end -lemma le_min_of_le_forall (l : list (order_dual M)) (n : (order_dual M)) - (h : ∀ (x ∈ l), n ≤ x) : - n ≤ l.foldr min ⊤ := +lemma le_min_of_le_forall (l : list Mᵒᵈ) (n : Mᵒᵈ) (h : ∀ (x ∈ l), n ≤ x) : n ≤ l.foldr min ⊤ := max_le_of_forall_le l n h lemma max_nat_le_of_forall_le (l : list ℕ) (n : ℕ) (h : ∀ (x ∈ l), x ≤ n) : diff --git a/src/data/nat/lattice.lean b/src/data/nat/lattice.lean index f6341a98e5162..2e0578ba30d21 100644 --- a/src/data/nat/lattice.lean +++ b/src/data/nat/lattice.lean @@ -152,10 +152,10 @@ lemma supr_lt_succ' (u : ℕ → α) (n : ℕ) : (⨆ k < n + 1, u k) = u 0 ⊔ by { rw ← sup_supr_nat_succ, simp } lemma infi_lt_succ (u : ℕ → α) (n : ℕ) : (⨅ k < n + 1, u k) = (⨅ k < n, u k) ⊓ u n := -@supr_lt_succ (order_dual α) _ _ _ +@supr_lt_succ αᵒᵈ _ _ _ lemma infi_lt_succ' (u : ℕ → α) (n : ℕ) : (⨅ k < n + 1, u k) = u 0 ⊓ (⨅ k < n, u (k + 1)) := -@supr_lt_succ' (order_dual α) _ _ _ +@supr_lt_succ' αᵒᵈ _ _ _ end diff --git a/src/data/nat/pairing.lean b/src/data/nat/pairing.lean index 365d0baaa34f4..3b61b8826f839 100644 --- a/src/data/nat/pairing.lean +++ b/src/data/nat/pairing.lean @@ -140,7 +140,7 @@ by rw [← (supr_prod : (⨆ i : ℕ × ℕ, f i.1 i.2) = _), ← nat.surjective lemma infi_unpair {α} [complete_lattice α] (f : ℕ → ℕ → α) : (⨅ n : ℕ, f n.unpair.1 n.unpair.2) = ⨅ i j : ℕ, f i j := -supr_unpair (show ℕ → ℕ → order_dual α, from f) +supr_unpair (show ℕ → ℕ → αᵒᵈ, from f) end complete_lattice diff --git a/src/data/ordmap/ordset.lean b/src/data/ordmap/ordset.lean index a19fea78f2afa..b0aacf4dd1846 100644 --- a/src/data/ordmap/ordset.lean +++ b/src/data/ordmap/ordset.lean @@ -540,7 +540,7 @@ theorem find_max'_all {P : α → Prop} : ∀ (x : α) t, P x → all P t → P /-! ### `insert` -/ theorem dual_insert [preorder α] [is_total α (≤)] [@decidable_rel α (≤)] (x : α) : - ∀ t : ordnode α, dual (ordnode.insert x t) = @ordnode.insert (order_dual α) _ _ x (dual t) + ∀ t : ordnode α, dual (ordnode.insert x t) = @ordnode.insert αᵒᵈ _ _ x (dual t) | nil := rfl | (node _ l y r) := begin rw [ordnode.insert, dual, ordnode.insert, order_dual.cmp_le_flip, ← cmp_le_swap x y], @@ -796,12 +796,12 @@ def bounded : ordnode α → with_bot α → with_top α → Prop | (node _ l x r) o₁ o₂ := bounded l o₁ ↑x ∧ bounded r ↑x o₂ theorem bounded.dual : ∀ {t : ordnode α} {o₁ o₂} (h : bounded t o₁ o₂), - @bounded (order_dual α) _ (dual t) o₂ o₁ + @bounded αᵒᵈ _ (dual t) o₂ o₁ | nil o₁ o₂ h := by cases o₁; cases o₂; try {trivial}; exact h | (node s l x r) _ _ ⟨ol, or⟩ := ⟨or.dual, ol.dual⟩ -theorem bounded.dual_iff {t : ordnode α} {o₁ o₂} : bounded t o₁ o₂ ↔ - @bounded (order_dual α) _ (dual t) o₂ o₁ := +theorem bounded.dual_iff {t : ordnode α} {o₁ o₂} : + bounded t o₁ o₂ ↔ @bounded αᵒᵈ _ (dual t) o₂ o₁ := ⟨bounded.dual, λ h, by have := bounded.dual h; rwa [dual_dual, order_dual.preorder.dual_dual] at this⟩ @@ -927,8 +927,7 @@ theorem valid'.node {s l x r o₁ o₂} valid' o₁ (@node α s l x r) o₂ := ⟨⟨hl.1, hr.1⟩, ⟨hs, hl.2, hr.2⟩, ⟨H, hl.3, hr.3⟩⟩ -theorem valid'.dual : ∀ {t : ordnode α} {o₁ o₂} (h : valid' o₁ t o₂), - @valid' (order_dual α) _ o₂ (dual t) o₁ +theorem valid'.dual : ∀ {t : ordnode α} {o₁ o₂} (h : valid' o₁ t o₂), @valid' αᵒᵈ _ o₂ (dual t) o₁ | nil o₁ o₂ h := valid'_nil h.1.dual | (node s l x r) o₁ o₂ ⟨⟨ol, or⟩, ⟨rfl, sl, sr⟩, ⟨b, bl, br⟩⟩ := let ⟨ol', sl', bl'⟩ := valid'.dual ⟨ol, sl, bl⟩, @@ -937,16 +936,13 @@ theorem valid'.dual : ∀ {t : ordnode α} {o₁ o₂} (h : valid' o₁ t o₂), ⟨by simp [size_dual, add_comm], sr', sl'⟩, ⟨by rw [size_dual, size_dual]; exact b.symm, br', bl'⟩⟩ -theorem valid'.dual_iff {t : ordnode α} {o₁ o₂} : valid' o₁ t o₂ ↔ - @valid' (order_dual α) _ o₂ (dual t) o₁ := +theorem valid'.dual_iff {t : ordnode α} {o₁ o₂} : valid' o₁ t o₂ ↔ @valid' αᵒᵈ _ o₂ (dual t) o₁ := ⟨valid'.dual, λ h, by have := valid'.dual h; rwa [dual_dual, order_dual.preorder.dual_dual] at this⟩ -theorem valid.dual {t : ordnode α} : valid t → - @valid (order_dual α) _ (dual t) := valid'.dual +theorem valid.dual {t : ordnode α} : valid t → @valid αᵒᵈ _ (dual t) := valid'.dual -theorem valid.dual_iff {t : ordnode α} : valid t ↔ - @valid (order_dual α) _ (dual t) := valid'.dual_iff +theorem valid.dual_iff {t : ordnode α} : valid t ↔ @valid αᵒᵈ _ (dual t) := valid'.dual_iff theorem valid'.left {s l x r o₁ o₂} (H : valid' o₁ (@node α s l x r) o₂) : valid' o₁ l x := ⟨H.1.1, H.2.2.1, H.3.2.1⟩ diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 001e3031e1c04..db0e3a5ab3140 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -253,10 +253,10 @@ lemma infi_ne_top [complete_lattice α] (f : ℝ≥0∞ → α) : (⨅ x ≠ ∞ by rw [infi_subtype', cinfi_ne_top] lemma csupr_ne_top [has_Sup α] (f : ℝ≥0∞ → α) : (⨆ x : {x // x ≠ ∞}, f x) = ⨆ x : ℝ≥0, f x := -@cinfi_ne_top (order_dual α) _ _ +@cinfi_ne_top αᵒᵈ _ _ lemma supr_ne_top [complete_lattice α] (f : ℝ≥0∞ → α) : (⨆ x ≠ ∞, f x) = ⨆ x : ℝ≥0, f x := -@infi_ne_top (order_dual α) _ _ +@infi_ne_top αᵒᵈ _ _ lemma infi_ennreal {α : Type*} [complete_lattice α] {f : ℝ≥0∞ → α} : (⨅ n, f n) = (⨅ n : ℝ≥0, f n) ⊓ f ∞ := @@ -266,7 +266,7 @@ le_antisymm lemma supr_ennreal {α : Type*} [complete_lattice α] {f : ℝ≥0∞ → α} : (⨆ n, f n) = (⨆ n : ℝ≥0, f n) ⊔ f ∞ := -@infi_ennreal (order_dual α) _ _ +@infi_ennreal αᵒᵈ _ _ @[simp] lemma add_top : a + ∞ = ∞ := add_top _ @[simp] lemma top_add : ∞ + a = ∞ := top_add _ @@ -1068,7 +1068,7 @@ inv_lt_iff_inv_lt.trans $ by rw [inv_one] /-- The inverse map `λ x, x⁻¹` is an order isomorphism between `ℝ≥0∞` and its `order_dual` -/ @[simps apply] -def _root_.order_iso.inv_ennreal : ℝ≥0∞ ≃o order_dual ℝ≥0∞ := +def _root_.order_iso.inv_ennreal : ℝ≥0∞ ≃o ℝ≥0∞ᵒᵈ := { to_fun := λ x, x⁻¹, inv_fun := λ x, x⁻¹, map_rel_iff' := λ a b, ennreal.inv_le_inv, diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index aec1b4bd4bbcc..c1c818caae5b0 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -415,7 +415,7 @@ by conv_lhs { rw [ereal.neg_le, neg_neg] } @[simp, norm_cast] lemma coe_neg (x : ℝ) : ((- x : ℝ) : ereal) = - (x : ereal) := rfl /-- Negation as an order reversing isomorphism on `ereal`. -/ -def neg_order_iso : ereal ≃o (order_dual ereal) := +def neg_order_iso : ereal ≃o erealᵒᵈ := { to_fun := λ x, order_dual.to_dual (-x), inv_fun := λ x, -x.of_dual, map_rel_iff' := λ x y, neg_le_neg_iff, diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 01f7b87e09b23..de6dfea7e0e91 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -845,13 +845,12 @@ section variables [semilattice_inf α] [nonempty α] {s : set α} /--A finite set is bounded below.-/ -protected lemma finite.bdd_below (hs : finite s) : bdd_below s := -@finite.bdd_above (order_dual α) _ _ _ hs +protected lemma finite.bdd_below (hs : finite s) : bdd_below s := @finite.bdd_above αᵒᵈ _ _ _ hs /--A finite union of sets which are all bounded below is still bounded below.-/ lemma finite.bdd_below_bUnion {I : set β} {S : β → set α} (H : finite I) : - (bdd_below (⋃i∈I, S i)) ↔ (∀i ∈ I, bdd_below (S i)) := -@finite.bdd_above_bUnion (order_dual α) _ _ _ _ _ H + bdd_below (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, bdd_below (S i) := +@finite.bdd_above_bUnion αᵒᵈ _ _ _ _ _ H lemma infinite_of_not_bdd_below : ¬ bdd_below s → s.infinite := begin diff --git a/src/data/set/function.lean b/src/data/set/function.lean index 40a99087faaa1..7497ee422bfcb 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -895,7 +895,7 @@ lemma piecewise_le {δ : α → Type*} [Π i, preorder (δ i)] {s : set α} [Π lemma le_piecewise {δ : α → Type*} [Π i, preorder (δ i)] {s : set α} [Π j, decidable (j ∈ s)] {f₁ f₂ g : Π i, δ i} (h₁ : ∀ i ∈ s, g i ≤ f₁ i) (h₂ : ∀ i ∉ s, g i ≤ f₂ i) : g ≤ s.piecewise f₁ f₂ := -@piecewise_le α (λ i, order_dual (δ i)) _ s _ _ _ _ h₁ h₂ +@piecewise_le α (λ i, (δ i)ᵒᵈ) _ s _ _ _ _ h₁ h₂ lemma piecewise_le_piecewise {δ : α → Type*} [Π i, preorder (δ i)] {s : set α} [Π j, decidable (j ∈ s)] {f₁ f₂ g₁ g₂ : Π i, δ i} (h₁ : ∀ i ∈ s, f₁ i ≤ g₁ i) @@ -1015,7 +1015,7 @@ lemma strict_mono_on.inj_on [linear_order α] [preorder β] {f : α → β} {s : lemma strict_anti_on.inj_on [linear_order α] [preorder β] {f : α → β} {s : set α} (H : strict_anti_on f s) : s.inj_on f := -@strict_mono_on.inj_on α (order_dual β) _ _ f s H +@strict_mono_on.inj_on α βᵒᵈ _ _ f s H lemma strict_mono_on.comp [preorder α] [preorder β] [preorder γ] {g : β → γ} {f : α → β} {s : set α} {t : set β} (hg : strict_mono_on g t) diff --git a/src/data/set/intervals/basic.lean b/src/data/set/intervals/basic.lean index c4a7a2815a8a5..2db4ebb1c677c 100644 --- a/src/data/set/intervals/basic.lean +++ b/src/data/set/intervals/basic.lean @@ -186,8 +186,7 @@ Ioo_eq_empty h.not_lt lemma Ici_subset_Ici : Ici a ⊆ Ici b ↔ b ≤ a := ⟨λ h, h $ left_mem_Ici, λ h x hx, h.trans hx⟩ -lemma Iic_subset_Iic : Iic a ⊆ Iic b ↔ a ≤ b := -@Ici_subset_Ici (order_dual α) _ _ _ +lemma Iic_subset_Iic : Iic a ⊆ Iic b ↔ a ≤ b := @Ici_subset_Ici αᵒᵈ _ _ _ lemma Ici_subset_Ioi : Ici a ⊆ Ioi b ↔ b < a := ⟨λ h, h left_mem_Ici, λ h x hx, h.trans_le hx⟩ @@ -485,7 +484,7 @@ classical.by_cases lemma mem_Iic_Iio_of_subset_of_subset {s : set α} (ho : Iio a ⊆ s) (hc : s ⊆ Iic a) : s ∈ ({Iic a, Iio a} : set (set α)) := -@mem_Ici_Ioi_of_subset_of_subset (order_dual α) _ a s ho hc +@mem_Ici_Ioi_of_subset_of_subset αᵒᵈ _ a s ho hc lemma mem_Icc_Ico_Ioc_Ioo_of_subset_of_subset {s : set α} (ho : Ioo a b ⊆ s) (hc : s ⊆ Icc a b) : s ∈ ({Icc a b, Ico a b, Ioc a b, Ioo a b} : set (set α)) := @@ -628,7 +627,7 @@ lemma Ico_subset_Ico_iff (h₁ : a₁ < b₁) : lemma Ioc_subset_Ioc_iff (h₁ : a₁ < b₁) : Ioc a₁ b₁ ⊆ Ioc a₂ b₂ ↔ b₁ ≤ b₂ ∧ a₂ ≤ a₁ := -by { convert @Ico_subset_Ico_iff (order_dual α) _ b₁ b₂ a₁ a₂ h₁; exact (@dual_Ico α _ _ _).symm } +by { convert @Ico_subset_Ico_iff αᵒᵈ _ b₁ b₂ a₁ a₂ h₁; exact (@dual_Ico α _ _ _).symm } lemma Ioo_subset_Ioo_iff [densely_ordered α] (h₁ : a₁ < b₁) : Ioo a₁ b₁ ⊆ Ioo a₂ b₂ ↔ a₂ ≤ a₁ ∧ b₁ ≤ b₂ := diff --git a/src/data/set/intervals/infinite.lean b/src/data/set/intervals/infinite.lean index 70b92219d666e..cdca43d3557a7 100644 --- a/src/data/set/intervals/infinite.lean +++ b/src/data/set/intervals/infinite.lean @@ -62,8 +62,7 @@ section unbounded_above variables [no_max_order α] -lemma Ioi.infinite {a : α} : infinite (Ioi a) := -by apply @Iio.infinite (order_dual α) +lemma Ioi.infinite {a : α} : infinite (Ioi a) := @Iio.infinite αᵒᵈ _ _ _ lemma Ici.infinite {a : α} : infinite (Ici a) := Ioi.infinite.mono Ioi_subset_Ici_self diff --git a/src/data/set/intervals/pi.lean b/src/data/set/intervals/pi.lean index 4ded6da0e95fd..575d49afb813c 100644 --- a/src/data/set/intervals/pi.lean +++ b/src/data/set/intervals/pi.lean @@ -51,7 +51,7 @@ lemma pi_univ_Ioi_subset : pi univ (λ i, Ioi (x i)) ⊆ Ioi x := λ h, nonempty.elim ‹nonempty ι› $ λ i, (h i).not_lt (hz i trivial)⟩ lemma pi_univ_Iio_subset : pi univ (λ i, Iio (x i)) ⊆ Iio x := -@pi_univ_Ioi_subset ι (λ i, order_dual (α i)) _ x _ +@pi_univ_Ioi_subset ι (λ i, (α i)ᵒᵈ) _ x _ lemma pi_univ_Ioo_subset : pi univ (λ i, Ioo (x i) (y i)) ⊆ Ioo x y := λ x hx, ⟨pi_univ_Ioi_subset _ $ λ i hi, (hx i hi).1, pi_univ_Iio_subset _ $ λ i hi, (hx i hi).2⟩ diff --git a/src/data/set/intervals/with_bot_top.lean b/src/data/set/intervals/with_bot_top.lean index 04f6644b2f3de..6f4c2ee095456 100644 --- a/src/data/set/intervals/with_bot_top.lean +++ b/src/data/set/intervals/with_bot_top.lean @@ -95,11 +95,11 @@ end with_top namespace with_bot @[simp] lemma preimage_coe_bot : (coe : α → with_bot α) ⁻¹' {⊥} = (∅ : set α) := -@with_top.preimage_coe_top (order_dual α) +@with_top.preimage_coe_top αᵒᵈ variables [partial_order α] {a b : α} -lemma range_coe : range (coe : α → with_bot α) = Ioi ⊥ := @with_top.range_coe (order_dual α) _ +lemma range_coe : range (coe : α → with_bot α) = Ioi ⊥ := @with_top.range_coe αᵒᵈ _ @[simp] lemma preimage_coe_Ioi : (coe : α → with_bot α) ⁻¹' (Ioi a) = Ioi a := ext $ λ x, coe_lt_coe @[simp] lemma preimage_coe_Ici : (coe : α → with_bot α) ⁻¹' (Ici a) = Ici a := ext $ λ x, coe_le_coe diff --git a/src/data/sum/order.lean b/src/data/sum/order.lean index f6f26e9e4637e..c6324afa960ac 100644 --- a/src/data/sum/order.lean +++ b/src/data/sum/order.lean @@ -438,8 +438,7 @@ rfl @[simp] lemma sum_assoc_symm_apply_inr_inr : (sum_assoc α β γ).symm (inr (inr c)) = inr c := rfl /-- `order_dual` is distributive over `⊕` up to an order isomorphism. -/ -def sum_dual_distrib (α β : Type*) [has_le α] [has_le β] : - order_dual (α ⊕ β) ≃o order_dual α ⊕ order_dual β := +def sum_dual_distrib (α β : Type*) [has_le α] [has_le β] : (α ⊕ β)ᵒᵈ ≃o αᵒᵈ ⊕ βᵒᵈ := { map_rel_iff' := begin rintro (a | a) (b | b), { change inl (to_dual a) ≤ inl (to_dual b) ↔ to_dual (inl a) ≤ to_dual (inl b), @@ -501,8 +500,7 @@ def sum_lex_assoc (α β γ : Type*) [has_le α] [has_le β] [has_le γ] : (α (sum_lex_assoc α β γ).symm (inr (inr c)) = inr c := rfl /-- `order_dual` is antidistributive over `⊕ₗ` up to an order isomorphism. -/ -def sum_lex_dual_antidistrib (α β : Type*) [has_le α] [has_le β] : - order_dual (α ⊕ₗ β) ≃o order_dual β ⊕ₗ order_dual α := +def sum_lex_dual_antidistrib (α β : Type*) [has_le α] [has_le β] : (α ⊕ₗ β)ᵒᵈ ≃o βᵒᵈ ⊕ₗ αᵒᵈ := { map_rel_iff' := begin rintro (a | a) (b | b), simp, { change to_lex (inr $ to_dual a) ≤ to_lex (inr $ to_dual b) ↔ diff --git a/src/field_theory/galois.lean b/src/field_theory/galois.lean index 22f5c16445623..6447db5cb2dac 100644 --- a/src/field_theory/galois.lean +++ b/src/field_theory/galois.lean @@ -251,7 +251,7 @@ by conv { to_rhs, rw [←fixed_field_fixing_subgroup K, /-- The Galois correspondence from intermediate fields to subgroups -/ def intermediate_field_equiv_subgroup [finite_dimensional F E] [is_galois F E] : - intermediate_field F E ≃o order_dual (subgroup (E ≃ₐ[F] E)) := + intermediate_field F E ≃o (subgroup (E ≃ₐ[F] E))ᵒᵈ := { to_fun := intermediate_field.fixing_subgroup, inv_fun := intermediate_field.fixed_field, left_inv := λ K, fixed_field_fixing_subgroup K, diff --git a/src/linear_algebra/affine_space/ordered.lean b/src/linear_algebra/affine_space/ordered.lean index 4fea2a3879290..833ef73fbcd54 100644 --- a/src/linear_algebra/affine_space/ordered.lean +++ b/src/linear_algebra/affine_space/ordered.lean @@ -96,13 +96,13 @@ lemma left_lt_line_map_iff_lt (h : 0 < r) : a < line_map a b r ↔ a < b := iff.trans (by rw line_map_apply_zero) (line_map_lt_line_map_iff_of_lt h) lemma line_map_lt_left_iff_lt (h : 0 < r) : line_map a b r < a ↔ b < a := -@left_lt_line_map_iff_lt k (order_dual E) _ _ _ _ _ _ _ h +@left_lt_line_map_iff_lt k Eᵒᵈ _ _ _ _ _ _ _ h lemma line_map_lt_right_iff_lt (h : r < 1) : line_map a b r < b ↔ a < b := iff.trans (by rw line_map_apply_one) (line_map_lt_line_map_iff_of_lt h) lemma right_lt_line_map_iff_lt (h : r < 1) : b < line_map a b r ↔ b < a := -@line_map_lt_right_iff_lt k (order_dual E) _ _ _ _ _ _ _ h +@line_map_lt_right_iff_lt k Eᵒᵈ _ _ _ _ _ _ _ h end ordered_ring @@ -143,7 +143,7 @@ iff.trans (by rw line_map_apply_zero) (line_map_le_line_map_iff_of_lt h) left_le_line_map_iff_le $ inv_pos.2 zero_lt_two lemma line_map_le_left_iff_le (h : 0 < r) : line_map a b r ≤ a ↔ b ≤ a := -@left_le_line_map_iff_le k (order_dual E) _ _ _ _ _ _ _ h +@left_le_line_map_iff_le k Eᵒᵈ _ _ _ _ _ _ _ h @[simp] lemma midpoint_le_left : midpoint k a b ≤ a ↔ b ≤ a := line_map_le_left_iff_le $ inv_pos.2 zero_lt_two @@ -155,7 +155,7 @@ iff.trans (by rw line_map_apply_one) (line_map_le_line_map_iff_of_lt h) line_map_le_right_iff_le $ inv_lt_one one_lt_two lemma right_le_line_map_iff_le (h : r < 1) : b ≤ line_map a b r ↔ b ≤ a := -@line_map_le_right_iff_le k (order_dual E) _ _ _ _ _ _ _ h +@line_map_le_right_iff_le k Eᵒᵈ _ _ _ _ _ _ _ h @[simp] lemma right_le_midpoint : b ≤ midpoint k a b ↔ b ≤ a := right_le_line_map_iff_le $ inv_lt_one one_lt_two @@ -213,7 +213,7 @@ end segment `[(a, f a), (b, f b)]` if and only if `slope f a b ≤ slope f a c`. -/ lemma line_map_le_map_iff_slope_le_slope_left (h : 0 < r * (b - a)) : line_map (f a) (f b) r ≤ f c ↔ slope f a b ≤ slope f a c := -@map_le_line_map_iff_slope_le_slope_left k (order_dual E) _ _ _ _ f a b r h +@map_le_line_map_iff_slope_le_slope_left k Eᵒᵈ _ _ _ _ f a b r h /-- Given `c = line_map a b r`, `a < c`, the point `(c, f c)` is strictly below the segment `[(a, f a), (b, f b)]` if and only if `slope f a c < slope f a b`. -/ @@ -226,7 +226,7 @@ lt_iff_lt_of_le_iff_le' (line_map_le_map_iff_slope_le_slope_left h) segment `[(a, f a), (b, f b)]` if and only if `slope f a b < slope f a c`. -/ lemma line_map_lt_map_iff_slope_lt_slope_left (h : 0 < r * (b - a)) : line_map (f a) (f b) r < f c ↔ slope f a b < slope f a c := -@map_lt_line_map_iff_slope_lt_slope_left k (order_dual E) _ _ _ _ f a b r h +@map_lt_line_map_iff_slope_lt_slope_left k Eᵒᵈ _ _ _ _ f a b r h /-- Given `c = line_map a b r`, `c < b`, the point `(c, f c)` is non-strictly below the segment `[(a, f a), (b, f b)]` if and only if `slope f a b ≤ slope f c b`. -/ @@ -247,7 +247,7 @@ end segment `[(a, f a), (b, f b)]` if and only if `slope f c b ≤ slope f a b`. -/ lemma line_map_le_map_iff_slope_le_slope_right (h : 0 < (1 - r) * (b - a)) : line_map (f a) (f b) r ≤ f c ↔ slope f c b ≤ slope f a b := -@map_le_line_map_iff_slope_le_slope_right k (order_dual E) _ _ _ _ f a b r h +@map_le_line_map_iff_slope_le_slope_right k Eᵒᵈ _ _ _ _ f a b r h /-- Given `c = line_map a b r`, `c < b`, the point `(c, f c)` is strictly below the segment `[(a, f a), (b, f b)]` if and only if `slope f a b < slope f c b`. -/ @@ -260,7 +260,7 @@ lt_iff_lt_of_le_iff_le' (line_map_le_map_iff_slope_le_slope_right h) segment `[(a, f a), (b, f b)]` if and only if `slope f c b < slope f a b`. -/ lemma line_map_lt_map_iff_slope_lt_slope_right (h : 0 < (1 - r) * (b - a)) : line_map (f a) (f b) r < f c ↔ slope f c b < slope f a b := -@map_lt_line_map_iff_slope_lt_slope_right k (order_dual E) _ _ _ _ f a b r h +@map_lt_line_map_iff_slope_lt_slope_right k Eᵒᵈ _ _ _ _ f a b r h /-- Given `c = line_map a b r`, `a < c < b`, the point `(c, f c)` is non-strictly below the segment `[(a, f a), (b, f b)]` if and only if `slope f a c ≤ slope f c b`. -/ @@ -277,7 +277,7 @@ end segment `[(a, f a), (b, f b)]` if and only if `slope f c b ≤ slope f a c`. -/ lemma line_map_le_map_iff_slope_le_slope (hab : a < b) (h₀ : 0 < r) (h₁ : r < 1) : line_map (f a) (f b) r ≤ f c ↔ slope f c b ≤ slope f a c := -@map_le_line_map_iff_slope_le_slope k (order_dual E) _ _ _ _ _ _ _ _ hab h₀ h₁ +@map_le_line_map_iff_slope_le_slope k Eᵒᵈ _ _ _ _ _ _ _ _ hab h₀ h₁ /-- Given `c = line_map a b r`, `a < c < b`, the point `(c, f c)` is strictly below the segment `[(a, f a), (b, f b)]` if and only if `slope f a c < slope f c b`. -/ @@ -290,6 +290,6 @@ lt_iff_lt_of_le_iff_le' (line_map_le_map_iff_slope_le_slope hab h₀ h₁) segment `[(a, f a), (b, f b)]` if and only if `slope f c b < slope f a c`. -/ lemma line_map_lt_map_iff_slope_lt_slope (hab : a < b) (h₀ : 0 < r) (h₁ : r < 1) : line_map (f a) (f b) r < f c ↔ slope f c b < slope f a c := -@map_lt_line_map_iff_slope_lt_slope k (order_dual E) _ _ _ _ _ _ _ _ hab h₀ h₁ +@map_lt_line_map_iff_slope_lt_slope k Eᵒᵈ _ _ _ _ _ _ _ _ hab h₀ h₁ end linear_ordered_field diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 43fde04ad6c3d..558762474a163 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -971,7 +971,7 @@ end The decreasing sequence of submodules consisting of the ranges of the iterates of a linear map. -/ @[simps] -def iterate_range (f : M →ₗ[R] M) : ℕ →o order_dual (submodule R M) := +def iterate_range (f : M →ₗ[R] M) : ℕ →o (submodule R M)ᵒᵈ := ⟨λ n, (f ^ n).range, λ n m w x h, begin obtain ⟨c, rfl⟩ := le_iff_exists_add.mp w, rw linear_map.mem_range at h, diff --git a/src/linear_algebra/prod.lean b/src/linear_algebra/prod.lean index b2ca23c5c34c1..a2898792568a7 100644 --- a/src/linear_algebra/prod.lean +++ b/src/linear_algebra/prod.lean @@ -686,7 +686,7 @@ noncomputable def tunnel' (f : M × N →ₗ[R] M) (i : injective f) : Give an injective map `f : M × N →ₗ[R] M` we can find a nested sequence of submodules all isomorphic to `M`. -/ -def tunnel (f : M × N →ₗ[R] M) (i : injective f) : ℕ →o order_dual (submodule R M) := +def tunnel (f : M × N →ₗ[R] M) (i : injective f) : ℕ →o (submodule R M)ᵒᵈ := ⟨λ n, (tunnel' f i n).1, monotone_nat_of_le_succ (λ n, begin dsimp [tunnel', tunnel_aux], rw [submodule.map_comp, submodule.map_comp], diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index d00dd27457ff2..fb54f747b226f 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -147,7 +147,7 @@ begin end lemma borel_eq_generate_from_Ioi : borel α = generate_from (range Ioi) := -@borel_eq_generate_from_Iio (order_dual α) _ (by apply_instance : second_countable_topology α) _ _ +@borel_eq_generate_from_Iio αᵒᵈ _ (by apply_instance : second_countable_topology α) _ _ end order_topology @@ -232,13 +232,13 @@ end tactic @[priority 100] instance order_dual.opens_measurable_space {α : Type*} [topological_space α] [measurable_space α] [h : opens_measurable_space α] : - opens_measurable_space (order_dual α) := + opens_measurable_space αᵒᵈ := { borel_le := h.borel_le } @[priority 100] instance order_dual.borel_space {α : Type*} [topological_space α] [measurable_space α] [h : borel_space α] : - borel_space (order_dual α) := + borel_space αᵒᵈ := { measurable_eq := h.measurable_eq } /-- In a `borel_space` all open sets are measurable. -/ @@ -619,7 +619,7 @@ lemma ext_of_Ioc_finite {α : Type*} [topological_space α] {m : measurable_spac [borel_space α] (μ ν : measure α) [is_finite_measure μ] (hμν : μ univ = ν univ) (h : ∀ ⦃a b⦄, a < b → μ (Ioc a b) = ν (Ioc a b)) : μ = ν := begin - refine @ext_of_Ico_finite (order_dual α) _ _ _ _ _ ‹_› μ ν _ hμν (λ a b hab, _), + refine @ext_of_Ico_finite αᵒᵈ _ _ _ _ _ ‹_› μ ν _ hμν (λ a b hab, _), erw dual_Ico, exact h hab end @@ -655,7 +655,7 @@ lemma ext_of_Ioc' {α : Type*} [topological_space α] {m : measurable_space α} [no_min_order α] (μ ν : measure α) (hμ : ∀ ⦃a b⦄, a < b → μ (Ioc a b) ≠ ∞) (h : ∀ ⦃a b⦄, a < b → μ (Ioc a b) = ν (Ioc a b)) : μ = ν := begin - refine @ext_of_Ico' (order_dual α) _ _ _ _ _ ‹_› _ μ ν _ _; + refine @ext_of_Ico' αᵒᵈ _ _ _ _ _ ‹_› _ μ ν _ _; intros a b hab; erw dual_Ico, exacts [hμ hab, h hab] end @@ -697,7 +697,7 @@ intervals. -/ lemma ext_of_Ici {α : Type*} [topological_space α] {m : measurable_space α} [second_countable_topology α] [linear_order α] [order_topology α] [borel_space α] (μ ν : measure α) [is_finite_measure μ] (h : ∀ a, μ (Ici a) = ν (Ici a)) : μ = ν := -@ext_of_Iic (order_dual α) _ _ _ _ _ ‹_› _ _ _ h +@ext_of_Iic αᵒᵈ _ _ _ _ _ ‹_› _ _ _ h end measure_theory.measure @@ -1129,12 +1129,12 @@ ae_measurable_restrict_of_measurable_subtype hs this.measurable protected lemma antitone.measurable [linear_order β] [order_closed_topology β] {f : β → α} (hf : antitone f) : measurable f := -@monotone.measurable (order_dual α) β _ _ ‹_› _ _ _ _ _ ‹_› _ _ _ hf +@monotone.measurable αᵒᵈ β _ _ ‹_› _ _ _ _ _ ‹_› _ _ _ hf lemma ae_measurable_restrict_of_antitone_on [linear_order β] [order_closed_topology β] {μ : measure β} {s : set β} (hs : measurable_set s) {f : β → α} (hf : antitone_on f s) : ae_measurable f (μ.restrict s) := -@ae_measurable_restrict_of_monotone_on (order_dual α) β _ _ ‹_› _ _ _ _ _ ‹_› _ _ _ _ hs _ hf +@ae_measurable_restrict_of_monotone_on αᵒᵈ β _ _ ‹_› _ _ _ _ _ ‹_› _ _ _ _ hs _ hf end linear_order diff --git a/src/measure_theory/function/ess_sup.lean b/src/measure_theory/function/ess_sup.lean index 5c3d3879924d6..469004444c966 100644 --- a/src/measure_theory/function/ess_sup.lean +++ b/src/measure_theory/function/ess_sup.lean @@ -46,7 +46,7 @@ lemma ess_sup_congr_ae {f g : α → β} (hfg : f =ᵐ[μ] g) : ess_sup f μ = e limsup_congr hfg lemma ess_inf_congr_ae {f g : α → β} (hfg : f =ᵐ[μ] g) : ess_inf f μ = ess_inf g μ := -@ess_sup_congr_ae α (order_dual β) _ _ _ _ _ hfg +@ess_sup_congr_ae α βᵒᵈ _ _ _ _ _ hfg end conditionally_complete_lattice @@ -73,7 +73,7 @@ le_bot_iff.mp (Inf_le (by simp [set.mem_set_of_eq, eventually_le, ae_iff])) @[simp] lemma ess_inf_measure_zero {m : measurable_space α} {f : α → β} : ess_inf f (0 : measure α) = ⊤ := -@ess_sup_measure_zero α (order_dual β) _ _ _ +@ess_sup_measure_zero α βᵒᵈ _ _ _ lemma ess_sup_mono_ae {f g : α → β} (hfg : f ≤ᵐ[μ] g) : ess_sup f μ ≤ ess_sup g μ := limsup_le_limsup hfg @@ -96,10 +96,10 @@ begin end lemma ess_inf_const (c : β) (hμ : μ ≠ 0) : ess_inf (λ x : α, c) μ = c := -@ess_sup_const α (order_dual β) _ _ _ _ hμ +@ess_sup_const α βᵒᵈ _ _ _ _ hμ lemma le_ess_inf_of_ae_le {f : α → β} (c : β) (hf : (λ _, c) ≤ᵐ[μ] f) : c ≤ ess_inf f μ := -@ess_sup_le_of_ae_le α (order_dual β) _ _ _ _ c hf +@ess_sup_le_of_ae_le α βᵒᵈ _ _ _ _ c hf lemma ess_sup_const_bot : ess_sup (λ x : α, (⊥ : β)) μ = (⊥ : β) := limsup_const_bot @@ -118,7 +118,7 @@ end lemma order_iso.ess_inf_apply {m : measurable_space α} {γ} [complete_lattice γ] (f : α → β) (μ : measure α) (g : β ≃o γ) : g (ess_inf f μ) = ess_inf (λ x, g (f x)) μ := -@order_iso.ess_sup_apply α (order_dual β) _ _ (order_dual γ) _ _ _ g.dual +@order_iso.ess_sup_apply α βᵒᵈ _ _ γᵒᵈ _ _ _ g.dual lemma ess_sup_mono_measure {f : α → β} (hμν : ν ≪ μ) : ess_sup f ν ≤ ess_sup f μ := begin @@ -207,7 +207,7 @@ lemma ae_lt_of_ess_sup_lt {f : α → β} {x : β} (hf : ess_sup f μ < x) : ∀ filter.eventually_lt_of_limsup_lt hf lemma ae_lt_of_lt_ess_inf {f : α → β} {x : β} (hf : x < ess_inf f μ) : ∀ᵐ y ∂μ, x < f y := -@ae_lt_of_ess_sup_lt α (order_dual β) _ _ _ _ _ hf +@ae_lt_of_ess_sup_lt α βᵒᵈ _ _ _ _ _ hf lemma ess_sup_indicator_eq_ess_sup_restrict [has_zero β] {s : set α} {f : α → β} (hf : 0 ≤ᵐ[μ.restrict s] f) (hs : measurable_set s) (hs_not_null : μ s ≠ 0) : diff --git a/src/measure_theory/function/locally_integrable.lean b/src/measure_theory/function/locally_integrable.lean index b2cd7339b04e7..e8d07e8aeebf3 100644 --- a/src/measure_theory/function/locally_integrable.lean +++ b/src/measure_theory/function/locally_integrable.lean @@ -180,12 +180,12 @@ end lemma antitone_on.integrable_on_compact (hs : is_compact s) (hanti : antitone_on f s) : integrable_on f s μ := -@monotone_on.integrable_on_compact X (order_dual E) _ _ _ _ _ _ _ _ _ _ _ _ _ _ hs hanti +hanti.dual_right.integrable_on_compact hs lemma monotone.locally_integrable (hmono : monotone f) : locally_integrable f μ := -λ s hs, monotone_on.integrable_on_compact hs (λ x y _ _ hxy, hmono hxy) +λ s hs, (hmono.monotone_on _).integrable_on_compact hs lemma antitone.locally_integrable (hanti : antitone f) : locally_integrable f μ := -@monotone.locally_integrable X (order_dual E) _ _ _ _ _ _ _ _ _ _ _ _ _ hanti +hanti.dual_right.locally_integrable end monotone diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index bdd9626da14f9..af85dc5380c35 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -358,7 +358,7 @@ end lemma antitone_on.interval_integrable {u : ℝ → E} {a b : ℝ} (hu : antitone_on u (interval a b)) : interval_integrable u μ a b := -@monotone_on.interval_integrable (order_dual E) _ _ _ _ _ _ _ _ _ hu +hu.dual_right.interval_integrable lemma monotone.interval_integrable {u : ℝ → E} {a b : ℝ} (hu : monotone u) : interval_integrable u μ a b := diff --git a/src/measure_theory/lattice.lean b/src/measure_theory/lattice.lean index 069c10d0c9adb..87b9da7e7679c 100644 --- a/src/measure_theory/lattice.lean +++ b/src/measure_theory/lattice.lean @@ -63,23 +63,19 @@ variables {M : Type*} [measurable_space M] section order_dual @[priority 100] -instance order_dual.has_measurable_sup [has_inf M] [has_measurable_inf M] : - has_measurable_sup (order_dual M) := +instance [has_inf M] [has_measurable_inf M] : has_measurable_sup Mᵒᵈ := ⟨@measurable_const_inf M _ _ _, @measurable_inf_const M _ _ _⟩ @[priority 100] -instance order_dual.has_measurable_inf [has_sup M] [has_measurable_sup M] : - has_measurable_inf (order_dual M) := +instance [has_sup M] [has_measurable_sup M] : has_measurable_inf Mᵒᵈ := ⟨@measurable_const_sup M _ _ _, @measurable_sup_const M _ _ _⟩ @[priority 100] -instance order_dual.has_measurable_sup₂ [has_inf M] [has_measurable_inf₂ M] : - has_measurable_sup₂ (order_dual M) := +instance [has_inf M] [has_measurable_inf₂ M] : has_measurable_sup₂ Mᵒᵈ := ⟨@measurable_inf M _ _ _⟩ @[priority 100] -instance order_dual.has_measurable_inf₂ [has_sup M] [has_measurable_sup₂ M] : - has_measurable_inf₂ (order_dual M) := +instance [has_sup M] [has_measurable_sup₂ M] : has_measurable_inf₂ Mᵒᵈ := ⟨@measurable_sup M _ _ _⟩ end order_dual diff --git a/src/measure_theory/measurable_space_def.lean b/src/measure_theory/measurable_space_def.lean index c136d5f854c29..c7cf0139715b6 100644 --- a/src/measure_theory/measurable_space_def.lean +++ b/src/measure_theory/measurable_space_def.lean @@ -55,7 +55,7 @@ structure measurable_space (α : Type*) := attribute [class] measurable_space -instance [h : measurable_space α] : measurable_space (order_dual α) := h +instance [h : measurable_space α] : measurable_space αᵒᵈ := h section diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 2cf0c911f4c6a..f86260b4fabfc 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -2046,8 +2046,7 @@ variables [partial_order α] {a b : α} lemma Iio_ae_eq_Iic' (ha : μ {a} = 0) : Iio a =ᵐ[μ] Iic a := by rw [←Iic_diff_right, diff_ae_eq_self, measure_mono_null (set.inter_subset_right _ _) ha] -lemma Ioi_ae_eq_Ici' (ha : μ {a} = 0) : Ioi a =ᵐ[μ] Ici a := -@Iio_ae_eq_Iic' (order_dual α) ‹_› ‹_› _ _ ha +lemma Ioi_ae_eq_Ici' (ha : μ {a} = 0) : Ioi a =ᵐ[μ] Ici a := @Iio_ae_eq_Iic' αᵒᵈ ‹_› ‹_› _ _ ha lemma Ioo_ae_eq_Ioc' (hb : μ {b} = 0) : Ioo a b =ᵐ[μ] Ioc a b := (ae_eq_refl _).inter (Iio_ae_eq_Iic' hb) diff --git a/src/measure_theory/pi_system.lean b/src/measure_theory/pi_system.lean index 485382125312d..efaaa9aa0f3cb 100644 --- a/src/measure_theory/pi_system.lean +++ b/src/measure_theory/pi_system.lean @@ -87,7 +87,7 @@ lemma is_pi_system_Iio : is_pi_system (range Iio : set (set α)) := @image_univ α _ Iio ▸ is_pi_system_image_Iio univ lemma is_pi_system_image_Ioi (s : set α) : is_pi_system (Ioi '' s) := -@is_pi_system_image_Iio (order_dual α) _ s +@is_pi_system_image_Iio αᵒᵈ _ s lemma is_pi_system_Ioi : is_pi_system (range Ioi : set (set α)) := @image_univ α _ Ioi ▸ is_pi_system_image_Ioi univ diff --git a/src/order/antisymmetrization.lean b/src/order/antisymmetrization.lean index c6ca0a435e52e..4e8098fb5a811 100644 --- a/src/order/antisymmetrization.lean +++ b/src/order/antisymmetrization.lean @@ -169,7 +169,7 @@ variables (α) /-- `antisymmetrization` and `order_dual` commute. -/ def order_iso.dual_antisymmetrization : - order_dual (antisymmetrization α (≤)) ≃o antisymmetrization (order_dual α) (≤) := + (antisymmetrization α (≤))ᵒᵈ ≃o antisymmetrization αᵒᵈ (≤) := { to_fun := quotient.map' id $ λ _ _, and.symm, inv_fun := quotient.map' id $ λ _ _, and.symm, left_inv := λ a, quotient.induction_on' a $ λ a, by simp_rw [quotient.map'_mk', id], diff --git a/src/order/atoms.lean b/src/order/atoms.lean index 052dcd0a24a51..bce44154fb480 100644 --- a/src/order/atoms.lean +++ b/src/order/atoms.lean @@ -155,20 +155,17 @@ export is_atomic (eq_bot_or_exists_atom_le) is_coatomic (eq_top_or_exists_le_coa variable {α} -@[simp] theorem is_coatomic_dual_iff_is_atomic [order_bot α] : - is_coatomic (order_dual α) ↔ is_atomic α := +@[simp] lemma is_coatomic_dual_iff_is_atomic [order_bot α] : is_coatomic αᵒᵈ ↔ is_atomic α := ⟨λ h, ⟨λ b, by apply h.eq_top_or_exists_le_coatom⟩, λ h, ⟨λ b, by apply h.eq_bot_or_exists_atom_le⟩⟩ -@[simp] theorem is_atomic_dual_iff_is_coatomic [order_top α] : - is_atomic (order_dual α) ↔ is_coatomic α := +@[simp] lemma is_atomic_dual_iff_is_coatomic [order_top α] : is_atomic αᵒᵈ ↔ is_coatomic α := ⟨λ h, ⟨λ b, by apply h.eq_bot_or_exists_atom_le⟩, λ h, ⟨λ b, by apply h.eq_top_or_exists_le_coatom⟩⟩ namespace is_atomic variables [order_bot α] [is_atomic α] -instance is_coatomic_dual : is_coatomic (order_dual α) := -is_coatomic_dual_iff_is_atomic.2 ‹is_atomic α› +instance is_coatomic_dual : is_coatomic αᵒᵈ := is_coatomic_dual_iff_is_atomic.2 ‹is_atomic α› instance {x : α} : is_atomic (set.Iic x) := ⟨λ ⟨y, hy⟩, (eq_bot_or_exists_atom_le y).imp subtype.mk_eq_mk.2 @@ -180,8 +177,7 @@ namespace is_coatomic variables [order_top α] [is_coatomic α] -instance is_coatomic : is_atomic (order_dual α) := -is_atomic_dual_iff_is_coatomic.2 ‹is_coatomic α› +instance is_coatomic : is_atomic αᵒᵈ := is_atomic_dual_iff_is_coatomic.2 ‹is_coatomic α› instance {x : α} : is_coatomic (set.Ici x) := ⟨λ ⟨y, hy⟩, (eq_top_or_exists_le_coatom y).imp subtype.mk_eq_mk.2 @@ -219,16 +215,16 @@ export is_atomistic (eq_Sup_atoms) is_coatomistic (eq_Inf_coatoms) variable {α} @[simp] -theorem is_coatomistic_dual_iff_is_atomistic : is_coatomistic (order_dual α) ↔ is_atomistic α := +theorem is_coatomistic_dual_iff_is_atomistic : is_coatomistic αᵒᵈ ↔ is_atomistic α := ⟨λ h, ⟨λ b, by apply h.eq_Inf_coatoms⟩, λ h, ⟨λ b, by apply h.eq_Sup_atoms⟩⟩ @[simp] -theorem is_atomistic_dual_iff_is_coatomistic : is_atomistic (order_dual α) ↔ is_coatomistic α := +theorem is_atomistic_dual_iff_is_coatomistic : is_atomistic αᵒᵈ ↔ is_coatomistic α := ⟨λ h, ⟨λ b, by apply h.eq_Sup_atoms⟩, λ h, ⟨λ b, by apply h.eq_Inf_coatoms⟩⟩ namespace is_atomistic -instance is_coatomistic_dual [h : is_atomistic α] : is_coatomistic (order_dual α) := +instance is_coatomistic_dual [h : is_atomistic α] : is_coatomistic αᵒᵈ := is_coatomistic_dual_iff_is_atomistic.2 h variable [is_atomistic α] @@ -270,7 +266,7 @@ end is_atomistic namespace is_coatomistic -instance is_atomistic_dual [h : is_coatomistic α] : is_atomistic (order_dual α) := +instance is_atomistic_dual [h : is_coatomistic α] : is_atomistic αᵒᵈ := is_atomistic_dual_iff_is_coatomistic.2 h variable [is_coatomistic α] @@ -292,12 +288,12 @@ class is_simple_order (α : Type*) [has_le α] [bounded_order α] extends nontri export is_simple_order (eq_bot_or_eq_top) theorem is_simple_order_iff_is_simple_order_order_dual [has_le α] [bounded_order α] : - is_simple_order α ↔ is_simple_order (order_dual α) := + is_simple_order α ↔ is_simple_order αᵒᵈ := begin split; intro i; haveI := i, { exact { exists_pair_ne := @exists_pair_ne α _, eq_bot_or_eq_top := λ a, or.symm (eq_bot_or_eq_top ((order_dual.of_dual a)) : _ ∨ _) } }, - { exact { exists_pair_ne := @exists_pair_ne (order_dual α) _, + { exact { exists_pair_ne := @exists_pair_ne αᵒᵈ _, eq_bot_or_eq_top := λ a, or.symm (eq_bot_or_eq_top (order_dual.to_dual a)) } } end @@ -314,7 +310,7 @@ section is_simple_order variables [partial_order α] [bounded_order α] [is_simple_order α] -instance {α} [has_le α] [bounded_order α] [is_simple_order α] : is_simple_order (order_dual α) := +instance {α} [has_le α] [bounded_order α] [is_simple_order α] : is_simple_order αᵒᵈ := is_simple_order_iff_is_simple_order_order_dual.1 (by apply_instance) /-- A simple `bounded_order` induces a preorder. This is not an instance to prevent loops. -/ diff --git a/src/order/basic.lean b/src/order/basic.lean index 48a73f4fa6044..d5ed39b3f15f1 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -14,7 +14,7 @@ classes and allows to transfer order instances. ## Type synonyms -* `order_dual α` : A type synonym reversing the meaning of all inequalities. +* `order_dual α` : A type synonym reversing the meaning of all inequalities, with notation `αᵒᵈ`. * `as_linear_order α`: A type synonym to promote `partial_order α` to `linear_order α` using `is_total α (≤)`. @@ -389,37 +389,40 @@ instance order.preimage.decidable {α β} (f : α → β) (s : β → β → Pro /-! ### Order dual -/ -/-- Type synonym to equip a type with the dual order: `≤` means `≥` and `<` means `>`. -/ +/-- Type synonym to equip a type with the dual order: `≤` means `≥` and `<` means `>`. `αᵒᵈ` is +notation for `order_dual α`. -/ def order_dual (α : Type*) : Type* := α +notation α `ᵒᵈ`:std.prec.max_plus := order_dual α + namespace order_dual -instance (α : Type*) [h : nonempty α] : nonempty (order_dual α) := h -instance (α : Type*) [h : subsingleton α] : subsingleton (order_dual α) := h -instance (α : Type*) [has_le α] : has_le (order_dual α) := ⟨λ x y : α, y ≤ x⟩ -instance (α : Type*) [has_lt α] : has_lt (order_dual α) := ⟨λ x y : α, y < x⟩ -instance (α : Type*) [has_zero α] : has_zero (order_dual α) := ⟨(0 : α)⟩ +instance (α : Type*) [h : nonempty α] : nonempty αᵒᵈ := h +instance (α : Type*) [h : subsingleton α] : subsingleton αᵒᵈ := h +instance (α : Type*) [has_le α] : has_le αᵒᵈ := ⟨λ x y : α, y ≤ x⟩ +instance (α : Type*) [has_lt α] : has_lt αᵒᵈ := ⟨λ x y : α, y < x⟩ +instance (α : Type*) [has_zero α] : has_zero αᵒᵈ := ⟨(0 : α)⟩ -- `dual_le` and `dual_lt` should not be simp lemmas: --- they cause a loop since `α` and `order_dual α` are definitionally equal +-- they cause a loop since `α` and `αᵒᵈ` are definitionally equal lemma dual_le [has_le α] {a b : α} : - @has_le.le (order_dual α) _ a b ↔ @has_le.le α _ b a := iff.rfl + @has_le.le αᵒᵈ _ a b ↔ @has_le.le α _ b a := iff.rfl lemma dual_lt [has_lt α] {a b : α} : - @has_lt.lt (order_dual α) _ a b ↔ @has_lt.lt α _ b a := iff.rfl + @has_lt.lt αᵒᵈ _ a b ↔ @has_lt.lt α _ b a := iff.rfl -instance (α : Type*) [preorder α] : preorder (order_dual α) := +instance (α : Type*) [preorder α] : preorder αᵒᵈ := { le_refl := le_refl, le_trans := λ a b c hab hbc, hbc.trans hab, lt_iff_le_not_le := λ _ _, lt_iff_le_not_le, .. order_dual.has_le α, .. order_dual.has_lt α } -instance (α : Type*) [partial_order α] : partial_order (order_dual α) := +instance (α : Type*) [partial_order α] : partial_order αᵒᵈ := { le_antisymm := λ a b hab hba, @le_antisymm α _ a b hba hab, .. order_dual.preorder α } -instance (α : Type*) [linear_order α] : linear_order (order_dual α) := +instance (α : Type*) [linear_order α] : linear_order αᵒᵈ := { le_total := λ a b : α, le_total b a, decidable_le := (infer_instance : decidable_rel (λ a b : α, b ≤ a)), decidable_lt := (infer_instance : decidable_rel (λ a b : α, b < a)), @@ -429,18 +432,18 @@ instance (α : Type*) [linear_order α] : linear_order (order_dual α) := max_def := @linear_order.min_def α _, .. order_dual.partial_order α } -instance : Π [inhabited α], inhabited (order_dual α) := id +instance : Π [inhabited α], inhabited αᵒᵈ := id theorem preorder.dual_dual (α : Type*) [H : preorder α] : - order_dual.preorder (order_dual α) = H := + order_dual.preorder αᵒᵈ = H := preorder.ext $ λ _ _, iff.rfl theorem partial_order.dual_dual (α : Type*) [H : partial_order α] : - order_dual.partial_order (order_dual α) = H := + order_dual.partial_order αᵒᵈ = H := partial_order.ext $ λ _ _, iff.rfl theorem linear_order.dual_dual (α : Type*) [H : linear_order α] : - order_dual.linear_order (order_dual α) = H := + order_dual.linear_order αᵒᵈ = H := linear_order.ext $ λ _ _, iff.rfl end order_dual @@ -626,7 +629,7 @@ lemma exists_between [has_lt α] [densely_ordered α] : densely_ordered.dense instance order_dual.densely_ordered (α : Type u) [has_lt α] [densely_ordered α] : - densely_ordered (order_dual α) := + densely_ordered αᵒᵈ := ⟨λ a₁ a₂ ha, (@exists_between α _ _ _ _ ha).imp $ λ a, and.symm⟩ lemma le_of_forall_le_of_dense [linear_order α] [densely_ordered α] {a₁ a₂ : α} diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index 85efade06d5a9..beaa2b5a72bc3 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -761,7 +761,7 @@ section boolean_algebra variables [boolean_algebra α] --TODO@Yaël: Once we have co-Heyting algebras, we won't need to go through `boolean_algebra.of_core` -instance : boolean_algebra (order_dual α) := +instance : boolean_algebra αᵒᵈ := boolean_algebra.of_core { compl := λ a, to_dual (of_dual a)ᶜ, inf_compl_le_bot := λ _, sup_compl_eq_top.ge, diff --git a/src/order/bounded.lean b/src/order/bounded.lean index a2ee26dd25ee9..4c29d5138dae0 100644 --- a/src/order/bounded.lean +++ b/src/order/bounded.lean @@ -41,7 +41,7 @@ lemma unbounded_lt_iff [linear_order α] : unbounded (<) s ↔ ∀ a, ∃ b ∈ by simp only [unbounded, not_lt] lemma unbounded_ge_of_forall_exists_gt [preorder α] (h : ∀ a, ∃ b ∈ s, b < a) : unbounded (≥) s := -@unbounded_le_of_forall_exists_lt (order_dual α) _ _ h +@unbounded_le_of_forall_exists_lt αᵒᵈ _ _ h lemma unbounded_ge_iff [linear_order α] : unbounded (≥) s ↔ ∀ a, ∃ b ∈ s, b < a := ⟨λ h a, let ⟨b, hb, hba⟩ := h a in ⟨b, hb, lt_of_not_ge hba⟩, unbounded_ge_of_forall_exists_gt⟩ @@ -90,11 +90,11 @@ lemma unbounded_gt_of_unbounded_ge [preorder α] (h : unbounded (≥) s) : unbou λ a, let ⟨b, hb, hba⟩ := h a in ⟨b, hb, λ hba', hba (le_of_lt hba')⟩ lemma bounded_ge_iff_bounded_gt [preorder α] [no_min_order α] : bounded (≥) s ↔ bounded (>) s := -@bounded_le_iff_bounded_lt (order_dual α) _ _ _ +@bounded_le_iff_bounded_lt αᵒᵈ _ _ _ lemma unbounded_gt_iff_unbounded_ge [preorder α] [no_min_order α] : unbounded (>) s ↔ unbounded (≥) s := -@unbounded_lt_iff_unbounded_le (order_dual α) _ _ _ +@unbounded_lt_iff_unbounded_le αᵒᵈ _ _ _ /-! ### The universal set -/ @@ -296,52 +296,52 @@ end theorem bounded_ge_inter_not_ge [semilattice_inf α] (a : α) : bounded (≥) (s ∩ {b | ¬ a ≤ b}) ↔ bounded (≥) s := -@bounded_le_inter_not_le (order_dual α) s _ a +@bounded_le_inter_not_le αᵒᵈ s _ a theorem unbounded_ge_inter_not_ge [semilattice_inf α] (a : α) : unbounded (≥) (s ∩ {b | ¬ a ≤ b}) ↔ unbounded (≥) s := -@unbounded_le_inter_not_le (order_dual α) s _ a +@unbounded_le_inter_not_le αᵒᵈ s _ a theorem bounded_ge_inter_gt [linear_order α] (a : α) : bounded (≥) (s ∩ {b | b < a}) ↔ bounded (≥) s := -@bounded_le_inter_lt (order_dual α) s _ a +@bounded_le_inter_lt αᵒᵈ s _ a theorem unbounded_ge_inter_gt [linear_order α] (a : α) : unbounded (≥) (s ∩ {b | b < a}) ↔ unbounded (≥) s := -@unbounded_le_inter_lt (order_dual α) s _ a +@unbounded_le_inter_lt αᵒᵈ s _ a theorem bounded_ge_inter_ge [linear_order α] (a : α) : bounded (≥) (s ∩ {b | b ≤ a}) ↔ bounded (≥) s := -@bounded_le_inter_le (order_dual α) s _ a +@bounded_le_inter_le αᵒᵈ s _ a theorem unbounded_ge_iff_unbounded_inter_ge [linear_order α] (a : α) : unbounded (≥) (s ∩ {b | b ≤ a}) ↔ unbounded (≥) s := -@unbounded_le_inter_le (order_dual α) s _ a +@unbounded_le_inter_le αᵒᵈ s _ a /-! #### Greater than -/ theorem bounded_gt_inter_not_gt [semilattice_inf α] (a : α) : bounded (>) (s ∩ {b | ¬ a < b}) ↔ bounded (>) s := -@bounded_lt_inter_not_lt (order_dual α) s _ a +@bounded_lt_inter_not_lt αᵒᵈ s _ a theorem unbounded_gt_inter_not_gt [semilattice_inf α] (a : α) : unbounded (>) (s ∩ {b | ¬ a < b}) ↔ unbounded (>) s := -@unbounded_lt_inter_not_lt (order_dual α) s _ a +@unbounded_lt_inter_not_lt αᵒᵈ s _ a theorem bounded_gt_inter_ge [linear_order α] (a : α) : bounded (>) (s ∩ {b | b ≤ a}) ↔ bounded (>) s := -@bounded_lt_inter_le (order_dual α) s _ a +@bounded_lt_inter_le αᵒᵈ s _ a theorem unbounded_inter_ge [linear_order α] (a : α) : unbounded (>) (s ∩ {b | b ≤ a}) ↔ unbounded (>) s := -@unbounded_lt_inter_le (order_dual α) s _ a +@unbounded_lt_inter_le αᵒᵈ s _ a theorem bounded_gt_inter_gt [linear_order α] [no_min_order α] (a : α) : bounded (>) (s ∩ {b | b < a}) ↔ bounded (>) s := -@bounded_lt_inter_lt (order_dual α) s _ _ a +@bounded_lt_inter_lt αᵒᵈ s _ _ a theorem unbounded_gt_inter_gt [linear_order α] [no_min_order α] (a : α) : unbounded (>) (s ∩ {b | b < a}) ↔ unbounded (>) s := -@unbounded_lt_inter_lt (order_dual α) s _ _ a +@unbounded_lt_inter_lt αᵒᵈ s _ _ a end set diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index dc463a734e37e..1b937d2324ae8 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -958,10 +958,10 @@ instance [lattice α] : lattice (with_top α) := { ..with_top.semilattice_sup, ..with_top.semilattice_inf } instance decidable_le [has_le α] [@decidable_rel α (≤)] : @decidable_rel (with_top α) (≤) := -λ x y, @with_bot.decidable_le (order_dual α) _ _ y x +λ x y, @with_bot.decidable_le αᵒᵈ _ _ y x instance decidable_lt [has_lt α] [@decidable_rel α (<)] : @decidable_rel (with_top α) (<) := -λ x y, @with_bot.decidable_lt (order_dual α) _ _ y x +λ x y, @with_bot.decidable_lt αᵒᵈ _ _ y x instance is_total_le [has_le α] [is_total α (≤)] : is_total (with_top α) (≤) := ⟨λ a b, match a, b with @@ -989,11 +989,11 @@ have acc_some : ∀ a : α, acc ((<) : with_top α → with_top α → Prop) (so (λ _ _, acc_some _))) acc_some⟩ lemma well_founded_gt [preorder α] (h : @well_founded α (>)) : @well_founded (with_top α) (>) := -@with_bot.well_founded_lt (order_dual α) _ h +@with_bot.well_founded_lt αᵒᵈ _ h lemma _root_.with_bot.well_founded_gt [preorder α] (h : @well_founded α (>)) : @well_founded (with_bot α) (>) := -@with_top.well_founded_lt (order_dual α) _ h +@with_top.well_founded_lt αᵒᵈ _ h instance [has_lt α] [densely_ordered α] [no_max_order α] : densely_ordered (with_top α) := ⟨ λ a b, @@ -1087,18 +1087,18 @@ end subtype namespace order_dual variable (α) -instance [has_bot α] : has_top (order_dual α) := ⟨(⊥ : α)⟩ -instance [has_top α] : has_bot (order_dual α) := ⟨(⊤ : α)⟩ +instance [has_bot α] : has_top αᵒᵈ := ⟨(⊥ : α)⟩ +instance [has_top α] : has_bot αᵒᵈ := ⟨(⊤ : α)⟩ -instance [has_le α] [order_bot α] : order_top (order_dual α) := +instance [has_le α] [order_bot α] : order_top αᵒᵈ := { le_top := @bot_le α _ _, .. order_dual.has_top α } -instance [has_le α] [order_top α] : order_bot (order_dual α) := +instance [has_le α] [order_top α] : order_bot αᵒᵈ := { bot_le := @le_top α _ _, .. order_dual.has_bot α } -instance [has_le α] [bounded_order α] : bounded_order (order_dual α) := +instance [has_le α] [bounded_order α] : bounded_order αᵒᵈ := { .. order_dual.order_top α, .. order_dual.order_bot α } end order_dual @@ -1221,7 +1221,7 @@ lemma max_top_right [order_top α] (a : α) : max a ⊤ = ⊤ := max_eq_right le by { symmetry, cases le_total a b; simpa [*, min_eq_left, min_eq_right] using eq_bot_mono h } @[simp] lemma max_eq_top [order_top α] {a b : α} : max a b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := -@min_eq_bot (order_dual α) _ _ a b +@min_eq_bot αᵒᵈ _ _ a b @[simp] lemma max_eq_bot [order_bot α] {a b : α} : max a b = ⊥ ↔ a = ⊥ ∧ b = ⊥ := sup_eq_bot_iff @[simp] lemma min_eq_top [order_top α] {a b : α} : min a b = ⊤ ↔ a = ⊤ ∧ b = ⊤ := inf_eq_top_iff @@ -1408,7 +1408,7 @@ export is_complemented (exists_is_compl) namespace is_complemented variables [lattice α] [bounded_order α] [is_complemented α] -instance : is_complemented (order_dual α) := +instance : is_complemented αᵒᵈ := ⟨λ a, let ⟨b, hb⟩ := exists_is_compl (show α, from a) in ⟨b, hb.to_order_dual⟩⟩ end is_complemented diff --git a/src/order/bounds.lean b/src/order/bounds.lean index 8e432601bbf1a..d9972992fe65d 100644 --- a/src/order/bounds.lean +++ b/src/order/bounds.lean @@ -78,8 +78,7 @@ by simp [bdd_above, upper_bounds, set.nonempty] /-- A set `s` is not bounded below if and only if for each `x` there exists `y ∈ s` such that `x` is not less than or equal to `y`. This version only assumes `preorder` structure and uses `¬(x ≤ y)`. A version for linear orders is called `not_bdd_below_iff`. -/ -lemma not_bdd_below_iff' : ¬bdd_below s ↔ ∀ x, ∃ y ∈ s, ¬(x ≤ y) := -@not_bdd_above_iff' (order_dual α) _ _ +lemma not_bdd_below_iff' : ¬bdd_below s ↔ ∀ x, ∃ y ∈ s, ¬ x ≤ y := @not_bdd_above_iff' αᵒᵈ _ _ /-- A set `s` is not bounded above if and only if for each `x` there exists `y ∈ s` that is greater than `x`. A version for preorders is called `not_bdd_above_iff'`. -/ @@ -91,7 +90,7 @@ by simp only [not_bdd_above_iff', not_le] than `x`. A version for preorders is called `not_bdd_below_iff'`. -/ lemma not_bdd_below_iff {α : Type*} [linear_order α] {s : set α} : ¬bdd_below s ↔ ∀ x, ∃ y ∈ s, y < x := -@not_bdd_above_iff (order_dual α) _ _ +@not_bdd_above_iff αᵒᵈ _ _ lemma bdd_above.dual (h : bdd_above s) : bdd_below (of_dual ⁻¹' s) := h @@ -214,7 +213,7 @@ lemma is_lub_iff_le_iff : is_lub s a ↔ ∀ b, a ≤ b ↔ b ∈ upper_bounds s ⟨λ h b, is_lub_le_iff h, λ H, ⟨(H _).1 le_rfl, λ b hb, (H b).2 hb⟩⟩ lemma is_glb_iff_le_iff : is_glb s a ↔ ∀ b, b ≤ a ↔ b ∈ lower_bounds s := -@is_lub_iff_le_iff (order_dual α) _ _ _ +@is_lub_iff_le_iff αᵒᵈ _ _ _ /-- If `s` has a least upper bound, then it is bounded above. -/ lemma is_lub.bdd_above (h : is_lub s a) : bdd_above s := ⟨a, h.1⟩ @@ -242,7 +241,7 @@ subset.antisymm (λ b hb x hx, hx.elim (λ hs, hb.1 hs) (λ ht, hb.2 ht)) @[simp] lemma lower_bounds_union : lower_bounds (s ∪ t) = lower_bounds s ∩ lower_bounds t := -@upper_bounds_union (order_dual α) _ s t +@upper_bounds_union αᵒᵈ _ s t lemma union_upper_bounds_subset_upper_bounds_inter : upper_bounds s ∪ upper_bounds t ⊆ upper_bounds (s ∩ t) := @@ -252,7 +251,7 @@ union_subset lemma union_lower_bounds_subset_lower_bounds_inter : lower_bounds s ∪ lower_bounds t ⊆ lower_bounds (s ∩ t) := -@union_upper_bounds_subset_upper_bounds_inter (order_dual α) _ s t +@union_upper_bounds_subset_upper_bounds_inter αᵒᵈ _ s t lemma is_least_union_iff {a : α} {s t : set α} : is_least (s ∪ t) a ↔ (is_least s a ∧ a ∈ lower_bounds t ∨ a ∈ lower_bounds s ∧ is_least t a) := @@ -261,7 +260,7 @@ by simp [is_least, lower_bounds_union, or_and_distrib_right, and_comm (a ∈ t), lemma is_greatest_union_iff : is_greatest (s ∪ t) a ↔ (is_greatest s a ∧ a ∈ upper_bounds t ∨ a ∈ upper_bounds s ∧ is_greatest t a) := -@is_least_union_iff (order_dual α) _ a s t +@is_least_union_iff αᵒᵈ _ a s t /-- If `s` is bounded, then so is `s ∩ t` -/ lemma bdd_above.inter_of_left (h : bdd_above s) : bdd_above (s ∩ t) := @@ -298,12 +297,12 @@ lemma bdd_above_union [semilattice_sup γ] {s t : set γ} : lemma bdd_below.union [semilattice_inf γ] {s t : set γ} : bdd_below s → bdd_below t → bdd_below (s ∪ t) := -@bdd_above.union (order_dual γ) _ s t +@bdd_above.union γᵒᵈ _ s t /--The union of two sets is bounded above if and only if each of the sets is.-/ lemma bdd_below_union [semilattice_inf γ] {s t : set γ} : bdd_below (s ∪ t) ↔ bdd_below s ∧ bdd_below t := -@bdd_above_union (order_dual γ) _ s t +@bdd_above_union γᵒᵈ _ s t /-- If `a` is the least upper bound of `s` and `b` is the least upper bound of `t`, then `a ⊔ b` is the least upper bound of `s ∪ t`. -/ @@ -389,8 +388,7 @@ lemma bdd_below_Ioi : bdd_below (Ioi a) := ⟨a, λ x hx, le_of_lt hx⟩ lemma lub_Iio_le (a : α) (hb : is_lub (set.Iio a) b) : b ≤ a := (is_lub_le_iff hb).mpr $ λ k hk, le_of_lt hk -lemma le_glb_Ioi (a : α) (hb : is_glb (set.Ioi a) b) : a ≤ b := -@lub_Iio_le (order_dual α) _ _ a hb +lemma le_glb_Ioi (a : α) (hb : is_glb (set.Ioi a) b) : a ≤ b := @lub_Iio_le αᵒᵈ _ _ a hb lemma lub_Iio_eq_self_or_Iio_eq_Iic [partial_order γ] {j : γ} (i : γ) (hj : is_lub (set.Iio i) j) : j = i ∨ set.Iio i = set.Iic j := @@ -403,7 +401,7 @@ end lemma glb_Ioi_eq_self_or_Ioi_eq_Ici [partial_order γ] {j : γ} (i : γ) (hj : is_glb (set.Ioi i) j) : j = i ∨ set.Ioi i = set.Ici j := -@lub_Iio_eq_self_or_Iio_eq_Iic (order_dual γ) _ j i hj +@lub_Iio_eq_self_or_Iio_eq_Iic γᵒᵈ _ j i hj section @@ -422,15 +420,14 @@ begin exact h, }, end -lemma exists_glb_Ioi (i : γ) : ∃ j, is_glb (set.Ioi i) j := -@exists_lub_Iio (order_dual γ) _ i +lemma exists_glb_Ioi (i : γ) : ∃ j, is_glb (set.Ioi i) j := @exists_lub_Iio γᵒᵈ _ i variables [densely_ordered γ] lemma is_lub_Iio {a : γ} : is_lub (Iio a) a := ⟨λ x hx, le_of_lt hx, λ y hy, le_of_forall_ge_of_dense hy⟩ -lemma is_glb_Ioi {a : γ} : is_glb (Ioi a) a := @is_lub_Iio (order_dual γ) _ _ a +lemma is_glb_Ioi {a : γ} : is_glb (Ioi a) a := @is_lub_Iio γᵒᵈ _ _ a lemma upper_bounds_Iio {a : γ} : upper_bounds (Iio a) = Ici a := is_lub_Iio.upper_bounds_eq @@ -445,8 +442,7 @@ end lemma is_greatest_singleton : is_greatest {a} a := ⟨mem_singleton a, λ x hx, le_of_eq $ eq_of_mem_singleton hx⟩ -lemma is_least_singleton : is_least {a} a := -@is_greatest_singleton (order_dual α) _ a +lemma is_least_singleton : is_least {a} a := @is_greatest_singleton αᵒᵈ _ a lemma is_lub_singleton : is_lub {a} a := is_greatest_singleton.is_lub @@ -583,10 +579,10 @@ is_greatest_univ.is_lub @[simp] lemma order_bot.lower_bounds_univ [partial_order γ] [order_bot γ] : lower_bounds (univ : set γ) = {⊥} := -@order_top.upper_bounds_univ (order_dual γ) _ _ +@order_top.upper_bounds_univ γᵒᵈ _ _ lemma is_least_univ [preorder γ] [order_bot γ] : is_least (univ : set γ) ⊥ := -@is_greatest_univ (order_dual γ) _ _ +@is_greatest_univ γᵒᵈ _ _ lemma is_glb_univ [preorder γ] [order_bot γ] : is_glb (univ : set γ) ⊥ := is_least_univ.is_glb @@ -596,13 +592,13 @@ eq_empty_of_subset_empty $ λ b hb, let ⟨x, hx⟩ := exists_gt b in not_le_of_lt hx (hb trivial) @[simp] lemma no_min_order.lower_bounds_univ [no_min_order α] : lower_bounds (univ : set α) = ∅ := -@no_max_order.upper_bounds_univ (order_dual α) _ _ +@no_max_order.upper_bounds_univ αᵒᵈ _ _ @[simp] lemma not_bdd_above_univ [no_max_order α] : ¬bdd_above (univ : set α) := by simp [bdd_above] @[simp] lemma not_bdd_below_univ [no_min_order α] : ¬bdd_below (univ : set α) := -@not_bdd_above_univ (order_dual α) _ _ +@not_bdd_above_univ αᵒᵈ _ _ /-! #### Empty set @@ -611,8 +607,7 @@ by simp [bdd_above] @[simp] lemma upper_bounds_empty : upper_bounds (∅ : set α) = univ := by simp only [upper_bounds, eq_univ_iff_forall, mem_set_of_eq, ball_empty_iff, forall_true_iff] -@[simp] lemma lower_bounds_empty : lower_bounds (∅ : set α) = univ := -@upper_bounds_empty (order_dual α) _ +@[simp] lemma lower_bounds_empty : lower_bounds (∅ : set α) = univ := @upper_bounds_empty αᵒᵈ _ @[simp] lemma bdd_above_empty [nonempty α] : bdd_above (∅ : set α) := by simp only [bdd_above, upper_bounds_empty, univ_nonempty] @@ -623,8 +618,7 @@ by simp only [bdd_below, lower_bounds_empty, univ_nonempty] lemma is_glb_empty [preorder γ] [order_top γ] : is_glb ∅ (⊤:γ) := by simp only [is_glb, lower_bounds_empty, is_greatest_univ] -lemma is_lub_empty [preorder γ] [order_bot γ] : is_lub ∅ (⊥:γ) := -@is_glb_empty (order_dual γ) _ _ +lemma is_lub_empty [preorder γ] [order_bot γ] : is_lub ∅ (⊥:γ) := @is_glb_empty γᵒᵈ _ _ lemma is_lub.nonempty [no_min_order α] (hs : is_lub s a) : s.nonempty := let ⟨a', ha'⟩ := exists_lt a in @@ -638,7 +632,7 @@ lemma nonempty_of_not_bdd_above [ha : nonempty α] (h : ¬bdd_above s) : s.nonem nonempty.elim ha $ λ x, (not_bdd_above_iff'.1 h x).imp $ λ a ha, ha.fst lemma nonempty_of_not_bdd_below [ha : nonempty α] (h : ¬bdd_below s) : s.nonempty := -@nonempty_of_not_bdd_above (order_dual α) _ _ _ h +@nonempty_of_not_bdd_above αᵒᵈ _ _ _ h /-! #### insert @@ -718,7 +712,7 @@ is_greatest_singleton.insert _ ⟨λ H, ⟨λ x hx, H.2 $ subset_upper_bounds_lower_bounds s hx, H.1⟩, is_greatest.is_lub⟩ @[simp] lemma is_glb_upper_bounds : is_glb (upper_bounds s) a ↔ is_lub s a := -@is_lub_lower_bounds (order_dual α) _ _ _ +@is_lub_lower_bounds αᵒᵈ _ _ _ end @@ -1038,7 +1032,7 @@ lemma is_glb.of_image [preorder α] [preorder β] {f : α → β} (hf : ∀ {x y lemma is_lub.of_image [preorder α] [preorder β] {f : α → β} (hf : ∀ {x y}, f x ≤ f y ↔ x ≤ y) {s : set α} {x : α} (hx : is_lub (f '' s) (f x)) : is_lub s x := -@is_glb.of_image (order_dual α) (order_dual β) _ _ f (λ x y, hf) _ _ hx +@is_glb.of_image αᵒᵈ βᵒᵈ _ _ f (λ x y, hf) _ _ hx lemma is_lub_pi {π : α → Type*} [Π a, preorder (π a)] {s : set (Π a, π a)} {f : Π a, π a} : is_lub s f ↔ ∀ a, is_lub (function.eval a '' s) (f a) := @@ -1054,7 +1048,7 @@ end lemma is_glb_pi {π : α → Type*} [Π a, preorder (π a)] {s : set (Π a, π a)} {f : Π a, π a} : is_glb s f ↔ ∀ a, is_glb (function.eval a '' s) (f a) := -@is_lub_pi α (λ a, order_dual (π a)) _ s f +@is_lub_pi α (λ a, (π a)ᵒᵈ) _ s f lemma is_lub_prod [preorder α] [preorder β] {s : set (α × β)} (p : α × β) : is_lub s p ↔ is_lub (prod.fst '' s) p.1 ∧ is_lub (prod.snd '' s) p.2 := @@ -1072,7 +1066,7 @@ end lemma is_glb_prod [preorder α] [preorder β] {s : set (α × β)} (p : α × β) : is_glb s p ↔ is_glb (prod.fst '' s) p.1 ∧ is_glb (prod.snd '' s) p.2 := -@is_lub_prod (order_dual α) (order_dual β) _ _ _ _ +@is_lub_prod αᵒᵈ βᵒᵈ _ _ _ _ namespace order_iso @@ -1084,9 +1078,8 @@ subset.antisymm (λ x hx, ⟨f.symm x, λ y hy, f.le_symm_apply.2 (hx $ mem_image_of_mem _ hy), f.apply_symm_apply x⟩) f.monotone.image_upper_bounds_subset_upper_bounds_image -lemma lower_bounds_image {s : set α} : - lower_bounds (f '' s) = f '' lower_bounds s := -@upper_bounds_image (order_dual α) (order_dual β) _ _ f.dual _ +lemma lower_bounds_image {s : set α} : lower_bounds (f '' s) = f '' lower_bounds s := +@upper_bounds_image αᵒᵈ βᵒᵈ _ _ f.dual _ @[simp] lemma is_lub_image {s : set α} {x : β} : is_lub (f '' s) x ↔ is_lub s (f.symm x) := diff --git a/src/order/category/BoolAlg.lean b/src/order/category/BoolAlg.lean index 13894871ed017..7301e685995a9 100644 --- a/src/order/category/BoolAlg.lean +++ b/src/order/category/BoolAlg.lean @@ -53,7 +53,7 @@ between them. -/ /-- `order_dual` as a functor. -/ @[simps] def dual : BoolAlg ⥤ BoolAlg := -{ obj := λ X, of (order_dual X), map := λ X Y, bounded_lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } /-- The equivalence between `BoolAlg` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : BoolAlg ≌ BoolAlg := diff --git a/src/order/category/BoundedDistribLattice.lean b/src/order/category/BoundedDistribLattice.lean index 58e342ecbddf9..ecac2f1ba4af5 100644 --- a/src/order/category/BoundedDistribLattice.lean +++ b/src/order/category/BoundedDistribLattice.lean @@ -68,7 +68,7 @@ between them. -/ /-- `order_dual` as a functor. -/ @[simps] def dual : BoundedDistribLattice ⥤ BoundedDistribLattice := -{ obj := λ X, of (order_dual X), map := λ X Y, bounded_lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } /-- The equivalence between `BoundedDistribLattice` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : BoundedDistribLattice ≌ BoundedDistribLattice := diff --git a/src/order/category/BoundedLattice.lean b/src/order/category/BoundedLattice.lean index 01043558a12a1..0201a28d1a191 100644 --- a/src/order/category/BoundedLattice.lean +++ b/src/order/category/BoundedLattice.lean @@ -98,7 +98,7 @@ between them. -/ /-- `order_dual` as a functor. -/ @[simps] def dual : BoundedLattice ⥤ BoundedLattice := -{ obj := λ X, of (order_dual X), map := λ X Y, bounded_lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } /-- The equivalence between `BoundedLattice` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : BoundedLattice ≌ BoundedLattice := diff --git a/src/order/category/BoundedOrder.lean b/src/order/category/BoundedOrder.lean index 5db92d0feda2b..8c359e2282b31 100644 --- a/src/order/category/BoundedOrder.lean +++ b/src/order/category/BoundedOrder.lean @@ -57,7 +57,7 @@ instance has_forget_to_Bipointed : has_forget₂ BoundedOrder Bipointed := /-- `order_dual` as a functor. -/ @[simps] def dual : BoundedOrder ⥤ BoundedOrder := -{ obj := λ X, of (order_dual X), map := λ X Y, bounded_order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_order_hom.dual } /-- Constructs an equivalence between bounded orders from an order isomorphism between them. -/ @[simps] def iso.mk {α β : BoundedOrder.{u}} (e : α ≃o β) : α ≅ β := diff --git a/src/order/category/CompleteLattice.lean b/src/order/category/CompleteLattice.lean index ffc4930c60e76..03ce1f1839b3a 100644 --- a/src/order/category/CompleteLattice.lean +++ b/src/order/category/CompleteLattice.lean @@ -53,7 +53,7 @@ instance has_forget_to_BoundedLattice : has_forget₂ CompleteLattice BoundedLat /-- `order_dual` as a functor. -/ @[simps] def dual : CompleteLattice ⥤ CompleteLattice := -{ obj := λ X, of (order_dual X), map := λ X Y, complete_lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, complete_lattice_hom.dual } /-- The equivalence between `CompleteLattice` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : CompleteLattice ≌ CompleteLattice := diff --git a/src/order/category/DistribLattice.lean b/src/order/category/DistribLattice.lean index b966c9f4c3099..7b7220477b2b9 100644 --- a/src/order/category/DistribLattice.lean +++ b/src/order/category/DistribLattice.lean @@ -50,7 +50,7 @@ instance has_forget_to_Lattice : has_forget₂ DistribLattice Lattice := bundled /-- `order_dual` as a functor. -/ @[simps] def dual : DistribLattice ⥤ DistribLattice := -{ obj := λ X, of (order_dual X), map := λ X Y, lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, lattice_hom.dual } /-- The equivalence between `DistribLattice` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : DistribLattice ≌ DistribLattice := diff --git a/src/order/category/FinBoolAlg.lean b/src/order/category/FinBoolAlg.lean index 26981b8b6003d..a6a1960a6ed5d 100644 --- a/src/order/category/FinBoolAlg.lean +++ b/src/order/category/FinBoolAlg.lean @@ -76,7 +76,7 @@ them. -/ /-- `order_dual` as a functor. -/ @[simps] def dual : FinBoolAlg ⥤ FinBoolAlg := -{ obj := λ X, of (order_dual X), map := λ X Y, bounded_lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } /-- The equivalence between `FinBoolAlg` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : FinBoolAlg ≌ FinBoolAlg := diff --git a/src/order/category/FinPartialOrder.lean b/src/order/category/FinPartialOrder.lean index 13186942f9072..78b7eae8b2fd9 100644 --- a/src/order/category/FinPartialOrder.lean +++ b/src/order/category/FinPartialOrder.lean @@ -64,7 +64,7 @@ instance has_forget_to_Fintype : has_forget₂ FinPartialOrder Fintype := /-- `order_dual` as a functor. -/ @[simps] def dual : FinPartialOrder ⥤ FinPartialOrder := -{ obj := λ X, of (order_dual X), map := λ X Y, order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } /-- The equivalence between `FinPartialOrder` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : FinPartialOrder ≌ FinPartialOrder := diff --git a/src/order/category/Lattice.lean b/src/order/category/Lattice.lean index ff1641b860e3d..3fa08568307f4 100644 --- a/src/order/category/Lattice.lean +++ b/src/order/category/Lattice.lean @@ -60,8 +60,7 @@ instance has_forget_to_PartialOrder : has_forget₂ Lattice PartialOrder := inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } /-- `order_dual` as a functor. -/ -@[simps] def dual : Lattice ⥤ Lattice := -{ obj := λ X, of (order_dual X), map := λ X Y, lattice_hom.dual } +@[simps] def dual : Lattice ⥤ Lattice := { obj := λ X, of Xᵒᵈ, map := λ X Y, lattice_hom.dual } /-- The equivalence between `Lattice` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : Lattice ≌ Lattice := diff --git a/src/order/category/LinearOrder.lean b/src/order/category/LinearOrder.lean index 97107e1468b93..e8d0bda96df25 100644 --- a/src/order/category/LinearOrder.lean +++ b/src/order/category/LinearOrder.lean @@ -49,7 +49,7 @@ instance has_forget_to_Lattice : has_forget₂ LinearOrder Lattice := /-- `order_dual` as a functor. -/ @[simps] def dual : LinearOrder ⥤ LinearOrder := -{ obj := λ X, of (order_dual X), map := λ X Y, order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } /-- The equivalence between `LinearOrder` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : LinearOrder ≌ LinearOrder := diff --git a/src/order/category/NonemptyFinLinOrd.lean b/src/order/category/NonemptyFinLinOrd.lean index 01d1e8e7b8b32..d281dff7e3125 100644 --- a/src/order/category/NonemptyFinLinOrd.lean +++ b/src/order/category/NonemptyFinLinOrd.lean @@ -42,7 +42,7 @@ instance ulift.nonempty_fin_lin_ord (α : Type u) [nonempty_fin_lin_ord α] : .. linear_order.lift equiv.ulift (equiv.injective _), .. ulift.fintype _ } -instance (α : Type*) [nonempty_fin_lin_ord α] : nonempty_fin_lin_ord (order_dual α) := +instance (α : Type*) [nonempty_fin_lin_ord α] : nonempty_fin_lin_ord αᵒᵈ := { ..order_dual.fintype α } /-- The category of nonempty finite linear orders. -/ @@ -78,7 +78,7 @@ between them. -/ /-- `order_dual` as a functor. -/ @[simps] def dual : NonemptyFinLinOrd ⥤ NonemptyFinLinOrd := -{ obj := λ X, of (order_dual X), map := λ X Y, order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } /-- The equivalence between `FinPartialOrder` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : NonemptyFinLinOrd ≌ NonemptyFinLinOrd := diff --git a/src/order/category/PartialOrder.lean b/src/order/category/PartialOrder.lean index 601fef3ed4de4..793f355aa681f 100644 --- a/src/order/category/PartialOrder.lean +++ b/src/order/category/PartialOrder.lean @@ -47,7 +47,7 @@ instance has_forget_to_Preorder : has_forget₂ PartialOrder Preorder := bundled /-- `order_dual` as a functor. -/ @[simps] def dual : PartialOrder ⥤ PartialOrder := -{ obj := λ X, of (order_dual X), map := λ X Y, order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } /-- The equivalence between `PartialOrder` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : PartialOrder ≌ PartialOrder := diff --git a/src/order/category/Preorder.lean b/src/order/category/Preorder.lean index 5517cfe0644f8..b28dd988e00b2 100644 --- a/src/order/category/Preorder.lean +++ b/src/order/category/Preorder.lean @@ -52,7 +52,7 @@ instance (α : Preorder) : preorder α := α.str /-- `order_dual` as a functor. -/ @[simps] def dual : Preorder ⥤ Preorder := -{ obj := λ X, of (order_dual X), map := λ X Y, order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } /-- The equivalence between `Preorder` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : Preorder ≌ Preorder := diff --git a/src/order/category/Semilattice.lean b/src/order/category/Semilattice.lean index 743f255f090eb..234a24b254b47 100644 --- a/src/order/category/Semilattice.lean +++ b/src/order/category/Semilattice.lean @@ -112,7 +112,7 @@ namespace SemilatticeSup /-- `order_dual` as a functor. -/ @[simps] def dual : SemilatticeSup ⥤ SemilatticeInf := -{ obj := λ X, SemilatticeInf.of (order_dual X), map := λ X Y, sup_bot_hom.dual } +{ obj := λ X, SemilatticeInf.of Xᵒᵈ, map := λ X Y, sup_bot_hom.dual } end SemilatticeSup @@ -127,7 +127,7 @@ namespace SemilatticeInf /-- `order_dual` as a functor. -/ @[simps] def dual : SemilatticeInf ⥤ SemilatticeSup := -{ obj := λ X, SemilatticeSup.of (order_dual X), map := λ X Y, inf_top_hom.dual } +{ obj := λ X, SemilatticeSup.of Xᵒᵈ, map := λ X Y, inf_top_hom.dual } end SemilatticeInf diff --git a/src/order/circular.lean b/src/order/circular.lean index 823316bc3056b..5255e3f546214 100644 --- a/src/order/circular.lean +++ b/src/order/circular.lean @@ -48,12 +48,12 @@ Some concrete circular orders one encounters in the wild are `zmod n` for `0 < n ## Notes -There's an unsolved diamond here. The instances `has_le α → has_btw (order_dual α)` and -`has_lt α → has_sbtw (order_dual α)` can each be inferred in two ways: -* `has_le α` → `has_btw α` → `has_btw (order_dual α)` vs - `has_le α` → `has_le (order_dual α)` → `has_btw (order_dual α)` -* `has_lt α` → `has_sbtw α` → `has_sbtw (order_dual α)` vs - `has_lt α` → `has_lt (order_dual α)` → `has_sbtw (order_dual α)` +There's an unsolved diamond on `order_dual α` here. The instances `has_le α → has_btw αᵒᵈ` and +`has_lt α → has_sbtw αᵒᵈ` can each be inferred in two ways: +* `has_le α` → `has_btw α` → `has_btw αᵒᵈ` vs + `has_le α` → `has_le αᵒᵈ` → `has_btw αᵒᵈ` +* `has_lt α` → `has_sbtw α` → `has_sbtw αᵒᵈ` vs + `has_lt α` → `has_lt αᵒᵈ` → `has_sbtw αᵒᵈ` The fields are propeq, but not defeq. It is temporarily fixed by turning the circularizing instances into definitions. @@ -385,10 +385,10 @@ def linear_order.to_circular_order (α : Type*) [linear_order α] : section order_dual -instance (α : Type*) [has_btw α] : has_btw (order_dual α) := ⟨λ a b c : α, btw c b a⟩ -instance (α : Type*) [has_sbtw α] : has_sbtw (order_dual α) := ⟨λ a b c : α, sbtw c b a⟩ +instance (α : Type*) [has_btw α] : has_btw αᵒᵈ := ⟨λ a b c : α, btw c b a⟩ +instance (α : Type*) [has_sbtw α] : has_sbtw αᵒᵈ := ⟨λ a b c : α, sbtw c b a⟩ -instance (α : Type*) [h : circular_preorder α] : circular_preorder (order_dual α) := +instance (α : Type*) [h : circular_preorder α] : circular_preorder αᵒᵈ := { btw_refl := btw_refl, btw_cyclic_left := λ a b c, btw_cyclic_right, sbtw_trans_left := λ a b c d habc hbdc, hbdc.trans_right habc, @@ -396,11 +396,11 @@ instance (α : Type*) [h : circular_preorder α] : circular_preorder (order_dual .. order_dual.has_btw α, .. order_dual.has_sbtw α } -instance (α : Type*) [circular_partial_order α] : circular_partial_order (order_dual α) := +instance (α : Type*) [circular_partial_order α] : circular_partial_order αᵒᵈ := { btw_antisymm := λ a b c habc hcba, @btw_antisymm α _ _ _ _ hcba habc, .. order_dual.circular_preorder α } -instance (α : Type*) [circular_order α] : circular_order (order_dual α) := +instance (α : Type*) [circular_order α] : circular_order αᵒᵈ := { btw_total := λ a b c, btw_total c b a, .. order_dual.circular_partial_order α } end order_dual diff --git a/src/order/compare.lean b/src/order/compare.lean index b792a468854a3..a9cefd8331663 100644 --- a/src/order/compare.lean +++ b/src/order/compare.lean @@ -127,7 +127,7 @@ by cases o₁; cases o₂; exact dec_trivial end ordering lemma order_dual.dual_compares [has_lt α] {a b : α} {o : ordering} : - @ordering.compares (order_dual α) _ o a b ↔ @ordering.compares α _ o b a := + @ordering.compares αᵒᵈ _ o a b ↔ @ordering.compares α _ o b a := by { cases o, exacts [iff.rfl, eq_comm, iff.rfl] } lemma cmp_compares [linear_order α] (a b : α) : (cmp a b).compares a b := @@ -146,7 +146,7 @@ begin end lemma order_dual.cmp_le_flip {α} [has_le α] [@decidable_rel α (≤)] (x y : α) : - @cmp_le (order_dual α) _ _ x y = cmp_le y x := rfl + @cmp_le αᵒᵈ _ _ x y = cmp_le y x := rfl /-- Generate a linear order structure from a preorder and `cmp` function. -/ def linear_order_of_compares [preorder α] (cmp : α → α → ordering) diff --git a/src/order/complete_boolean_algebra.lean b/src/order/complete_boolean_algebra.lean index 5270899009733..d084d45c86a3a 100644 --- a/src/order/complete_boolean_algebra.lean +++ b/src/order/complete_boolean_algebra.lean @@ -64,7 +64,7 @@ instance complete_distrib_lattice.to_coframe [complete_distrib_lattice α] : cof section frame variables [frame α] {s t : set α} {a b : α} -instance order_dual.coframe : coframe (order_dual α) := +instance order_dual.coframe : coframe αᵒᵈ := { infi_sup_le_sup_Inf := frame.inf_Sup_le_supr_inf, ..order_dual.complete_lattice α } lemma inf_Sup_eq : a ⊓ Sup s = ⨆ b ∈ s, a ⊓ b := @@ -122,37 +122,31 @@ end frame section coframe variables [coframe α] {s t : set α} {a b : α} -instance order_dual.frame : frame (order_dual α) := +instance order_dual.frame : frame αᵒᵈ := { inf_Sup_le_supr_inf := coframe.infi_sup_le_sup_Inf, ..order_dual.complete_lattice α } -theorem sup_Inf_eq : a ⊔ Inf s = (⨅ b ∈ s, a ⊔ b) := -@inf_Sup_eq (order_dual α) _ _ _ +lemma sup_Inf_eq : a ⊔ Inf s = ⨅ b ∈ s, a ⊔ b := @inf_Sup_eq αᵒᵈ _ _ _ +lemma Inf_sup_eq : Inf s ⊔ b = ⨅ a ∈ s, a ⊔ b := @Sup_inf_eq αᵒᵈ _ _ _ -theorem Inf_sup_eq : Inf s ⊔ b = (⨅ a ∈ s, a ⊔ b) := -@Sup_inf_eq (order_dual α) _ _ _ - -theorem infi_sup_eq (f : ι → α) (a : α) : (⨅ i, f i) ⊔ a = ⨅ i, f i ⊔ a := -@supr_inf_eq (order_dual α) _ _ _ _ - -theorem sup_infi_eq (a : α) (f : ι → α) : a ⊔ (⨅ i, f i) = ⨅ i, a ⊔ f i := -@inf_supr_eq (order_dual α) _ _ _ _ +lemma infi_sup_eq (f : ι → α) (a : α) : (⨅ i, f i) ⊔ a = ⨅ i, f i ⊔ a := @supr_inf_eq αᵒᵈ _ _ _ _ +lemma sup_infi_eq (a : α) (f : ι → α) : a ⊔ (⨅ i, f i) = ⨅ i, a ⊔ f i := @inf_supr_eq αᵒᵈ _ _ _ _ lemma binfi_sup_eq {f : Π i, κ i → α} (a : α) : (⨅ i j, f i j) ⊔ a = ⨅ i j, f i j ⊔ a := -@bsupr_inf_eq (order_dual α) _ _ _ _ _ +@bsupr_inf_eq αᵒᵈ _ _ _ _ _ lemma sup_binfi_eq {f : Π i, κ i → α} (a : α) : a ⊔ (⨅ i j, f i j) = ⨅ i j, a ⊔ f i j := -@inf_bsupr_eq (order_dual α) _ _ _ _ _ +@inf_bsupr_eq αᵒᵈ _ _ _ _ _ lemma infi_sup_infi {ι ι' : Type*} {f : ι → α} {g : ι' → α} : (⨅ i, f i) ⊔ (⨅ i, g i) = ⨅ i : ι × ι', f i.1 ⊔ g i.2 := -@supr_inf_supr (order_dual α) _ _ _ _ _ +@supr_inf_supr αᵒᵈ _ _ _ _ _ lemma binfi_sup_binfi {ι ι' : Type*} {f : ι → α} {g : ι' → α} {s : set ι} {t : set ι'} : (⨅ i ∈ s, f i) ⊔ (⨅ j ∈ t, g j) = ⨅ p ∈ s ×ˢ t, f (p : ι × ι').1 ⊔ g p.2 := -@bsupr_inf_bsupr (order_dual α) _ _ _ _ _ _ _ +@bsupr_inf_bsupr αᵒᵈ _ _ _ _ _ _ _ theorem Inf_sup_Inf : Inf s ⊔ Inf t = (⨅ p ∈ s ×ˢ t, (p : α × α).1 ⊔ p.2) := -@Sup_inf_Sup (order_dual α) _ _ _ +@Sup_inf_Sup αᵒᵈ _ _ _ instance pi.coframe {ι : Type*} {π : ι → Type*} [Π i, coframe (π i)] : coframe (Π i, π i) := { Inf := Inf, @@ -165,7 +159,7 @@ end coframe section complete_distrib_lattice variables [complete_distrib_lattice α] {a b : α} {s t : set α} -instance : complete_distrib_lattice (order_dual α) := { ..order_dual.frame, ..order_dual.coframe } +instance : complete_distrib_lattice αᵒᵈ := { ..order_dual.frame, ..order_dual.coframe } instance pi.complete_distrib_lattice {ι : Type*} {π : ι → Type*} [Π i, complete_distrib_lattice (π i)] : complete_distrib_lattice (Π i, π i) := diff --git a/src/order/complete_lattice.lean b/src/order/complete_lattice.lean index c299d80d2ebb6..e67a577f6f81f 100644 --- a/src/order/complete_lattice.lean +++ b/src/order/complete_lattice.lean @@ -67,8 +67,8 @@ def infi [has_Inf α] {ι} (s : ι → α) : α := Inf (range s) notation `⨆` binders `, ` r:(scoped f, supr f) := r notation `⨅` binders `, ` r:(scoped f, infi f) := r -instance (α) [has_Inf α] : has_Sup (order_dual α) := ⟨(Inf : set α → α)⟩ -instance (α) [has_Sup α] : has_Inf (order_dual α) := ⟨(Sup : set α → α)⟩ +instance (α) [has_Inf α] : has_Sup αᵒᵈ := ⟨(Inf : set α → α)⟩ +instance (α) [has_Sup α] : has_Inf αᵒᵈ := ⟨(Sup : set α → α)⟩ /-- Note that we rarely use `complete_semilattice_Sup` @@ -280,7 +280,7 @@ class complete_linear_order (α : Type*) extends complete_lattice α, namespace order_dual variable (α) -instance [complete_lattice α] : complete_lattice (order_dual α) := +instance [complete_lattice α] : complete_lattice αᵒᵈ := { le_Sup := @complete_lattice.Inf_le α _, Sup_le := @complete_lattice.le_Inf α _, Inf_le := @complete_lattice.le_Sup α _, @@ -288,7 +288,7 @@ instance [complete_lattice α] : complete_lattice (order_dual α) := .. order_dual.lattice α, ..order_dual.has_Sup α, ..order_dual.has_Inf α, .. order_dual.bounded_order α } -instance [complete_linear_order α] : complete_linear_order (order_dual α) := +instance [complete_linear_order α] : complete_linear_order αᵒᵈ := { .. order_dual.complete_lattice α, .. order_dual.linear_order α } end order_dual @@ -311,8 +311,7 @@ Sup_le $ λ b hb, le_inf (le_Sup hb.1) (le_Sup hb.2) theorem Inf_union {s t : set α} : Inf (s ∪ t) = Inf s ⊓ Inf t := ((is_glb_Inf s).union (is_glb_Inf t)).Inf_eq -theorem le_Inf_inter {s t : set α} : Inf s ⊔ Inf t ≤ Inf (s ∩ t) := -@Sup_inter_le (order_dual α) _ _ _ +theorem le_Inf_inter {s t : set α} : Inf s ⊔ Inf t ≤ Inf (s ∩ t) := @Sup_inter_le αᵒᵈ _ _ _ @[simp] theorem Sup_empty : Sup ∅ = (⊥ : α) := (@is_lub_empty α _ _).Sup_eq @@ -354,8 +353,7 @@ lemma eq_singleton_top_of_Inf_eq_top_of_nonempty {s : set α} (h_inf : Inf s = ⊤) (hne : s.nonempty) : s = {⊤} := by { rw set.eq_singleton_iff_nonempty_unique_mem, rw Inf_eq_top at h_inf, exact ⟨hne, h_inf⟩, } -@[simp] theorem Sup_eq_bot : Sup s = ⊥ ↔ (∀ a ∈ s, a = ⊥) := -@Inf_eq_top (order_dual α) _ _ +@[simp] lemma Sup_eq_bot : Sup s = ⊥ ↔ ∀ a ∈ s, a = ⊥ := @Inf_eq_top αᵒᵈ _ _ lemma eq_singleton_bot_of_Sup_eq_bot_of_nonempty {s : set α} (h_sup : Sup s = ⊥) (hne : s.nonempty) : s = {⊥} := @@ -380,7 +378,7 @@ See `cInf_eq_of_forall_ge_of_forall_gt_exists_lt` for a version in conditionally lattices. -/ theorem Inf_eq_of_forall_ge_of_forall_gt_exists_lt : (∀ a ∈ s, b ≤ a) → (∀ w, b < w → (∃ a ∈ s, a < w)) → Inf s = b := -@Sup_eq_of_forall_le_of_forall_lt_exists_gt (order_dual α) _ _ ‹_› +@Sup_eq_of_forall_le_of_forall_lt_exists_gt αᵒᵈ _ _ ‹_› end @@ -397,8 +395,7 @@ iff.intro let ⟨a, ha, h⟩ := h _ h' in lt_irrefl a $ lt_of_le_of_lt (le_Sup ha) h) -lemma Inf_eq_bot : Inf s = ⊥ ↔ (∀ b > ⊥, ∃ a ∈ s, a < b) := -@Sup_eq_top (order_dual α) _ _ +lemma Inf_eq_bot : Inf s = ⊥ ↔ ∀ b > ⊥, ∃ a ∈ s, a < b := @Sup_eq_top αᵒᵈ _ _ lemma lt_supr_iff {f : ι → α} : a < supr f ↔ ∃ i, a < f i := lt_Sup_iff.trans exists_range_iff lemma infi_lt_iff {f : ι → α} : infi f < a ↔ ∃ i, f i < a := Inf_lt_iff.trans exists_range_iff @@ -445,27 +442,26 @@ section has_Inf variables [has_Inf α] {f g : ι → α} lemma Inf_range : Inf (range f) = infi f := rfl -lemma Inf_eq_infi' (s : set α) : Inf s = ⨅ a : s, a := @Sup_eq_supr' (order_dual α) _ _ +lemma Inf_eq_infi' (s : set α) : Inf s = ⨅ a : s, a := @Sup_eq_supr' αᵒᵈ _ _ lemma infi_congr (h : ∀ i, f i = g i) : (⨅ i, f i) = ⨅ i, g i := congr_arg _ $ funext h lemma function.surjective.infi_comp {f : ι → ι'} (hf : surjective f) (g : ι' → α) : (⨅ x, g (f x)) = ⨅ y, g y := -@function.surjective.supr_comp (order_dual α) _ _ _ f hf g +@function.surjective.supr_comp αᵒᵈ _ _ _ f hf g lemma function.surjective.infi_congr {g : ι' → α} (h : ι → ι') (h1 : surjective h) (h2 : ∀ x, g (h x) = f x) : (⨅ x, f x) = ⨅ y, g y := -@function.surjective.supr_congr (order_dual α) _ _ _ _ _ h h1 h2 +@function.surjective.supr_congr αᵒᵈ _ _ _ _ _ h h1 h2 @[congr]lemma infi_congr_Prop {p q : Prop} {f₁ : p → α} {f₂ : q → α} (pq : p ↔ q) (f : ∀ x, f₁ (pq.mpr x) = f₂ x) : infi f₁ = infi f₂ := -@supr_congr_Prop (order_dual α) _ p q f₁ f₂ pq f +@supr_congr_Prop αᵒᵈ _ p q f₁ f₂ pq f lemma infi_range' (g : β → α) (f : ι → β) : (⨅ b : range f, g b) = ⨅ i, g (f i) := -@supr_range' (order_dual α) _ _ _ _ _ +@supr_range' αᵒᵈ _ _ _ _ _ -lemma Inf_image' {s : set β} {f : β → α} : Inf (f '' s) = (⨅ a : s, f a) := -@Sup_image' (order_dual α) _ _ _ _ +lemma Inf_image' {s : set β} {f : β → α} : Inf (f '' s) = ⨅ a : s, f a := @Sup_image' αᵒᵈ _ _ _ _ end has_Inf @@ -577,8 +573,7 @@ lemma lt_infi_iff : a < infi f ↔ ∃ b, a < b ∧ ∀ i, b ≤ f i := lemma Sup_eq_supr {s : set α} : Sup s = ⨆ a ∈ s, a := le_antisymm (Sup_le le_supr₂) (supr₂_le $ λ b, le_Sup) -lemma Inf_eq_infi {s : set α} : Inf s = ⨅ a ∈ s, a := -@Sup_eq_supr (order_dual α) _ _ +lemma Inf_eq_infi {s : set α} : Inf s = ⨅ a ∈ s, a := @Sup_eq_supr αᵒᵈ _ _ lemma Sup_sUnion (s : set (set α)) : Sup (⋃₀ s) = ⨆ t ∈ s, Sup t := begin @@ -591,7 +586,7 @@ begin exact supr_le (λ ts, Sup_le_Sup (λ x xt, ⟨t, ts, xt⟩)) } end -lemma Inf_sUnion (s : set (set α)) : Inf (⋃₀ s) = ⨅ t ∈ s, Inf t := @Sup_sUnion (order_dual α) _ _ +lemma Inf_sUnion (s : set (set α)) : Inf (⋃₀ s) = ⨅ t ∈ s, Inf t := @Sup_sUnion αᵒᵈ _ _ lemma monotone.le_map_supr [complete_lattice β] {f : α → β} (hf : monotone f) : (⨆ i, f (s i)) ≤ f (supr s) := @@ -685,14 +680,10 @@ theorem infi_const [nonempty ι] {a : α} : (⨅ b : ι, a) = a := by rw [infi, range_const, Inf_singleton] -- We will generalize this to conditionally complete lattices in `csupr_const`. -theorem supr_const [nonempty ι] {a : α} : (⨆ b : ι, a) = a := -@infi_const (order_dual α) _ _ _ _ +theorem supr_const [nonempty ι] {a : α} : (⨆ b : ι, a) = a := @infi_const αᵒᵈ _ _ _ _ -@[simp] lemma infi_top : (⨅ i:ι, ⊤ : α) = ⊤ := -top_unique $ le_infi $ λ i, le_rfl - -@[simp] lemma supr_bot : (⨆ i:ι, ⊥ : α) = ⊥ := -@infi_top (order_dual α) _ _ +@[simp] lemma infi_top : (⨅ i:ι, ⊤ : α) = ⊤ := top_unique $ le_infi $ λ i, le_rfl +@[simp] lemma supr_bot : (⨆ i:ι, ⊥ : α) = ⊥ := @infi_top αᵒᵈ _ _ @[simp] lemma supr_eq_bot : supr s = ⊥ ↔ ∀ i, s i = ⊥ := Sup_eq_bot.trans forall_range_iff @[simp] lemma infi_eq_top : infi s = ⊤ ↔ ∀ i, s i = ⊤ := Inf_eq_top.trans forall_range_iff @@ -730,7 +721,7 @@ See `cinfi_eq_of_forall_ge_of_forall_gt_exists_lt` for a version in conditionall lattices. -/ theorem infi_eq_of_forall_ge_of_forall_gt_exists_lt {f : ι → α} (h₁ : ∀ i, b ≤ f i) (h₂ : ∀ w, b < w → (∃ i, f i < w)) : (⨅ (i : ι), f i) = b := -@supr_eq_of_forall_le_of_forall_lt_exists_gt (order_dual α) _ _ _ ‹_› ‹_› ‹_› +@supr_eq_of_forall_le_of_forall_lt_exists_gt αᵒᵈ _ _ _ ‹_› ‹_› ‹_› lemma supr_eq_dif {p : Prop} [decidable p] (a : p → α) : (⨆ h : p, a h) = if h : p then a h else ⊥ := @@ -742,7 +733,7 @@ supr_eq_dif (λ _, a) lemma infi_eq_dif {p : Prop} [decidable p] (a : p → α) : (⨅ h : p, a h) = if h : p then a h else ⊤ := -@supr_eq_dif (order_dual α) _ _ _ _ +@supr_eq_dif αᵒᵈ _ _ _ _ lemma infi_eq_if {p : Prop} [decidable p] (a : α) : (⨅ h : p, a) = if p then a else ⊤ := @@ -755,8 +746,7 @@ le_antisymm (supr_le $ λ j, supr_mono $ λ i, le_supr _ _) -- TODO: should this be @[simp]? -lemma infi_comm {f : ι → ι' → α} : (⨅ i j, f i j) = ⨅ j i, f i j := -@supr_comm (order_dual α) _ _ _ _ +lemma infi_comm {f : ι → ι' → α} : (⨅ i j, f i j) = ⨅ j i, f i j := @supr_comm αᵒᵈ _ _ _ _ /- TODO: this is strange. In the proof below, we get exactly the desired among the equalities, but close does not get it. @@ -784,11 +774,11 @@ le_antisymm @[simp] theorem supr_supr_eq_left {b : β} {f : Π x : β, x = b → α} : (⨆ x, ⨆ h : x = b, f x h) = f b rfl := -@infi_infi_eq_left (order_dual α) _ _ _ _ +@infi_infi_eq_left αᵒᵈ _ _ _ _ @[simp] theorem supr_supr_eq_right {b : β} {f : Π x : β, b = x → α} : (⨆ x, ⨆ h : b = x, f x h) = f b rfl := -@infi_infi_eq_right (order_dual α) _ _ _ _ +@infi_infi_eq_right αᵒᵈ _ _ _ _ attribute [ematch] le_refl @@ -833,13 +823,13 @@ lemma inf_binfi {p : ι → Prop} {f : Π i (hi : p i), α} {a : α} (h : ∃ i, by simpa only [inf_comm] using binfi_inf h theorem supr_sup_eq {f g : ι → α} : (⨆ x, f x ⊔ g x) = (⨆ x, f x) ⊔ (⨆ x, g x) := -@infi_inf_eq (order_dual α) ι _ _ _ +@infi_inf_eq αᵒᵈ ι _ _ _ lemma supr_sup [h : nonempty ι] {f : ι → α} {a : α} : (⨆ x, f x) ⊔ a = (⨆ x, f x ⊔ a) := -@infi_inf (order_dual α) _ _ _ _ _ +@infi_inf αᵒᵈ _ _ _ _ _ lemma sup_supr [nonempty ι] {f : ι → α} {a : α} : a ⊔ (⨆ x, f x) = (⨆ x, a ⊔ f x) := -@inf_infi (order_dual α) _ _ _ _ _ +@inf_infi αᵒᵈ _ _ _ _ _ /-! ### `supr` and `infi` under `Prop` -/ @@ -856,7 +846,7 @@ lemma infi_true {s : true → α} : infi s = s trivial := infi_pos trivial le_antisymm (le_infi₂ $ λ i h, infi_le _ _) (le_infi $ λ ⟨i, h⟩, infi₂_le _ _) @[simp] lemma supr_exists {p : ι → Prop} {f : Exists p → α} : (⨆ x, f x) = ⨆ i h, f ⟨i, h⟩ := -@infi_exists (order_dual α) _ _ _ _ +@infi_exists αᵒᵈ _ _ _ _ lemma infi_and {p q : Prop} {s : p ∧ q → α} : infi s = ⨅ h₁ h₂, s ⟨h₁, h₂⟩ := le_antisymm (le_infi₂ $ λ i j, infi_le _ _) (le_infi $ λ ⟨i, h⟩, infi₂_le _ _) @@ -866,8 +856,7 @@ lemma infi_and' {p q : Prop} {s : p → q → α} : (⨅ (h₁ : p) (h₂ : q), s h₁ h₂) = ⨅ (h : p ∧ q), s h.1 h.2 := by { symmetry, exact infi_and } -lemma supr_and {p q : Prop} {s : p ∧ q → α} : supr s = ⨆ h₁ h₂, s ⟨h₁, h₂⟩ := -@infi_and (order_dual α) _ _ _ _ +lemma supr_and {p q : Prop} {s : p ∧ q → α} : supr s = ⨆ h₁ h₂, s ⟨h₁, h₂⟩ := @infi_and αᵒᵈ _ _ _ _ /-- The symmetric case of `supr_and`, useful for rewriting into a supremum over a conjunction -/ lemma supr_and' {p q : Prop} {s : p → q → α} : @@ -885,7 +874,7 @@ le_antisymm theorem supr_or {p q : Prop} {s : p ∨ q → α} : (⨆ x, s x) = (⨆ i, s (or.inl i)) ⊔ (⨆ j, s (or.inr j)) := -@infi_or (order_dual α) _ _ _ _ +@infi_or αᵒᵈ _ _ _ _ section @@ -906,7 +895,7 @@ supr_dite _ _ _ lemma infi_dite (f : Π i, p i → α) (g : Π i, ¬p i → α) : (⨅ i, if h : p i then f i h else g i h) = (⨅ i (h : p i), f i h) ⊓ (⨅ i (h : ¬ p i), g i h) := -supr_dite p (show Π i, p i → order_dual α, from f) g +supr_dite p (show Π i, p i → αᵒᵈ, from f) g lemma infi_ite (f g : ι → α) : (⨅ i, if p i then f i else g i) = (⨅ i (h : p i), f i) ⊓ (⨅ i (h : ¬ p i), g i) := @@ -918,13 +907,12 @@ lemma infi_range {g : β → α} {f : ι → β} : (⨅ b ∈ range f, g b) = by rw [← infi_subtype'', infi_range'] lemma supr_range {g : β → α} {f : ι → β} : (⨆ b ∈ range f, g b) = ⨆ i, g (f i) := -@infi_range (order_dual α) _ _ _ _ _ +@infi_range αᵒᵈ _ _ _ _ _ theorem Inf_image {s : set β} {f : β → α} : Inf (f '' s) = ⨅ a ∈ s, f a := by rw [← infi_subtype'', Inf_image'] -theorem Sup_image {s : set β} {f : β → α} : Sup (f '' s) = ⨆ a ∈ s, f a := -@Inf_image (order_dual α) _ _ _ _ +theorem Sup_image {s : set β} {f : β → α} : Sup (f '' s) = ⨆ a ∈ s, f a := @Inf_image αᵒᵈ _ _ _ _ /- ### supr and infi under set constructions @@ -959,19 +947,19 @@ by rw [(union_eq_self_of_subset_left h).symm, infi_union]; exact inf_le_left theorem supr_union {f : β → α} {s t : set β} : (⨆ x ∈ s ∪ t, f x) = (⨆ x ∈ s, f x) ⊔ (⨆ x ∈ t, f x) := -@infi_union (order_dual α) _ _ _ _ _ +@infi_union αᵒᵈ _ _ _ _ _ lemma supr_split (f : β → α) (p : β → Prop) : (⨆ i, f i) = (⨆ i (h : p i), f i) ⊔ (⨆ i (h : ¬ p i), f i) := -@infi_split (order_dual α) _ _ _ _ +@infi_split αᵒᵈ _ _ _ _ lemma supr_split_single (f : β → α) (i₀ : β) : (⨆ i, f i) = f i₀ ⊔ (⨆ i (h : i ≠ i₀), f i) := -@infi_split_single (order_dual α) _ _ _ _ +@infi_split_single αᵒᵈ _ _ _ _ theorem supr_le_supr_of_subset {f : β → α} {s t : set β} (h : s ⊆ t) : (⨆ x ∈ s, f x) ≤ (⨆ x ∈ t, f x) := -@infi_le_infi_of_subset (order_dual α) _ _ _ _ _ h +@infi_le_infi_of_subset αᵒᵈ _ _ _ _ _ h theorem infi_insert {f : β → α} {s : set β} {b : β} : (⨅ x ∈ insert b s, f x) = f b ⊓ (⨅ x ∈ s, f x) := @@ -988,7 +976,7 @@ theorem infi_pair {f : β → α} {a b : β} : (⨅ x ∈ ({a, b} : set β), f x by rw [infi_insert, infi_singleton] theorem supr_singleton {f : β → α} {b : β} : (⨆ x ∈ (singleton b : set β), f x) = f b := -@infi_singleton (order_dual α) _ _ _ _ +@infi_singleton αᵒᵈ _ _ _ _ theorem supr_pair {f : β → α} {a b : β} : (⨆ x ∈ ({a, b} : set β), f x) = f a ⊔ f b := by rw [supr_insert, supr_singleton] @@ -999,7 +987,7 @@ by rw [← Inf_image, ← Inf_image, ← image_comp] lemma supr_image {γ} {f : β → γ} {g : γ → α} {t : set β} : (⨆ c ∈ f '' t, g c) = (⨆ b ∈ t, g (f b)) := -@infi_image (order_dual α) _ _ _ _ _ _ +@infi_image αᵒᵈ _ _ _ _ _ _ theorem supr_extend_bot {e : ι → β} (he : injective e) (f : ι → α) : (⨆ j, extend e f ⊥ j) = ⨆ i, f i := @@ -1023,26 +1011,23 @@ theorem infi_of_empty' {α ι} [has_Inf α] [is_empty ι] (f : ι → α) : infi f = Inf (∅ : set α) := congr_arg Inf (range_eq_empty f) -theorem infi_of_empty [is_empty ι] (f : ι → α) : infi f = ⊤ := -@supr_of_empty (order_dual α) _ _ _ f +theorem infi_of_empty [is_empty ι] (f : ι → α) : infi f = ⊤ := @supr_of_empty αᵒᵈ _ _ _ f lemma supr_bool_eq {f : bool → α} : (⨆b:bool, f b) = f tt ⊔ f ff := by rw [supr, bool.range_eq, Sup_pair, sup_comm] -lemma infi_bool_eq {f : bool → α} : (⨅b:bool, f b) = f tt ⊓ f ff := -@supr_bool_eq (order_dual α) _ _ +lemma infi_bool_eq {f : bool → α} : (⨅b:bool, f b) = f tt ⊓ f ff := @supr_bool_eq αᵒᵈ _ _ lemma sup_eq_supr (x y : α) : x ⊔ y = ⨆ b : bool, cond b x y := by rw [supr_bool_eq, bool.cond_tt, bool.cond_ff] -lemma inf_eq_infi (x y : α) : x ⊓ y = ⨅ b : bool, cond b x y := -@sup_eq_supr (order_dual α) _ _ _ +lemma inf_eq_infi (x y : α) : x ⊓ y = ⨅ b : bool, cond b x y := @sup_eq_supr αᵒᵈ _ _ _ lemma is_glb_binfi {s : set β} {f : β → α} : is_glb (f '' s) (⨅ x ∈ s, f x) := by simpa only [range_comp, subtype.range_coe, infi_subtype'] using @is_glb_infi α s _ (f ∘ coe) theorem supr_subtype {p : ι → Prop} {f : subtype p → α} : (⨆ x, f x) = (⨆ i (h:p i), f ⟨i, h⟩) := -@infi_subtype (order_dual α) _ _ _ _ +@infi_subtype αᵒᵈ _ _ _ _ lemma supr_subtype' {p : ι → Prop} {f : ∀ i, p i → α} : (⨆ i (h : p i), f i h) = (⨆ x : subtype p, f x x.property) := @@ -1059,13 +1044,13 @@ theorem infi_sigma {p : β → Type*} {f : sigma p → α} : (⨅ x, f x) = (⨅ eq_of_forall_le_iff $ λ c, by simp only [le_infi_iff, sigma.forall] theorem supr_sigma {p : β → Type*} {f : sigma p → α} : (⨆ x, f x) = (⨆ i (h : p i), f ⟨i, h⟩) := -@infi_sigma (order_dual α) _ _ _ _ +@infi_sigma αᵒᵈ _ _ _ _ theorem infi_prod {γ : Type*} {f : β × γ → α} : (⨅ x, f x) = (⨅ i j, f (i, j)) := eq_of_forall_le_iff $ λ c, by simp only [le_infi_iff, prod.forall] theorem supr_prod {γ : Type*} {f : β × γ → α} : (⨆ x, f x) = (⨆ i j, f (i, j)) := -@infi_prod (order_dual α) _ _ _ _ +@infi_prod αᵒᵈ _ _ _ _ theorem infi_sum {γ : Type*} {f : β ⊕ γ → α} : (⨅ x, f x) = (⨅ i, f (sum.inl i)) ⊓ (⨅ j, f (sum.inr j)) := @@ -1073,7 +1058,7 @@ eq_of_forall_le_iff $ λ c, by simp only [le_inf_iff, le_infi_iff, sum.forall] theorem supr_sum {γ : Type*} {f : β ⊕ γ → α} : (⨆ x, f x) = (⨆ i, f (sum.inl i)) ⊔ (⨆ j, f (sum.inr j)) := -@infi_sum (order_dual α) _ _ _ _ +@infi_sum αᵒᵈ _ _ _ _ theorem supr_option (f : option β → α) : (⨆ o, f o) = f none ⊔ ⨆ b, f (option.some b) := @@ -1081,7 +1066,7 @@ eq_of_forall_ge_iff $ λ c, by simp only [supr_le_iff, sup_le_iff, option.forall theorem infi_option (f : option β → α) : (⨅ o, f o) = f none ⊓ ⨅ b, f (option.some b) := -@supr_option (order_dual α) _ _ _ +@supr_option αᵒᵈ _ _ _ /-- A version of `supr_option` useful for rewriting right-to-left. -/ lemma supr_option_elim (a : α) (f : β → α) : (⨆ o : option β, o.elim a f) = a ⊔ ⨆ b, f b := @@ -1089,7 +1074,7 @@ by simp [supr_option] /-- A version of `infi_option` useful for rewriting right-to-left. -/ lemma infi_option_elim (a : α) (f : β → α) : (⨅ o : option β, o.elim a f) = a ⊓ ⨅ b, f b := -@supr_option_elim (order_dual α) _ _ _ _ +@supr_option_elim αᵒᵈ _ _ _ _ /-- When taking the supremum of `f : ι → α`, the elements of `ι` on which `f` gives `⊥` can be dropped, without changing the result. -/ @@ -1109,7 +1094,7 @@ end /-- When taking the infimum of `f : ι → α`, the elements of `ι` on which `f` gives `⊤` can be dropped, without changing the result. -/ lemma infi_ne_top_subtype (f : ι → α) : (⨅ i : {i // f i ≠ ⊤}, f i) = ⨅ i, f i := -@supr_ne_bot_subtype (order_dual α) ι _ f +@supr_ne_bot_subtype αᵒᵈ ι _ f /-! ### `supr` and `infi` under `ℕ` @@ -1124,7 +1109,7 @@ begin end lemma infi_ge_eq_infi_nat_add {u : ℕ → α} (n : ℕ) : (⨅ i ≥ n, u i) = ⨅ i, u (i + n) := -@supr_ge_eq_supr_nat_add (order_dual α) _ _ _ +@supr_ge_eq_supr_nat_add αᵒᵈ _ _ _ lemma monotone.supr_nat_add {f : ℕ → α} (hf : monotone f) (k : ℕ) : (⨆ n, f (n + k)) = ⨆ n, f n := @@ -1148,7 +1133,7 @@ begin end lemma inf_infi_nat_succ (u : ℕ → α) : u 0 ⊓ (⨅ i, u (i + 1)) = ⨅ i, u i := -@sup_supr_nat_succ (order_dual α) _ u +@sup_supr_nat_succ αᵒᵈ _ u end @@ -1228,7 +1213,7 @@ by { change (∃ _, _) ↔ _, simp [-eq_iff_iff] } @[simp] lemma supr_apply {α : Type*} {β : α → Type*} {ι : Sort*} [Π i, has_Sup (β i)] {f : ι → Π a, β a} {a : α} : (⨆ i, f i) a = (⨆ i, f i a) := -@infi_apply α (λ i, order_dual (β i)) _ _ f a +@infi_apply α (λ i, (β i)ᵒᵈ) _ _ f a section complete_lattice variables [preorder α] [complete_lattice β] diff --git a/src/order/concept.lean b/src/order/concept.lean index c8d112fc01e15..0327b0d58dc01 100644 --- a/src/order/concept.lean +++ b/src/order/concept.lean @@ -255,7 +255,7 @@ instance : inhabited (concept α β r) := ⟨⊥⟩ @[simp] lemma swap_lt_swap_iff : c.swap < d.swap ↔ d < c := snd_ssubset_snd_iff /-- The dual of a concept lattice is isomorphic to the concept lattice of the dual context. -/ -@[simps] def swap_equiv : order_dual (concept α β r) ≃o concept β α (function.swap r) := +@[simps] def swap_equiv : (concept α β r)ᵒᵈ ≃o concept β α (function.swap r) := { to_fun := swap ∘ of_dual, inv_fun := to_dual ∘ swap, left_inv := swap_swap, diff --git a/src/order/conditionally_complete_lattice.lean b/src/order/conditionally_complete_lattice.lean index cbf266460a395..dc46cae9c8bf2 100644 --- a/src/order/conditionally_complete_lattice.lean +++ b/src/order/conditionally_complete_lattice.lean @@ -50,10 +50,10 @@ noncomputable instance {α : Type*} [has_Inf α] : has_Inf (with_top α) := ⟨λ S, if S ⊆ {⊤} then ⊤ else ↑(Inf (coe ⁻¹' S : set α))⟩ noncomputable instance {α : Type*} [has_Sup α] : has_Sup (with_bot α) := -⟨(@with_top.has_Inf (order_dual α) _).Inf⟩ +⟨(@with_top.has_Inf αᵒᵈ _).Inf⟩ noncomputable instance {α : Type*} [preorder α] [has_Inf α] : has_Inf (with_bot α) := -⟨(@with_top.has_Sup (order_dual α) _ _).Sup⟩ +⟨(@with_top.has_Sup αᵒᵈ _ _).Sup⟩ @[simp] theorem with_top.cInf_empty {α : Type*} [has_Inf α] : Inf (∅ : set (with_top α)) = ⊤ := @@ -182,8 +182,7 @@ end section order_dual -instance (α : Type*) [conditionally_complete_lattice α] : - conditionally_complete_lattice (order_dual α) := +instance (α : Type*) [conditionally_complete_lattice α] : conditionally_complete_lattice αᵒᵈ := { le_cSup := @conditionally_complete_lattice.cInf_le α _, cSup_le := @conditionally_complete_lattice.le_cInf α _, le_cInf := @conditionally_complete_lattice.cSup_le α _, @@ -193,7 +192,7 @@ instance (α : Type*) [conditionally_complete_lattice α] : ..order_dual.lattice α } instance (α : Type*) [conditionally_complete_linear_order α] : - conditionally_complete_linear_order (order_dual α) := + conditionally_complete_linear_order αᵒᵈ := { ..order_dual.conditionally_complete_lattice α, ..order_dual.linear_order α } @@ -254,7 +253,7 @@ is_glb_cInf (range_nonempty f) H lemma is_glb_cinfi_set {f : β → α} {s : set β} (H : bdd_below (f '' s)) (Hne : s.nonempty) : is_glb (f '' s) (⨅ i : s, f i) := -@is_lub_csupr_set (order_dual α) _ _ _ _ H Hne +@is_lub_csupr_set αᵒᵈ _ _ _ _ H Hne lemma is_lub.cSup_eq (H : is_lub s a) (ne : s.nonempty) : Sup s = a := (is_lub_cSup ne ⟨a, H.1⟩).unique H @@ -312,7 +311,7 @@ lemma not_mem_of_lt_cInf {x : α} {s : set α} (h : x < Inf s) (hs : bdd_below s λ hx, lt_irrefl _ (h.trans_le (cInf_le hs hx)) lemma not_mem_of_cSup_lt {x : α} {s : set α} (h : Sup s < x) (hs : bdd_above s) : x ∉ s := -@not_mem_of_lt_cInf (order_dual α) _ x s h hs +@not_mem_of_lt_cInf αᵒᵈ _ x s h hs /--Introduction rule to prove that `b` is the supremum of `s`: it suffices to check that `b` is larger than all elements of `s`, and that this is not the case of any `wb`. See `Inf_eq_of_forall_ge_of_forall_gt_exists_lt` for a version in complete lattices. -/ theorem cInf_eq_of_forall_ge_of_forall_gt_exists_lt (_ : s.nonempty) (_ : ∀a∈s, b ≤ a) (H : ∀w, b < w → (∃a∈s, a < w)) : Inf s = b := -@cSup_eq_of_forall_le_of_forall_lt_exists_gt (order_dual α) _ _ _ ‹_› ‹_› ‹_› +@cSup_eq_of_forall_le_of_forall_lt_exists_gt αᵒᵈ _ _ _ ‹_› ‹_› ‹_› /--b < Sup s when there is an element a in s with b < a, when s is bounded above. This is essentially an iff, except that the assumptions for the two implications are @@ -349,7 +348,7 @@ slightly different (one needs boundedness below for one direction, nonemptiness order for the other one), so we formulate separately the two implications, contrary to the complete_lattice case.-/ lemma cInf_lt_of_lt (_ : bdd_below s) (_ : a ∈ s) (_ : a < b) : Inf s < b := -@lt_cSup_of_lt (order_dual α) _ _ _ _ ‹_› ‹_› ‹_› +@lt_cSup_of_lt αᵒᵈ _ _ _ _ ‹_› ‹_› ‹_› /-- If all elements of a nonempty set `s` are less than or equal to all elements of a nonempty set `t`, then there exists an element between these sets. -/ @@ -387,7 +386,7 @@ theorem cSup_union (hs : bdd_above s) (sne : s.nonempty) (ht : bdd_above t) (tne that all sets are bounded below and nonempty.-/ theorem cInf_union (hs : bdd_below s) (sne : s.nonempty) (ht : bdd_below t) (tne : t.nonempty) : Inf (s ∪ t) = Inf s ⊓ Inf t := -@cSup_union (order_dual α) _ _ _ hs sne ht tne +@cSup_union αᵒᵈ _ _ _ hs sne ht tne /--The supremum of an intersection of two sets is bounded by the minimum of the suprema of each set, if all sets are bounded above and nonempty.-/ @@ -403,7 +402,7 @@ end infima of each set, if all sets are bounded below and nonempty.-/ theorem le_cInf_inter (_ : bdd_below s) (_ : bdd_below t) (hst : (s ∩ t).nonempty) : Inf s ⊔ Inf t ≤ Inf (s ∩ t) := -@cSup_inter_le (order_dual α) _ _ _ ‹_› ‹_› hst +@cSup_inter_le αᵒᵈ _ _ _ ‹_› ‹_› hst /-- The supremum of insert a s is the maximum of a and the supremum of s, if s is nonempty and bounded above.-/ @@ -413,7 +412,7 @@ theorem cSup_insert (hs : bdd_above s) (sne : s.nonempty) : Sup (insert a s) = a /-- The infimum of insert a s is the minimum of a and the infimum of s, if s is nonempty and bounded below.-/ theorem cInf_insert (hs : bdd_below s) (sne : s.nonempty) : Inf (insert a s) = a ⊓ Inf s := -@cSup_insert (order_dual α) _ _ _ hs sne +@cSup_insert αᵒᵈ _ _ _ hs sne @[simp] lemma cInf_Icc (h : a ≤ b) : Inf (Icc a b) = a := (is_glb_Icc h).cInf_eq (nonempty_Icc.2 h) @@ -478,41 +477,40 @@ lemma le_csupr_set {f : β → α} {s : set β} /--The indexed infimum of two functions are comparable if the functions are pointwise comparable-/ lemma cinfi_mono {f g : ι → α} (B : bdd_below (range f)) (H : ∀ x, f x ≤ g x) : infi f ≤ infi g := -@csupr_mono (order_dual α) _ _ _ _ B H +@csupr_mono αᵒᵈ _ _ _ _ B H /--The indexed minimum of a function is bounded below by a uniform lower bound-/ lemma le_cinfi [nonempty ι] {f : ι → α} {c : α} (H : ∀x, c ≤ f x) : c ≤ infi f := -@csupr_le (order_dual α) _ _ _ _ _ H +@csupr_le αᵒᵈ _ _ _ _ _ H /--The indexed infimum of a function is bounded above by the value taken at one point-/ lemma cinfi_le {f : ι → α} (H : bdd_below (range f)) (c : ι) : infi f ≤ f c := -@le_csupr (order_dual α) _ _ _ H c +@le_csupr αᵒᵈ _ _ _ H c lemma cinfi_le_of_le {f : ι → α} (H : bdd_below (range f)) (c : ι) (h : f c ≤ a) : infi f ≤ a := -@le_csupr_of_le (order_dual α) _ _ _ _ H c h +@le_csupr_of_le αᵒᵈ _ _ _ _ H c h lemma cinfi_set_le {f : β → α} {s : set β} (H : bdd_below (f '' s)) {c : β} (hc : c ∈ s) : (⨅ i : s, f i) ≤ f c := -@le_csupr_set (order_dual α) _ _ _ _ H _ hc +@le_csupr_set αᵒᵈ _ _ _ _ H _ hc @[simp] theorem csupr_const [hι : nonempty ι] {a : α} : (⨆ b:ι, a) = a := by rw [supr, range_const, cSup_singleton] -@[simp] theorem cinfi_const [hι : nonempty ι] {a : α} : (⨅ b:ι, a) = a := -@csupr_const (order_dual α) _ _ _ _ +@[simp] theorem cinfi_const [hι : nonempty ι] {a : α} : (⨅ b:ι, a) = a := @csupr_const αᵒᵈ _ _ _ _ @[simp] theorem supr_unique [unique ι] {s : ι → α} : (⨆ i, s i) = s default := have ∀ i, s i = s default := λ i, congr_arg s (unique.eq_default i), by simp only [this, csupr_const] @[simp] theorem infi_unique [unique ι] {s : ι → α} : (⨅ i, s i) = s default := -@supr_unique (order_dual α) _ _ _ _ +@supr_unique αᵒᵈ _ _ _ _ @[simp] lemma csupr_pos {p : Prop} {f : p → α} (hp : p) : (⨆ h : p, f h) = f hp := by haveI := unique_prop hp; exact supr_unique @[simp] lemma cinfi_pos {p : Prop} {f : p → α} (hp : p) : (⨅ h : p, f h) = f hp := -@csupr_pos (order_dual α) _ _ _ hp +@csupr_pos αᵒᵈ _ _ _ hp lemma csupr_set {s : set β} {f : β → α} : (⨆ x : s, f x) = Sup (f '' s) := begin @@ -523,8 +521,7 @@ begin simp_rw [subtype.coe_mk, exists_prop], end -lemma cinfi_set {s : set β} {f : β → α} : (⨅ x : s, f x) = Inf (f '' s) := -@csupr_set (order_dual α) _ _ _ _ +lemma cinfi_set {s : set β} {f : β → α} : (⨅ x : s, f x) = Inf (f '' s) := @csupr_set αᵒᵈ _ _ _ _ /--Introduction rule to prove that `b` is the supremum of `f`: it suffices to check that `b` is larger than `f i` for all `i`, and that this is not the case of any `wb`. See `infi_eq_of_forall_ge_of_forall_gt_exists_lt` for a version in complete lattices. -/ theorem cinfi_eq_of_forall_ge_of_forall_gt_exists_lt [nonempty ι] {f : ι → α} (h₁ : ∀ i, b ≤ f i) (h₂ : ∀ w, b < w → (∃ i, f i < w)) : (⨅ (i : ι), f i) = b := -@csupr_eq_of_forall_le_of_forall_lt_exists_gt (order_dual α) _ _ _ _ ‹_› ‹_› ‹_› +@csupr_eq_of_forall_le_of_forall_lt_exists_gt αᵒᵈ _ _ _ _ ‹_› ‹_› ‹_› /-- Nested intervals lemma: if `f` is a monotone sequence, `g` is an antitone sequence, and `f n ≤ g n` for all `n`, then `⨆ n, f n` belongs to all the intervals `[f n, g n]`. -/ @@ -601,25 +598,25 @@ lemma finset.nonempty.cSup_eq_max' {s : finset α} (h : s.nonempty) : Sup ↑s = eq_of_forall_ge_iff $ λ a, (cSup_le_iff s.bdd_above h.to_set).trans (s.max'_le_iff h).symm lemma finset.nonempty.cInf_eq_min' {s : finset α} (h : s.nonempty) : Inf ↑s = s.min' h := -@finset.nonempty.cSup_eq_max' (order_dual α) _ s h +@finset.nonempty.cSup_eq_max' αᵒᵈ _ s h lemma finset.nonempty.cSup_mem {s : finset α} (h : s.nonempty) : Sup (s : set α) ∈ s := by { rw h.cSup_eq_max', exact s.max'_mem _ } lemma finset.nonempty.cInf_mem {s : finset α} (h : s.nonempty) : Inf (s : set α) ∈ s := -@finset.nonempty.cSup_mem (order_dual α) _ _ h +@finset.nonempty.cSup_mem αᵒᵈ _ _ h lemma set.nonempty.cSup_mem (h : s.nonempty) (hs : finite s) : Sup s ∈ s := by { lift s to finset α using hs, exact finset.nonempty.cSup_mem h } lemma set.nonempty.cInf_mem (h : s.nonempty) (hs : finite s) : Inf s ∈ s := -@set.nonempty.cSup_mem (order_dual α) _ _ h hs +@set.nonempty.cSup_mem αᵒᵈ _ _ h hs lemma set.finite.cSup_lt_iff (hs : finite s) (h : s.nonempty) : Sup s < a ↔ ∀ x ∈ s, x < a := ⟨λ h x hx, (le_cSup hs.bdd_above hx).trans_lt h, λ H, H _ $ h.cSup_mem hs⟩ lemma set.finite.lt_cInf_iff (hs : finite s) (h : s.nonempty) : a < Inf s ↔ ∀ x ∈ s, a < x := -@set.finite.cSup_lt_iff (order_dual α) _ _ _ hs h +@set.finite.cSup_lt_iff αᵒᵈ _ _ _ hs h /-- When b < Sup s, there is an element a in s with b < a, if s is nonempty and the order is a linear order. -/ @@ -640,7 +637,7 @@ let ⟨_, ⟨i, rfl⟩, h⟩ := exists_lt_of_lt_cSup (range_nonempty f) h in ⟨ /--When Inf s < b, there is an element a in s with a < b, if s is nonempty and the order is a linear order.-/ lemma exists_lt_of_cInf_lt (hs : s.nonempty) (hb : Inf s < b) : ∃a∈s, a < b := -@exists_lt_of_lt_cSup (order_dual α) _ _ _ hs hb +@exists_lt_of_lt_cSup αᵒᵈ _ _ _ hs hb /-- Indexed version of the above lemma `exists_lt_of_cInf_lt` @@ -648,7 +645,7 @@ When `infi f < a`, there is an element `i` such that `f i < a`. -/ lemma exists_lt_of_cinfi_lt [nonempty ι] {f : ι → α} (h : infi f < a) : (∃i, f i < a) := -@exists_lt_of_lt_csupr (order_dual α) _ _ _ _ _ h +@exists_lt_of_lt_csupr αᵒᵈ _ _ _ _ _ h open function variables [is_well_order α (<)] @@ -859,11 +856,11 @@ cSup_le (nonempty.image f hs) (h_mono.mem_upper_bounds_image hB) lemma cInf_image_le {s : set α} {c : α} (hcs : c ∈ s) (h_bdd : bdd_below s) : Inf (f '' s) ≤ f c := -@le_cSup_image (order_dual α) (order_dual β) _ _ _ (λ x y hxy, h_mono hxy) _ _ hcs h_bdd +@le_cSup_image αᵒᵈ βᵒᵈ _ _ _ (λ x y hxy, h_mono hxy) _ _ hcs h_bdd lemma le_cInf_image {s : set α} (hs : s.nonempty) {B : α} (hB: B ∈ lower_bounds s) : f B ≤ Inf (f '' s) := -@cSup_image_le (order_dual α) (order_dual β) _ _ _ (λ x y hxy, h_mono hxy) _ hs _ hB +@cSup_image_le αᵒᵈ βᵒᵈ _ _ _ (λ x y hxy, h_mono hxy) _ hs _ hB end monotone @@ -977,7 +974,7 @@ end lemma inf'_eq_cInf_image [conditionally_complete_lattice β] (s : finset α) (H) (f : α → β) : s.inf' H f = Inf (f '' s) := -@sup'_eq_cSup_image _ (order_dual β) _ _ _ _ +@sup'_eq_cSup_image _ βᵒᵈ _ _ _ _ lemma sup'_id_eq_cSup [conditionally_complete_lattice α] (s : finset α) (H) : s.sup' H id = Sup s := @@ -985,7 +982,7 @@ by rw [sup'_eq_cSup_image s H, set.image_id] lemma inf'_id_eq_cInf [conditionally_complete_lattice α] (s : finset α) (H) : s.inf' H id = Inf s := -@sup'_id_eq_cSup (order_dual α) _ _ _ +@sup'_id_eq_cSup αᵒᵈ _ _ _ end finset @@ -1025,10 +1022,10 @@ gives a conditionally complete lattice -/ noncomputable instance with_bot.conditionally_complete_lattice {α : Type*} [conditionally_complete_lattice α] : conditionally_complete_lattice (with_bot α) := -{ le_cSup := (@with_top.conditionally_complete_lattice (order_dual α) _).cInf_le, - cSup_le := (@with_top.conditionally_complete_lattice (order_dual α) _).le_cInf, - cInf_le := (@with_top.conditionally_complete_lattice (order_dual α) _).le_cSup, - le_cInf := (@with_top.conditionally_complete_lattice (order_dual α) _).cSup_le, +{ le_cSup := (@with_top.conditionally_complete_lattice αᵒᵈ _).cInf_le, + cSup_le := (@with_top.conditionally_complete_lattice αᵒᵈ _).le_cInf, + cInf_le := (@with_top.conditionally_complete_lattice αᵒᵈ _).le_cSup, + le_cInf := (@with_top.conditionally_complete_lattice αᵒᵈ _).cSup_le, ..with_bot.lattice, ..with_bot.has_Sup, ..with_bot.has_Inf } @@ -1081,7 +1078,7 @@ inv_mul_le_iff_le_mul.mp $ le_cinfi $ λ hi, inv_mul_le_iff_le_mul.mpr $ H _ @[to_additive] lemma mul_csupr_le [covariant_class α α (*) (≤)] {a : α} {g : α} {h : ι → α} (H : ∀ j, g * h j ≤ a) : g * supr h ≤ a := -@le_mul_cinfi (order_dual α) _ _ _ _ _ _ _ _ H +@le_mul_cinfi αᵒᵈ _ _ _ _ _ _ _ _ H @[to_additive] lemma le_cinfi_mul [covariant_class α α (function.swap (*)) (≤)] {a : α} {g : ι → α} {h : α} @@ -1091,7 +1088,7 @@ mul_inv_le_iff_le_mul.mp $ le_cinfi $ λ gi, mul_inv_le_iff_le_mul.mpr $ H _ @[to_additive] lemma csupr_mul_le [covariant_class α α (function.swap (*)) (≤)] {a : α} {g : ι → α} {h : α} (H : ∀ i, g i * h ≤ a) : supr g * h ≤ a := -@le_cinfi_mul (order_dual α) _ _ _ _ _ _ _ _ H +@le_cinfi_mul αᵒᵈ _ _ _ _ _ _ _ _ H @[to_additive] lemma le_cinfi_mul_cinfi [covariant_class α α (*) (≤)] [covariant_class α α (function.swap (*)) (≤)] diff --git a/src/order/cover.lean b/src/order/cover.lean index 0fd3a7fa85f0b..bcb6c46c22f2a 100644 --- a/src/order/cover.lean +++ b/src/order/cover.lean @@ -84,7 +84,7 @@ lemma set.ord_connected.apply_wcovby_apply_iff (f : α ↪o β) (h : (range f).o @[simp] lemma to_dual_wcovby_to_dual_iff : to_dual b ⩿ to_dual a ↔ a ⩿ b := and_congr_right' $ forall_congr $ λ c, forall_swap -@[simp] lemma of_dual_wcovby_of_dual_iff {a b : order_dual α} : +@[simp] lemma of_dual_wcovby_of_dual_iff {a b : αᵒᵈ} : of_dual a ⩿ of_dual b ↔ b ⩿ a := and_congr_right' $ forall_congr $ λ c, forall_swap @@ -148,8 +148,7 @@ lemma densely_ordered_iff_forall_not_covby : densely_ordered α ↔ ∀ a b : α @[simp] lemma to_dual_covby_to_dual_iff : to_dual b ⋖ to_dual a ↔ a ⋖ b := and_congr_right' $ forall_congr $ λ c, forall_swap -@[simp] lemma of_dual_covby_of_dual_iff {a b : order_dual α} : - of_dual a ⋖ of_dual b ↔ b ⋖ a := +@[simp] lemma of_dual_covby_of_dual_iff {a b : αᵒᵈ} : of_dual a ⋖ of_dual b ↔ b ⋖ a := and_congr_right' $ forall_congr $ λ c, forall_swap alias to_dual_covby_to_dual_iff ↔ _ covby.to_dual diff --git a/src/order/directed.lean b/src/order/directed.lean index c7ace46f1073b..11f44fa57d220 100644 --- a/src/order/directed.lean +++ b/src/order/directed.lean @@ -122,12 +122,10 @@ directed_of (≤) a b lemma exists_le_le [has_le α] [is_directed α (swap (≤))] (a b : α) : ∃ c, c ≤ a ∧ c ≤ b := directed_of (swap (≤)) a b -instance order_dual.is_directed_ge [has_le α] [is_directed α (≤)] : - is_directed (order_dual α) (swap (≤)) := +instance order_dual.is_directed_ge [has_le α] [is_directed α (≤)] : is_directed αᵒᵈ (swap (≤)) := by assumption -instance order_dual.is_directed_le [has_le α] [is_directed α (swap (≤))] : - is_directed (order_dual α) (≤) := +instance order_dual.is_directed_le [has_le α] [is_directed α (swap (≤))] : is_directed αᵒᵈ (≤) := by assumption section preorder diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index 1c6947f521a19..f28d73a903a54 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -56,7 +56,7 @@ lemma disjoint_at_bot_principal_Ioi [preorder α] (x : α) : disjoint at_bot ( disjoint_of_disjoint_of_mem (Iic_disjoint_Ioi le_rfl) (Iic_mem_at_bot x) (mem_principal_self _) lemma disjoint_at_top_principal_Iio [preorder α] (x : α) : disjoint at_top (𝓟 (Iio x)) := -@disjoint_at_bot_principal_Ioi (order_dual α) _ _ +@disjoint_at_bot_principal_Ioi αᵒᵈ _ _ lemma disjoint_at_top_principal_Iic [preorder α] [no_max_order α] (x : α) : disjoint at_top (𝓟 (Iic x)) := @@ -64,7 +64,7 @@ disjoint_of_disjoint_of_mem (Iic_disjoint_Ioi le_rfl).symm (Ioi_mem_at_top x) (m lemma disjoint_at_bot_principal_Ici [preorder α] [no_min_order α] (x : α) : disjoint at_bot (𝓟 (Ici x)) := -@disjoint_at_top_principal_Iic (order_dual α) _ _ _ +@disjoint_at_top_principal_Iic αᵒᵈ _ _ _ lemma disjoint_at_bot_at_top [partial_order α] [nontrivial α] : disjoint (at_bot : filter α) at_top := @@ -91,13 +91,11 @@ lemma at_top_basis' [semilattice_sup α] (a : α) : ⟨λ ⟨x, _, hx⟩, ⟨x ⊔ a, le_sup_right, λ y hy, hx (le_trans le_sup_left hy)⟩, λ ⟨x, _, hx⟩, ⟨x, trivial, hx⟩⟩⟩ -lemma at_bot_basis [nonempty α] [semilattice_inf α] : - (@at_bot α _).has_basis (λ _, true) Iic := -@at_top_basis (order_dual α) _ _ +lemma at_bot_basis [nonempty α] [semilattice_inf α] : (@at_bot α _).has_basis (λ _, true) Iic := +@at_top_basis αᵒᵈ _ _ -lemma at_bot_basis' [semilattice_inf α] (a : α) : - (@at_bot α _).has_basis (λ x, x ≤ a) Iic := -@at_top_basis' (order_dual α) _ _ +lemma at_bot_basis' [semilattice_inf α] (a : α) : (@at_bot α _).has_basis (λ x, x ≤ a) Iic := +@at_top_basis' αᵒᵈ _ _ @[instance] lemma at_top_ne_bot [nonempty α] [semilattice_sup α] : ne_bot (at_top : filter α) := @@ -105,7 +103,7 @@ at_top_basis.ne_bot_iff.2 $ λ a _, nonempty_Ici @[instance] lemma at_bot_ne_bot [nonempty α] [semilattice_inf α] : ne_bot (at_bot : filter α) := -@at_top_ne_bot (order_dual α) _ _ +@at_top_ne_bot αᵒᵈ _ _ @[simp] lemma mem_at_top_sets [nonempty α] [semilattice_sup α] {s : set α} : @@ -115,7 +113,7 @@ at_top_basis.mem_iff.trans $ exists_congr $ λ _, exists_const _ @[simp] lemma mem_at_bot_sets [nonempty α] [semilattice_inf α] {s : set α} : s ∈ (at_bot : filter α) ↔ ∃a:α, ∀b≤a, b ∈ s := -@mem_at_top_sets (order_dual α) _ _ _ +@mem_at_top_sets αᵒᵈ _ _ _ @[simp] lemma eventually_at_top [semilattice_sup α] [nonempty α] {p : α → Prop} : @@ -173,7 +171,7 @@ le_antisymm (le_pure_iff.2 $ (eventually_ge_at_top ⊤).mono $ λ b, top_unique) (le_infi $ λ b, le_principal_iff.2 le_top) lemma order_bot.at_bot_eq (α) [partial_order α] [order_bot α] : (at_bot : filter α) = pure ⊥ := -@order_top.at_top_eq (order_dual α) _ _ +@order_top.at_top_eq αᵒᵈ _ _ @[nontriviality] lemma subsingleton.at_top_eq (α) [subsingleton α] [preorder α] : (at_top : filter α) = ⊤ := @@ -186,7 +184,7 @@ end @[nontriviality] lemma subsingleton.at_bot_eq (α) [subsingleton α] [preorder α] : (at_bot : filter α) = ⊤ := -@subsingleton.at_top_eq (order_dual α) _ _ +@subsingleton.at_top_eq αᵒᵈ _ _ lemma tendsto_at_top_pure [partial_order α] [order_top α] (f : α → β) : tendsto f at_top (pure $ f ⊤) := @@ -194,7 +192,7 @@ lemma tendsto_at_top_pure [partial_order α] [order_top α] (f : α → β) : lemma tendsto_at_bot_pure [partial_order α] [order_bot α] (f : α → β) : tendsto f at_bot (pure $ f ⊥) := -@tendsto_at_top_pure (order_dual α) _ _ _ _ +@tendsto_at_top_pure αᵒᵈ _ _ _ _ lemma eventually.exists_forall_of_at_top [semilattice_sup α] [nonempty α] {p : α → Prop} (h : ∀ᶠ x in at_top, p x) : ∃ a, ∀ b ≥ a, p b := @@ -210,7 +208,7 @@ by simp [at_top_basis.frequently_iff] lemma frequently_at_bot [semilattice_inf α] [nonempty α] {p : α → Prop} : (∃ᶠ x in at_bot, p x) ↔ (∀ a, ∃ b ≤ a, p b) := -@frequently_at_top (order_dual α) _ _ _ +@frequently_at_top αᵒᵈ _ _ _ lemma frequently_at_top' [semilattice_sup α] [nonempty α] [no_max_order α] {p : α → Prop} : (∃ᶠ x in at_top, p x) ↔ (∀ a, ∃ b > a, p b) := @@ -218,7 +216,7 @@ by simp [at_top_basis_Ioi.frequently_iff] lemma frequently_at_bot' [semilattice_inf α] [nonempty α] [no_min_order α] {p : α → Prop} : (∃ᶠ x in at_bot, p x) ↔ (∀ a, ∃ b < a, p b) := -@frequently_at_top' (order_dual α) _ _ _ _ +@frequently_at_top' αᵒᵈ _ _ _ _ lemma frequently.forall_exists_of_at_top [semilattice_sup α] [nonempty α] {p : α → Prop} (h : ∃ᶠ x in at_top, p x) : ∀ a, ∃ b ≥ a, p b := @@ -234,7 +232,7 @@ lemma map_at_top_eq [nonempty α] [semilattice_sup α] {f : α → β} : lemma map_at_bot_eq [nonempty α] [semilattice_inf α] {f : α → β} : at_bot.map f = (⨅a, 𝓟 $ f '' {a' | a' ≤ a}) := -@map_at_top_eq (order_dual α) _ _ _ _ +@map_at_top_eq αᵒᵈ _ _ _ _ lemma tendsto_at_top [preorder β] {m : α → β} {f : filter α} : tendsto m f at_top ↔ (∀b, ∀ᶠ a in f, b ≤ m a) := @@ -242,7 +240,7 @@ by simp only [at_top, tendsto_infi, tendsto_principal, mem_Ici] lemma tendsto_at_bot [preorder β] {m : α → β} {f : filter α} : tendsto m f at_bot ↔ (∀b, ∀ᶠ a in f, m a ≤ b) := -@tendsto_at_top α (order_dual β) _ m f +@tendsto_at_top α βᵒᵈ _ m f lemma tendsto_at_top_mono' [preorder β] (l : filter α) ⦃f₁ f₂ : α → β⦄ (h : f₁ ≤ᶠ[l] f₂) : tendsto f₁ l at_top → tendsto f₂ l at_top := @@ -251,7 +249,7 @@ assume h₁, tendsto_at_top.2 $ λ b, mp_mem (tendsto_at_top.1 h₁ b) lemma tendsto_at_bot_mono' [preorder β] (l : filter α) ⦃f₁ f₂ : α → β⦄ (h : f₁ ≤ᶠ[l] f₂) : tendsto f₂ l at_bot → tendsto f₁ l at_bot := -@tendsto_at_top_mono' _ (order_dual β) _ _ _ _ h +@tendsto_at_top_mono' _ βᵒᵈ _ _ _ _ h lemma tendsto_at_top_mono [preorder β] {l : filter α} {f g : α → β} (h : ∀ n, f n ≤ g n) : tendsto f l at_top → tendsto g l at_top := @@ -259,7 +257,7 @@ tendsto_at_top_mono' l $ eventually_of_forall h lemma tendsto_at_bot_mono [preorder β] {l : filter α} {f g : α → β} (h : ∀ n, f n ≤ g n) : tendsto g l at_bot → tendsto f l at_bot := -@tendsto_at_top_mono _ (order_dual β) _ _ _ _ h +@tendsto_at_top_mono _ βᵒᵈ _ _ _ _ h /-! ### Sequences @@ -271,7 +269,7 @@ by simp_rw [inf_ne_bot_iff_frequently_left, frequently_map, frequently_at_top]; lemma inf_map_at_bot_ne_bot_iff [semilattice_inf α] [nonempty α] {F : filter β} {u : α → β} : ne_bot (F ⊓ (map u at_bot)) ↔ ∀ U ∈ F, ∀ N, ∃ n ≤ N, u n ∈ U := -@inf_map_at_top_ne_bot_iff (order_dual α) _ _ _ _ _ +@inf_map_at_top_ne_bot_iff αᵒᵈ _ _ _ _ _ lemma extraction_of_frequently_at_top' {P : ℕ → Prop} (h : ∀ N, ∃ n > N, P n) : ∃ φ : ℕ → ℕ, strict_mono φ ∧ ∀ n, P (φ n) := @@ -327,7 +325,7 @@ end @[nolint ge_or_gt] -- see Note [nolint_ge] lemma exists_le_of_tendsto_at_bot [semilattice_sup α] [preorder β] {u : α → β} (h : tendsto u at_top at_bot) : ∀ a b, ∃ a' ≥ a, u a' ≤ b := -@exists_le_of_tendsto_at_top _ (order_dual β) _ _ _ h +@exists_le_of_tendsto_at_top _ βᵒᵈ _ _ _ h lemma exists_lt_of_tendsto_at_top [semilattice_sup α] [preorder β] [no_max_order β] {u : α → β} (h : tendsto u at_top at_top) (a : α) (b : β) : ∃ a' ≥ a, b < u a' := @@ -340,7 +338,7 @@ end @[nolint ge_or_gt] -- see Note [nolint_ge] lemma exists_lt_of_tendsto_at_bot [semilattice_sup α] [preorder β] [no_min_order β] {u : α → β} (h : tendsto u at_top at_bot) : ∀ a b, ∃ a' ≥ a, u a' < b := -@exists_lt_of_tendsto_at_top _ (order_dual β) _ _ _ _ h +@exists_lt_of_tendsto_at_top _ βᵒᵈ _ _ _ _ h /-- If `u` is a sequence which is unbounded above, @@ -376,7 +374,7 @@ then after any point, it reaches a value strictly smaller than all previous valu @[nolint ge_or_gt] -- see Note [nolint_ge] lemma low_scores [linear_order β] [no_min_order β] {u : ℕ → β} (hu : tendsto u at_top at_bot) : ∀ N, ∃ n ≥ N, ∀ k < n, u n < u k := -@high_scores (order_dual β) _ _ _ hu +@high_scores βᵒᵈ _ _ _ hu /-- If `u` is a sequence which is unbounded above, @@ -392,7 +390,7 @@ then it `frequently` reaches a value strictly smaller than all previous values. -/ lemma frequently_low_scores [linear_order β] [no_min_order β] {u : ℕ → β} (hu : tendsto u at_top at_bot) : ∃ᶠ n in at_top, ∀ k < n, u n < u k := -@frequently_high_scores (order_dual β) _ _ _ hu +@frequently_high_scores βᵒᵈ _ _ _ hu lemma strict_mono_subseq_of_tendsto_at_top {β : Type*} [linear_order β] [no_max_order β] @@ -419,7 +417,7 @@ tendsto_at_top_mono' l (hf.mono (λ x, le_add_of_nonneg_left)) hg lemma tendsto_at_bot_add_nonpos_left' (hf : ∀ᶠ x in l, f x ≤ 0) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_nonneg_left' _ (order_dual β) _ _ _ _ hf hg +@tendsto_at_top_add_nonneg_left' _ βᵒᵈ _ _ _ _ hf hg lemma tendsto_at_top_add_nonneg_left (hf : ∀ x, 0 ≤ f x) (hg : tendsto g l at_top) : tendsto (λ x, f x + g x) l at_top := @@ -427,7 +425,7 @@ tendsto_at_top_add_nonneg_left' (eventually_of_forall hf) hg lemma tendsto_at_bot_add_nonpos_left (hf : ∀ x, f x ≤ 0) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_nonneg_left _ (order_dual β) _ _ _ _ hf hg +@tendsto_at_top_add_nonneg_left _ βᵒᵈ _ _ _ _ hf hg lemma tendsto_at_top_add_nonneg_right' (hf : tendsto f l at_top) (hg : ∀ᶠ x in l, 0 ≤ g x) : tendsto (λ x, f x + g x) l at_top := @@ -435,7 +433,7 @@ tendsto_at_top_mono' l (monotone_mem (λ x, le_add_of_nonneg_right) hg) hf lemma tendsto_at_bot_add_nonpos_right' (hf : tendsto f l at_bot) (hg : ∀ᶠ x in l, g x ≤ 0) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_nonneg_right' _ (order_dual β) _ _ _ _ hf hg +@tendsto_at_top_add_nonneg_right' _ βᵒᵈ _ _ _ _ hf hg lemma tendsto_at_top_add_nonneg_right (hf : tendsto f l at_top) (hg : ∀ x, 0 ≤ g x) : tendsto (λ x, f x + g x) l at_top := @@ -443,7 +441,7 @@ tendsto_at_top_add_nonneg_right' hf (eventually_of_forall hg) lemma tendsto_at_bot_add_nonpos_right (hf : tendsto f l at_bot) (hg : ∀ x, g x ≤ 0) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_nonneg_right _ (order_dual β) _ _ _ _ hf hg +@tendsto_at_top_add_nonneg_right _ βᵒᵈ _ _ _ _ hf hg lemma tendsto_at_top_add (hf : tendsto f l at_top) (hg : tendsto g l at_top) : tendsto (λ x, f x + g x) l at_top := @@ -451,7 +449,7 @@ tendsto_at_top_add_nonneg_left' (tendsto_at_top.mp hf 0) hg lemma tendsto_at_bot_add (hf : tendsto f l at_bot) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add _ (order_dual β) _ _ _ _ hf hg +@tendsto_at_top_add _ βᵒᵈ _ _ _ _ hf hg lemma tendsto.nsmul_at_top (hf : tendsto f l at_top) {n : ℕ} (hn : 0 < n) : tendsto (λ x, n • f x) l at_top := @@ -462,7 +460,7 @@ calc y ≤ f x : hy lemma tendsto.nsmul_at_bot (hf : tendsto f l at_bot) {n : ℕ} (hn : 0 < n) : tendsto (λ x, n • f x) l at_bot := -@tendsto.nsmul_at_top α (order_dual β) _ l f hf n hn +@tendsto.nsmul_at_top α βᵒᵈ _ l f hf n hn lemma tendsto_bit0_at_top : tendsto bit0 (at_top : filter β) at_top := tendsto_at_top_add tendsto_id tendsto_id @@ -482,7 +480,7 @@ tendsto_at_top.2 $ assume b, (tendsto_at_top.1 hf (C + b)).mono (λ x, le_of_add lemma tendsto_at_bot_of_add_const_left (C : β) (hf : tendsto (λ x, C + f x) l at_bot) : tendsto f l at_bot := -@tendsto_at_top_of_add_const_left _ (order_dual β) _ _ _ C hf +@tendsto_at_top_of_add_const_left _ βᵒᵈ _ _ _ C hf lemma tendsto_at_top_of_add_const_right (C : β) (hf : tendsto (λ x, f x + C) l at_top) : tendsto f l at_top := @@ -490,7 +488,7 @@ tendsto_at_top.2 $ assume b, (tendsto_at_top.1 hf (b + C)).mono (λ x, le_of_add lemma tendsto_at_bot_of_add_const_right (C : β) (hf : tendsto (λ x, f x + C) l at_bot) : tendsto f l at_bot := -@tendsto_at_top_of_add_const_right _ (order_dual β) _ _ _ C hf +@tendsto_at_top_of_add_const_right _ βᵒᵈ _ _ _ C hf lemma tendsto_at_top_of_add_bdd_above_left' (C) (hC : ∀ᶠ x in l, f x ≤ C) (h : tendsto (λ x, f x + g x) l at_top) : @@ -501,7 +499,7 @@ tendsto_at_top_of_add_const_left C lemma tendsto_at_bot_of_add_bdd_below_left' (C) (hC : ∀ᶠ x in l, C ≤ f x) (h : tendsto (λ x, f x + g x) l at_bot) : tendsto g l at_bot := -@tendsto_at_top_of_add_bdd_above_left' _ (order_dual β) _ _ _ _ C hC h +@tendsto_at_top_of_add_bdd_above_left' _ βᵒᵈ _ _ _ _ C hC h lemma tendsto_at_top_of_add_bdd_above_left (C) (hC : ∀ x, f x ≤ C) : tendsto (λ x, f x + g x) l at_top → tendsto g l at_top := @@ -509,7 +507,7 @@ tendsto_at_top_of_add_bdd_above_left' C (univ_mem' hC) lemma tendsto_at_bot_of_add_bdd_below_left (C) (hC : ∀ x, C ≤ f x) : tendsto (λ x, f x + g x) l at_bot → tendsto g l at_bot := -@tendsto_at_top_of_add_bdd_above_left _ (order_dual β) _ _ _ _ C hC +@tendsto_at_top_of_add_bdd_above_left _ βᵒᵈ _ _ _ _ C hC lemma tendsto_at_top_of_add_bdd_above_right' (C) (hC : ∀ᶠ x in l, g x ≤ C) (h : tendsto (λ x, f x + g x) l at_top) : @@ -520,7 +518,7 @@ tendsto_at_top_of_add_const_right C lemma tendsto_at_bot_of_add_bdd_below_right' (C) (hC : ∀ᶠ x in l, C ≤ g x) (h : tendsto (λ x, f x + g x) l at_bot) : tendsto f l at_bot := -@tendsto_at_top_of_add_bdd_above_right' _ (order_dual β) _ _ _ _ C hC h +@tendsto_at_top_of_add_bdd_above_right' _ βᵒᵈ _ _ _ _ C hC h lemma tendsto_at_top_of_add_bdd_above_right (C) (hC : ∀ x, g x ≤ C) : tendsto (λ x, f x + g x) l at_top → tendsto f l at_top := @@ -528,7 +526,7 @@ tendsto_at_top_of_add_bdd_above_right' C (univ_mem' hC) lemma tendsto_at_bot_of_add_bdd_below_right (C) (hC : ∀ x, C ≤ g x) : tendsto (λ x, f x + g x) l at_bot → tendsto f l at_bot := -@tendsto_at_top_of_add_bdd_above_right _ (order_dual β) _ _ _ _ C hC +@tendsto_at_top_of_add_bdd_above_right _ βᵒᵈ _ _ _ _ C hC end ordered_cancel_add_comm_monoid @@ -543,7 +541,7 @@ lemma tendsto_at_top_add_left_of_le' (C : β) (hf : ∀ᶠ x in l, C ≤ f x) (h lemma tendsto_at_bot_add_left_of_ge' (C : β) (hf : ∀ᶠ x in l, f x ≤ C) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_left_of_le' _ (order_dual β) _ _ _ _ C hf hg +@tendsto_at_top_add_left_of_le' _ βᵒᵈ _ _ _ _ C hf hg lemma tendsto_at_top_add_left_of_le (C : β) (hf : ∀ x, C ≤ f x) (hg : tendsto g l at_top) : tendsto (λ x, f x + g x) l at_top := @@ -551,7 +549,7 @@ tendsto_at_top_add_left_of_le' l C (univ_mem' hf) hg lemma tendsto_at_bot_add_left_of_ge (C : β) (hf : ∀ x, f x ≤ C) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_left_of_le _ (order_dual β) _ _ _ _ C hf hg +@tendsto_at_top_add_left_of_le _ βᵒᵈ _ _ _ _ C hf hg lemma tendsto_at_top_add_right_of_le' (C : β) (hf : tendsto f l at_top) (hg : ∀ᶠ x in l, C ≤ g x) : tendsto (λ x, f x + g x) l at_top := @@ -560,7 +558,7 @@ lemma tendsto_at_top_add_right_of_le' (C : β) (hf : tendsto f l at_top) (hg : lemma tendsto_at_bot_add_right_of_ge' (C : β) (hf : tendsto f l at_bot) (hg : ∀ᶠ x in l, g x ≤ C) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_right_of_le' _ (order_dual β) _ _ _ _ C hf hg +@tendsto_at_top_add_right_of_le' _ βᵒᵈ _ _ _ _ C hf hg lemma tendsto_at_top_add_right_of_le (C : β) (hf : tendsto f l at_top) (hg : ∀ x, C ≤ g x) : tendsto (λ x, f x + g x) l at_top := @@ -568,7 +566,7 @@ tendsto_at_top_add_right_of_le' l C hf (univ_mem' hg) lemma tendsto_at_bot_add_right_of_ge (C : β) (hf : tendsto f l at_bot) (hg : ∀ x, g x ≤ C) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_right_of_le _ (order_dual β) _ _ _ _ C hf hg +@tendsto_at_top_add_right_of_le _ βᵒᵈ _ _ _ _ C hf hg lemma tendsto_at_top_add_const_left (C : β) (hf : tendsto f l at_top) : tendsto (λ x, C + f x) l at_top := @@ -576,7 +574,7 @@ tendsto_at_top_add_left_of_le' l C (univ_mem' $ λ _, le_refl C) hf lemma tendsto_at_bot_add_const_left (C : β) (hf : tendsto f l at_bot) : tendsto (λ x, C + f x) l at_bot := -@tendsto_at_top_add_const_left _ (order_dual β) _ _ _ C hf +@tendsto_at_top_add_const_left _ βᵒᵈ _ _ _ C hf lemma tendsto_at_top_add_const_right (C : β) (hf : tendsto f l at_top) : tendsto (λ x, f x + C) l at_top := @@ -584,7 +582,7 @@ tendsto_at_top_add_right_of_le' l C hf (univ_mem' $ λ _, le_refl C) lemma tendsto_at_bot_add_const_right (C : β) (hf : tendsto f l at_bot) : tendsto (λ x, f x + C) l at_bot := -@tendsto_at_top_add_const_right _ (order_dual β) _ _ _ C hf +@tendsto_at_top_add_const_right _ βᵒᵈ _ _ _ C hf lemma tendsto_neg_at_top_at_bot : tendsto (has_neg.neg : β → β) at_top at_bot := begin @@ -593,14 +591,14 @@ begin end lemma tendsto_neg_at_bot_at_top : tendsto (has_neg.neg : β → β) at_bot at_top := -@tendsto_neg_at_top_at_bot (order_dual β) _ +@tendsto_neg_at_top_at_bot βᵒᵈ _ lemma tendsto_at_top_iff_tends_to_neg_at_bot : tendsto f l at_top ↔ tendsto (-f) l at_bot := have hf : f = has_neg.neg ∘ -f, { ext, simp, }, ⟨tendsto_neg_at_top_at_bot.comp, λ h, hf.symm ▸ tendsto_neg_at_bot_at_top.comp h⟩ lemma tendsto_at_bot_iff_tends_to_neg_at_top : tendsto f l at_bot ↔ tendsto (-f) l at_top := -@tendsto_at_top_iff_tends_to_neg_at_bot α (order_dual β) _ l f +@tendsto_at_top_iff_tends_to_neg_at_bot α βᵒᵈ _ l f end ordered_group @@ -827,7 +825,7 @@ by simp only [tendsto_def, mem_at_top_sets]; refl lemma tendsto_at_bot' [nonempty α] [semilattice_inf α] {f : α → β} {l : filter β} : tendsto f at_bot l ↔ (∀s ∈ l, ∃a, ∀b≤a, f b ∈ s) := -@tendsto_at_top' (order_dual α) _ _ _ _ _ +@tendsto_at_top' αᵒᵈ _ _ _ _ _ theorem tendsto_at_top_principal [nonempty β] [semilattice_sup β] {f : β → α} {s : set α} : tendsto f at_top (𝓟 s) ↔ ∃N, ∀n≥N, f n ∈ s := @@ -835,7 +833,7 @@ by rw [tendsto_iff_comap, comap_principal, le_principal_iff, mem_at_top_sets]; r theorem tendsto_at_bot_principal [nonempty β] [semilattice_inf β] {f : β → α} {s : set α} : tendsto f at_bot (𝓟 s) ↔ ∃N, ∀n≤N, f n ∈ s := -@tendsto_at_top_principal _ (order_dual β) _ _ _ _ +@tendsto_at_top_principal _ βᵒᵈ _ _ _ _ /-- A function `f` grows to `+∞` independent of an order-preserving embedding `e`. -/ lemma tendsto_at_top_at_top [nonempty α] [semilattice_sup α] [preorder β] {f : α → β} : @@ -844,15 +842,15 @@ iff.trans tendsto_infi $ forall_congr $ assume b, tendsto_at_top_principal lemma tendsto_at_top_at_bot [nonempty α] [semilattice_sup α] [preorder β] {f : α → β} : tendsto f at_top at_bot ↔ ∀ (b : β), ∃ (i : α), ∀ (a : α), i ≤ a → f a ≤ b := -@tendsto_at_top_at_top α (order_dual β) _ _ _ f +@tendsto_at_top_at_top α βᵒᵈ _ _ _ f lemma tendsto_at_bot_at_top [nonempty α] [semilattice_inf α] [preorder β] {f : α → β} : tendsto f at_bot at_top ↔ ∀ (b : β), ∃ (i : α), ∀ (a : α), a ≤ i → b ≤ f a := -@tendsto_at_top_at_top (order_dual α) β _ _ _ f +@tendsto_at_top_at_top αᵒᵈ β _ _ _ f lemma tendsto_at_bot_at_bot [nonempty α] [semilattice_inf α] [preorder β] {f : α → β} : tendsto f at_bot at_bot ↔ ∀ (b : β), ∃ (i : α), ∀ (a : α), a ≤ i → f a ≤ b := -@tendsto_at_top_at_top (order_dual α) (order_dual β) _ _ _ f +@tendsto_at_top_at_top αᵒᵈ βᵒᵈ _ _ _ f lemma tendsto_at_top_at_top_of_monotone [preorder α] [preorder β] {f : α → β} (hf : monotone f) (h : ∀ b, ∃ a, b ≤ f a) : @@ -893,7 +891,7 @@ le_antisymm lemma comap_embedding_at_bot [preorder β] [preorder γ] {e : β → γ} (hm : ∀ b₁ b₂, e b₁ ≤ e b₂ ↔ b₁ ≤ b₂) (hu : ∀c, ∃b, e b ≤ c) : comap e at_bot = at_bot := -@comap_embedding_at_top (order_dual β) (order_dual γ) _ _ e (function.swap hm) hu +@comap_embedding_at_top βᵒᵈ γᵒᵈ _ _ e (function.swap hm) hu lemma tendsto_at_top_embedding [preorder β] [preorder γ] {f : α → β} {e : β → γ} {l : filter α} @@ -906,7 +904,7 @@ lemma tendsto_at_bot_embedding [preorder β] [preorder γ] {f : α → β} {e : β → γ} {l : filter α} (hm : ∀b₁ b₂, e b₁ ≤ e b₂ ↔ b₁ ≤ b₂) (hu : ∀c, ∃b, e b ≤ c) : tendsto (e ∘ f) l at_bot ↔ tendsto f l at_bot := -@tendsto_at_top_embedding α (order_dual β) (order_dual γ) _ _ f e l (function.swap hm) hu +@tendsto_at_top_embedding α βᵒᵈ γᵒᵈ _ _ f e l (function.swap hm) hu lemma tendsto_finset_range : tendsto finset.range at_top at_top := finset.range_mono.tendsto_at_top_at_top finset.exists_nat_subset_range @@ -959,7 +957,7 @@ end lemma prod_at_bot_at_bot_eq {β₁ β₂ : Type*} [semilattice_inf β₁] [semilattice_inf β₂] : (at_bot : filter β₁) ×ᶠ (at_bot : filter β₂) = (at_bot : filter (β₁ × β₂)) := -@prod_at_top_at_top_eq (order_dual β₁) (order_dual β₂) _ _ +@prod_at_top_at_top_eq β₁ᵒᵈ β₂ᵒᵈ _ _ lemma prod_map_at_top_eq {α₁ α₂ β₁ β₂ : Type*} [semilattice_sup β₁] [semilattice_sup β₂] (u₁ : β₁ → α₁) (u₂ : β₂ → α₂) : @@ -969,7 +967,7 @@ by rw [prod_map_map_eq, prod_at_top_at_top_eq, prod.map_def] lemma prod_map_at_bot_eq {α₁ α₂ β₁ β₂ : Type*} [semilattice_inf β₁] [semilattice_inf β₂] (u₁ : β₁ → α₁) (u₂ : β₂ → α₂) : (map u₁ at_bot) ×ᶠ (map u₂ at_bot) = map (prod.map u₁ u₂) at_bot := -@prod_map_at_top_eq _ _ (order_dual β₁) (order_dual β₂) _ _ _ _ +@prod_map_at_top_eq _ _ β₁ᵒᵈ β₂ᵒᵈ _ _ _ _ lemma tendsto.subseq_mem {F : filter α} {V : ℕ → set α} (h : ∀ n, V n ∈ F) {u : ℕ → α} (hu : tendsto u at_top F) : ∃ φ : ℕ → ℕ, strict_mono φ ∧ ∀ n, u (φ n) ∈ V n := @@ -1036,7 +1034,7 @@ end lemma eventually_at_bot_curry [semilattice_inf α] [semilattice_inf β] {p : α × β → Prop} (hp : ∀ᶠ (x : α × β) in filter.at_bot, p x) : ∀ᶠ k in at_bot, ∀ᶠ l in at_bot, p (k, l) := -@eventually_at_top_curry (order_dual α) (order_dual β) _ _ _ hp +@eventually_at_top_curry αᵒᵈ βᵒᵈ _ _ _ hp /-- A function `f` maps upwards closed sets (at_top sets) to upwards closed sets when it is a Galois insertion. The Galois "insertion" and "connection" is weakened to only require it to be an @@ -1056,7 +1054,7 @@ end lemma map_at_bot_eq_of_gc [semilattice_inf α] [semilattice_inf β] {f : α → β} (g : β → α) (b' : β) (hf : monotone f) (gc : ∀a, ∀b≤b', b ≤ f a ↔ g b ≤ a) (hgi : ∀b≤b', f (g b) ≤ b) : map f at_bot = at_bot := -@map_at_top_eq_of_gc (order_dual α) (order_dual β) _ _ _ _ _ hf.dual gc hgi +@map_at_top_eq_of_gc αᵒᵈ βᵒᵈ _ _ _ _ _ hf.dual gc hgi lemma map_coe_at_top_of_Ici_subset [semilattice_sup α] {a : α} {s : set α} (h : Ici a ⊆ s) : map (coe : s → α) at_top = at_top := @@ -1109,25 +1107,24 @@ by rw [← map_coe_Ici_at_top a, comap_map subtype.coe_injective] order. -/ @[simp] lemma map_coe_Iio_at_bot [semilattice_inf α] [no_min_order α] (a : α) : map (coe : Iio a → α) at_bot = at_bot := -@map_coe_Ioi_at_top (order_dual α) _ _ _ +@map_coe_Ioi_at_top αᵒᵈ _ _ _ /-- The `at_bot` filter for an open interval `Iio a` comes from the `at_bot` filter in the ambient order. -/ -lemma at_bot_Iio_eq [semilattice_inf α] (a : α) : - at_bot = comap (coe : Iio a → α) at_bot := -@at_top_Ioi_eq (order_dual α) _ _ +lemma at_bot_Iio_eq [semilattice_inf α] (a : α) : at_bot = comap (coe : Iio a → α) at_bot := +@at_top_Ioi_eq αᵒᵈ _ _ /-- The `at_bot` filter for an open interval `Iic a` comes from the `at_bot` filter in the ambient order. -/ @[simp] lemma map_coe_Iic_at_bot [semilattice_inf α] (a : α) : map (coe : Iic a → α) at_bot = at_bot := -@map_coe_Ici_at_top (order_dual α) _ _ +@map_coe_Ici_at_top αᵒᵈ _ _ /-- The `at_bot` filter for an open interval `Iic a` comes from the `at_bot` filter in the ambient order. -/ lemma at_bot_Iic_eq [semilattice_inf α] (a : α) : at_bot = comap (coe : Iic a → α) at_bot := -@at_top_Ici_eq (order_dual α) _ _ +@at_top_Ici_eq αᵒᵈ _ _ lemma tendsto_Ioi_at_top [semilattice_sup α] {a : α} {f : β → Ioi a} {l : filter β} : @@ -1224,7 +1221,7 @@ below, then `tendsto u at_bot at_bot`. -/ lemma tendsto_at_bot_at_bot_of_monotone' [preorder ι] [linear_order α] {u : ι → α} (h : monotone u) (H : ¬bdd_below (range u)) : tendsto u at_bot at_bot := -@tendsto_at_top_at_top_of_monotone' (order_dual ι) (order_dual α) _ _ _ h.dual H +@tendsto_at_top_at_top_of_monotone' ιᵒᵈ αᵒᵈ _ _ _ h.dual H lemma unbounded_of_tendsto_at_top [nonempty α] [semilattice_sup α] [preorder β] [no_max_order β] {f : α → β} (h : tendsto f at_top at_top) : @@ -1241,17 +1238,17 @@ end lemma unbounded_of_tendsto_at_bot [nonempty α] [semilattice_sup α] [preorder β] [no_min_order β] {f : α → β} (h : tendsto f at_top at_bot) : ¬ bdd_below (range f) := -@unbounded_of_tendsto_at_top _ (order_dual β) _ _ _ _ _ h +@unbounded_of_tendsto_at_top _ βᵒᵈ _ _ _ _ _ h lemma unbounded_of_tendsto_at_top' [nonempty α] [semilattice_inf α] [preorder β] [no_max_order β] {f : α → β} (h : tendsto f at_bot at_top) : ¬ bdd_above (range f) := -@unbounded_of_tendsto_at_top (order_dual α) _ _ _ _ _ _ h +@unbounded_of_tendsto_at_top αᵒᵈ _ _ _ _ _ _ h lemma unbounded_of_tendsto_at_bot' [nonempty α] [semilattice_inf α] [preorder β] [no_min_order β] {f : α → β} (h : tendsto f at_bot at_bot) : ¬ bdd_below (range f) := -@unbounded_of_tendsto_at_top (order_dual α) (order_dual β) _ _ _ _ _ h +@unbounded_of_tendsto_at_top αᵒᵈ βᵒᵈ _ _ _ _ _ h /-- If a monotone function `u : ι → α` tends to `at_top` along *some* non-trivial filter `l`, then it tends to `at_top` along `at_top`. -/ @@ -1265,7 +1262,7 @@ it tends to `at_bot` along `at_bot`. -/ lemma tendsto_at_bot_of_monotone_of_filter [preorder ι] [preorder α] {l : filter ι} {u : ι → α} (h : monotone u) [ne_bot l] (hu : tendsto u l at_bot) : tendsto u at_bot at_bot := -@tendsto_at_top_of_monotone_of_filter (order_dual ι) (order_dual α) _ _ _ _ h.dual _ hu +@tendsto_at_top_of_monotone_of_filter ιᵒᵈ αᵒᵈ _ _ _ _ h.dual _ hu lemma tendsto_at_top_of_monotone_of_subseq [preorder ι] [preorder α] {u : ι → α} {φ : ι' → ι} (h : monotone u) {l : filter ι'} [ne_bot l] diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 68c9530e6aa07..748cf89a8400b 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -351,7 +351,7 @@ show u ∈ (filter.mk_of_closure s hs).sets ↔ u ∈ (generate s).sets, from hs /-- Galois insertion from sets of sets into filters. -/ def gi_generate (α : Type*) : - @galois_insertion (set (set α)) (order_dual (filter α)) _ _ filter.generate filter.sets := + @galois_insertion (set (set α)) (filter α)ᵒᵈ _ _ filter.generate filter.sets := { gc := λ s f, sets_iff_generate, le_l_u := λ f u h, generate_sets.basic h, choice := λ s hs, filter.mk_of_closure s (le_antisymm hs $ sets_iff_generate.1 $ le_rfl), diff --git a/src/order/filter/cofinite.lean b/src/order/filter/cofinite.lean index 33f72e0f88abd..38f8908f2dec6 100644 --- a/src/order/filter/cofinite.lean +++ b/src/order/filter/cofinite.lean @@ -143,12 +143,12 @@ let ⟨a₀, _, ha₀⟩ := hf.exists_within_forall_le univ_nonempty in ⟨a₀, lemma filter.tendsto.exists_within_forall_ge [linear_order β] {s : set α} (hs : s.nonempty) {f : α → β} (hf : filter.tendsto f filter.cofinite filter.at_bot) : ∃ a₀ ∈ s, ∀ a ∈ s, f a ≤ f a₀ := -@filter.tendsto.exists_within_forall_le _ (order_dual β) _ _ hs _ hf +@filter.tendsto.exists_within_forall_le _ βᵒᵈ _ _ hs _ hf lemma filter.tendsto.exists_forall_ge [nonempty α] [linear_order β] {f : α → β} (hf : tendsto f cofinite at_bot) : ∃ a₀, ∀ a, f a ≤ f a₀ := -@filter.tendsto.exists_forall_le _ (order_dual β) _ _ _ hf +@filter.tendsto.exists_forall_le _ βᵒᵈ _ _ _ hf /-- For an injective function `f`, inverse images of finite sets are finite. See also `filter.comap_cofinite_le` and `function.injective.comap_cofinite_eq`. -/ diff --git a/src/order/filter/extr.lean b/src/order/filter/extr.lean index 160c0bec9ac22..937e748bcbce8 100644 --- a/src/order/filter/extr.lean +++ b/src/order/filter/extr.lean @@ -529,7 +529,7 @@ lemma filter.eventually_eq.is_max_filter_iff {α β : Type*} [preorder β] {f g lemma filter.eventually_le.is_min_filter {α β : Type*} [preorder β] {f g : α → β} {a : α} {l : filter α} (hle : f ≤ᶠ[l] g) (hfga : f a = g a) (h : is_min_filter f l a) : is_min_filter g l a := -@filter.eventually_le.is_max_filter _ (order_dual β) _ _ _ _ _ hle hfga h +@filter.eventually_le.is_max_filter _ βᵒᵈ _ _ _ _ _ hle hfga h lemma is_min_filter.congr {α β : Type*} [preorder β] {f g : α → β} {a : α} {l : filter α} (h : is_min_filter f l a) (heq : f =ᶠ[l] g) (hfga : f a = g a) : @@ -568,8 +568,7 @@ begin exact csupr_eq_of_forall_le_of_forall_lt_exists_gt (λ x, h x.prop) (λ w hw, ⟨⟨x₀, hx₀⟩, hw⟩), end -lemma is_min_on.infi_eq (hx₀ : x₀ ∈ s) (h : is_min_on f s x₀) : - (⨅ x : s, f x) = f x₀ := -@is_max_on.supr_eq (order_dual α) β _ _ _ _ hx₀ h +lemma is_min_on.infi_eq (hx₀ : x₀ ∈ s) (h : is_min_on f s x₀) : (⨅ x : s, f x) = f x₀ := +@is_max_on.supr_eq αᵒᵈ β _ _ _ _ hx₀ h end conditionally_complete_linear_order diff --git a/src/order/fixed_points.lean b/src/order/fixed_points.lean index cb7fecc3fda1c..5ce5fa03dd506 100644 --- a/src/order/fixed_points.lean +++ b/src/order/fixed_points.lean @@ -137,10 +137,8 @@ begin ... = a : ha end -lemma gfp_gfp (h : α →o α →o α) : - gfp (gfp.comp h) = gfp h.on_diag := -@lfp_lfp (order_dual α) _ $ (order_hom.dual_iso (order_dual α) - (order_dual α)).symm.to_order_embedding.to_order_hom.comp h.dual +lemma gfp_gfp (h : α →o α →o α) : gfp (gfp.comp h) = gfp h.on_diag := +@lfp_lfp αᵒᵈ _ $ (order_hom.dual_iso αᵒᵈ αᵒᵈ).symm.to_order_embedding.to_order_hom.comp h.dual end eqn diff --git a/src/order/grade.lean b/src/order/grade.lean index 1aa5a116184b6..9a8f467c7cb7f 100644 --- a/src/order/grade.lean +++ b/src/order/grade.lean @@ -14,7 +14,7 @@ This file defines graded orders, also known as ranked orders. A `𝕆`-graded order is an order `α` equipped with a distinguished "grade" function `α → 𝕆` which should be understood as giving the "height" of the elements. Usual graded orders are `ℕ`-graded, -cograded orders are `order_dual ℕ`-graded, but we can also grade by `ℤ`, and polytopes are naturally +cograded orders are `ℕᵒᵈ`-graded, but we can also grade by `ℤ`, and polytopes are naturally `fin n`-graded. Visually, `grade ℕ a` is the height of `a` in the Hasse diagram of `α`. @@ -178,26 +178,26 @@ instance preorder.to_grade_bounded_order : grade_bounded_order α α := /-! #### Dual -/ -instance [grade_order 𝕆 α] : grade_order (order_dual 𝕆) (order_dual α) := +instance [grade_order 𝕆 α] : grade_order 𝕆ᵒᵈ αᵒᵈ := { grade := to_dual ∘ grade 𝕆 ∘ of_dual, grade_strict_mono := grade_strict_mono.dual, covby_grade := λ a b h, (h.of_dual.grade _).to_dual } -instance [grade_max_order 𝕆 α] : grade_min_order (order_dual 𝕆) (order_dual α) := +instance [grade_max_order 𝕆 α] : grade_min_order 𝕆ᵒᵈ αᵒᵈ := { is_min_grade := λ _, is_max.grade _, ..order_dual.grade_order } -instance [grade_min_order 𝕆 α] : grade_max_order (order_dual 𝕆) (order_dual α) := +instance [grade_min_order 𝕆 α] : grade_max_order 𝕆ᵒᵈ αᵒᵈ := { is_max_grade := λ _, is_min.grade _, ..order_dual.grade_order } -instance [grade_bounded_order 𝕆 α] : grade_bounded_order (order_dual 𝕆) (order_dual α) := +instance [grade_bounded_order 𝕆 α] : grade_bounded_order 𝕆ᵒᵈ αᵒᵈ := { ..order_dual.grade_min_order, ..order_dual.grade_max_order } @[simp] lemma grade_to_dual [grade_order 𝕆 α] (a : α) : - grade (order_dual 𝕆) (to_dual a) = to_dual (grade 𝕆 a) := rfl -@[simp] lemma grade_of_dual [grade_order 𝕆 α] (a : order_dual α) : - grade 𝕆 (of_dual a) = of_dual (grade (order_dual 𝕆) a) := rfl + grade 𝕆ᵒᵈ (to_dual a) = to_dual (grade 𝕆 a) := rfl +@[simp] lemma grade_of_dual [grade_order 𝕆 α] (a : αᵒᵈ) : + grade 𝕆 (of_dual a) = of_dual (grade 𝕆ᵒᵈ a) := rfl /-! #### Lifting a graded order -/ diff --git a/src/order/hom/basic.lean b/src/order/hom/basic.lean index 1c646e6b3e071..5db3beb839b61 100644 --- a/src/order/hom/basic.lean +++ b/src/order/hom/basic.lean @@ -51,11 +51,9 @@ because the more bundled version usually does not work with dot notation. `α →o Π i, π i`; * `order_hom.pi_iso`: order isomorphism between `α →o Π i, π i` and `Π i, α →o π i`; * `order_hom.subtyle.val`: embedding `subtype.val : subtype p → α` as a bundled monotone map; -* `order_hom.dual`: reinterpret a monotone map `α →o β` as a monotone map - `order_dual α →o order_dual β`; -* `order_hom.dual_iso`: order isomorphism between `α →o β` and - `order_dual (order_dual α →o order_dual β)`; -* `order_iso.compl`: order isomorphism `α ≃o order_dual α` given by taking complements in a +* `order_hom.dual`: reinterpret a monotone map `α →o β` as a monotone map `αᵒᵈ →o βᵒᵈ`; +* `order_hom.dual_iso`: order isomorphism between `α →o β` and `(αᵒᵈ →o βᵒᵈ)ᵒᵈ`; +* `order_iso.compl`: order isomorphism `α ≃o αᵒᵈ` given by taking complements in a boolean algebra; We also define two functions to convert other bundled maps to `α →o β`: @@ -347,7 +345,7 @@ lemma order_hom_eq_id [subsingleton α] (g : α →o α) : g = order_hom.id := subsingleton.elim _ _ /-- Reinterpret a bundled monotone function as a monotone function between dual orders. -/ -@[simps] protected def dual : (α →o β) ≃ (order_dual α →o order_dual β) := +@[simps] protected def dual : (α →o β) ≃ (αᵒᵈ →o βᵒᵈ) := { to_fun := λ f, ⟨order_dual.to_dual ∘ f ∘ order_dual.of_dual, f.mono.dual⟩, inv_fun := λ f, ⟨order_dual.of_dual ∘ f ∘ order_dual.to_dual, f.mono.dual⟩, left_inv := λ f, ext _ _ rfl, @@ -357,13 +355,11 @@ subsingleton.elim _ _ @[simp] lemma dual_comp (g : β →o γ) (f : α →o β) : (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : order_hom.dual.symm order_hom.id = (order_hom.id : α →o α) := rfl -@[simp] lemma symm_dual_comp (g : order_dual β →o order_dual γ) - (f : order_dual α →o order_dual β) : +@[simp] lemma symm_dual_comp (g : βᵒᵈ →o γᵒᵈ) (f : αᵒᵈ →o βᵒᵈ) : order_hom.dual.symm (g.comp f) = (order_hom.dual.symm g).comp (order_hom.dual.symm f) := rfl /-- `order_hom.dual` as an order isomorphism. -/ -def dual_iso (α β : Type*) [preorder α] [preorder β] : - (α →o β) ≃o order_dual (order_dual α →o order_dual β) := +def dual_iso (α β : Type*) [preorder α] [preorder β] : (α →o β) ≃o (αᵒᵈ →o βᵒᵈ)ᵒᵈ := { to_equiv := order_hom.dual.trans order_dual.to_dual, map_rel_iff' := λ f g, iff.rfl } @@ -412,7 +408,7 @@ protected theorem is_well_order [is_well_order β (<)] : is_well_order α (<) := f.lt_embedding.is_well_order /-- An order embedding is also an order embedding between dual orders. -/ -protected def dual : order_dual α ↪o order_dual β := +protected def dual : αᵒᵈ ↪o βᵒᵈ := ⟨f.to_embedding, λ a b, f.map_rel_iff⟩ /-- @@ -579,7 +575,7 @@ def prod_comm : (α × β) ≃o (β × α) := variables (α) /-- The order isomorphism between a type and its double dual. -/ -def dual_dual : α ≃o order_dual (order_dual α) := refl α +def dual_dual : α ≃o αᵒᵈᵒᵈ := refl α @[simp] lemma coe_dual_dual : ⇑(dual_dual α) = to_dual ∘ to_dual := rfl @[simp] lemma coe_dual_dual_symm : ⇑(dual_dual α).symm = of_dual ∘ of_dual := rfl @@ -587,9 +583,7 @@ def dual_dual : α ≃o order_dual (order_dual α) := refl α variables {α} @[simp] lemma dual_dual_apply (a : α) : dual_dual α a = to_dual (to_dual a) := rfl - -@[simp] lemma dual_dual_symm_apply (a : order_dual (order_dual α)) : - (dual_dual α).symm a = of_dual (of_dual a) := rfl +@[simp] lemma dual_dual_symm_apply (a : αᵒᵈᵒᵈ) : (dual_dual α).symm a = of_dual (of_dual a) := rfl end has_le @@ -715,8 +709,8 @@ lemma order_iso_of_surjective_self_symm_apply (b : β) : end strict_mono /-- An order isomorphism is also an order isomorphism between dual orders. -/ -protected def order_iso.dual [has_le α] [has_le β] (f : α ≃o β) : - order_dual α ≃o order_dual β := ⟨f.to_equiv, λ _ _, f.map_rel_iff⟩ +protected def order_iso.dual [has_le α] [has_le β] (f : α ≃o β) : αᵒᵈ ≃o βᵒᵈ := +⟨f.to_equiv, λ _ _, f.map_rel_iff⟩ section lattice_isos @@ -774,8 +768,7 @@ f.dual.map_inf x y namespace with_bot /-- Taking the dual then adding `⊥` is the same as adding `⊤` then taking the dual. -/ -protected def to_dual_top [has_le α] : with_bot (order_dual α) ≃o order_dual (with_top α) := -order_iso.refl _ +protected def to_dual_top [has_le α] : with_bot αᵒᵈ ≃o (with_top α)ᵒᵈ := order_iso.refl _ @[simp] lemma to_dual_top_coe [has_le α] (a : α) : with_bot.to_dual_top ↑(to_dual a) = to_dual (a : with_top α) := rfl @@ -787,8 +780,7 @@ end with_bot namespace with_top /-- Taking the dual then adding `⊤` is the same as adding `⊥` then taking the dual. -/ -protected def to_dual_bot [has_le α] : with_top (order_dual α) ≃o order_dual (with_bot α) := -order_iso.refl _ +protected def to_dual_bot [has_le α] : with_top αᵒᵈ ≃o (with_bot α)ᵒᵈ := order_iso.refl _ @[simp] lemma to_dual_bot_coe [has_le α] (a : α) : with_top.to_dual_bot ↑(to_dual a) = to_dual (a : with_bot α) := rfl @@ -873,7 +865,7 @@ variables (α) [boolean_algebra α] /-- Taking complements as an order isomorphism to the order dual. -/ @[simps] -def order_iso.compl : α ≃o order_dual α := +def order_iso.compl : α ≃o αᵒᵈ := { to_fun := order_dual.to_dual ∘ compl, inv_fun := compl ∘ order_dual.of_dual, left_inv := compl_compl, diff --git a/src/order/hom/bounded.lean b/src/order/hom/bounded.lean index d3c8d3bf916f9..e35a40a93b5e4 100644 --- a/src/order/hom/bounded.lean +++ b/src/order/hom/bounded.lean @@ -436,7 +436,7 @@ namespace top_hom variables [has_le α] [order_top α] [has_le β] [order_top β] [has_le γ] [order_top γ] /-- Reinterpret a top homomorphism as a bot homomorphism between the dual lattices. -/ -@[simps] protected def dual : top_hom α β ≃ bot_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : top_hom α β ≃ bot_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f, f.map_top'⟩, inv_fun := λ f, ⟨f, f.map_bot'⟩, left_inv := λ f, top_hom.ext $ λ _, rfl, @@ -447,8 +447,7 @@ variables [has_le α] [order_top α] [has_le β] [order_top β] [has_le γ] [ord (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : top_hom.dual.symm (bot_hom.id _) = top_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : bot_hom (order_dual β) (order_dual γ)) - (f : bot_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : bot_hom βᵒᵈ γᵒᵈ) (f : bot_hom αᵒᵈ βᵒᵈ) : top_hom.dual.symm (g.comp f) = (top_hom.dual.symm g).comp (top_hom.dual.symm f) := rfl end top_hom @@ -457,7 +456,7 @@ namespace bot_hom variables [has_le α] [order_bot α] [has_le β] [order_bot β] [has_le γ] [order_bot γ] /-- Reinterpret a bot homomorphism as a top homomorphism between the dual lattices. -/ -@[simps] protected def dual : bot_hom α β ≃ top_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : bot_hom α β ≃ top_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f, f.map_bot'⟩, inv_fun := λ f, ⟨f, f.map_top'⟩, left_inv := λ f, bot_hom.ext $ λ _, rfl, @@ -468,8 +467,7 @@ variables [has_le α] [order_bot α] [has_le β] [order_bot β] [has_le γ] [ord (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : bot_hom.dual.symm (top_hom.id _) = bot_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : top_hom (order_dual β) (order_dual γ)) - (f : top_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : top_hom βᵒᵈ γᵒᵈ) (f : top_hom αᵒᵈ βᵒᵈ) : bot_hom.dual.symm (g.comp f) = (bot_hom.dual.symm g).comp (bot_hom.dual.symm f) := rfl end bot_hom @@ -480,8 +478,7 @@ variables [preorder α] [bounded_order α] [preorder β] [bounded_order β] [pre /-- Reinterpret a bounded order homomorphism as a bounded order homomorphism between the dual orders. -/ -@[simps] protected def dual : - bounded_order_hom α β ≃ bounded_order_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : bounded_order_hom α β ≃ bounded_order_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_order_hom.dual, f.map_bot', f.map_top'⟩, inv_fun := λ f, ⟨order_hom.dual.symm f.to_order_hom, f.map_bot', f.map_top'⟩, left_inv := λ f, ext $ λ a, rfl, @@ -493,8 +490,7 @@ orders. -/ @[simp] lemma symm_dual_id : bounded_order_hom.dual.symm (bounded_order_hom.id _) = bounded_order_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : bounded_order_hom (order_dual β) (order_dual γ)) - (f : bounded_order_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : bounded_order_hom βᵒᵈ γᵒᵈ) (f : bounded_order_hom αᵒᵈ βᵒᵈ) : bounded_order_hom.dual.symm (g.comp f) = (bounded_order_hom.dual.symm g).comp (bounded_order_hom.dual.symm f) := rfl diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index ccec7e6f74eba..596d043ba064f 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -484,7 +484,7 @@ namespace Sup_hom variables [has_Sup α] [has_Sup β] [has_Sup γ] /-- Reinterpret a `⨆`-homomorphism as an `⨅`-homomorphism between the dual orders. -/ -@[simps] protected def dual : Sup_hom α β ≃ Inf_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : Sup_hom α β ≃ Inf_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨to_dual ∘ f ∘ of_dual, f.map_Sup'⟩, inv_fun := λ f, ⟨of_dual ∘ f ∘ to_dual, f.map_Inf'⟩, left_inv := λ f, Sup_hom.ext $ λ a, rfl, @@ -495,8 +495,7 @@ variables [has_Sup α] [has_Sup β] [has_Sup γ] (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : Sup_hom.dual.symm (Inf_hom.id _) = Sup_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : Inf_hom (order_dual β) (order_dual γ)) - (f : Inf_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : Inf_hom βᵒᵈ γᵒᵈ) (f : Inf_hom αᵒᵈ βᵒᵈ) : Sup_hom.dual.symm (g.comp f) = (Sup_hom.dual.symm g).comp (Sup_hom.dual.symm f) := rfl end Sup_hom @@ -505,7 +504,7 @@ namespace Inf_hom variables [has_Inf α] [has_Inf β] [has_Inf γ] /-- Reinterpret an `⨅`-homomorphism as a `⨆`-homomorphism between the dual orders. -/ -@[simps] protected def dual : Inf_hom α β ≃ Sup_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : Inf_hom α β ≃ Sup_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, { to_fun := to_dual ∘ f ∘ of_dual, map_Sup' := λ _, congr_arg to_dual (map_Inf f _) }, inv_fun := λ f, { to_fun := of_dual ∘ f ∘ to_dual, @@ -518,8 +517,7 @@ variables [has_Inf α] [has_Inf β] [has_Inf γ] (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : Inf_hom.dual.symm (Sup_hom.id _) = Inf_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : Sup_hom (order_dual β) (order_dual γ)) - (f : Sup_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : Sup_hom βᵒᵈ γᵒᵈ) (f : Sup_hom αᵒᵈ βᵒᵈ) : Inf_hom.dual.symm (g.comp f) = (Inf_hom.dual.symm g).comp (Inf_hom.dual.symm f) := rfl end Inf_hom @@ -529,8 +527,7 @@ variables [complete_lattice α] [complete_lattice β] [complete_lattice γ] /-- Reinterpret a complete lattice homomorphism as a complete lattice homomorphism between the dual lattices. -/ -@[simps] protected def dual : - complete_lattice_hom α β ≃ complete_lattice_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : complete_lattice_hom α β ≃ complete_lattice_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_Sup_hom.dual, f.map_Inf'⟩, inv_fun := λ f, ⟨f.to_Sup_hom.dual, f.map_Inf'⟩, left_inv := λ f, ext $ λ a, rfl, @@ -542,8 +539,7 @@ lattices. -/ @[simp] lemma symm_dual_id : complete_lattice_hom.dual.symm (complete_lattice_hom.id _) = complete_lattice_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : complete_lattice_hom (order_dual β) (order_dual γ)) - (f : complete_lattice_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : complete_lattice_hom βᵒᵈ γᵒᵈ) (f : complete_lattice_hom αᵒᵈ βᵒᵈ) : complete_lattice_hom.dual.symm (g.comp f) = (complete_lattice_hom.dual.symm g).comp (complete_lattice_hom.dual.symm f) := rfl diff --git a/src/order/hom/lattice.lean b/src/order/hom/lattice.lean index 1721e0128c3c9..214fb509677d1 100644 --- a/src/order/hom/lattice.lean +++ b/src/order/hom/lattice.lean @@ -821,7 +821,7 @@ namespace sup_hom variables [has_sup α] [has_sup β] [has_sup γ] /-- Reinterpret a supremum homomorphism as an infimum homomorphism between the dual lattices. -/ -@[simps] protected def dual : sup_hom α β ≃ inf_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : sup_hom α β ≃ inf_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f, f.map_sup'⟩, inv_fun := λ f, ⟨f, f.map_inf'⟩, left_inv := λ f, sup_hom.ext $ λ _, rfl, @@ -832,8 +832,7 @@ variables [has_sup α] [has_sup β] [has_sup γ] (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : sup_hom.dual.symm (inf_hom.id _) = sup_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : inf_hom (order_dual β) (order_dual γ)) - (f : inf_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : inf_hom βᵒᵈ γᵒᵈ) (f : inf_hom αᵒᵈ βᵒᵈ) : sup_hom.dual.symm (g.comp f) = (sup_hom.dual.symm g).comp (sup_hom.dual.symm f) := rfl end sup_hom @@ -842,7 +841,7 @@ namespace inf_hom variables [has_inf α] [has_inf β] [has_inf γ] /-- Reinterpret an infimum homomorphism as a supremum homomorphism between the dual lattices. -/ -@[simps] protected def dual : inf_hom α β ≃ sup_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : inf_hom α β ≃ sup_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f, f.map_inf'⟩, inv_fun := λ f, ⟨f, f.map_sup'⟩, left_inv := λ f, inf_hom.ext $ λ _, rfl, @@ -853,8 +852,7 @@ variables [has_inf α] [has_inf β] [has_inf γ] (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : inf_hom.dual.symm (sup_hom.id _) = inf_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : sup_hom (order_dual β) (order_dual γ)) - (f : sup_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : sup_hom βᵒᵈ γᵒᵈ) (f : sup_hom αᵒᵈ βᵒᵈ) : inf_hom.dual.symm (g.comp f) = (inf_hom.dual.symm g).comp (inf_hom.dual.symm f) := rfl end inf_hom @@ -864,7 +862,7 @@ variables [has_sup α] [has_bot α] [has_sup β] [has_bot β] [has_sup γ] [has_ /-- Reinterpret a finitary supremum homomorphism as a finitary infimum homomorphism between the dual lattices. -/ -def dual : sup_bot_hom α β ≃ inf_top_hom (order_dual α) (order_dual β) := +def dual : sup_bot_hom α β ≃ inf_top_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_sup_hom.dual, f.map_bot'⟩, inv_fun := λ f, ⟨sup_hom.dual.symm f.to_inf_hom, f.map_top'⟩, left_inv := λ f, sup_bot_hom.ext $ λ _, rfl, @@ -875,8 +873,7 @@ def dual : sup_bot_hom α β ≃ inf_top_hom (order_dual α) (order_dual β) := (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : sup_bot_hom.dual.symm (inf_top_hom.id _) = sup_bot_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : inf_top_hom (order_dual β) (order_dual γ)) - (f : inf_top_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : inf_top_hom βᵒᵈ γᵒᵈ) (f : inf_top_hom αᵒᵈ βᵒᵈ) : sup_bot_hom.dual.symm (g.comp f) = (sup_bot_hom.dual.symm g).comp (sup_bot_hom.dual.symm f) := rfl end sup_bot_hom @@ -886,7 +883,7 @@ variables [has_inf α] [has_top α] [has_inf β] [has_top β] [has_inf γ] [has_ /-- Reinterpret a finitary infimum homomorphism as a finitary supremum homomorphism between the dual lattices. -/ -@[simps] protected def dual : inf_top_hom α β ≃ sup_bot_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : inf_top_hom α β ≃ sup_bot_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_inf_hom.dual, f.map_top'⟩, inv_fun := λ f, ⟨inf_hom.dual.symm f.to_sup_hom, f.map_bot'⟩, left_inv := λ f, inf_top_hom.ext $ λ _, rfl, @@ -897,8 +894,7 @@ lattices. -/ (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : inf_top_hom.dual.symm (sup_bot_hom.id _) = inf_top_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : sup_bot_hom (order_dual β) (order_dual γ)) - (f : sup_bot_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : sup_bot_hom βᵒᵈ γᵒᵈ) (f : sup_bot_hom αᵒᵈ βᵒᵈ) : inf_top_hom.dual.symm (g.comp f) = (inf_top_hom.dual.symm g).comp (inf_top_hom.dual.symm f) := rfl end inf_top_hom @@ -907,7 +903,7 @@ namespace lattice_hom variables [lattice α] [lattice β] [lattice γ] /-- Reinterpret a lattice homomorphism as a lattice homomorphism between the dual lattices. -/ -@[simps] protected def dual : lattice_hom α β ≃ lattice_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : lattice_hom α β ≃ lattice_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_inf_hom.dual, f.map_sup'⟩, inv_fun := λ f, ⟨f.to_inf_hom.dual, f.map_sup'⟩, left_inv := λ f, ext $ λ a, rfl, @@ -918,8 +914,7 @@ variables [lattice α] [lattice β] [lattice γ] (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : lattice_hom.dual.symm (lattice_hom.id _) = lattice_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : lattice_hom (order_dual β) (order_dual γ)) - (f : lattice_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : lattice_hom βᵒᵈ γᵒᵈ) (f : lattice_hom αᵒᵈ βᵒᵈ) : lattice_hom.dual.symm (g.comp f) = (lattice_hom.dual.symm g).comp (lattice_hom.dual.symm f) := rfl end lattice_hom @@ -929,8 +924,7 @@ variables [lattice α] [bounded_order α] [lattice β] [bounded_order β] [latti /-- Reinterpret a bounded lattice homomorphism as a bounded lattice homomorphism between the dual bounded lattices. -/ -@[simps] protected def dual : - bounded_lattice_hom α β ≃ bounded_lattice_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : bounded_lattice_hom α β ≃ bounded_lattice_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_lattice_hom.dual, f.map_bot', f.map_top'⟩, inv_fun := λ f, ⟨lattice_hom.dual.symm f.to_lattice_hom, f.map_bot', f.map_top'⟩, left_inv := λ f, ext $ λ a, rfl, @@ -942,8 +936,7 @@ bounded lattices. -/ @[simp] lemma symm_dual_id : bounded_lattice_hom.dual.symm (bounded_lattice_hom.id _) = bounded_lattice_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : bounded_lattice_hom (order_dual β) (order_dual γ)) - (f : bounded_lattice_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : bounded_lattice_hom βᵒᵈ γᵒᵈ) (f : bounded_lattice_hom αᵒᵈ βᵒᵈ) : bounded_lattice_hom.dual.symm (g.comp f) = (bounded_lattice_hom.dual.symm g).comp (bounded_lattice_hom.dual.symm f) := rfl diff --git a/src/order/iterate.lean b/src/order/iterate.lean index bcd0164440a1a..2b0b78c82ce61 100644 --- a/src/order/iterate.lean +++ b/src/order/iterate.lean @@ -123,13 +123,13 @@ lemma id_le_iterate_of_id_le (h : id ≤ f) (n : ℕ) : id ≤ (f^[n]) := by simpa only [iterate_id] using monotone_id.iterate_le_of_le h n lemma iterate_le_id_of_le_id (h : f ≤ id) (n : ℕ) : (f^[n]) ≤ id := -@id_le_iterate_of_id_le (order_dual α) _ f h n +@id_le_iterate_of_id_le αᵒᵈ _ f h n lemma monotone_iterate_of_id_le (h : id ≤ f) : monotone (λ m, f^[m]) := monotone_nat_of_le_succ $ λ n x, by { rw iterate_succ_apply', exact h _ } lemma antitone_iterate_of_le_id (h : f ≤ id) : antitone (λ m, f^[m]) := -λ m n hmn, @monotone_iterate_of_id_le (order_dual α) _ f h m n hmn +λ m n hmn, @monotone_iterate_of_id_le αᵒᵈ _ f h m n hmn end preorder @@ -161,7 +161,7 @@ by refine hf.seq_pos_lt_seq_of_le_of_lt hn _ (λ k hk, _) (λ k hk, _); lemma iterate_pos_lt_of_map_lt' (h : commute f g) (hf : strict_mono f) (hg : monotone g) {x} (hx : f x < g x) {n} (hn : 0 < n) : f^[n] x < (g^[n]) x := -@iterate_pos_lt_of_map_lt (order_dual α) _ g f h.symm hg.dual hf.dual x hx n hn +@iterate_pos_lt_of_map_lt αᵒᵈ _ g f h.symm hg.dual hf.dual x hx n hn end preorder @@ -180,7 +180,7 @@ end lemma iterate_pos_lt_iff_map_lt' (h : commute f g) (hf : strict_mono f) (hg : monotone g) {x n} (hn : 0 < n) : f^[n] x < (g^[n]) x ↔ f x < g x := -@iterate_pos_lt_iff_map_lt (order_dual α) _ _ _ h.symm hg.dual hf.dual x n hn +@iterate_pos_lt_iff_map_lt αᵒᵈ _ _ _ h.symm hg.dual hf.dual x n hn lemma iterate_pos_le_iff_map_le (h : commute f g) (hf : monotone f) (hg : strict_mono g) {x n} (hn : 0 < n) : diff --git a/src/order/lattice.lean b/src/order/lattice.lean index a0aeec33efa9e..c59a20e18ac9a 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -119,8 +119,8 @@ def semilattice_sup.mk' {α : Type*} [has_sup α] rwa [sup_assoc, hbc], end } -instance (α : Type*) [has_inf α] : has_sup (order_dual α) := ⟨((⊓) : α → α → α)⟩ -instance (α : Type*) [has_sup α] : has_inf (order_dual α) := ⟨((⊔) : α → α → α)⟩ +instance (α : Type*) [has_inf α] : has_sup αᵒᵈ := ⟨((⊓) : α → α → α)⟩ +instance (α : Type*) [has_sup α] : has_inf αᵒᵈ := ⟨((⊔) : α → α → α)⟩ section semilattice_sup variables [semilattice_sup α] {a b c d : α} @@ -284,20 +284,20 @@ class semilattice_inf (α : Type u) extends has_inf α, partial_order α := (inf_le_right : ∀ a b : α, a ⊓ b ≤ b) (le_inf : ∀ a b c : α, a ≤ b → a ≤ c → a ≤ b ⊓ c) -instance (α) [semilattice_inf α] : semilattice_sup (order_dual α) := +instance (α) [semilattice_inf α] : semilattice_sup αᵒᵈ := { le_sup_left := semilattice_inf.inf_le_left, le_sup_right := semilattice_inf.inf_le_right, sup_le := assume a b c hca hcb, @semilattice_inf.le_inf α _ _ _ _ hca hcb, .. order_dual.partial_order α, .. order_dual.has_sup α } -instance (α) [semilattice_sup α] : semilattice_inf (order_dual α) := +instance (α) [semilattice_sup α] : semilattice_inf αᵒᵈ := { inf_le_left := @le_sup_left α _, inf_le_right := @le_sup_right α _, le_inf := assume a b c hca hcb, @sup_le α _ _ _ _ hca hcb, .. order_dual.partial_order α, .. order_dual.has_inf α } theorem semilattice_sup.dual_dual (α : Type*) [H : semilattice_sup α] : - order_dual.semilattice_sup (order_dual α) = H := + order_dual.semilattice_sup αᵒᵈ = H := semilattice_sup.ext $ λ _ _, iff.rfl section semilattice_inf @@ -330,8 +330,7 @@ lt_of_le_of_lt inf_le_left h theorem inf_lt_of_right_lt (h : b < c) : a ⊓ b < c := lt_of_le_of_lt inf_le_right h -@[simp] theorem le_inf_iff : a ≤ b ⊓ c ↔ a ≤ b ∧ a ≤ c := -@sup_le_iff (order_dual α) _ _ _ _ +@[simp] theorem le_inf_iff : a ≤ b ⊓ c ↔ a ≤ b ∧ a ≤ c := @sup_le_iff αᵒᵈ _ _ _ _ @[simp] theorem inf_eq_left : a ⊓ b = a ↔ a ≤ b := le_antisymm_iff.trans $ by simp [le_refl] @@ -363,18 +362,15 @@ inf_le_inf le_rfl h theorem le_of_inf_eq (h : a ⊓ b = a) : a ≤ b := by { rw ← h, simp } -@[simp] theorem inf_idem : a ⊓ a = a := -@sup_idem (order_dual α) _ _ +@[simp] lemma inf_idem : a ⊓ a = a := @sup_idem αᵒᵈ _ _ instance inf_is_idempotent : is_idempotent α (⊓) := ⟨@inf_idem _ _⟩ -theorem inf_comm : a ⊓ b = b ⊓ a := -@sup_comm (order_dual α) _ _ _ +lemma inf_comm : a ⊓ b = b ⊓ a := @sup_comm αᵒᵈ _ _ _ instance inf_is_commutative : is_commutative α (⊓) := ⟨@inf_comm _ _⟩ -theorem inf_assoc : a ⊓ b ⊓ c = a ⊓ (b ⊓ c) := -@sup_assoc (order_dual α) _ a b c +lemma inf_assoc : a ⊓ b ⊓ c = a ⊓ (b ⊓ c) := @sup_assoc αᵒᵈ _ a b c instance inf_is_associative : is_associative α (⊓) := ⟨@inf_assoc _ _⟩ @@ -382,28 +378,28 @@ lemma inf_left_right_swap (a b c : α) : a ⊓ b ⊓ c = c ⊓ b ⊓ a := by rw [inf_comm, @inf_comm _ _ a, inf_assoc] @[simp] lemma inf_left_idem : a ⊓ (a ⊓ b) = a ⊓ b := -@sup_left_idem (order_dual α) _ a b +@sup_left_idem αᵒᵈ _ a b @[simp] lemma inf_right_idem : (a ⊓ b) ⊓ b = a ⊓ b := -@sup_right_idem (order_dual α) _ a b +@sup_right_idem αᵒᵈ _ a b lemma inf_left_comm (a b c : α) : a ⊓ (b ⊓ c) = b ⊓ (a ⊓ c) := -@sup_left_comm (order_dual α) _ a b c +@sup_left_comm αᵒᵈ _ a b c lemma inf_right_comm (a b c : α) : a ⊓ b ⊓ c = a ⊓ c ⊓ b := -@sup_right_comm (order_dual α) _ a b c +@sup_right_comm αᵒᵈ _ a b c lemma inf_inf_inf_comm (a b c d : α) : a ⊓ b ⊓ (c ⊓ d) = a ⊓ c ⊓ (b ⊓ d) := -@sup_sup_sup_comm (order_dual α) _ _ _ _ _ +@sup_sup_sup_comm αᵒᵈ _ _ _ _ _ lemma inf_inf_distrib_left (a b c : α) : a ⊓ (b ⊓ c) = (a ⊓ b) ⊓ (a ⊓ c) := -@sup_sup_distrib_left (order_dual α) _ _ _ _ +@sup_sup_distrib_left αᵒᵈ _ _ _ _ lemma inf_inf_distrib_right (a b c : α) : (a ⊓ b) ⊓ c = (a ⊓ c) ⊓ (b ⊓ c) := -@sup_sup_distrib_right (order_dual α) _ _ _ _ +@sup_sup_distrib_right αᵒᵈ _ _ _ _ lemma forall_le_or_exists_lt_inf (a : α) : (∀b, a ≤ b) ∨ (∃b, b < a) := -@forall_le_or_exists_lt_sup (order_dual α) _ a +@forall_le_or_exists_lt_sup αᵒᵈ _ a theorem semilattice_inf.ext_inf {α} {A B : semilattice_inf α} (H : ∀ x y : α, (by haveI := A; exact x ≤ y) ↔ x ≤ y) @@ -421,12 +417,12 @@ begin end theorem semilattice_inf.dual_dual (α : Type*) [H : semilattice_inf α] : - order_dual.semilattice_inf (order_dual α) = H := + order_dual.semilattice_inf αᵒᵈ = H := semilattice_inf.ext $ λ _ _, iff.rfl theorem exists_lt_of_inf (α : Type*) [semilattice_inf α] [nontrivial α] : ∃ a b : α, a < b := -let ⟨a, b, h⟩ := exists_lt_of_sup (order_dual α) in ⟨b, a, h⟩ +let ⟨a, b, h⟩ := exists_lt_of_sup αᵒᵈ in ⟨b, a, h⟩ lemma inf_le_ite (s s' : α) (P : Prop) [decidable P] : s ⊓ s' ≤ ite P s s' := if h : P then inf_le_left.trans_eq (if_pos h).symm else inf_le_right.trans_eq (if_neg h).symm @@ -444,8 +440,8 @@ def semilattice_inf.mk' {α : Type*} [has_inf α] (inf_assoc : ∀ (a b c : α), a ⊓ b ⊓ c = a ⊓ (b ⊓ c)) (inf_idem : ∀ (a : α), a ⊓ a = a) : semilattice_inf α := begin - haveI : semilattice_sup (order_dual α) := semilattice_sup.mk' inf_comm inf_assoc inf_idem, - haveI i := order_dual.semilattice_inf (order_dual α), + haveI : semilattice_sup αᵒᵈ := semilattice_sup.mk' inf_comm inf_assoc inf_idem, + haveI i := order_dual.semilattice_inf αᵒᵈ, exact i, end @@ -456,7 +452,7 @@ end /-- A lattice is a join-semilattice which is also a meet-semilattice. -/ @[protect_proj] class lattice (α : Type u) extends semilattice_sup α, semilattice_inf α -instance (α) [lattice α] : lattice (order_dual α) := +instance (α) [lattice α] : lattice αᵒᵈ := { .. order_dual.semilattice_sup α, .. order_dual.semilattice_inf α } /-- The partial orders from `semilattice_sup_mk'` and `semilattice_inf_mk'` agree @@ -598,7 +594,7 @@ calc x ⊓ (y ⊔ z) = (x ⊓ (x ⊔ z)) ⊓ (y ⊔ z) : by rw [inf_sup_se ... = ((x ⊓ y) ⊔ x) ⊓ ((x ⊓ y) ⊔ z) : by rw [sup_comm] ... = (x ⊓ y) ⊔ (x ⊓ z) : by rw [sup_inf_left] -instance (α : Type*) [distrib_lattice α] : distrib_lattice (order_dual α) := +instance (α : Type*) [distrib_lattice α] : distrib_lattice αᵒᵈ := { le_sup_inf := assume x y z, le_of_eq inf_sup_left.symm, .. order_dual.lattice α } @@ -663,12 +659,11 @@ lemma sup_ind (a b : α) {p : α → Prop} (ha : p a) (hb : p b) : p (a ⊔ b) : @[simp] lemma sup_lt_iff : b ⊔ c < a ↔ b < a ∧ c < a := ⟨λ h, ⟨le_sup_left.trans_lt h, le_sup_right.trans_lt h⟩, λ h, sup_ind b c h.1 h.2⟩ -lemma inf_ind (a b : α) {p : α → Prop} (ha : p a) (hb : p b) : p (a ⊓ b) := -@sup_ind (order_dual α) _ _ _ _ ha hb +lemma inf_ind (a b : α) {p : α → Prop} : p a → p b → p (a ⊓ b) := @sup_ind αᵒᵈ _ _ _ _ -@[simp] lemma inf_le_iff : b ⊓ c ≤ a ↔ b ≤ a ∨ c ≤ a := @le_sup_iff (order_dual α) _ _ _ _ -@[simp] lemma inf_lt_iff : b ⊓ c < a ↔ b < a ∨ c < a := @lt_sup_iff (order_dual α) _ _ _ _ -@[simp] lemma lt_inf_iff : a < b ⊓ c ↔ a < b ∧ a < c := @sup_lt_iff (order_dual α) _ _ _ _ +@[simp] lemma inf_le_iff : b ⊓ c ≤ a ↔ b ≤ a ∨ c ≤ a := @le_sup_iff αᵒᵈ _ _ _ _ +@[simp] lemma inf_lt_iff : b ⊓ c < a ↔ b < a ∨ c < a := @lt_sup_iff αᵒᵈ _ _ _ _ +@[simp] lemma lt_inf_iff : a < b ⊓ c ↔ a < b ∧ a < c := @sup_lt_iff αᵒᵈ _ _ _ _ end linear_order @@ -683,7 +678,7 @@ end lemma inf_eq_min_default [semilattice_inf α] [decidable_rel ((≤) : α → α → Prop)] [is_total α (≤)] : (⊓) = (min_default : α → α → α) := -@sup_eq_max_default (order_dual α) _ _ _ +@sup_eq_max_default αᵒᵈ _ _ _ /-- A lattice with total order is a linear order. diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index 19fd5d35a6d4c..f80f61d5317e2 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -105,7 +105,7 @@ end lemma not_is_bounded_under_of_tendsto_at_bot [preorder β] [no_min_order β] {f : α → β} {l : filter α} [l.ne_bot](hf : tendsto f l at_bot) : ¬ is_bounded_under (≥) l f := -@not_is_bounded_under_of_tendsto_at_top α (order_dual β) _ _ _ _ _ hf +@not_is_bounded_under_of_tendsto_at_top α βᵒᵈ _ _ _ _ _ hf lemma is_bounded_under.bdd_above_range_of_cofinite [semilattice_sup β] {f : α → β} (hf : is_bounded_under (≤) cofinite f) : bdd_above (range f) := @@ -118,7 +118,7 @@ end lemma is_bounded_under.bdd_below_range_of_cofinite [semilattice_inf β] {f : α → β} (hf : is_bounded_under (≥) cofinite f) : bdd_below (range f) := -@is_bounded_under.bdd_above_range_of_cofinite α (order_dual β) _ _ hf +@is_bounded_under.bdd_above_range_of_cofinite α βᵒᵈ _ _ hf lemma is_bounded_under.bdd_above_range [semilattice_sup β] {f : ℕ → β} (hf : is_bounded_under (≤) at_top f) : bdd_above (range f) := @@ -126,7 +126,7 @@ by { rw ← nat.cofinite_eq_at_top at hf, exact hf.bdd_above_range_of_cofinite } lemma is_bounded_under.bdd_below_range [semilattice_inf β] {f : ℕ → β} (hf : is_bounded_under (≥) at_top f) : bdd_below (range f) := -@is_bounded_under.bdd_above_range (order_dual β) _ _ hf +@is_bounded_under.bdd_above_range βᵒᵈ _ _ hf /-- `is_cobounded (≺) f` states that the filter `f` does not tend to infinity w.r.t. `≺`. This is also called frequently bounded. Will be usually instantiated with `≤` or `≥`. @@ -299,7 +299,7 @@ lemma liminf_le_liminf {α : Type*} [conditionally_complete_lattice β] {f : fil (hu : f.is_bounded_under (≥) u . is_bounded_default) (hv : f.is_cobounded_under (≥) v . is_bounded_default) : f.liminf u ≤ f.liminf v := -@limsup_le_limsup (order_dual β) α _ _ _ _ h hv hu +@limsup_le_limsup βᵒᵈ α _ _ _ _ h hv hu lemma limsup_le_limsup_of_le {α β} [conditionally_complete_lattice β] {f g : filter α} (h : f ≤ g) {u : α → β} (hf : f.is_cobounded_under (≤) u . is_bounded_default) @@ -319,7 +319,7 @@ by simp [Limsup]; exact cInf_upper_bounds_eq_cSup h hs theorem Liminf_principal {s : set α} (h : bdd_below s) (hs : s.nonempty) : (𝓟 s).Liminf = Inf s := -@Limsup_principal (order_dual α) _ s h hs +@Limsup_principal αᵒᵈ _ s h hs lemma limsup_congr {α : Type*} [conditionally_complete_lattice β] {f : filter α} {u v : α → β} (h : ∀ᶠ a in f, u a = v a) : limsup f u = limsup f v := @@ -331,7 +331,7 @@ end lemma liminf_congr {α : Type*} [conditionally_complete_lattice β] {f : filter α} {u v : α → β} (h : ∀ᶠ a in f, u a = v a) : liminf f u = liminf f v := -@limsup_congr (order_dual β) _ _ _ _ _ h +@limsup_congr βᵒᵈ _ _ _ _ _ h lemma limsup_const {α : Type*} [conditionally_complete_lattice β] {f : filter α} [ne_bot f] (b : β) : limsup f (λ x, b) = b := @@ -339,7 +339,7 @@ by simpa only [limsup_eq, eventually_const] using cInf_Ici lemma liminf_const {α : Type*} [conditionally_complete_lattice β] {f : filter α} [ne_bot f] (b : β) : liminf f (λ x, b) = b := -@limsup_const (order_dual β) α _ f _ b +@limsup_const βᵒᵈ α _ f _ b lemma liminf_le_limsup {f : filter β} [ne_bot f] {u : β → α} (h : f.is_bounded_under (≤) u . is_bounded_default) @@ -375,7 +375,7 @@ end /-- Same as limsup_const applied to `⊤` but without the `ne_bot f` assumption -/ lemma liminf_const_top {f : filter β} : liminf f (λ x : β, (⊤ : α)) = (⊤ : α) := -@limsup_const_bot (order_dual α) β _ _ +@limsup_const_bot αᵒᵈ β _ _ theorem has_basis.Limsup_eq_infi_Sup {ι} {p : ι → Prop} {s} {f : filter α} (h : f.has_basis p s) : f.Limsup = ⨅ i (hi : p i), Sup (s i) := @@ -386,13 +386,13 @@ le_antisymm theorem has_basis.Liminf_eq_supr_Inf {p : ι → Prop} {s : ι → set α} {f : filter α} (h : f.has_basis p s) : f.Liminf = ⨆ i (hi : p i), Inf (s i) := -@has_basis.Limsup_eq_infi_Sup (order_dual α) _ _ _ _ _ h +@has_basis.Limsup_eq_infi_Sup αᵒᵈ _ _ _ _ _ h theorem Limsup_eq_infi_Sup {f : filter α} : f.Limsup = ⨅ s ∈ f, Sup s := f.basis_sets.Limsup_eq_infi_Sup theorem Liminf_eq_supr_Inf {f : filter α} : f.Liminf = ⨆ s ∈ f, Inf s := -@Limsup_eq_infi_Sup (order_dual α) _ _ +@Limsup_eq_infi_Sup αᵒᵈ _ _ /-- In a complete lattice, the limsup of a function is the infimum over sets `s` in the filter of the supremum of the function over `s` -/ @@ -414,17 +414,17 @@ theorem has_basis.limsup_eq_infi_supr {p : ι → Prop} {s : ι → set β} {f : /-- In a complete lattice, the liminf of a function is the infimum over sets `s` in the filter of the supremum of the function over `s` -/ theorem liminf_eq_supr_infi {f : filter β} {u : β → α} : f.liminf u = ⨆ s ∈ f, ⨅ a ∈ s, u a := -@limsup_eq_infi_supr (order_dual α) β _ _ _ +@limsup_eq_infi_supr αᵒᵈ β _ _ _ lemma liminf_eq_supr_infi_of_nat {u : ℕ → α} : liminf at_top u = ⨆ n : ℕ, ⨅ i ≥ n, u i := -@limsup_eq_infi_supr_of_nat (order_dual α) _ u +@limsup_eq_infi_supr_of_nat αᵒᵈ _ u lemma liminf_eq_supr_infi_of_nat' {u : ℕ → α} : liminf at_top u = ⨆ n : ℕ, ⨅ i : ℕ, u (i + n) := -@limsup_eq_infi_supr_of_nat' (order_dual α) _ _ +@limsup_eq_infi_supr_of_nat' αᵒᵈ _ _ theorem has_basis.liminf_eq_supr_infi {p : ι → Prop} {s : ι → set β} {f : filter β} {u : β → α} (h : f.has_basis p s) : f.liminf u = ⨆ i (hi : p i), ⨅ a ∈ s i, u a := -@has_basis.limsup_eq_infi_supr (order_dual α) _ _ _ _ _ _ _ h +@has_basis.limsup_eq_infi_supr αᵒᵈ _ _ _ _ _ _ _ h @[simp] lemma liminf_nat_add (f : ℕ → α) (k : ℕ) : at_top.liminf (λ i, f (i + k)) = at_top.liminf f := @@ -432,7 +432,7 @@ by { simp_rw liminf_eq_supr_infi_of_nat, exact supr_infi_ge_nat_add f k } @[simp] lemma limsup_nat_add (f : ℕ → α) (k : ℕ) : at_top.limsup (λ i, f (i + k)) = at_top.limsup f := -@liminf_nat_add (order_dual α) _ f k +@liminf_nat_add αᵒᵈ _ f k lemma liminf_le_of_frequently_le' {α β} [complete_lattice β] {f : filter α} {u : α → β} {x : β} (h : ∃ᶠ a in f, u a ≤ x) : @@ -450,7 +450,7 @@ end lemma le_limsup_of_frequently_le' {α β} [complete_lattice β] {f : filter α} {u : α → β} {x : β} (h : ∃ᶠ a in f, x ≤ u a) : x ≤ f.limsup u := -@liminf_le_of_frequently_le' _ (order_dual β) _ _ _ _ h +@liminf_le_of_frequently_le' _ βᵒᵈ _ _ _ _ h end complete_lattice @@ -468,7 +468,7 @@ end lemma eventually_lt_of_limsup_lt {f : filter α} [conditionally_complete_linear_order β] {u : α → β} {b : β} (h : limsup f u < b) (hu : f.is_bounded_under (≤) u . is_bounded_default) : ∀ᶠ a in f, u a < b := -@eventually_lt_of_lt_liminf _ (order_dual β) _ _ _ _ h hu +@eventually_lt_of_lt_liminf _ βᵒᵈ _ _ _ _ h hu lemma le_limsup_of_frequently_le {α β} [conditionally_complete_linear_order β] {f : filter α} {u : α → β} {b : β} (hu_le : ∃ᶠ x in f, b ≤ u x) @@ -485,7 +485,7 @@ lemma liminf_le_of_frequently_le {α β} [conditionally_complete_linear_order {u : α → β} {b : β} (hu_le : ∃ᶠ x in f, u x ≤ b) (hu : f.is_bounded_under (≥) u . is_bounded_default) : f.liminf u ≤ b := -@le_limsup_of_frequently_le _ (order_dual β) _ f u b hu_le hu +@le_limsup_of_frequently_le _ βᵒᵈ _ f u b hu_le hu lemma frequently_lt_of_lt_limsup {α β} [conditionally_complete_linear_order β] {f : filter α} {u : α → β} {b : β} @@ -501,7 +501,7 @@ lemma frequently_lt_of_liminf_lt {α β} [conditionally_complete_linear_order β {u : α → β} {b : β} (hu : f.is_cobounded_under (≥) u . is_bounded_default) (h : f.liminf u < b) : ∃ᶠ x in f, u x < b := -@frequently_lt_of_lt_limsup _ (order_dual β) _ f u b hu h +@frequently_lt_of_lt_limsup _ βᵒᵈ _ f u b hu h end conditionally_complete_linear_order @@ -548,6 +548,6 @@ lemma order_iso.liminf_apply {γ} [conditionally_complete_lattice β] (hgu : f.is_bounded_under (≥) (λ x, g (u x)) . is_bounded_default) (hgu_co : f.is_cobounded_under (≥) (λ x, g (u x)) . is_bounded_default) : g (f.liminf u) = f.liminf (λ x, g (u x)) := -@order_iso.limsup_apply α (order_dual β) (order_dual γ) _ _ f u g.dual hu hu_co hgu hgu_co +@order_iso.limsup_apply α βᵒᵈ γᵒᵈ _ _ f u g.dual hu hu_co hgu hgu_co end order diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 937ba1a08cef6..54634155c3da1 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -423,13 +423,13 @@ open order_dual variables [locally_finite_order α] (a b : α) /-- Note we define `Icc (to_dual a) (to_dual b)` as `Icc α _ _ b a` (which has type `finset α` not -`finset (order_dual α)`!) instead of `(Icc b a).map to_dual.to_embedding` as this means the +`finset αᵒᵈ`!) instead of `(Icc b a).map to_dual.to_embedding` as this means the following is defeq: ``` lemma this : (Icc (to_dual (to_dual a)) (to_dual (to_dual b)) : _) = (Icc a b : _) := rfl ``` -/ -instance : locally_finite_order (order_dual α) := +instance : locally_finite_order αᵒᵈ := { finset_Icc := λ a b, @Icc α _ _ (of_dual b) (of_dual a), finset_Ico := λ a b, @Ioc α _ _ (of_dual b) (of_dual a), finset_Ioc := λ a b, @Ico α _ _ (of_dual b) (of_dual a), @@ -570,7 +570,7 @@ namespace with_bot variables (α) [partial_order α] [order_bot α] [locally_finite_order α] instance : locally_finite_order (with_bot α) := -@order_dual.locally_finite_order (with_top (order_dual α)) _ _ +@order_dual.locally_finite_order (with_top αᵒᵈ) _ _ variables (a b : α) diff --git a/src/order/max.lean b/src/order/max.lean index 56c53cab2685a..38f2de28e046f 100644 --- a/src/order/max.lean +++ b/src/order/max.lean @@ -59,20 +59,16 @@ nonempty_subtype.2 (exists_lt a) instance nonempty_gt [has_lt α] [no_max_order α] (a : α) : nonempty {x // a < x} := nonempty_subtype.2 (exists_gt a) -instance order_dual.no_bot_order (α : Type*) [has_le α] [no_top_order α] : - no_bot_order (order_dual α) := +instance order_dual.no_bot_order (α : Type*) [has_le α] [no_top_order α] : no_bot_order αᵒᵈ := ⟨λ a, @exists_not_le α _ _ a⟩ -instance order_dual.no_top_order (α : Type*) [has_le α] [no_bot_order α] : - no_top_order (order_dual α) := +instance order_dual.no_top_order (α : Type*) [has_le α] [no_bot_order α] : no_top_order αᵒᵈ := ⟨λ a, @exists_not_ge α _ _ a⟩ -instance order_dual.no_min_order (α : Type*) [has_lt α] [no_max_order α] : - no_min_order (order_dual α) := +instance order_dual.no_min_order (α : Type*) [has_lt α] [no_max_order α] : no_min_order αᵒᵈ := ⟨λ a, @exists_gt α _ _ a⟩ -instance order_dual.no_max_order (α : Type*) [has_lt α] [no_min_order α] : - no_max_order (order_dual α) := +instance order_dual.no_max_order (α : Type*) [has_lt α] [no_min_order α] : no_max_order αᵒᵈ := ⟨λ a, @exists_lt α _ _ a⟩ @[priority 100] -- See note [lower instance priority] @@ -121,10 +117,10 @@ protected lemma is_top.is_max (h : is_top a) : is_max a := λ b _, h b @[simp] lemma is_top_to_dual_iff : is_top (to_dual a) ↔ is_bot a := iff.rfl @[simp] lemma is_min_to_dual_iff : is_min (to_dual a) ↔ is_max a := iff.rfl @[simp] lemma is_max_to_dual_iff : is_max (to_dual a) ↔ is_min a := iff.rfl -@[simp] lemma is_bot_of_dual_iff {a : order_dual α} : is_bot (of_dual a) ↔ is_top a := iff.rfl -@[simp] lemma is_top_of_dual_iff {a : order_dual α} : is_top (of_dual a) ↔ is_bot a := iff.rfl -@[simp] lemma is_min_of_dual_iff {a : order_dual α} : is_min (of_dual a) ↔ is_max a := iff.rfl -@[simp] lemma is_max_of_dual_iff {a : order_dual α} : is_max (of_dual a) ↔ is_min a := iff.rfl +@[simp] lemma is_bot_of_dual_iff {a : αᵒᵈ} : is_bot (of_dual a) ↔ is_top a := iff.rfl +@[simp] lemma is_top_of_dual_iff {a : αᵒᵈ} : is_top (of_dual a) ↔ is_bot a := iff.rfl +@[simp] lemma is_min_of_dual_iff {a : αᵒᵈ} : is_min (of_dual a) ↔ is_max a := iff.rfl +@[simp] lemma is_max_of_dual_iff {a : αᵒᵈ} : is_max (of_dual a) ↔ is_min a := iff.rfl alias is_bot_to_dual_iff ↔ _ is_top.to_dual alias is_top_to_dual_iff ↔ _ is_bot.to_dual @@ -196,6 +192,6 @@ variables [linear_order α] lemma is_top_or_exists_gt (a : α) : is_top a ∨ ∃ b, a < b := by simpa only [or_iff_not_imp_left, is_top, not_forall, not_le] using id -lemma is_bot_or_exists_lt (a : α) : is_bot a ∨ ∃ b, b < a := @is_top_or_exists_gt (order_dual α) _ a +lemma is_bot_or_exists_lt (a : α) : is_bot a ∨ ∃ b, b < a := @is_top_or_exists_gt αᵒᵈ _ a end linear_order diff --git a/src/order/min_max.lean b/src/order/min_max.lean index cbe72653f0ef4..46d96d9a40b4d 100644 --- a/src/order/min_max.lean +++ b/src/order/min_max.lean @@ -68,8 +68,7 @@ end /-- For elements `a` and `b` of a linear order, either `max a b = a` and `b ≤ a`, or `max a b = b` and `a < b`. Use cases on this lemma to automate linarith in inequalities -/ -lemma max_cases (a b : α) : max a b = a ∧ b ≤ a ∨ max a b = b ∧ a < b := -@min_cases (order_dual α) _ a b +lemma max_cases (a b : α) : max a b = a ∧ b ≤ a ∨ max a b = b ∧ a < b := @min_cases αᵒᵈ _ a b lemma min_eq_iff : min a b = c ↔ a = c ∧ a ≤ b ∨ b = c ∧ b ≤ a := begin @@ -81,8 +80,7 @@ begin simp [h] } end -lemma max_eq_iff : max a b = c ↔ a = c ∧ b ≤ a ∨ b = c ∧ a ≤ b := -@min_eq_iff (order_dual α) _ a b c +lemma max_eq_iff : max a b = c ↔ a = c ∧ b ≤ a ∨ b = c ∧ a ≤ b := @min_eq_iff αᵒᵈ _ a b c lemma min_lt_min_left_iff : min a c < min b c ↔ a < b ∧ a < c := by { simp_rw [lt_min_iff, min_lt_iff, or_iff_left (lt_irrefl _)], @@ -91,10 +89,8 @@ by { simp_rw [lt_min_iff, min_lt_iff, or_iff_left (lt_irrefl _)], lemma min_lt_min_right_iff : min a b < min a c ↔ b < c ∧ b < a := by simp_rw [min_comm a, min_lt_min_left_iff] -lemma max_lt_max_left_iff : max a c < max b c ↔ a < b ∧ c < b := -@min_lt_min_left_iff (order_dual α) _ _ _ _ -lemma max_lt_max_right_iff : max a b < max a c ↔ b < c ∧ a < c := -@min_lt_min_right_iff (order_dual α) _ _ _ _ +lemma max_lt_max_left_iff : max a c < max b c ↔ a < b ∧ c < b := @min_lt_min_left_iff αᵒᵈ _ _ _ _ +lemma max_lt_max_right_iff : max a b < max a c ↔ b < c ∧ a < c := @min_lt_min_right_iff αᵒᵈ _ _ _ _ /-- An instance asserting that `max a a = a` -/ instance max_idem : is_idempotent α max := by apply_instance -- short-circuit type class inference @@ -107,8 +103,7 @@ lemma min_lt_max : min a b < max a b ↔ a ≠ b := inf_lt_sup lemma max_lt_max (h₁ : a < c) (h₂ : b < d) : max a b < max c d := by simp [lt_max_iff, max_lt_iff, *] -lemma min_lt_min (h₁ : a < c) (h₂ : b < d) : min a b < min c d := -@max_lt_max (order_dual α) _ _ _ _ _ h₁ h₂ +lemma min_lt_min (h₁ : a < c) (h₂ : b < d) : min a b < min c d := @max_lt_max αᵒᵈ _ _ _ _ _ h₁ h₂ theorem min_right_comm (a b c : α) : min (min a b) c = min (min a c) b := right_comm min min_comm min_assoc a b c @@ -152,7 +147,7 @@ lemma min_rec {p : α → Prop} {x y : α} (hx : x ≤ y → p x) (hy : y ≤ x (λ h, (min_eq_right h).symm.subst (hy h)) lemma max_rec {p : α → Prop} {x y : α} (hx : y ≤ x → p x) (hy : x ≤ y → p y) : p (max x y) := -@min_rec (order_dual α) _ _ _ _ hx hy +@min_rec αᵒᵈ _ _ _ _ hx hy lemma min_rec' (p : α → Prop) {x y : α} (hx : p x) (hy : p y) : p (min x y) := min_rec (λ _, hx) (λ _, hy) @@ -164,7 +159,7 @@ theorem min_choice (a b : α) : min a b = a ∨ min a b = b := by cases le_total a b; simp * theorem max_choice (a b : α) : max a b = a ∨ max a b = b := -@min_choice (order_dual α) _ a b +@min_choice αᵒᵈ _ a b lemma le_of_max_le_left {a b c : α} (h : max a b ≤ c) : a ≤ c := le_trans (le_max_left _ _) h diff --git a/src/order/modular_lattice.lean b/src/order/modular_lattice.lean index 6901c7bc46e33..ce193145da41c 100644 --- a/src/order/modular_lattice.lean +++ b/src/order/modular_lattice.lean @@ -52,7 +52,7 @@ lemma inf_sup_assoc_of_le {x : α} (y : α) {z : α} (h : z ≤ x) : (x ⊓ y) ⊔ z = x ⊓ (y ⊔ z) := by rw [inf_comm, sup_comm, ← sup_inf_assoc_of_le y h, inf_comm, sup_comm] -instance : is_modular_lattice (order_dual α) := +instance : is_modular_lattice αᵒᵈ := ⟨λ x y z xz, le_of_eq (by { rw [inf_comm, sup_comm, eq_comm, inf_comm, sup_comm], convert sup_inf_assoc_of_le (order_dual.of_dual y) (order_dual.dual_le.2 xz) })⟩ @@ -60,7 +60,7 @@ variables {x y z : α} theorem is_modular_lattice.sup_inf_sup_assoc : (x ⊔ z) ⊓ (y ⊔ z) = ((x ⊔ z) ⊓ y) ⊔ z := -@is_modular_lattice.inf_sup_inf_assoc (order_dual α) _ _ _ _ _ +@is_modular_lattice.inf_sup_inf_assoc αᵒᵈ _ _ _ _ _ theorem eq_of_le_of_inf_le_of_sup_le (hxy : x ≤ y) (hinf : y ⊓ z ≤ x ⊓ z) (hsup : y ⊔ z ≤ x ⊔ z) : x = y := @@ -81,7 +81,7 @@ lt_of_le_of_ne (le_of_eq hsup.symm)) theorem inf_lt_inf_of_lt_of_sup_le_sup (hxy : x < y) (hinf : y ⊔ z ≤ x ⊔ z) : x ⊓ z < y ⊓ z := -@sup_lt_sup_of_lt_of_inf_le_inf (order_dual α) _ _ _ _ _ hxy hinf +@sup_lt_sup_of_lt_of_inf_le_inf αᵒᵈ _ _ _ _ _ hxy hinf /-- A generalization of the theorem that if `N` is a submodule of `M` and `N` and `M / N` are both Artinian, then `M` is Artinian. -/ @@ -119,9 +119,7 @@ theorem well_founded_gt_exact_sequence (hf : ∀ a, f₁ (f₂ a) = a ⊓ K) (hg : ∀ a, g₁ (g₂ a) = a ⊔ K) : well_founded ((>) : α → α → Prop) := -@well_founded_lt_exact_sequence - (order_dual α) _ _ (order_dual γ) (order_dual β) _ _ - h₂ h₁ K g₁ g₂ f₁ f₂ gi.dual gci.dual hg hf +@well_founded_lt_exact_sequence αᵒᵈ _ _ γᵒᵈ βᵒᵈ _ _ h₂ h₁ K g₁ g₂ f₁ f₂ gi.dual gci.dual hg hf /-- The diamond isomorphism between the intervals `[a ⊓ b, a]` and `[b, a ⊔ b]` -/ def inf_Icc_order_iso_Icc_sup (a b : α) : set.Icc (a ⊓ b) a ≃o set.Icc b (a ⊔ b) := diff --git a/src/order/monotone.lean b/src/order/monotone.lean index 8df550fd1916a..57f24e1e1f547 100644 --- a/src/order/monotone.lean +++ b/src/order/monotone.lean @@ -637,17 +637,14 @@ lemma monotone_nat_of_le_succ {f : ℕ → α} (hf : ∀ n, f n ≤ f (n + 1)) : monotone f := nat.rel_of_forall_rel_succ_of_le (≤) hf -lemma antitone_nat_of_succ_le {f : ℕ → α} (hf : ∀ n, f (n + 1) ≤ f n) : - antitone f := -@monotone_nat_of_le_succ (order_dual α) _ _ hf +lemma antitone_nat_of_succ_le {f : ℕ → α} (hf : ∀ n, f (n + 1) ≤ f n) : antitone f := +@monotone_nat_of_le_succ αᵒᵈ _ _ hf -lemma strict_mono_nat_of_lt_succ {f : ℕ → α} (hf : ∀ n, f n < f (n + 1)) : - strict_mono f := +lemma strict_mono_nat_of_lt_succ {f : ℕ → α} (hf : ∀ n, f n < f (n + 1)) : strict_mono f := nat.rel_of_forall_rel_succ_of_lt (<) hf -lemma strict_anti_nat_of_succ_lt {f : ℕ → α} (hf : ∀ n, f (n + 1) < f n) : - strict_anti f := -@strict_mono_nat_of_lt_succ (order_dual α) _ f hf +lemma strict_anti_nat_of_succ_lt {f : ℕ → α} (hf : ∀ n, f (n + 1) < f n) : strict_anti f := +@strict_mono_nat_of_lt_succ αᵒᵈ _ f hf lemma int.rel_of_forall_rel_succ_of_lt (r : β → β → Prop) [is_trans β r] {f : ℤ → β} (h : ∀ n, r (f n) (f (n + 1))) ⦃a b : ℤ⦄ (hab : a < b) : r (f a) (f b) := diff --git a/src/order/order_dual.lean b/src/order/order_dual.lean index a0a833081b658..b11185399d059 100644 --- a/src/order/order_dual.lean +++ b/src/order/order_dual.lean @@ -25,61 +25,38 @@ variables {α : Type u} {β : Type v} {γ : Type w} {r : α → α → Prop} namespace order_dual -instance [nontrivial α] : nontrivial (order_dual α) := by delta order_dual; assumption +instance [nontrivial α] : nontrivial αᵒᵈ := by delta order_dual; assumption /-- `to_dual` is the identity function to the `order_dual` of a linear order. -/ -def to_dual : α ≃ order_dual α := ⟨id, id, λ h, rfl, λ h, rfl⟩ +def to_dual : α ≃ αᵒᵈ := ⟨id, id, λ h, rfl, λ h, rfl⟩ /-- `of_dual` is the identity function from the `order_dual` of a linear order. -/ -def of_dual : order_dual α ≃ α := to_dual.symm +def of_dual : αᵒᵈ ≃ α := to_dual.symm @[simp] lemma to_dual_symm_eq : (@to_dual α).symm = of_dual := rfl - @[simp] lemma of_dual_symm_eq : (@of_dual α).symm = to_dual := rfl - -@[simp] lemma to_dual_of_dual (a : order_dual α) : to_dual (of_dual a) = a := rfl +@[simp] lemma to_dual_of_dual (a : αᵒᵈ) : to_dual (of_dual a) = a := rfl @[simp] lemma of_dual_to_dual (a : α) : of_dual (to_dual a) = a := rfl -@[simp] lemma to_dual_inj {a b : α} : - to_dual a = to_dual b ↔ a = b := iff.rfl - -@[simp] lemma to_dual_le_to_dual [has_le α] {a b : α} : - to_dual a ≤ to_dual b ↔ b ≤ a := iff.rfl - -@[simp] lemma to_dual_lt_to_dual [has_lt α] {a b : α} : - to_dual a < to_dual b ↔ b < a := iff.rfl - -@[simp] lemma of_dual_inj {a b : order_dual α} : - of_dual a = of_dual b ↔ a = b := iff.rfl +@[simp] lemma to_dual_inj {a b : α} : to_dual a = to_dual b ↔ a = b := iff.rfl +@[simp] lemma of_dual_inj {a b : αᵒᵈ} : of_dual a = of_dual b ↔ a = b := iff.rfl -@[simp] lemma of_dual_le_of_dual [has_le α] {a b : order_dual α} : - of_dual a ≤ of_dual b ↔ b ≤ a := iff.rfl +@[simp] lemma to_dual_le_to_dual [has_le α] {a b : α} : to_dual a ≤ to_dual b ↔ b ≤ a := iff.rfl +@[simp] lemma to_dual_lt_to_dual [has_lt α] {a b : α} : to_dual a < to_dual b ↔ b < a := iff.rfl +@[simp] lemma of_dual_le_of_dual [has_le α] {a b : αᵒᵈ} : of_dual a ≤ of_dual b ↔ b ≤ a := iff.rfl +@[simp] lemma of_dual_lt_of_dual [has_lt α] {a b : αᵒᵈ} : of_dual a < of_dual b ↔ b < a := iff.rfl -@[simp] lemma of_dual_lt_of_dual [has_lt α] {a b : order_dual α} : - of_dual a < of_dual b ↔ b < a := iff.rfl +lemma le_to_dual [has_le α] {a : αᵒᵈ} {b : α} : a ≤ to_dual b ↔ b ≤ of_dual a := iff.rfl +lemma lt_to_dual [has_lt α] {a : αᵒᵈ} {b : α} : a < to_dual b ↔ b < of_dual a := iff.rfl +lemma to_dual_le [has_le α] {a : α} {b : αᵒᵈ} : to_dual a ≤ b ↔ of_dual b ≤ a := iff.rfl +lemma to_dual_lt [has_lt α] {a : α} {b : αᵒᵈ} : to_dual a < b ↔ of_dual b < a := iff.rfl -lemma le_to_dual [has_le α] {a : order_dual α} {b : α} : - a ≤ to_dual b ↔ b ≤ of_dual a := iff.rfl - -lemma lt_to_dual [has_lt α] {a : order_dual α} {b : α} : - a < to_dual b ↔ b < of_dual a := iff.rfl - -lemma to_dual_le [has_le α] {a : α} {b : order_dual α} : - to_dual a ≤ b ↔ of_dual b ≤ a := iff.rfl - -lemma to_dual_lt [has_lt α] {a : α} {b : order_dual α} : - to_dual a < b ↔ of_dual b < a := iff.rfl - -/-- Recursor for `order_dual α`. -/ +/-- Recursor for `αᵒᵈ`. -/ @[elab_as_eliminator] -protected def rec {C : order_dual α → Sort*} (h₂ : Π (a : α), C (to_dual a)) : - Π (a : order_dual α), C a := h₂ - -@[simp] protected lemma «forall» {p : order_dual α → Prop} : (∀ a, p a) ↔ ∀ a, p (to_dual a) := -iff.rfl +protected def rec {C : αᵒᵈ → Sort*} (h₂ : Π (a : α), C (to_dual a)) : Π (a : αᵒᵈ), C a := h₂ -@[simp] protected lemma «exists» {p : order_dual α → Prop} : (∃ a, p a) ↔ ∃ a, p (to_dual a) := -iff.rfl +@[simp] protected lemma «forall» {p : αᵒᵈ → Prop} : (∀ a, p a) ↔ ∀ a, p (to_dual a) := iff.rfl +@[simp] protected lemma «exists» {p : αᵒᵈ → Prop} : (∃ a, p a) ↔ ∃ a, p (to_dual a) := iff.rfl end order_dual diff --git a/src/order/pfilter.lean b/src/order/pfilter.lean index 43b49eb367283..e421064d01b2a 100644 --- a/src/order/pfilter.lean +++ b/src/order/pfilter.lean @@ -16,7 +16,7 @@ a join-semilattice structure. - `order.pfilter P`: The type of nonempty, downward directed, upward closed subsets of `P`. This is dual to `order.ideal`, so it - simply wraps `order.ideal (order_dual P)`. + simply wraps `order.ideal Pᵒᵈ`. - `order.is_pfilter P`: a predicate for when a `set P` is a filter. @@ -43,11 +43,11 @@ variables {P : Type*} - downward directed - upward closed. -/ structure pfilter (P) [preorder P] := -(dual : ideal (order_dual P)) +(dual : ideal Pᵒᵈ) /-- A predicate for when a subset of `P` is a filter. -/ def is_pfilter [preorder P] (F : set P) : Prop := -@is_ideal (order_dual P) _ F +@is_ideal Pᵒᵈ _ F lemma is_pfilter.of_def [preorder P] {F : set P} (nonempty : F.nonempty) (directed : directed_on (≥) F) (mem_of_le : ∀ {x y : P}, x ≤ y → x ∈ F → y ∈ F) : is_pfilter F := diff --git a/src/order/rel_classes.lean b/src/order/rel_classes.lean index 8db3f3846d5bf..d2f5c9a52f45c 100644 --- a/src/order/rel_classes.lean +++ b/src/order/rel_classes.lean @@ -458,10 +458,10 @@ lemma transitive_lt [preorder α] : transitive (@has_lt.lt α _) := transitive_o lemma transitive_ge [preorder α] : transitive (@ge α _) := transitive_of_trans _ lemma transitive_gt [preorder α] : transitive (@gt α _) := transitive_of_trans _ -instance order_dual.is_total_le [has_le α] [is_total α (≤)] : is_total (order_dual α) (≤) := +instance order_dual.is_total_le [has_le α] [is_total α (≤)] : is_total αᵒᵈ (≤) := @is_total.swap α _ _ instance nat.lt.is_well_order : is_well_order ℕ (<) := ⟨nat.lt_wf⟩ -instance [linear_order α] [h : is_well_order α (<)] : is_well_order (order_dual α) (>) := h -instance [linear_order α] [h : is_well_order α (>)] : is_well_order (order_dual α) (<) := h +instance [linear_order α] [h : is_well_order α (<)] : is_well_order αᵒᵈ (>) := h +instance [linear_order α] [h : is_well_order α (>)] : is_well_order αᵒᵈ (<) := h diff --git a/src/order/rel_iso.lean b/src/order/rel_iso.lean index f92c19c5168cd..bbb073b837f78 100644 --- a/src/order/rel_iso.lean +++ b/src/order/rel_iso.lean @@ -74,7 +74,7 @@ lemma map_inf [semilattice_inf α] [linear_order β] lemma map_sup [semilattice_sup α] [linear_order β] [rel_hom_class F ((>) : β → β → Prop) ((>) : α → α → Prop)] (a : F) (m n : β) : a (m ⊔ n) = a m ⊔ a n := -@map_inf (order_dual α) (order_dual β) _ _ _ _ _ _ _ +@map_inf αᵒᵈ βᵒᵈ _ _ _ _ _ _ _ protected theorem is_irrefl [rel_hom_class F r s] (f : F) : ∀ [is_irrefl β s], is_irrefl α r | ⟨H⟩ := ⟨λ a h, H _ (map_rel f h)⟩ diff --git a/src/order/succ_pred/basic.lean b/src/order/succ_pred/basic.lean index 49640974bca81..a657e33acc0b5 100644 --- a/src/order/succ_pred/basic.lean +++ b/src/order/succ_pred/basic.lean @@ -72,14 +72,14 @@ variables {α : Type*} (le_pred_of_lt {a b} : a < b → a ≤ pred b) (le_of_pred_lt {a b} : pred a < b → a ≤ b) -instance [preorder α] [succ_order α] : pred_order (order_dual α) := +instance [preorder α] [succ_order α] : pred_order αᵒᵈ := { pred := to_dual ∘ succ_order.succ ∘ of_dual, pred_le := succ_order.le_succ, min_of_le_pred := λ _, succ_order.max_of_succ_le, le_pred_of_lt := λ a b h, succ_order.succ_le_of_lt h, le_of_pred_lt := λ a b, succ_order.le_of_lt_succ } -instance [preorder α] [pred_order α] : succ_order (order_dual α) := +instance [preorder α] [pred_order α] : succ_order αᵒᵈ := { succ := to_dual ∘ pred_order.pred ∘ of_dual, le_succ := pred_order.pred_le, max_of_succ_le := λ _, pred_order.min_of_le_pred, @@ -532,8 +532,8 @@ variables [order_bot α] @[simp] lemma pred_bot : pred (⊥ : α) = ⊥ := is_min_bot.pred_eq -@[simp] lemma le_pred_iff_eq_bot : a ≤ pred a ↔ a = ⊥ := @succ_le_iff_eq_top (order_dual α) _ _ _ _ -@[simp] lemma pred_lt_iff_ne_bot : pred a < a ↔ a ≠ ⊥ := @lt_succ_iff_ne_top (order_dual α) _ _ _ _ +@[simp] lemma le_pred_iff_eq_bot : a ≤ pred a ↔ a = ⊥ := @succ_le_iff_eq_top αᵒᵈ _ _ _ _ +@[simp] lemma pred_lt_iff_ne_bot : pred a < a ↔ a ≠ ⊥ := @lt_succ_iff_ne_top αᵒᵈ _ _ _ _ end order_bot @@ -880,8 +880,7 @@ variables [preorder α] section succ_order variables [succ_order α] [is_succ_archimedean α] {a b : α} -instance : is_pred_archimedean (order_dual α) := -⟨λ a b h, by convert exists_succ_iterate_of_le h.of_dual⟩ +instance : is_pred_archimedean αᵒᵈ := ⟨λ a b h, by convert exists_succ_iterate_of_le h.of_dual⟩ lemma has_le.le.exists_succ_iterate (h : a ≤ b) : ∃ n, succ^[n] a = b := exists_succ_iterate_of_le h @@ -915,23 +914,22 @@ end succ_order section pred_order variables [pred_order α] [is_pred_archimedean α] {a b : α} -instance : is_succ_archimedean (order_dual α) := -⟨λ a b h, by convert exists_pred_iterate_of_le h.of_dual⟩ +instance : is_succ_archimedean αᵒᵈ := ⟨λ a b h, by convert exists_pred_iterate_of_le h.of_dual⟩ lemma has_le.le.exists_pred_iterate (h : a ≤ b) : ∃ n, pred^[n] b = a := exists_pred_iterate_of_le h lemma exists_pred_iterate_iff_le : (∃ n, pred^[n] b = a) ↔ a ≤ b := -@exists_succ_iterate_iff_le (order_dual α) _ _ _ _ _ +@exists_succ_iterate_iff_le αᵒᵈ _ _ _ _ _ /-- Induction principle on a type with a `pred_order` for all elements below a given element `m`. -/ @[elab_as_eliminator] lemma pred.rec {P : α → Prop} {m : α} (h0 : P m) (h1 : ∀ n, n ≤ m → P n → P (pred n)) ⦃n : α⦄ (hmn : n ≤ m) : P n := -@succ.rec (order_dual α) _ _ _ _ _ h0 h1 _ hmn +@succ.rec αᵒᵈ _ _ _ _ _ h0 h1 _ hmn lemma pred.rec_iff {p : α → Prop} (hsucc : ∀ a, p a ↔ p (pred a)) {a b : α} (h : a ≤ b) : p a ↔ p b := -(@succ.rec_iff (order_dual α) _ _ _ _ hsucc _ _ h).symm +(@succ.rec_iff αᵒᵈ _ _ _ _ hsucc _ _ h).symm end pred_order end preorder @@ -983,7 +981,7 @@ end⟩ @[priority 100] instance is_well_order.to_is_succ_archimedean [h : is_well_order α (>)] [succ_order α] : is_succ_archimedean α := -by convert @order_dual.is_succ_archimedean (order_dual α) _ _ _ +by convert @order_dual.is_succ_archimedean αᵒᵈ _ _ _ end is_well_order diff --git a/src/order/succ_pred/relation.lean b/src/order/succ_pred/relation.lean index 0767f09851d15..15a145eeb1af0 100644 --- a/src/order/succ_pred/relation.lean +++ b/src/order/succ_pred/relation.lean @@ -86,25 +86,25 @@ variables {α : Type*} [partial_order α] [pred_order α] [is_pred_archimedean for all `i` between `n` and `m`. -/ lemma refl_trans_gen_of_pred_of_ge (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc m n, r i (pred i)) (hnm : m ≤ n) : refl_trans_gen r n m := -@refl_trans_gen_of_succ_of_le (order_dual α) _ _ _ r n m (λ x hx, h x ⟨hx.2, hx.1⟩) hnm +@refl_trans_gen_of_succ_of_le αᵒᵈ _ _ _ r n m (λ x hx, h x ⟨hx.2, hx.1⟩) hnm /-- For `n ≤ m`, `(n, m)` is in the reflexive-transitive closure of `~` if `pred i ~ i` for all `i` between `n` and `m`. -/ lemma refl_trans_gen_of_pred_of_le (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc n m, r (pred i) i) (hmn : n ≤ m) : refl_trans_gen r n m := -@refl_trans_gen_of_succ_of_ge (order_dual α) _ _ _ r n m (λ x hx, h x ⟨hx.2, hx.1⟩) hmn +@refl_trans_gen_of_succ_of_ge αᵒᵈ _ _ _ r n m (λ x hx, h x ⟨hx.2, hx.1⟩) hmn /-- For `m < n`, `(n, m)` is in the transitive closure of a relation `~` for `n ≠ m` if `i ~ pred i` for all `i` between `n` and `m`. -/ lemma trans_gen_of_pred_of_gt (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc m n, r i (pred i)) (hnm : m < n) : trans_gen r n m := -@trans_gen_of_succ_of_lt (order_dual α) _ _ _ r _ _ (λ x hx, h x ⟨hx.2, hx.1⟩) hnm +@trans_gen_of_succ_of_lt αᵒᵈ _ _ _ r _ _ (λ x hx, h x ⟨hx.2, hx.1⟩) hnm /-- For `n < m`, `(n, m)` is in the transitive closure of a relation `~` for `n ≠ m` if `pred i ~ i` for all `i` between `n` and `m`. -/ lemma trans_gen_of_pred_of_lt (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc n m, r (pred i) i) (hmn : n < m) : trans_gen r n m := -@trans_gen_of_succ_of_gt (order_dual α) _ _ _ r _ _ (λ x hx, h x ⟨hx.2, hx.1⟩) hmn +@trans_gen_of_succ_of_gt αᵒᵈ _ _ _ r _ _ (λ x hx, h x ⟨hx.2, hx.1⟩) hmn end partial_pred @@ -115,7 +115,7 @@ variables {α : Type*} [linear_order α] [pred_order α] [is_pred_archimedean α for all `i` between `n` and `m`. -/ lemma refl_trans_gen_of_pred (r : α → α → Prop) {n m : α} (h1 : ∀ i ∈ Ioc m n, r i (pred i)) (h2 : ∀ i ∈ Ioc n m, r (pred i) i) : refl_trans_gen r n m := -@refl_trans_gen_of_succ (order_dual α) _ _ _ r n m (λ x hx, h1 x ⟨hx.2, hx.1⟩) +@refl_trans_gen_of_succ αᵒᵈ _ _ _ r n m (λ x hx, h1 x ⟨hx.2, hx.1⟩) (λ x hx, h2 x ⟨hx.2, hx.1⟩) /-- For `n ≠ m`, `(n, m)` is in the transitive closure of a relation `~` if `i ~ pred i` and @@ -123,14 +123,14 @@ lemma refl_trans_gen_of_pred (r : α → α → Prop) {n m : α} lemma trans_gen_of_pred_of_ne (r : α → α → Prop) {n m : α} (h1 : ∀ i ∈ Ioc m n, r i (pred i)) (h2 : ∀ i ∈ Ioc n m, r (pred i) i) (hnm : n ≠ m) : trans_gen r n m := -@trans_gen_of_succ_of_ne (order_dual α) _ _ _ r n m (λ x hx, h1 x ⟨hx.2, hx.1⟩) +@trans_gen_of_succ_of_ne αᵒᵈ _ _ _ r n m (λ x hx, h1 x ⟨hx.2, hx.1⟩) (λ x hx, h2 x ⟨hx.2, hx.1⟩) hnm /-- `(n, m)` is in the transitive closure of a reflexive relation `~` if `i ~ pred i` and `pred i ~ i` for all `i` between `n` and `m`. -/ lemma trans_gen_of_pred_of_reflexive (r : α → α → Prop) {n m : α} (hr : reflexive r) (h1 : ∀ i ∈ Ioc m n, r i (pred i)) (h2 : ∀ i ∈ Ioc n m, r (pred i) i) : trans_gen r n m := -@trans_gen_of_succ_of_reflexive (order_dual α) _ _ _ r n m hr (λ x hx, h1 x ⟨hx.2, hx.1⟩) +@trans_gen_of_succ_of_reflexive αᵒᵈ _ _ _ r n m hr (λ x hx, h1 x ⟨hx.2, hx.1⟩) (λ x hx, h2 x ⟨hx.2, hx.1⟩) end linear_pred diff --git a/src/order/upper_lower.lean b/src/order/upper_lower.lean index 18ca128ad0824..9ece02f8ee12f 100644 --- a/src/order/upper_lower.lean +++ b/src/order/upper_lower.lean @@ -109,9 +109,9 @@ lemma is_lower_set_sInter {S : set (set α)} (hf : ∀ s ∈ S, is_lower_set s) iff.rfl @[simp] lemma is_upper_set_preimage_of_dual_iff : is_upper_set (of_dual ⁻¹' s) ↔ is_lower_set s := iff.rfl -@[simp] lemma is_lower_set_preimage_to_dual_iff {s : set (order_dual α)} : +@[simp] lemma is_lower_set_preimage_to_dual_iff {s : set αᵒᵈ} : is_lower_set (to_dual ⁻¹' s) ↔ is_upper_set s := iff.rfl -@[simp] lemma is_upper_set_preimage_to_dual_iff {s : set (order_dual α)} : +@[simp] lemma is_upper_set_preimage_to_dual_iff {s : set αᵒᵈ} : is_upper_set (to_dual ⁻¹' s) ↔ is_lower_set s := iff.rfl alias is_lower_set_preimage_of_dual_iff ↔ _ is_upper_set.of_dual @@ -383,7 +383,7 @@ variables [semilattice_sup α] @[simp] lemma Ici_sup (a b : α) : Ici (a ⊔ b) = Ici a ⊓ Ici b := ext Ici_inter_Ici.symm /-- `upper_set.Ici` as a `sup_hom`. -/ -def Ici_sup_hom : sup_hom α (order_dual $ upper_set α) := ⟨Ici, Ici_sup⟩ +def Ici_sup_hom : sup_hom α (upper_set α)ᵒᵈ := ⟨Ici, Ici_sup⟩ @[simp] lemma Ici_sup_hom_apply (a : α) : Ici_sup_hom a = to_dual (Ici a) := rfl @@ -402,7 +402,7 @@ set_like.ext $ λ c, by simp only [mem_Ici_iff, mem_infi_iff, supr_le_iff] by simp_rw Ici_supr /-- `upper_set.Ici` as a `Sup_hom`. -/ -def Ici_Sup_hom : Sup_hom α (order_dual $ upper_set α) := +def Ici_Sup_hom : Sup_hom α (upper_set α)ᵒᵈ := ⟨Ici, λ s, (Ici_Sup s).trans Inf_image.symm⟩ @[simp] lemma Ici_Sup_hom_apply (a : α) : Ici_Sup_hom a = to_dual (Ici a) := rfl diff --git a/src/order/well_founded.lean b/src/order/well_founded.lean index 2642c4ba58431..2323747cf163d 100644 --- a/src/order/well_founded.lean +++ b/src/order/well_founded.lean @@ -94,7 +94,7 @@ by simp only [eq_iff_not_lt_of_le, well_founded_iff_has_min] theorem well_founded_iff_has_min' [partial_order α] : (well_founded (has_lt.lt : α → α → Prop)) ↔ ∀ (p : set α), p.nonempty → ∃ m ∈ p, ∀ x ∈ p, x ≤ m → x = m := -@well_founded_iff_has_max' (order_dual α) _ +@well_founded_iff_has_max' αᵒᵈ _ open set /-- The supremum of a bounded, well-founded order -/ diff --git a/src/order/well_founded_set.lean b/src/order/well_founded_set.lean index c3ca197b9ac11..d425dd2f43372 100644 --- a/src/order/well_founded_set.lean +++ b/src/order/well_founded_set.lean @@ -131,8 +131,7 @@ end has_lt section partial_order variables [partial_order α] {s t : set α} {a : α} -theorem is_wf_iff_no_descending_seq : - is_wf s ↔ ∀ (f : (order_dual ℕ) ↪o α), ¬ (range f) ⊆ s := +theorem is_wf_iff_no_descending_seq : is_wf s ↔ ∀ f : ℕᵒᵈ ↪o α, ¬ (range f) ⊆ s := begin haveI : is_strict_order α (λ (a b : α), a < b ∧ a ∈ s ∧ b ∈ s) := { to_is_irrefl := ⟨λ x con, lt_irrefl x con.1⟩, diff --git a/src/order/zorn.lean b/src/order/zorn.lean index ac7aa1ffe64b5..2804ab41623b0 100644 --- a/src/order/zorn.lean +++ b/src/order/zorn.lean @@ -168,13 +168,12 @@ zorn_nonempty_partial_order₀ _ (λ c cS hc y yc, H _ cS hc ⟨y, yc⟩) _ hx lemma zorn_superset (S : set (set α)) (h : ∀ c ⊆ S, is_chain (⊆) c → ∃ lb ∈ S, ∀ s ∈ c, lb ⊆ s) : ∃ m ∈ S, ∀ a ∈ S, a ⊆ m → a = m := -@zorn_partial_order₀ (order_dual (set α)) _ S $ λ c cS hc, h c cS hc.symm +@zorn_partial_order₀ (set α)ᵒᵈ _ S $ λ c cS hc, h c cS hc.symm lemma zorn_superset_nonempty (S : set (set α)) (H : ∀ c ⊆ S, is_chain (⊆) c → c.nonempty → ∃ lb ∈ S, ∀ s ∈ c, lb ⊆ s) (x) (hx : x ∈ S) : ∃ m ∈ S, m ⊆ x ∧ ∀ a ∈ S, a ⊆ m → a = m := -@zorn_nonempty_partial_order₀ (order_dual (set α)) _ S (λ c cS hc y yc, H _ cS - hc.symm ⟨y, yc⟩) _ hx +@zorn_nonempty_partial_order₀ (set α)ᵒᵈ _ S (λ c cS hc y yc, H _ cS hc.symm ⟨y, yc⟩) _ hx /-- Every chain is contained in a maximal chain. This generalizes Hausdorff's maximality principle. -/ diff --git a/src/ring_theory/artinian.lean b/src/ring_theory/artinian.lean index 3add0c5b5d499..4f32cb64d8db4 100644 --- a/src/ring_theory/artinian.lean +++ b/src/ring_theory/artinian.lean @@ -201,12 +201,12 @@ set_has_minimal_iff_artinian.mpr ‹_› a ha /-- A module is Artinian iff every decreasing chain of submodules stabilizes. -/ theorem monotone_stabilizes_iff_artinian : - (∀ (f : ℕ →o order_dual (submodule R M)), ∃ n, ∀ m, n ≤ m → f n = f m) + (∀ (f : ℕ →o (submodule R M)ᵒᵈ), ∃ n, ∀ m, n ≤ m → f n = f m) ↔ is_artinian R M := by rw [is_artinian_iff_well_founded]; - exact (well_founded.monotone_chain_condition (order_dual (submodule R M))).symm + exact (well_founded.monotone_chain_condition (submodule R M)ᵒᵈ).symm -theorem is_artinian.monotone_stabilizes [is_artinian R M] (f : ℕ →o order_dual (submodule R M)) : +theorem is_artinian.monotone_stabilizes [is_artinian R M] (f : ℕ →o (submodule R M)ᵒᵈ) : ∃ n, ∀ m, n ≤ m → f n = f m := monotone_stabilizes_iff_artinian.mpr ‹_› f @@ -399,7 +399,7 @@ variables {R : Type*} [comm_ring R] [is_artinian_ring R] lemma is_nilpotent_jacobson_bot : is_nilpotent (ideal.jacobson (⊥ : ideal R)) := begin let Jac := ideal.jacobson (⊥ : ideal R), - let f : ℕ →o order_dual (ideal R) := ⟨λ n, Jac ^ n, λ _ _ h, ideal.pow_le_pow h⟩, + let f : ℕ →o (ideal R)ᵒᵈ := ⟨λ n, Jac ^ n, λ _ _ h, ideal.pow_le_pow h⟩, obtain ⟨n, hn⟩ : ∃ n, ∀ m, n ≤ m → Jac ^ n = Jac ^ m := is_artinian.monotone_stabilizes f, refine ⟨n, _⟩, let J : ideal R := annihilator (Jac ^ n), diff --git a/src/ring_theory/nullstellensatz.lean b/src/ring_theory/nullstellensatz.lean index e182530ddc107..91682491d504a 100644 --- a/src/ring_theory/nullstellensatz.lean +++ b/src/ring_theory/nullstellensatz.lean @@ -76,7 +76,7 @@ lemma zero_locus_vanishing_ideal_le (V : set (σ → k)) : λ V hV p hp, hp V hV theorem zero_locus_vanishing_ideal_galois_connection : - @galois_connection (ideal (mv_polynomial σ k)) (order_dual (set (σ → k))) _ _ + @galois_connection (ideal (mv_polynomial σ k)) (set (σ → k))ᵒᵈ _ _ zero_locus vanishing_ideal := λ I V, ⟨λ h, le_trans (le_vanishing_ideal_zero_locus I) (vanishing_ideal_anti_mono h), λ h, le_trans (zero_locus_anti_mono h) (zero_locus_vanishing_ideal_le V)⟩ diff --git a/src/ring_theory/valuation/basic.lean b/src/ring_theory/valuation/basic.lean index 4375641751d0c..e966286baa593 100644 --- a/src/ring_theory/valuation/basic.lean +++ b/src/ring_theory/valuation/basic.lean @@ -47,7 +47,7 @@ on R / J = `ideal.quotient J` is `on_quot v h`. ## Implementation Details -`add_valuation R Γ₀` is implemented as `valuation R (multiplicative (order_dual Γ₀))`. +`add_valuation R Γ₀` is implemented as `valuation R (multiplicative Γ₀)ᵒᵈ`. ## TODO @@ -504,7 +504,7 @@ variables (R) [ring R] (Γ₀ : Type*) [linear_ordered_add_comm_monoid_with_top /-- The type of `Γ₀`-valued additive valuations on `R`. -/ @[nolint has_inhabited_instance] -def add_valuation := valuation R (multiplicative (order_dual Γ₀)) +def add_valuation := valuation R (multiplicative Γ₀ᵒᵈ) end add_monoid @@ -529,8 +529,7 @@ section variables (f : R → Γ₀) (h0 : f 0 = ⊤) (h1 : f 1 = 0) variables (hadd : ∀ x y, min (f x) (f y) ≤ f (x + y)) (hmul : ∀ x y, f (x * y) = f x + f y) -/-- An alternate constructor of `add_valuation`, that doesn't reference - `multiplicative (order_dual Γ₀)` -/ +/-- An alternate constructor of `add_valuation`, that doesn't reference `multiplicative Γ₀ᵒᵈ` -/ def of : add_valuation R Γ₀ := { to_fun := f, map_one' := h1, @@ -545,7 +544,7 @@ theorem of_apply : (of f h0 h1 hadd hmul) r = f r := rfl /-- The `valuation` associated to an `add_valuation` (useful if the latter is constructed using `add_valuation.of`). -/ -def valuation : valuation R (multiplicative (order_dual Γ₀)) := v +def valuation : valuation R (multiplicative Γ₀ᵒᵈ) := v @[simp] lemma valuation_apply (r : R) : v.valuation r = multiplicative.of_add (order_dual.to_dual (v r)) := rfl diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index 83bfc7285a718..f814c27e93924 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -284,8 +284,7 @@ def prime_spectrum_equiv : /-- An ordered variant of `prime_spectrum_equiv`. -/ @[simps] -def prime_spectrum_order_equiv : - order_dual (prime_spectrum A) ≃o { S | A ≤ S } := +def prime_spectrum_order_equiv : (prime_spectrum A)ᵒᵈ ≃o {S | A ≤ S} := { map_rel_iff' := λ P Q, ⟨ λ h, begin have := ideal_of_le_le_of_le A _ _ _ _ h, diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index f1cb85b76af35..5894f7a29336b 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -90,17 +90,15 @@ generally, and suffices to derive many interesting properties relating order and class order_closed_topology (α : Type*) [topological_space α] [preorder α] : Prop := (is_closed_le' : is_closed {p:α×α | p.1 ≤ p.2}) -instance : Π [topological_space α], topological_space (order_dual α) := id +instance : Π [topological_space α], topological_space αᵒᵈ := id -instance [topological_space α] [h : first_countable_topology α] : - first_countable_topology (order_dual α) := h +instance [topological_space α] [h : first_countable_topology α] : first_countable_topology αᵒᵈ := h -instance [topological_space α] [h : second_countable_topology α] : - second_countable_topology (order_dual α) := h +instance [topological_space α] [h : second_countable_topology α] : second_countable_topology αᵒᵈ := +h @[to_additive] -instance [topological_space α] [has_mul α] [h : has_continuous_mul α] : - has_continuous_mul (order_dual α) := h +instance [topological_space α] [has_mul α] [h : has_continuous_mul α] : has_continuous_mul αᵒᵈ := h lemma dense.order_dual [topological_space α] {s : set α} (hs : dense s) : dense (order_dual.of_dual ⁻¹' s) := hs @@ -140,7 +138,7 @@ is_closed_le continuous_const continuous_id lemma is_closed_Ici {a : α} : is_closed (Ici a) := is_closed_ge' a -instance : order_closed_topology (order_dual α) := +instance : order_closed_topology αᵒᵈ := ⟨(@order_closed_topology.is_closed_le' α _ _ _).preimage continuous_swap⟩ lemma is_closed_Icc {a b : α} : is_closed (Icc a b) := @@ -534,8 +532,7 @@ end lemma frontier_Iic_subset (a : α) : frontier (Iic a) ⊆ {a} := frontier_le_subset_eq (@continuous_id α _) continuous_const -lemma frontier_Ici_subset (a : α) : frontier (Ici a) ⊆ {a} := -@frontier_Iic_subset (order_dual α) _ _ _ _ +lemma frontier_Ici_subset (a : α) : frontier (Ici a) ⊆ {a} := @frontier_Iic_subset αᵒᵈ _ _ _ _ lemma frontier_lt_subset_eq (hf : continuous f) (hg : continuous g) : frontier {b | f b < g b} ⊆ {b | f b = g b} := @@ -581,7 +578,7 @@ by { simp only [min_def], exact hf.if_le hg hf hg (λ x, id) } @[continuity] lemma continuous.max (hf : continuous f) (hg : continuous g) : continuous (λb, max (f b) (g b)) := -@continuous.min (order_dual α) _ _ _ _ _ _ _ hf hg +@continuous.min αᵒᵈ _ _ _ _ _ _ _ hf hg end @@ -646,7 +643,7 @@ end /-- A compact set is bounded above -/ lemma is_compact.bdd_above {s : set α} (hs : is_compact s) : bdd_above s := -@is_compact.bdd_below (order_dual α) _ _ _ _ _ hs +@is_compact.bdd_below αᵒᵈ _ _ _ _ _ hs /-- A continuous function is bounded below on a compact set. -/ lemma is_compact.bdd_below_image {f : β → α} {K : set β} @@ -656,7 +653,7 @@ lemma is_compact.bdd_below_image {f : β → α} {K : set β} /-- A continuous function is bounded above on a compact set. -/ lemma is_compact.bdd_above_image {f : β → α} {K : set β} (hK : is_compact K) (hf : continuous_on f K) : bdd_above (f '' K) := -@is_compact.bdd_below_image (order_dual α) _ _ _ _ _ _ _ _ hK hf +@is_compact.bdd_below_image αᵒᵈ _ _ _ _ _ _ _ _ hK hf /-- A continuous function with compact support is bounded below. -/ @[to_additive /-" A continuous function with compact support is bounded below. "-/] @@ -669,7 +666,7 @@ lemma continuous.bdd_below_range_of_has_compact_mul_support [has_one α] {f : β lemma continuous.bdd_above_range_of_has_compact_mul_support [has_one α] {f : β → α} (hf : continuous f) (h : has_compact_mul_support f) : bdd_above (range f) := -@continuous.bdd_below_range_of_has_compact_mul_support (order_dual α) _ _ _ _ _ _ _ _ hf h +@continuous.bdd_below_range_of_has_compact_mul_support αᵒᵈ _ _ _ _ _ _ _ _ hf h end linear_order @@ -711,7 +708,7 @@ generate_from {s : set α | ∃ (a : α), s = {b : α | a < b} ∨ s = {b : α | section order_topology instance {α : Type*} [topological_space α] [partial_order α] [order_topology α] : - order_topology (order_dual α) := + order_topology αᵒᵈ := ⟨by convert @order_topology.topology_eq_generate_intervals α _ _ _; conv in (_ ∨ _) { rw or.comm }; refl⟩ @@ -944,7 +941,7 @@ lemma nhds_top_basis [topological_space α] [linear_order α] [order_top α] [or lemma nhds_bot_basis [topological_space α] [linear_order α] [order_bot α] [order_topology α] [nontrivial α] : (𝓝 ⊥).has_basis (λ a : α, ⊥ < a) (λ a : α, Iio a) := -@nhds_top_basis (order_dual α) _ _ _ _ _ +@nhds_top_basis αᵒᵈ _ _ _ _ _ lemma nhds_top_basis_Ici [topological_space α] [linear_order α] [order_top α] [order_topology α] [nontrivial α] [densely_ordered α] : @@ -956,7 +953,7 @@ nhds_top_basis.to_has_basis lemma nhds_bot_basis_Iic [topological_space α] [linear_order α] [order_bot α] [order_topology α] [nontrivial α] [densely_ordered α] : (𝓝 ⊥).has_basis (λ a : α, ⊥ < a) Iic := -@nhds_top_basis_Ici (order_dual α) _ _ _ _ _ _ +@nhds_top_basis_Ici αᵒᵈ _ _ _ _ _ _ lemma tendsto_nhds_top_mono [topological_space β] [partial_order β] [order_top β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊤)) (hg : f ≤ᶠ[l] g) : @@ -970,7 +967,7 @@ end lemma tendsto_nhds_bot_mono [topological_space β] [partial_order β] [order_bot β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊥)) (hg : g ≤ᶠ[l] f) : tendsto g l (𝓝 ⊥) := -@tendsto_nhds_top_mono α (order_dual β) _ _ _ _ _ _ _ hf hg +@tendsto_nhds_top_mono α βᵒᵈ _ _ _ _ _ _ _ hf hg lemma tendsto_nhds_top_mono' [topological_space β] [partial_order β] [order_top β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊤)) (hg : f ≤ g) : @@ -1185,7 +1182,7 @@ lemma pi_Iio_mem_nhds' (ha : ∀ i, x' i < a' i) : Iio a' ∈ 𝓝 x' := pi_Iio_mem_nhds ha lemma pi_Ioi_mem_nhds (ha : ∀ i, a i < x i) : Ioi a ∈ 𝓝 x := -@pi_Iio_mem_nhds ι (λ i, order_dual (π i)) _ _ _ _ _ _ _ ha +@pi_Iio_mem_nhds ι (λ i, (π i)ᵒᵈ) _ _ _ _ _ _ _ ha lemma pi_Ioi_mem_nhds' (ha : ∀ i, a' i < x' i) : Ioi a' ∈ 𝓝 x' := pi_Ioi_mem_nhds ha @@ -1234,13 +1231,11 @@ end 𝓝 x ⊓ at_top = ⊥ := disjoint_iff.1 (disjoint_nhds_at_top x) -lemma disjoint_nhds_at_bot [no_min_order α] (x : α) : - disjoint (𝓝 x) at_bot := -@disjoint_nhds_at_top (order_dual α) _ _ _ _ x +lemma disjoint_nhds_at_bot [no_min_order α] (x : α) : disjoint (𝓝 x) at_bot := +@disjoint_nhds_at_top αᵒᵈ _ _ _ _ x -@[simp] lemma inf_nhds_at_bot [no_min_order α] (x : α) : - 𝓝 x ⊓ at_bot = ⊥ := -@inf_nhds_at_top (order_dual α) _ _ _ _ x +@[simp] lemma inf_nhds_at_bot [no_min_order α] (x : α) : 𝓝 x ⊓ at_bot = ⊥ := +@inf_nhds_at_top αᵒᵈ _ _ _ _ x lemma not_tendsto_nhds_of_tendsto_at_top [no_max_order α] {F : filter β} [ne_bot F] {f : β → α} (hf : tendsto f F at_top) (x : α) : @@ -1469,8 +1464,8 @@ with `l < a`. -/ lemma mem_nhds_within_Iic_iff_exists_Icc_subset' [no_min_order α] [densely_ordered α] {a : α} {s : set α} : s ∈ 𝓝[≤] a ↔ ∃l ∈ Iio a, Icc l a ⊆ s := begin - convert @mem_nhds_within_Ici_iff_exists_Icc_subset' (order_dual α) _ _ _ _ _ _ _, - simp_rw (show ∀ u : order_dual α, @Icc (order_dual α) _ a u = @Icc α _ u a, from λ u, dual_Icc), + convert @mem_nhds_within_Ici_iff_exists_Icc_subset' αᵒᵈ _ _ _ _ _ _ _, + simp_rw (show ∀ u : αᵒᵈ, @Icc αᵒᵈ _ a u = @Icc α _ u a, from λ u, dual_Icc), refl, end @@ -1661,7 +1656,7 @@ end and `g` tends to `at_bot` then `f + g` tends to `at_bot`. -/ lemma filter.tendsto.add_at_bot {C : α} (hf : tendsto f l (𝓝 C)) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@filter.tendsto.add_at_top (order_dual α) _ _ _ _ _ _ _ _ hf hg +@filter.tendsto.add_at_top αᵒᵈ _ _ _ _ _ _ _ _ hf hg /-- In a linearly ordered additive commutative group with the order topology, if `f` tends to `at_top` and `g` tends to `C` then `f + g` tends to `at_top`. -/ @@ -1984,7 +1979,7 @@ lemma is_lub.frequently_nhds_mem {a : α} {s : set α} (ha : is_lub s a) (hs : s lemma is_glb.frequently_mem {a : α} {s : set α} (ha : is_glb s a) (hs : s.nonempty) : ∃ᶠ x in 𝓝[≥] a, x ∈ s := -@is_lub.frequently_mem (order_dual α) _ _ _ _ _ ha hs +@is_lub.frequently_mem αᵒᵈ _ _ _ _ _ ha hs lemma is_glb.frequently_nhds_mem {a : α} {s : set α} (ha : is_glb s a) (hs : s.nonempty) : ∃ᶠ x in 𝓝 a, x ∈ s := @@ -2004,7 +1999,7 @@ mem_closure_iff_nhds_within_ne_bot.1 (ha.mem_closure hs) lemma is_glb.nhds_within_ne_bot : ∀ {a : α} {s : set α}, is_glb s a → s.nonempty → ne_bot (𝓝[s] a) := -@is_lub.nhds_within_ne_bot (order_dual α) _ _ _ +@is_lub.nhds_within_ne_bot αᵒᵈ _ _ _ lemma is_lub_of_mem_nhds {s : set α} {a : α} {f : filter α} (hsa : a ∈ upper_bounds s) (hsf : s ∈ f) [ne_bot (f ⊓ 𝓝 a)] : is_lub s a := @@ -2026,11 +2021,11 @@ end lemma is_glb_of_mem_nhds : ∀ {s : set α} {a : α} {f : filter α}, a ∈ lower_bounds s → s ∈ f → ne_bot (f ⊓ 𝓝 a) → is_glb s a := -@is_lub_of_mem_nhds (order_dual α) _ _ _ +@is_lub_of_mem_nhds αᵒᵈ _ _ _ lemma is_glb_of_mem_closure {s : set α} {a : α} (hsa : a ∈ lower_bounds s) (hsf : a ∈ closure s) : is_glb s a := -@is_lub_of_mem_closure (order_dual α) _ _ _ s a hsa hsf +@is_lub_of_mem_closure αᵒᵈ _ _ _ s a hsa hsf lemma is_lub.mem_upper_bounds_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] {f : α → γ} {s : set α} {a : α} {b : γ} @@ -2060,7 +2055,7 @@ lemma is_glb.mem_lower_bounds_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] {f : α → γ} {s : set α} {a : α} {b : γ} (hf : monotone_on f s) (ha : is_glb s a) (hb : tendsto f (𝓝[s] a) (𝓝 b)) : b ∈ lower_bounds (f '' s) := -@is_lub.mem_upper_bounds_of_tendsto (order_dual α) (order_dual γ) _ _ _ _ _ _ _ _ _ _ hf.dual ha hb +@is_lub.mem_upper_bounds_of_tendsto αᵒᵈ γᵒᵈ _ _ _ _ _ _ _ _ _ _ hf.dual ha hb -- For a version of this theorem in which the convergence considered on the domain `α` is as -- `x : α` tends to negative infinity, rather than tending to a point `x` in `α`, see @@ -2069,31 +2064,31 @@ lemma is_glb.is_glb_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] {f : α → γ} {s : set α} {a : α} {b : γ} (hf : monotone_on f s) : is_glb s a → s.nonempty → tendsto f (𝓝[s] a) (𝓝 b) → is_glb (f '' s) b := -@is_lub.is_lub_of_tendsto (order_dual α) (order_dual γ) _ _ _ _ _ _ f s a b hf.dual +@is_lub.is_lub_of_tendsto αᵒᵈ γᵒᵈ _ _ _ _ _ _ f s a b hf.dual lemma is_lub.mem_lower_bounds_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] {f : α → γ} {s : set α} {a : α} {b : γ} (hf : antitone_on f s) (ha : is_lub s a) (hb : tendsto f (𝓝[s] a) (𝓝 b)) : b ∈ lower_bounds (f '' s) := -@is_lub.mem_upper_bounds_of_tendsto α (order_dual γ) _ _ _ _ _ _ _ _ _ _ hf ha hb +@is_lub.mem_upper_bounds_of_tendsto α γᵒᵈ _ _ _ _ _ _ _ _ _ _ hf ha hb lemma is_lub.is_glb_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] : ∀ {f : α → γ} {s : set α} {a : α} {b : γ}, (antitone_on f s) → is_lub s a → s.nonempty → tendsto f (𝓝[s] a) (𝓝 b) → is_glb (f '' s) b := -@is_lub.is_lub_of_tendsto α (order_dual γ) _ _ _ _ _ _ +@is_lub.is_lub_of_tendsto α γᵒᵈ _ _ _ _ _ _ lemma is_glb.mem_upper_bounds_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] {f : α → γ} {s : set α} {a : α} {b : γ} (hf : antitone_on f s) (ha : is_glb s a) (hb : tendsto f (𝓝[s] a) (𝓝 b)) : b ∈ upper_bounds (f '' s) := -@is_glb.mem_lower_bounds_of_tendsto α (order_dual γ) _ _ _ _ _ _ _ _ _ _ hf ha hb +@is_glb.mem_lower_bounds_of_tendsto α γᵒᵈ _ _ _ _ _ _ _ _ _ _ hf ha hb lemma is_glb.is_lub_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] : ∀ {f : α → γ} {s : set α} {a : α} {b : γ}, (antitone_on f s) → is_glb s a → s.nonempty → tendsto f (𝓝[s] a) (𝓝 b) → is_lub (f '' s) b := -@is_glb.is_glb_of_tendsto α (order_dual γ) _ _ _ _ _ _ +@is_glb.is_glb_of_tendsto α γᵒᵈ _ _ _ _ _ _ lemma is_lub.mem_of_is_closed {a : α} {s : set α} (ha : is_lub s a) (hs : s.nonempty) (sc : is_closed s) : a ∈ s := @@ -2191,13 +2186,13 @@ lemma is_glb.exists_seq_strict_anti_tendsto_of_not_mem {t : set α} {x : α} [is_countably_generated (𝓝 x)] (htx : is_glb t x) (not_mem : x ∉ t) (ht : t.nonempty) : ∃ u : ℕ → α, strict_anti u ∧ (∀ n, x < u n) ∧ tendsto u at_top (𝓝 x) ∧ (∀ n, u n ∈ t) := -@is_lub.exists_seq_strict_mono_tendsto_of_not_mem (order_dual α) _ _ _ t x _ htx not_mem ht +@is_lub.exists_seq_strict_mono_tendsto_of_not_mem αᵒᵈ _ _ _ t x _ htx not_mem ht lemma is_glb.exists_seq_antitone_tendsto {t : set α} {x : α} [is_countably_generated (𝓝 x)] (htx : is_glb t x) (ht : t.nonempty) : ∃ u : ℕ → α, antitone u ∧ (∀ n, x ≤ u n) ∧ tendsto u at_top (𝓝 x) ∧ (∀ n, u n ∈ t) := -@is_lub.exists_seq_monotone_tendsto (order_dual α) _ _ _ t x _ htx ht +@is_lub.exists_seq_monotone_tendsto αᵒᵈ _ _ _ t x _ htx ht lemma exists_seq_strict_anti_tendsto' [densely_ordered α] [first_countable_topology α] {x y : α} (hy : x < y) : @@ -2207,7 +2202,7 @@ by simpa only [dual_Ioo] using exists_seq_strict_mono_tendsto' (order_dual.to_du lemma exists_seq_strict_anti_tendsto [densely_ordered α] [no_max_order α] [first_countable_topology α] (x : α) : ∃ u : ℕ → α, strict_anti u ∧ (∀ n, x < u n) ∧ tendsto u at_top (𝓝 x) := -@exists_seq_strict_mono_tendsto (order_dual α) _ _ _ _ _ _ x +@exists_seq_strict_mono_tendsto αᵒᵈ _ _ _ _ _ _ x lemma exists_seq_strict_anti_strict_mono_tendsto [densely_ordered α] [first_countable_topology α] {x y : α} (h : x < y) : @@ -2224,7 +2219,7 @@ lemma exists_seq_tendsto_Inf {α : Type*} [conditionally_complete_linear_order [topological_space α] [order_topology α] [first_countable_topology α] {S : set α} (hS : S.nonempty) (hS' : bdd_below S) : ∃ (u : ℕ → α), antitone u ∧ tendsto u at_top (𝓝 (Inf S)) ∧ (∀ n, u n ∈ S) := -@exists_seq_tendsto_Sup (order_dual α) _ _ _ _ S hS hS' +@exists_seq_tendsto_Sup αᵒᵈ _ _ _ _ S hS hS' end order_topology @@ -2251,9 +2246,7 @@ closure_Ioi' nonempty_Ioi /-- The closure of the interval `(-∞, a)` is the closed interval `(-∞, a]`, unless `a` is a bottom element. -/ -lemma closure_Iio' {a : α} (h : (Iio a).nonempty) : - closure (Iio a) = Iic a := -@closure_Ioi' (order_dual α) _ _ _ _ _ h +lemma closure_Iio' (h : (Iio a).nonempty) : closure (Iio a) = Iic a := @closure_Ioi' αᵒᵈ _ _ _ _ _ h /-- The closure of the interval `(-∞, a)` is the interval `(-∞, a]`. -/ @[simp] lemma closure_Iio (a : α) [no_min_order α] : @@ -2301,7 +2294,7 @@ lemma interior_Ici [no_min_order α] {a : α} : interior (Ici a) = Ioi a := interior_Ici' nonempty_Iio @[simp] lemma interior_Iic' {a : α} (ha : (Ioi a).nonempty) : interior (Iic a) = Iio a := -@interior_Ici' (order_dual α) _ _ _ _ _ ha +@interior_Ici' αᵒᵈ _ _ _ _ _ ha lemma interior_Iic [no_max_order α] {a : α} : interior (Iic a) = Iio a := interior_Iic' nonempty_Ioi @@ -2413,7 +2406,7 @@ nhds_within_Iio_ne_bot (le_refl a) lemma filter.eventually.exists_lt [no_min_order α] {a : α} {p : α → Prop} (h : ∀ᶠ x in 𝓝 a, p x) : ∃ b < a, p b := -@filter.eventually.exists_gt (order_dual α) _ _ _ _ _ _ _ h +@filter.eventually.exists_gt αᵒᵈ _ _ _ _ _ _ _ h lemma right_nhds_within_Ico_ne_bot {a b : α} (H : a < b) : ne_bot (𝓝[Ico a b] b) := (is_lub_Ico H).nhds_within_ne_bot (nonempty_Ico.2 H) @@ -2496,7 +2489,7 @@ comap_coe_nhds_within_Ioi_of_Ioo_subset (subset.refl _) $ lemma comap_coe_Iio_nhds_within_Iio (a : α) : comap (coe : Iio a → α) (𝓝[<] a) = at_top := -@comap_coe_Ioi_nhds_within_Ioi (order_dual α) _ _ _ _ a +@comap_coe_Ioi_nhds_within_Ioi αᵒᵈ _ _ _ _ a @[simp] lemma map_coe_Ioo_at_top {a b : α} (h : a < b) : map (coe : Ioo a b → α) at_top = 𝓝[<] b := @@ -2512,7 +2505,7 @@ map_coe_at_bot_of_Ioo_subset (subset.refl _) $ λ b hb, ⟨b, hb, Ioo_subset_Ioi @[simp] lemma map_coe_Iio_at_top (a : α) : map (coe : Iio a → α) at_top = 𝓝[<] a := -@map_coe_Ioi_at_bot (order_dual α) _ _ _ _ _ +@map_coe_Ioi_at_bot αᵒᵈ _ _ _ _ _ variables {l : filter β} {f : α → β} @@ -2667,32 +2660,28 @@ the infimum of the image of this set. -/ lemma map_Inf_of_continuous_at_of_monotone' {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) (Mf : monotone f) (hs : s.nonempty) : f (Inf s) = Inf (f '' s) := -@map_Sup_of_continuous_at_of_monotone' (order_dual α) (order_dual β) _ _ _ _ _ _ f s Cf - Mf.dual hs +@map_Sup_of_continuous_at_of_monotone' αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual hs /-- A monotone function `s` sending `top` to `top` and continuous at the infimum of a set sends this infimum to the infimum of the image of this set. -/ lemma map_Inf_of_continuous_at_of_monotone {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) (Mf : monotone f) (ftop : f ⊤ = ⊤) : f (Inf s) = Inf (f '' s) := -@map_Sup_of_continuous_at_of_monotone (order_dual α) (order_dual β) _ _ _ _ _ _ f s Cf - Mf.dual ftop +@map_Sup_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual ftop /-- A monotone function continuous at the indexed infimum over a nonempty `Sort` sends this indexed infimum to the indexed infimum of the composition. -/ lemma map_infi_of_continuous_at_of_monotone' {ι : Sort*} [nonempty ι] {f : α → β} {g : ι → α} (Cf : continuous_at f (infi g)) (Mf : monotone f) : f (⨅ i, g i) = ⨅ i, f (g i) := -@map_supr_of_continuous_at_of_monotone' (order_dual α) (order_dual β) _ _ _ _ _ _ ι _ f g Cf - Mf.dual +@map_supr_of_continuous_at_of_monotone' αᵒᵈ βᵒᵈ _ _ _ _ _ _ ι _ f g Cf Mf.dual /-- If a monotone function sending `top` to `top` is continuous at the indexed infimum over a `Sort`, then it sends this indexed infimum to the indexed infimum of the composition. -/ lemma map_infi_of_continuous_at_of_monotone {ι : Sort*} {f : α → β} {g : ι → α} (Cf : continuous_at f (infi g)) (Mf : monotone f) (ftop : f ⊤ = ⊤) : f (infi g) = infi (f ∘ g) := -@map_supr_of_continuous_at_of_monotone (order_dual α) (order_dual β) _ _ _ _ _ _ ι f g Cf - Mf.dual ftop +@map_supr_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ ι f g Cf Mf.dual ftop end complete_linear_order @@ -2738,16 +2727,14 @@ then it sends this infimum to the infimum of the image of `s`. -/ lemma map_cInf_of_continuous_at_of_monotone {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) (Mf : monotone f) (ne : s.nonempty) (H : bdd_below s) : f (Inf s) = Inf (f '' s) := -@map_cSup_of_continuous_at_of_monotone (order_dual α) (order_dual β) _ _ _ _ _ _ f s Cf - Mf.dual ne H +@map_cSup_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual ne H /-- A continuous monotone function sends indexed infimum to indexed infimum in conditionally complete linear order, under a boundedness assumption. -/ lemma map_cinfi_of_continuous_at_of_monotone {f : α → β} {g : γ → α} (Cf : continuous_at f (⨅ i, g i)) (Mf : monotone f) (H : bdd_below (range g)) : f (⨅ i, g i) = ⨅ i, f (g i) := -@map_csupr_of_continuous_at_of_monotone (order_dual α) (order_dual β) _ _ _ _ _ _ _ _ _ _ - Cf Mf.dual H +@map_csupr_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ _ _ _ _ Cf Mf.dual H /-- A monotone map has a limit to the left of any point `x`, equal to `Sup (f '' (Iio x))`. -/ lemma monotone.tendsto_nhds_within_Iio @@ -2772,7 +2759,7 @@ lemma monotone.tendsto_nhds_within_Ioi {α : Type*} [linear_order α] [topological_space α] [order_topology α] {f : α → β} (Mf : monotone f) (x : α) : tendsto f (𝓝[>] x) (𝓝 (Inf (f '' (Ioi x)))) := -@monotone.tendsto_nhds_within_Iio (order_dual β) _ _ _ (order_dual α) _ _ _ f Mf.dual x +@monotone.tendsto_nhds_within_Iio βᵒᵈ _ _ _ αᵒᵈ _ _ _ f Mf.dual x end conditionally_complete_linear_order diff --git a/src/topology/algebra/order/compact.lean b/src/topology/algebra/order/compact.lean index 995c7f12478dd..469fa59f229be 100644 --- a/src/topology/algebra/order/compact.lean +++ b/src/topology/algebra/order/compact.lean @@ -142,9 +142,8 @@ lemma is_compact.Inf_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : Inf s ∈ s := hs.is_closed.cInf_mem ne_s hs.bdd_below -lemma is_compact.Sup_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : - Sup s ∈ s := -@is_compact.Inf_mem (order_dual α) _ _ _ _ hs ne_s +lemma is_compact.Sup_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : Sup s ∈ s := +@is_compact.Inf_mem αᵒᵈ _ _ _ _ hs ne_s lemma is_compact.is_glb_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : is_glb s (Inf s) := @@ -152,7 +151,7 @@ is_glb_cInf ne_s hs.bdd_below lemma is_compact.is_lub_Sup {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : is_lub s (Sup s) := -@is_compact.is_glb_Inf (order_dual α) _ _ _ _ hs ne_s +@is_compact.is_glb_Inf αᵒᵈ _ _ _ _ hs ne_s lemma is_compact.is_least_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : is_least s (Inf s) := @@ -160,7 +159,7 @@ lemma is_compact.is_least_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempt lemma is_compact.is_greatest_Sup {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : is_greatest s (Sup s) := -@is_compact.is_least_Inf (order_dual α) _ _ _ _ hs ne_s +@is_compact.is_least_Inf αᵒᵈ _ _ _ _ hs ne_s lemma is_compact.exists_is_least {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : ∃ x, is_least s x := @@ -188,7 +187,7 @@ in ⟨x, hxs, hx.symm, λ y hy, lemma is_compact.exists_Sup_image_eq_and_ge {s : set β} (hs : is_compact s) (ne_s : s.nonempty) {f : β → α} (hf : continuous_on f s) : ∃ x ∈ s, Sup (f '' s) = f x ∧ ∀ y ∈ s, f y ≤ f x := -@is_compact.exists_Inf_image_eq_and_le (order_dual α) _ _ _ _ _ _ hs ne_s _ hf +@is_compact.exists_Inf_image_eq_and_le αᵒᵈ _ _ _ _ _ _ hs ne_s _ hf lemma is_compact.exists_Inf_image_eq {s : set β} (hs : is_compact s) (ne_s : s.nonempty) {f : β → α} (hf : continuous_on f s) : @@ -198,7 +197,7 @@ let ⟨x, hxs, hx, _⟩ := hs.exists_Inf_image_eq_and_le ne_s hf in ⟨x, hxs, h lemma is_compact.exists_Sup_image_eq : ∀ {s : set β}, is_compact s → s.nonempty → ∀ {f : β → α}, continuous_on f s → ∃ x ∈ s, Sup (f '' s) = f x := -@is_compact.exists_Inf_image_eq (order_dual α) _ _ _ _ _ +@is_compact.exists_Inf_image_eq αᵒᵈ _ _ _ _ _ lemma eq_Icc_of_connected_compact {s : set α} (h₁ : is_connected s) (h₂ : is_compact s) : s = Icc (Inf s) (Sup s) := @@ -222,7 +221,7 @@ end lemma is_compact.exists_forall_ge : ∀ {s : set β}, is_compact s → s.nonempty → ∀ {f : β → α}, continuous_on f s → ∃x∈s, ∀y∈s, f y ≤ f x := -@is_compact.exists_forall_le (order_dual α) _ _ _ _ _ +@is_compact.exists_forall_le αᵒᵈ _ _ _ _ _ /-- The **extreme value theorem**: if a function `f` is continuous on a closed set `s` and it is larger than a value in its image away from compact sets, then it has a minimum on this set. -/ @@ -244,7 +243,7 @@ smaller than a value in its image away from compact sets, then it has a maximum lemma continuous_on.exists_forall_ge' {s : set β} {f : β → α} (hf : continuous_on f s) (hsc : is_closed s) {x₀ : β} (h₀ : x₀ ∈ s) (hc : ∀ᶠ x in cocompact β ⊓ 𝓟 s, f x ≤ f x₀) : ∃ x ∈ s, ∀ y ∈ s, f y ≤ f x := -@continuous_on.exists_forall_le' (order_dual α) _ _ _ _ _ _ _ hf hsc _ h₀ hc +@continuous_on.exists_forall_le' αᵒᵈ _ _ _ _ _ _ _ hf hsc _ h₀ hc /-- The **extreme value theorem**: if a continuous function `f` is larger than a value in its range away from compact sets, then it has a global minimum. -/ @@ -258,7 +257,7 @@ in ⟨x, λ y, hx y (mem_univ y)⟩ away from compact sets, then it has a global maximum. -/ lemma _root_.continuous.exists_forall_ge' {f : β → α} (hf : continuous f) (x₀ : β) (h : ∀ᶠ x in cocompact β, f x ≤ f x₀) : ∃ (x : β), ∀ (y : β), f y ≤ f x := -@continuous.exists_forall_le' (order_dual α) _ _ _ _ _ _ hf x₀ h +@continuous.exists_forall_le' αᵒᵈ _ _ _ _ _ _ hf x₀ h /-- The **extreme value theorem**: if a continuous function `f` tends to infinity away from compact sets, then it has a global minimum. -/ @@ -272,7 +271,7 @@ compact sets, then it has a global maximum. -/ lemma continuous.exists_forall_ge [nonempty β] {f : β → α} (hf : continuous f) (hlim : tendsto f (cocompact β) at_bot) : ∃ x, ∀ y, f y ≤ f x := -@continuous.exists_forall_le (order_dual α) _ _ _ _ _ _ _ hf hlim +@continuous.exists_forall_le αᵒᵈ _ _ _ _ _ _ _ hf hlim lemma is_compact.Sup_lt_iff_of_continuous {f : β → α} {K : set β} (hK : is_compact K) (h0K : K.nonempty) (hf : continuous_on f K) (y : α) : @@ -289,7 +288,7 @@ lemma is_compact.lt_Inf_iff_of_continuous {α β : Type*} [order_topology α] [topological_space β] {f : β → α} {K : set β} (hK : is_compact K) (h0K : K.nonempty) (hf : continuous_on f K) (y : α) : y < Inf (f '' K) ↔ ∀ x ∈ K, y < f x := -@is_compact.Sup_lt_iff_of_continuous (order_dual α) β _ _ _ _ _ _ hK h0K hf y +@is_compact.Sup_lt_iff_of_continuous αᵒᵈ β _ _ _ _ _ _ hK h0K hf y /-- A continuous function with compact support has a global minimum. -/ @[to_additive] @@ -307,7 +306,7 @@ end lemma continuous.exists_forall_ge_of_has_compact_mul_support [nonempty β] [has_one α] {f : β → α} (hf : continuous f) (h : has_compact_mul_support f) : ∃ (x : β), ∀ (y : β), f y ≤ f x := -@continuous.exists_forall_le_of_has_compact_mul_support (order_dual α) _ _ _ _ _ _ _ _ hf h +@continuous.exists_forall_le_of_has_compact_mul_support αᵒᵈ _ _ _ _ _ _ _ _ hf h lemma is_compact.continuous_Sup {f : γ → β → α} {K : set β} (hK : is_compact K) (hf : continuous ↿f) : @@ -339,7 +338,7 @@ end lemma is_compact.continuous_Inf {f : γ → β → α} {K : set β} (hK : is_compact K) (hf : continuous ↿f) : continuous (λ x, Inf (f x '' K)) := -@is_compact.continuous_Sup (order_dual α) β γ _ _ _ _ _ _ _ hK hf +@is_compact.continuous_Sup αᵒᵈ β γ _ _ _ _ _ _ _ hK hf namespace continuous_on /-! diff --git a/src/topology/algebra/order/extend_from.lean b/src/topology/algebra/order/extend_from.lean index 9078a10044091..9578d72199b64 100644 --- a/src/topology/algebra/order/extend_from.lean +++ b/src/topology/algebra/order/extend_from.lean @@ -74,7 +74,7 @@ lemma continuous_on_Ioc_extend_from_Ioo [topological_space α] (hb : tendsto f (𝓝[<] b) (𝓝 lb)) : continuous_on (extend_from (Ioo a b) f) (Ioc a b) := begin - have := @continuous_on_Ico_extend_from_Ioo (order_dual α) _ _ _ _ _ _ _ f _ _ _ hab, + have := @continuous_on_Ico_extend_from_Ioo αᵒᵈ _ _ _ _ _ _ _ f _ _ _ hab, erw [dual_Ico, dual_Ioi, dual_Ioo] at this, exact this hf hb end diff --git a/src/topology/algebra/order/intermediate_value.lean b/src/topology/algebra/order/intermediate_value.lean index 18bce4b1a6321..83e6f74b73028 100644 --- a/src/topology/algebra/order/intermediate_value.lean +++ b/src/topology/algebra/order/intermediate_value.lean @@ -267,7 +267,7 @@ end lemma is_preconnected.Iio_cSup_subset {s : set α} (hs : is_preconnected s) (hb : ¬bdd_below s) (ha : bdd_above s) : Iio (Sup s) ⊆ s := -@is_preconnected.Ioi_cInf_subset (order_dual α) _ _ _ s hs ha hb +@is_preconnected.Ioi_cInf_subset αᵒᵈ _ _ _ s hs ha hb /-- A preconnected set in a conditionally complete linear order is either one of the intervals `[Inf s, Sup s]`, `[Inf s, Sup s)`, `(Inf s, Sup s]`, `(Inf s, Sup s)`, `[Inf s, +∞)`, @@ -547,7 +547,7 @@ lemma continuous.surjective {f : α → δ} (hf : continuous f) (h_top : tendsto lemma continuous.surjective' {f : α → δ} (hf : continuous f) (h_top : tendsto f at_bot at_top) (h_bot : tendsto f at_top at_bot) : function.surjective f := -@continuous.surjective (order_dual α) _ _ _ _ _ _ _ _ _ hf h_top h_bot +@continuous.surjective αᵒᵈ _ _ _ _ _ _ _ _ _ hf h_top h_bot /-- If a function `f : α → β` is continuous on a nonempty interval `s`, its restriction to `s` tends to `at_bot : filter β` along `at_bot : filter ↥s` and tends to `at_top : filter β` along @@ -569,4 +569,4 @@ lemma continuous_on.surj_on_of_tendsto' {f : α → δ} {s : set α} [ord_connec (hs : s.nonempty) (hf : continuous_on f s) (hbot : tendsto (λ x : s, f x) at_bot at_top) (htop : tendsto (λ x : s, f x) at_top at_bot) : surj_on f s univ := -@continuous_on.surj_on_of_tendsto α _ _ _ _ (order_dual δ) _ _ _ _ _ _ hs hf hbot htop +@continuous_on.surj_on_of_tendsto α _ _ _ _ δᵒᵈ _ _ _ _ _ _ hs hf hbot htop diff --git a/src/topology/algebra/order/left_right.lean b/src/topology/algebra/order/left_right.lean index 5540e488258f4..21e4190f327e4 100644 --- a/src/topology/algebra/order/left_right.lean +++ b/src/topology/algebra/order/left_right.lean @@ -35,7 +35,7 @@ by simp only [← Ici_diff_left, continuous_within_at_diff_self] lemma continuous_within_at_Iio_iff_Iic {a : α} {f : α → β} : continuous_within_at f (Iio a) a ↔ continuous_within_at f (Iic a) a := -@continuous_within_at_Ioi_iff_Ici (order_dual α) _ ‹topological_space α› _ _ _ f +@continuous_within_at_Ioi_iff_Ici αᵒᵈ _ ‹topological_space α› _ _ _ f end partial_order diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index e37326d87f201..30d7484a0e540 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -51,8 +51,7 @@ end order_closed_topology section order_closed_topology variables [semilattice_inf α] [topological_space α] [order_topology α] -lemma is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥) := -@is_bounded_le_nhds (order_dual α) _ _ _ a +lemma is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥) := @is_bounded_le_nhds αᵒᵈ _ _ _ a lemma filter.tendsto.is_bounded_under_ge {f : filter β} {u : β → α} {a : α} (h : tendsto u f (𝓝 a)) : f.is_bounded_under (≥) u := @@ -85,7 +84,7 @@ mem_of_superset h $ assume a hac, lt_of_le_of_lt hac hcb theorem gt_mem_sets_of_Liminf_gt : ∀ {f : filter α} {b}, f.is_bounded (≥) → b < f.Liminf → ∀ᶠ a in f, b < a := -@lt_mem_sets_of_Limsup_lt (order_dual α) _ +@lt_mem_sets_of_Limsup_lt αᵒᵈ _ variables [topological_space α] [order_topology α] @@ -107,8 +106,7 @@ cInf_eq_of_forall_ge_of_forall_gt_exists_lt (is_bounded_le_nhds a) | or.inr ⟨_, h⟩ := ⟨a, (𝓝 a).sets_of_superset (gt_mem_nhds hba) h, hba⟩ end) -theorem Liminf_nhds : ∀ (a : α), Liminf (𝓝 a) = a := -@Limsup_nhds (order_dual α) _ _ _ +theorem Liminf_nhds : ∀ (a : α), Liminf (𝓝 a) = a := @Limsup_nhds αᵒᵈ _ _ _ /-- If a filter is converging, its limsup coincides with its limit. -/ theorem Liminf_eq_of_le_nhds {f : filter α} {a : α} [ne_bot f] (h : f ≤ 𝓝 a) : f.Liminf = a := @@ -125,7 +123,7 @@ le_antisymm /-- If a filter is converging, its liminf coincides with its limit. -/ theorem Limsup_eq_of_le_nhds : ∀ {f : filter α} {a : α} [ne_bot f], f ≤ 𝓝 a → f.Limsup = a := -@Liminf_eq_of_le_nhds (order_dual α) _ _ _ +@Liminf_eq_of_le_nhds αᵒᵈ _ _ _ /-- If a function has a limit, then its limsup coincides with its limit. -/ theorem filter.tendsto.limsup_eq {f : filter β} {u : β → α} {a : α} [ne_bot f] diff --git a/src/topology/algebra/order/monotone_continuity.lean b/src/topology/algebra/order/monotone_continuity.lean index 9a199fdabf655..957caae659ca6 100644 --- a/src/topology/algebra/order/monotone_continuity.lean +++ b/src/topology/algebra/order/monotone_continuity.lean @@ -152,8 +152,7 @@ lemma continuous_at_left_of_monotone_on_of_exists_between {f : α → β} {s : s (hf : monotone_on f s) (hs : s ∈ 𝓝[≤] a) (hfs : ∀ b < f a, ∃ c ∈ s, f c ∈ Ioo b (f a)) : continuous_within_at f (Iic a) a := -@continuous_at_right_of_monotone_on_of_exists_between (order_dual α) (order_dual β) _ _ _ _ _ _ - f s a hf.dual hs $ +@continuous_at_right_of_monotone_on_of_exists_between αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s a hf.dual hs $ λ b hb, let ⟨c, hcs, hcb, hca⟩ := hfs b hb in ⟨c, hcs, hca, hcb⟩ /-- If a function `f` with a densely ordered codomain is monotone on a left neighborhood of `a` and @@ -163,8 +162,8 @@ lemma continuous_at_left_of_monotone_on_of_closure_image_mem_nhds_within [densel {f : α → β} {s : set α} {a : α} (hf : monotone_on f s) (hs : s ∈ 𝓝[≤] a) (hfs : closure (f '' s) ∈ 𝓝[≤] (f a)) : continuous_within_at f (Iic a) a := -@continuous_at_right_of_monotone_on_of_closure_image_mem_nhds_within (order_dual α) (order_dual β) - _ _ _ _ _ _ _ f s a hf.dual hs hfs +@continuous_at_right_of_monotone_on_of_closure_image_mem_nhds_within αᵒᵈ βᵒᵈ _ _ _ _ _ _ _ + f s a hf.dual hs hfs /-- If a function `f` with a densely ordered codomain is monotone on a left neighborhood of `a` and the image of this neighborhood under `f` is a left neighborhood of `f a`, then `f` is continuous at diff --git a/src/topology/algebra/order/monotone_convergence.lean b/src/topology/algebra/order/monotone_convergence.lean index 1b72764230fe6..fa813e949b10c 100644 --- a/src/topology/algebra/order/monotone_convergence.lean +++ b/src/topology/algebra/order/monotone_convergence.lean @@ -49,11 +49,11 @@ class Inf_convergence_class (α : Type*) [preorder α] [topological_space α] : (tendsto_coe_at_bot_is_glb : ∀ (a : α) (s : set α), is_glb s a → tendsto (coe : s → α) at_bot (𝓝 a)) instance order_dual.Sup_convergence_class [preorder α] [topological_space α] - [Inf_convergence_class α] : Sup_convergence_class (order_dual α) := + [Inf_convergence_class α] : Sup_convergence_class αᵒᵈ := ⟨‹Inf_convergence_class α›.1⟩ instance order_dual.Inf_convergence_class [preorder α] [topological_space α] - [Sup_convergence_class α] : Inf_convergence_class (order_dual α) := + [Sup_convergence_class α] : Inf_convergence_class αᵒᵈ := ⟨‹Sup_convergence_class α›.1⟩ @[priority 100] -- see Note [lower instance priority] @@ -70,7 +70,7 @@ end @[priority 100] -- see Note [lower instance priority] instance linear_order.Inf_convergence_class [topological_space α] [linear_order α] [order_topology α] : Inf_convergence_class α := -show Inf_convergence_class (order_dual $ order_dual α), from order_dual.Inf_convergence_class +show Inf_convergence_class αᵒᵈᵒᵈ, from order_dual.Inf_convergence_class section @@ -88,9 +88,9 @@ begin exact h_mono.range_factorization.tendsto_at_top_at_top (λ b, b.2.imp $ λ a ha, ha.ge) end -lemma tendsto_at_bot_is_lub (h_anti : antitone f) - (ha : is_lub (set.range f) a) : tendsto f at_bot (𝓝 a) := -@tendsto_at_top_is_lub α (order_dual ι) _ _ _ _ f a h_anti.dual ha +lemma tendsto_at_bot_is_lub (h_anti : antitone f) (ha : is_lub (set.range f) a) : + tendsto f at_bot (𝓝 a) := +by convert tendsto_at_top_is_lub h_anti.dual_left ha end is_lub @@ -100,12 +100,11 @@ variables [preorder α] [Inf_convergence_class α] {f : ι → α} {a : α} lemma tendsto_at_bot_is_glb (h_mono : monotone f) (ha : is_glb (set.range f) a) : tendsto f at_bot (𝓝 a) := -@tendsto_at_top_is_lub (order_dual α) (order_dual ι) _ _ _ _ f a h_mono.dual ha +by convert tendsto_at_top_is_lub h_mono.dual ha.dual -lemma tendsto_at_top_is_glb (h_anti : antitone f) - (ha : is_glb (set.range f) a) : +lemma tendsto_at_top_is_glb (h_anti : antitone f) (ha : is_glb (set.range f) a) : tendsto f at_top (𝓝 a) := -@tendsto_at_top_is_lub (order_dual α) ι _ _ _ _ f a h_anti ha +by convert tendsto_at_bot_is_lub h_anti.dual ha.dual end is_glb @@ -120,10 +119,9 @@ begin exacts [tendsto_of_is_empty, tendsto_at_top_is_lub h_mono (is_lub_csupr hbdd)] end -lemma tendsto_at_bot_csupr (h_anti : antitone f) - (hbdd : bdd_above $ range f) : - tendsto f at_bot (𝓝 (⨆i, f i)) := -@tendsto_at_top_csupr α (order_dual ι) _ _ _ _ _ h_anti.dual hbdd +lemma tendsto_at_bot_csupr (h_anti : antitone f) (hbdd : bdd_above $ range f) : + tendsto f at_bot (𝓝 (⨆ i, f i)) := +by convert tendsto_at_top_csupr h_anti.dual hbdd.dual end csupr @@ -132,13 +130,12 @@ section cinfi variables [conditionally_complete_lattice α] [Inf_convergence_class α] {f : ι → α} {a : α} lemma tendsto_at_bot_cinfi (h_mono : monotone f) (hbdd : bdd_below $ range f) : - tendsto f at_bot (𝓝 (⨅i, f i)) := -@tendsto_at_top_csupr (order_dual α) (order_dual ι) _ _ _ _ _ h_mono.dual hbdd + tendsto f at_bot (𝓝 (⨅ i, f i)) := +by convert tendsto_at_top_csupr h_mono.dual hbdd.dual -lemma tendsto_at_top_cinfi (h_anti : antitone f) - (hbdd : bdd_below $ range f) : - tendsto f at_top (𝓝 (⨅i, f i)) := -@tendsto_at_top_csupr (order_dual α) ι _ _ _ _ _ h_anti hbdd +lemma tendsto_at_top_cinfi (h_anti : antitone f) (hbdd : bdd_below $ range f) : + tendsto f at_top (𝓝 (⨅ i, f i)) := +by convert tendsto_at_bot_csupr h_anti.dual hbdd.dual end cinfi @@ -186,8 +183,7 @@ end instance [preorder α] [preorder β] [topological_space α] [topological_space β] [Inf_convergence_class α] [Inf_convergence_class β] : Inf_convergence_class (α × β) := -show Inf_convergence_class (order_dual $ (order_dual α × order_dual β)), - from order_dual.Inf_convergence_class +show Inf_convergence_class (αᵒᵈ × βᵒᵈ)ᵒᵈ, from order_dual.Inf_convergence_class instance {ι : Type*} {α : ι → Type*} [Π i, preorder (α i)] [Π i, topological_space (α i)] [Π i, Sup_convergence_class (α i)] : Sup_convergence_class (Π i, α i) := @@ -199,8 +195,7 @@ end instance {ι : Type*} {α : ι → Type*} [Π i, preorder (α i)] [Π i, topological_space (α i)] [Π i, Inf_convergence_class (α i)] : Inf_convergence_class (Π i, α i) := -show Inf_convergence_class (order_dual $ Π i, order_dual (α i)), - from order_dual.Inf_convergence_class +show Inf_convergence_class (Π i, (α i)ᵒᵈ)ᵒᵈ, from order_dual.Inf_convergence_class instance pi.Sup_convergence_class' {ι : Type*} [preorder α] [topological_space α] [Sup_convergence_class α] : Sup_convergence_class (ι → α) := @@ -279,19 +274,19 @@ lemma is_glb_of_tendsto_at_bot [topological_space α] [preorder α] [order_close [nonempty β] [semilattice_inf β] {f : β → α} {a : α} (hf : monotone f) (ha : tendsto f at_bot (𝓝 a)) : is_glb (set.range f) a := -@is_lub_of_tendsto_at_top (order_dual α) (order_dual β) _ _ _ _ _ _ _ hf.dual ha +@is_lub_of_tendsto_at_top αᵒᵈ βᵒᵈ _ _ _ _ _ _ _ hf.dual ha lemma is_lub_of_tendsto_at_bot [topological_space α] [preorder α] [order_closed_topology α] [nonempty β] [semilattice_inf β] {f : β → α} {a : α} (hf : antitone f) (ha : tendsto f at_bot (𝓝 a)) : is_lub (set.range f) a := -@is_lub_of_tendsto_at_top α (order_dual β) _ _ _ _ _ _ _ hf.dual_left ha +@is_lub_of_tendsto_at_top α βᵒᵈ _ _ _ _ _ _ _ hf.dual_left ha lemma is_glb_of_tendsto_at_top [topological_space α] [preorder α] [order_closed_topology α] [nonempty β] [semilattice_sup β] {f : β → α} {a : α} (hf : antitone f) (ha : tendsto f at_top (𝓝 a)) : is_glb (set.range f) a := -@is_glb_of_tendsto_at_bot α (order_dual β) _ _ _ _ _ _ _ hf.dual_left ha +@is_glb_of_tendsto_at_bot α βᵒᵈ _ _ _ _ _ _ _ hf.dual_left ha lemma supr_eq_of_tendsto {α β} [topological_space α] [complete_linear_order α] [order_topology α] [nonempty β] [semilattice_sup β] {f : β → α} {a : α} (hf : monotone f) : diff --git a/src/topology/continuous_function/ordered.lean b/src/topology/continuous_function/ordered.lean index 6961a685200de..ce6e88b9f1c1a 100644 --- a/src/topology/continuous_function/ordered.lean +++ b/src/topology/continuous_function/ordered.lean @@ -110,12 +110,12 @@ variables [linear_order γ] [order_closed_topology γ] lemma inf'_apply {ι : Type*} {s : finset ι} (H : s.nonempty) (f : ι → C(β, γ)) (b : β) : s.inf' H f b = s.inf' H (λ a, f a b) := -@sup'_apply _ (order_dual γ) _ _ _ _ _ _ H f b +@sup'_apply _ γᵒᵈ _ _ _ _ _ _ H f b @[simp, norm_cast] lemma inf'_coe {ι : Type*} {s : finset ι} (H : s.nonempty) (f : ι → C(β, γ)) : ((s.inf' H f : C(β, γ)) : ι → β) = s.inf' H (λ a, (f a : β → γ)) := -@sup'_coe _ (order_dual γ) _ _ _ _ _ _ H f +@sup'_coe _ γᵒᵈ _ _ _ _ _ _ H f end inf' diff --git a/src/topology/local_extr.lean b/src/topology/local_extr.lean index bdbbe6da9b0e8..ed73d7aeb1c31 100644 --- a/src/topology/local_extr.lean +++ b/src/topology/local_extr.lean @@ -143,7 +143,7 @@ let ⟨y, hy⟩ := (this.and self_mem_nhds_within).exists in hy.1.not_lt hy.2 lemma is_local_max_on.not_nhds_le_map [topological_space β] (hf : is_local_max_on f s a) [ne_bot (𝓝[>] (f a))] : ¬𝓝 (f a) ≤ map f (𝓝[s] a) := -@is_local_min_on.not_nhds_le_map α (order_dual β) _ _ _ _ _ ‹_› hf ‹_› +@is_local_min_on.not_nhds_le_map α βᵒᵈ _ _ _ _ _ ‹_› hf ‹_› lemma is_local_extr_on.not_nhds_le_map [topological_space β] (hf : is_local_extr_on f s a) [ne_bot (𝓝[<] (f a))] [ne_bot (𝓝[>] (f a))] : diff --git a/src/topology/order.lean b/src/topology/order.lean index ec2b2d47ce9f7..ede00f242234c 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -808,56 +808,54 @@ variables {α : Type u} {ι : Sort v} lemma generate_from_union (a₁ a₂ : set (set α)) : topological_space.generate_from (a₁ ∪ a₂) = topological_space.generate_from a₁ ⊓ topological_space.generate_from a₂ := -@galois_connection.l_sup _ (order_dual (topological_space α)) a₁ a₂ _ _ _ _ +@galois_connection.l_sup _ (topological_space α)ᵒᵈ a₁ a₂ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) lemma set_of_is_open_sup (t₁ t₂ : topological_space α) : {s | (t₁ ⊔ t₂).is_open s} = {s | t₁.is_open s} ∩ {s | t₂.is_open s} := -@galois_connection.u_inf _ (order_dual (topological_space α)) t₁ t₂ _ _ _ _ +@galois_connection.u_inf _ (topological_space α)ᵒᵈ t₁ t₂ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) lemma generate_from_Union {f : ι → set (set α)} : topological_space.generate_from (⋃ i, f i) = (⨅ i, topological_space.generate_from (f i)) := -@galois_connection.l_supr _ (order_dual (topological_space α)) _ _ _ _ _ +@galois_connection.l_supr _ (topological_space α)ᵒᵈ _ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) f lemma set_of_is_open_supr {t : ι → topological_space α} : {s | (⨆ i, t i).is_open s} = ⋂ i, {s | (t i).is_open s} := -@galois_connection.u_infi _ (order_dual (topological_space α)) _ _ _ _ _ +@galois_connection.u_infi _ (topological_space α)ᵒᵈ _ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) t lemma generate_from_sUnion {S : set (set (set α))} : topological_space.generate_from (⋃₀ S) = (⨅ s ∈ S, topological_space.generate_from s) := -@galois_connection.l_Sup _ (order_dual (topological_space α)) _ _ _ _ +@galois_connection.l_Sup _ (topological_space α)ᵒᵈ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) S lemma set_of_is_open_Sup {T : set (topological_space α)} : {s | (Sup T).is_open s} = ⋂ t ∈ T, {s | (t : topological_space α).is_open s} := -@galois_connection.u_Inf _ (order_dual (topological_space α)) _ _ _ _ +@galois_connection.u_Inf _ (topological_space α)ᵒᵈ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) T lemma generate_from_union_is_open (a b : topological_space α) : topological_space.generate_from ({s | a.is_open s} ∪ {s | b.is_open s}) = a ⊓ b := -@galois_insertion.l_sup_u _ (order_dual (topological_space α)) _ _ _ _ (gi_generate_from α) a b +@galois_insertion.l_sup_u _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) a b lemma generate_from_Union_is_open (f : ι → topological_space α) : topological_space.generate_from (⋃ i, {s | (f i).is_open s}) = ⨅ i, (f i) := -@galois_insertion.l_supr_u _ (order_dual (topological_space α)) _ _ _ _ (gi_generate_from α) _ f +@galois_insertion.l_supr_u _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) _ f lemma generate_from_inter (a b : topological_space α) : topological_space.generate_from ({s | a.is_open s} ∩ {s | b.is_open s}) = a ⊔ b := -@galois_insertion.l_inf_u _ (order_dual (topological_space α)) _ _ _ _ - (gi_generate_from α) a b +@galois_insertion.l_inf_u _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) a b lemma generate_from_Inter (f : ι → topological_space α) : topological_space.generate_from (⋂ i, {s | (f i).is_open s}) = ⨆ i, (f i) := -@galois_insertion.l_infi_u _ (order_dual (topological_space α)) _ _ _ _ (gi_generate_from α) _ f +@galois_insertion.l_infi_u _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) _ f lemma generate_from_Inter_of_generate_from_eq_self (f : ι → set (set α)) (hf : ∀ i, {s | (topological_space.generate_from (f i)).is_open s} = f i) : topological_space.generate_from (⋂ i, (f i)) = ⨆ i, topological_space.generate_from (f i) := -@galois_insertion.l_infi_of_ul_eq_self _ (order_dual (topological_space α)) _ _ _ _ - (gi_generate_from α) _ f hf +@galois_insertion.l_infi_of_ul_eq_self _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) _ f hf variables {t : ι → topological_space α} diff --git a/src/topology/order/lattice.lean b/src/topology/order/lattice.lean index 71156f5249437..e42a6d380704e 100644 --- a/src/topology/order/lattice.lean +++ b/src/topology/order/lattice.lean @@ -43,14 +43,12 @@ class has_continuous_sup (L : Type*) [topological_space L] [has_sup L] : Prop := @[priority 100] -- see Note [lower instance priority] instance order_dual.has_continuous_sup - (L : Type*) [topological_space L] [has_inf L] [has_continuous_inf L] : - has_continuous_sup (order_dual L) := + (L : Type*) [topological_space L] [has_inf L] [has_continuous_inf L] : has_continuous_sup Lᵒᵈ := { continuous_sup := @has_continuous_inf.continuous_inf L _ _ _ } @[priority 100] -- see Note [lower instance priority] instance order_dual.has_continuous_inf - (L : Type*) [topological_space L] [has_sup L] [has_continuous_sup L] : - has_continuous_inf (order_dual L) := + (L : Type*) [topological_space L] [has_sup L] [has_continuous_sup L] : has_continuous_inf Lᵒᵈ := { continuous_inf := @has_continuous_sup.continuous_sup L _ _ _ } /-- @@ -63,7 +61,7 @@ class topological_lattice (L : Type*) [topological_space L] [lattice L] @[priority 100] -- see Note [lower instance priority] instance order_dual.topological_lattice (L : Type*) [topological_space L] [lattice L] [topological_lattice L] : - topological_lattice (order_dual L) := {} + topological_lattice Lᵒᵈ := {} variables {L : Type*} [topological_space L] variables {X : Type*} [topological_space X] diff --git a/src/topology/semicontinuous.lean b/src/topology/semicontinuous.lean index d62e6cd694c96..bec1d2dd141c2 100644 --- a/src/topology/semicontinuous.lean +++ b/src/topology/semicontinuous.lean @@ -297,13 +297,12 @@ lemma continuous.comp_lower_semicontinuous lemma continuous_at.comp_lower_semicontinuous_within_at_antitone {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : lower_semicontinuous_within_at f s x) (gmon : antitone g) : upper_semicontinuous_within_at (g ∘ f) s x := -@continuous_at.comp_lower_semicontinuous_within_at α _ x s γ _ _ _ (order_dual δ) _ _ _ - g f hg hf gmon +@continuous_at.comp_lower_semicontinuous_within_at α _ x s γ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon lemma continuous_at.comp_lower_semicontinuous_at_antitone {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : lower_semicontinuous_at f x) (gmon : antitone g) : upper_semicontinuous_at (g ∘ f) x := -@continuous_at.comp_lower_semicontinuous_at α _ x γ _ _ _ (order_dual δ) _ _ _ g f hg hf gmon +@continuous_at.comp_lower_semicontinuous_at α _ x γ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon lemma continuous.comp_lower_semicontinuous_on_antitone {g : γ → δ} {f : α → γ} (hg : continuous g) (hf : lower_semicontinuous_on f s) @@ -632,7 +631,7 @@ variables [has_zero β] lemma is_open.upper_semicontinuous_indicator (hs : is_open s) (hy : y ≤ 0) : upper_semicontinuous (indicator s (λ x, y)) := -@is_open.lower_semicontinuous_indicator α _ (order_dual β) _ s y _ hs hy +@is_open.lower_semicontinuous_indicator α _ βᵒᵈ _ s y _ hs hy lemma is_open.upper_semicontinuous_on_indicator (hs : is_open s) (hy : y ≤ 0) : upper_semicontinuous_on (indicator s (λ x, y)) t := @@ -648,7 +647,7 @@ lemma is_open.upper_semicontinuous_within_at_indicator (hs : is_open s) (hy : y lemma is_closed.upper_semicontinuous_indicator (hs : is_closed s) (hy : 0 ≤ y) : upper_semicontinuous (indicator s (λ x, y)) := -@is_closed.lower_semicontinuous_indicator α _ (order_dual β) _ s y _ hs hy +@is_closed.lower_semicontinuous_indicator α _ βᵒᵈ _ s y _ hs hy lemma is_closed.upper_semicontinuous_on_indicator (hs : is_closed s) (hy : 0 ≤ y) : upper_semicontinuous_on (indicator s (λ x, y)) t := @@ -704,14 +703,12 @@ variables {δ : Type*} [linear_order δ] [topological_space δ] [order_topology lemma continuous_at.comp_upper_semicontinuous_within_at {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : upper_semicontinuous_within_at f s x) (gmon : monotone g) : upper_semicontinuous_within_at (g ∘ f) s x := -@continuous_at.comp_lower_semicontinuous_within_at α _ x s (order_dual γ) _ _ _ - (order_dual δ) _ _ _ g f hg hf (λ x y hxy, gmon hxy) +@continuous_at.comp_lower_semicontinuous_within_at α _ x s γᵒᵈ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon.dual lemma continuous_at.comp_upper_semicontinuous_at {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : upper_semicontinuous_at f x) (gmon : monotone g) : upper_semicontinuous_at (g ∘ f) x := -@continuous_at.comp_lower_semicontinuous_at α _ x (order_dual γ) _ _ _ - (order_dual δ) _ _ _ g f hg hf (λ x y hxy, gmon hxy) +@continuous_at.comp_lower_semicontinuous_at α _ x γᵒᵈ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon.dual lemma continuous.comp_upper_semicontinuous_on {g : γ → δ} {f : α → γ} (hg : continuous g) (hf : upper_semicontinuous_on f s) @@ -726,13 +723,12 @@ lemma continuous.comp_upper_semicontinuous lemma continuous_at.comp_upper_semicontinuous_within_at_antitone {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : upper_semicontinuous_within_at f s x) (gmon : antitone g) : lower_semicontinuous_within_at (g ∘ f) s x := -@continuous_at.comp_upper_semicontinuous_within_at α _ x s γ _ _ _ (order_dual δ) _ _ _ - g f hg hf gmon +@continuous_at.comp_upper_semicontinuous_within_at α _ x s γ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon lemma continuous_at.comp_upper_semicontinuous_at_antitone {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : upper_semicontinuous_at f x) (gmon : antitone g) : lower_semicontinuous_at (g ∘ f) x := -@continuous_at.comp_upper_semicontinuous_at α _ x γ _ _ _ (order_dual δ) _ _ _ g f hg hf gmon +@continuous_at.comp_upper_semicontinuous_at α _ x γ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon lemma continuous.comp_upper_semicontinuous_on_antitone {g : γ → δ} {f : α → γ} (hg : continuous g) (hf : upper_semicontinuous_on f s) @@ -759,7 +755,7 @@ lemma upper_semicontinuous_within_at.add' {f g : α → γ} (hf : upper_semicontinuous_within_at f s x) (hg : upper_semicontinuous_within_at g s x) (hcont : continuous_at (λ (p : γ × γ), p.1 + p.2) (f x, g x)) : upper_semicontinuous_within_at (λ z, f z + g z) s x := -@lower_semicontinuous_within_at.add' α _ x s (order_dual γ) _ _ _ _ _ hf hg hcont +@lower_semicontinuous_within_at.add' α _ x s γᵒᵈ _ _ _ _ _ hf hg hcont /-- The sum of two upper semicontinuous functions is upper semicontinuous. Formulated with an explicit continuity assumption on addition, for application to `ereal`. The unprimed version of @@ -825,7 +821,7 @@ hf.add' hg (λ x, continuous_add.continuous_at) lemma upper_semicontinuous_within_at_sum {f : ι → α → γ} {a : finset ι} (ha : ∀ i ∈ a, upper_semicontinuous_within_at (f i) s x) : upper_semicontinuous_within_at (λ z, (∑ i in a, f i z)) s x := -@lower_semicontinuous_within_at_sum α _ x s ι (order_dual γ) _ _ _ _ f a ha +@lower_semicontinuous_within_at_sum α _ x s ι γᵒᵈ _ _ _ _ f a ha lemma upper_semicontinuous_at_sum {f : ι → α → γ} {a : finset ι} (ha : ∀ i ∈ a, upper_semicontinuous_at (f i) x) : @@ -855,7 +851,7 @@ variables {ι : Sort*} {δ : Type*} [complete_linear_order δ] lemma upper_semicontinuous_within_at_infi {f : ι → α → δ} (h : ∀ i, upper_semicontinuous_within_at (f i) s x) : upper_semicontinuous_within_at (λ x', ⨅ i, f i x') s x := -@lower_semicontinuous_within_at_supr α _ x s ι (order_dual δ) _ f h +@lower_semicontinuous_within_at_supr α _ x s ι δᵒᵈ _ f h lemma upper_semicontinuous_within_at_binfi {p : ι → Prop} {f : Π i (h : p i), α → δ} (h : ∀ i hi, upper_semicontinuous_within_at (f i hi) s x) : @@ -865,7 +861,7 @@ upper_semicontinuous_within_at_infi $ λ i, upper_semicontinuous_within_at_infi lemma upper_semicontinuous_at_infi {f : ι → α → δ} (h : ∀ i, upper_semicontinuous_at (f i) x) : upper_semicontinuous_at (λ x', ⨅ i, f i x') x := -@lower_semicontinuous_at_supr α _ x ι (order_dual δ) _ f h +@lower_semicontinuous_at_supr α _ x ι δᵒᵈ _ f h lemma upper_semicontinuous_at_binfi {p : ι → Prop} {f : Π i (h : p i), α → δ} (h : ∀ i hi, upper_semicontinuous_at (f i hi) x) :