Skip to content
This repository has been archived by the owner on Jul 24, 2024. It is now read-only.

Commit

Permalink
Merge branch 'YK-cont-alt-new' into YK-cont-alternating
Browse files Browse the repository at this point in the history
  • Loading branch information
urkud committed Jun 22, 2023
2 parents 47a9075 + 56e63e6 commit 0d4a886
Show file tree
Hide file tree
Showing 11 changed files with 275 additions and 159 deletions.
Original file line number Diff line number Diff line change
Expand Up @@ -24,7 +24,6 @@ sequences, increasing, decreasing, Ramsey, Erdos-Szekeres, Erdős–Szekeres, Er
variables {α : Type*} [linear_order α] {β : Type*}

open function finset
open_locale classical

namespace theorems_100

Expand Down
21 changes: 0 additions & 21 deletions src/analysis/normed_space/continuous_linear_map.lean
Original file line number Diff line number Diff line change
Expand Up @@ -192,27 +192,6 @@ lemma to_span_singleton_homothety (x : E) (c : 𝕜) :
‖linear_map.to_span_singleton 𝕜 E x c‖ = ‖x‖ * ‖c‖ :=
by {rw mul_comm, exact norm_smul _ _}

/-- Given an element `x` of a normed space `E` over a field `𝕜`, the natural continuous
linear map from `𝕜` to `E` by taking multiples of `x`.-/
def to_span_singleton (x : E) : 𝕜 →L[𝕜] E :=
of_homothety (linear_map.to_span_singleton 𝕜 E x) ‖x‖ (to_span_singleton_homothety 𝕜 x)

lemma to_span_singleton_apply (x : E) (r : 𝕜) : to_span_singleton 𝕜 x r = r • x :=
by simp [to_span_singleton, of_homothety, linear_map.to_span_singleton]

lemma to_span_singleton_add (x y : E) :
to_span_singleton 𝕜 (x + y) = to_span_singleton 𝕜 x + to_span_singleton 𝕜 y :=
by { ext1, simp [to_span_singleton_apply], }

lemma to_span_singleton_smul' (𝕜') [normed_field 𝕜'] [normed_space 𝕜' E]
[smul_comm_class 𝕜 𝕜' E] (c : 𝕜') (x : E) :
to_span_singleton 𝕜 (c • x) = c • to_span_singleton 𝕜 x :=
by { ext1, rw [to_span_singleton_apply, smul_apply, to_span_singleton_apply, smul_comm], }

lemma to_span_singleton_smul (c : 𝕜) (x : E) :
to_span_singleton 𝕜 (c • x) = c • to_span_singleton 𝕜 x :=
to_span_singleton_smul' 𝕜 𝕜 c x

end continuous_linear_map

section
Expand Down
8 changes: 8 additions & 0 deletions src/data/finset/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -570,6 +570,10 @@ lemma forall_mem_cons (h : a ∉ s) (p : α → Prop) :
(∀ x, x ∈ cons a s h → p x) ↔ p a ∧ ∀ x, x ∈ s → p x :=
by simp only [mem_cons, or_imp_distrib, forall_and_distrib, forall_eq]

/-- Useful in proofs by induction. -/
lemma forall_of_forall_cons {p : α → Prop} {h : a ∉ s} (H : ∀ x, x ∈ cons a s h → p x) (x)
(h : x ∈ s) : p x := H _ $ mem_cons.2 $ or.inr h

@[simp] lemma mk_cons {s : multiset α} (h : (a ::ₘ s).nodup) :
(⟨a ::ₘ s, h⟩ : finset α) = cons a ⟨s, (nodup_cons.1 h).2⟩ (nodup_cons.1 h).1 := rfl

Expand Down Expand Up @@ -2084,6 +2088,10 @@ lemma forall_mem_insert [decidable_eq α] (a : α) (s : finset α) (p : α → P
(∀ x, x ∈ insert a s → p x) ↔ p a ∧ ∀ x, x ∈ s → p x :=
by simp only [mem_insert, or_imp_distrib, forall_and_distrib, forall_eq]

/-- Useful in proofs by induction. -/
lemma forall_of_forall_insert [decidable_eq α] {p : α → Prop} {a : α} {s : finset α}
(H : ∀ x, x ∈ insert a s → p x) (x) (h : x ∈ s) : p x := H _ $ mem_insert_of_mem h

end finset

/-- Equivalence between the set of natural numbers which are `≥ k` and `ℕ`, given by `n → n - k`. -/
Expand Down
31 changes: 31 additions & 0 deletions src/data/finset/n_ary.lean
Original file line number Diff line number Diff line change
Expand Up @@ -358,6 +358,37 @@ lemma image₂_right_identity {f : γ → β → γ} {b : β} (h : ∀ a, f a b
image₂ f s {b} = s :=
by rw [image₂_singleton_right, funext h, image_id']

/-- If each partial application of `f` is injective, and images of `s` under those partial
applications are disjoint (but not necessarily distinct!), then the size of `t` divides the size of
`finset.image₂ f s t`. -/
lemma card_dvd_card_image₂_right (hf : ∀ a ∈ s, injective (f a))
(hs : ((λ a, t.image $ f a) '' s).pairwise_disjoint id) :
t.card ∣ (image₂ f s t).card :=
begin
classical,
induction s using finset.induction with a s ha ih,
{ simp },
specialize ih (forall_of_forall_insert hf)
(hs.subset $ set.image_subset _ $ coe_subset.2 $ subset_insert _ _),
rw image₂_insert_left,
by_cases h : disjoint (image (f a) t) (image₂ f s t),
{ rw card_union_eq h,
exact (card_image_of_injective _ $ hf _ $ mem_insert_self _ _).symm.dvd.add ih },
simp_rw [←bUnion_image_left, disjoint_bUnion_right, not_forall] at h,
obtain ⟨b, hb, h⟩ := h,
rwa union_eq_right_iff_subset.2,
exact (hs.eq (set.mem_image_of_mem _ $ mem_insert_self _ _)
(set.mem_image_of_mem _ $ mem_insert_of_mem hb) h).trans_subset (image_subset_image₂_right hb),
end

/-- If each partial application of `f` is injective, and images of `t` under those partial
applications are disjoint (but not necessarily distinct!), then the size of `s` divides the size of
`finset.image₂ f s t`. -/
lemma card_dvd_card_image₂_left (hf : ∀ b ∈ t, injective (λ a, f a b))
(ht : ((λ b, s.image $ λ a, f a b) '' t).pairwise_disjoint id) :
s.card ∣ (image₂ f s t).card :=
by { rw ←image₂_swap, exact card_dvd_card_image₂_right hf ht }

variables [decidable_eq α] [decidable_eq β]

lemma image₂_inter_union_subset_union :
Expand Down
18 changes: 18 additions & 0 deletions src/data/finset/pointwise.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1078,6 +1078,24 @@ coe_injective $ by { push_cast, exact set.smul_univ hs }
@[simp, to_additive] lemma card_smul_finset (a : α) (s : finset β) : (a • s).card = s.card :=
card_image_of_injective _ $ mul_action.injective _

/-- If the left cosets of `t` by elements of `s` are disjoint (but not necessarily distinct!), then
the size of `t` divides the size of `s * t`. -/
@[to_additive "If the left cosets of `t` by elements of `s` are disjoint (but not necessarily
distinct!), then the size of `t` divides the size of `s + t`."]
lemma card_dvd_card_smul_right {s : finset α} :
((• t) '' (s : set α)).pairwise_disjoint id → t.card ∣ (s • t).card :=
card_dvd_card_image₂_right (λ _ _, mul_action.injective _)

variables [decidable_eq α]

/-- If the right cosets of `s` by elements of `t` are disjoint (but not necessarily distinct!), then
the size of `s` divides the size of `s * t`. -/
@[to_additive "If the right cosets of `s` by elements of `t` are disjoint (but not necessarily
distinct!), then the size of `s` divides the size of `s + t`."]
lemma card_dvd_card_mul_left {s t : finset α} :
((λ b, s.image $ λ a, a * b) '' (t : set α)).pairwise_disjoint id → s.card ∣ (s * t).card :=
card_dvd_card_image₂_left (λ _ _, mul_left_injective _)

end group

section group_with_zero
Expand Down
6 changes: 5 additions & 1 deletion src/order/directed.lean
Original file line number Diff line number Diff line change
Expand Up @@ -131,6 +131,10 @@ lemma directed_on_of_inf_mem [semilattice_inf α] {S : set α}
(H : ∀ ⦃i j⦄, i ∈ S → j ∈ S → i ⊓ j ∈ S) : directed_on (≥) S :=
λ a ha b hb, ⟨a ⊓ b, H ha hb, inf_le_left, inf_le_right⟩

lemma is_total.directed [is_total α r] (f : ι → α) :
directed r f :=
λ i j, or.cases_on (total_of r (f i) (f j)) (λ h, ⟨j, h, refl _⟩) (λ h, ⟨i, refl _, h⟩)

/-- `is_directed α r` states that for any elements `a`, `b` there exists an element `c` such that
`r a c` and `r b c`. -/
class is_directed (α : Type*) (r : α → α → Prop) : Prop :=
Expand All @@ -150,7 +154,7 @@ lemma directed_on_univ_iff : directed_on r set.univ ↔ is_directed α r :=

@[priority 100] -- see Note [lower instance priority]
instance is_total.to_is_directed [is_total α r] : is_directed α r :=
⟨λ a b, or.cases_on (total_of r a b) (λ h, ⟨b, h, refl _⟩) (λ h, ⟨a, refl _, h⟩)⟩
by rw ← directed_id_iff; exact is_total.directed _

lemma is_directed_mono [is_directed α r] (h : ∀ ⦃a b⦄, r a b → s a b) : is_directed α s :=
⟨λ a b, let ⟨c, ha, hb⟩ := is_directed.directed a b in ⟨c, h ha, h hb⟩⟩
Expand Down
15 changes: 15 additions & 0 deletions src/order/monotone/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -102,6 +102,21 @@ def strict_anti_on (f : α → β) (s : set α) : Prop :=

end monotone_def

section decidable

variables [preorder α] [preorder β] {f : α → β} {s : set α}

instance [i : decidable (∀ a b, a ≤ b → f a ≤ f b)] : decidable (monotone f) := i
instance [i : decidable (∀ a b, a ≤ b → f b ≤ f a)] : decidable (antitone f) := i
instance [i : decidable (∀ a b ∈ s, a ≤ b → f a ≤ f b)] : decidable (monotone_on f s) := i
instance [i : decidable (∀ a b ∈ s, a ≤ b → f b ≤ f a)] : decidable (antitone_on f s) := i
instance [i : decidable (∀ a b, a < b → f a < f b)] : decidable (strict_mono f) := i
instance [i : decidable (∀ a b, a < b → f b < f a)] : decidable (strict_anti f) := i
instance [i : decidable (∀ a b ∈ s, a < b → f a < f b)] : decidable (strict_mono_on f s) := i
instance [i : decidable (∀ a b ∈ s, a < b → f b < f a)] : decidable (strict_anti_on f s) := i

end decidable

/-! ### Monotonicity on the dual order
Strictly, many of the `*_on.dual` lemmas in this section should use `of_dual ⁻¹' s` instead of `s`,
Expand Down
31 changes: 31 additions & 0 deletions src/topology/algebra/module/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1003,6 +1003,37 @@ lemma smul_right_comp [has_continuous_mul R₁] {x : M₂} {c : R₁} :
smul_right (1 : R₁ →L[R₁] R₁) (c • x) :=
by { ext, simp [mul_smul] }

section to_span_singleton
variables (R₁)
variables [has_continuous_smul R₁ M₁]

/-- Given an element `x` of a topological space `M` over a semiring `R`, the natural continuous
linear map from `R` to `M` by taking multiples of `x`.-/
def to_span_singleton (x : M₁) : R₁ →L[R₁] M₁ :=
{ to_linear_map := linear_map.to_span_singleton R₁ M₁ x,
cont := continuous_id.smul continuous_const }

lemma to_span_singleton_apply (x : M₁) (r : R₁) : to_span_singleton R₁ x r = r • x :=
rfl

lemma to_span_singleton_add [has_continuous_add M₁] (x y : M₁) :
to_span_singleton R₁ (x + y) = to_span_singleton R₁ x + to_span_singleton R₁ y :=
by { ext1, simp [to_span_singleton_apply], }

lemma to_span_singleton_smul' {α} [monoid α] [distrib_mul_action α M₁]
[has_continuous_const_smul α M₁]
[smul_comm_class R₁ α M₁] (c : α) (x : M₁) :
to_span_singleton R₁ (c • x) = c • to_span_singleton R₁ x :=
by { ext1, rw [to_span_singleton_apply, smul_apply, to_span_singleton_apply, smul_comm], }

/-- A special case of `to_span_singleton_smul'` for when `R` is commutative. -/
lemma to_span_singleton_smul (R) {M₁} [comm_semiring R] [add_comm_monoid M₁] [module R M₁]
[topological_space R] [topological_space M₁] [has_continuous_smul R M₁] (c : R) (x : M₁) :
to_span_singleton R (c • x) = c • to_span_singleton R x :=
to_span_singleton_smul' R c x

end to_span_singleton

end semiring

section pi
Expand Down
Loading

0 comments on commit 0d4a886

Please sign in to comment.