Skip to content

Commit

Permalink
chore(analysis/normed_space/normed_group_hom): remove bound_by (#7860)
Browse files Browse the repository at this point in the history
`bound_by f C` is the same as `∥f∥ ≤ C` and it is therefore useless now that we have `∥f∥`.
  • Loading branch information
riccardobrasca committed Jun 10, 2021
1 parent 46760cf commit 13e02fa
Show file tree
Hide file tree
Showing 2 changed files with 18 additions and 27 deletions.
43 changes: 17 additions & 26 deletions src/analysis/normed_space/normed_group_hom.lean
Expand Up @@ -120,36 +120,18 @@ f.to_add_monoid_hom.map_sum _ _

@[simp] lemma map_neg (x) : f (-x) = -(f x) := f.to_add_monoid_hom.map_neg _

/-- Predicate asserting a norm bound on a normed group hom. -/
def bound_by (f : normed_group_hom V₁ V₂) (C : ℝ≥0) : Prop := ∀ x, ∥f x∥ ≤ C * ∥x∥

lemma mk_normed_group_hom'_bound_by (f : V₁ →+ V₂) (C) (hC) :
(f.mk_normed_group_hom' C hC).bound_by C := hC

lemma bound : ∃ C, 0 < C ∧ f.bound_by C :=
lemma bound : ∃ C, 0 < C ∧ ∀ x, ∥f x∥ ≤ C * ∥x∥ :=
begin
obtain ⟨C, hC⟩ := f.bound',
rcases exists_pos_bound_of_bound _ hC with ⟨C', C'pos, hC'⟩,
exact ⟨⟨C', C'pos.le⟩, C'pos, hC'⟩,
exact ⟨C', C'pos, hC'⟩,
end

lemma lipschitz_of_bound_by (C : ℝ≥0) (h : f.bound_by C) :
lipschitz_with (real.to_nnreal C) f :=
lipschitz_with.of_dist_le' $ λ x y, by simpa only [dist_eq_norm, f.map_sub] using h (x - y)

theorem antilipschitz_of_bound_by {K : ℝ≥0} (h : ∀ x, ∥x∥ ≤ K * ∥f x∥) :
theorem antilipschitz_of_norm_ge {K : ℝ≥0} (h : ∀ x, ∥x∥ ≤ K * ∥f x∥) :
antilipschitz_with K f :=
antilipschitz_with.of_le_mul_dist $
λ x y, by simpa only [dist_eq_norm, f.map_sub] using h (x - y)

protected lemma uniform_continuous (f : normed_group_hom V₁ V₂) :
uniform_continuous f :=
let ⟨C, C_pos, hC⟩ := f.bound in (lipschitz_of_bound_by f C hC).uniform_continuous

@[continuity]
protected lemma continuous (f : normed_group_hom V₁ V₂) : continuous f :=
f.uniform_continuous.continuous

/-! ### The operator norm -/

/-- The operator norm of a seminormed group homomorphism is the inf of all its bounds. -/
Expand Down Expand Up @@ -194,6 +176,13 @@ theorem lipschitz : lipschitz_with ⟨∥f∥, op_norm_nonneg f⟩ f :=
lipschitz_with.of_dist_le_mul $ λ x y,
by { rw [dist_eq_norm, dist_eq_norm, ←map_sub], apply le_op_norm }

protected lemma uniform_continuous (f : normed_group_hom V₁ V₂) :
uniform_continuous f := f.lipschitz.uniform_continuous

@[continuity]
protected lemma continuous (f : normed_group_hom V₁ V₂) : continuous f :=
f.uniform_continuous.continuous

lemma ratio_le_op_norm (x : V₁) : ∥f x∥ / ∥x∥ ≤ ∥f∥ :=
div_le_of_nonneg_of_le_mul (norm_nonneg _) f.op_norm_nonneg (le_op_norm _ _)

Expand Down Expand Up @@ -481,8 +470,13 @@ def norm_noninc (f : normed_group_hom V W) : Prop :=

namespace norm_noninc

lemma bound_by_one (hf : f.norm_noninc) : f.bound_by 1 :=
λ v, by simpa only [one_mul, nnreal.coe_one] using hf v
lemma norm_noninc_iff_norm_le_one : f.norm_noninc ↔ ∥f∥ ≤ 1 :=
begin
refine ⟨λ h, _, λ h, λ v, _⟩,
{ refine op_norm_le_bound _ (zero_le_one) (λ v, _),
simpa [one_mul] using h v },
{ simpa using le_of_op_norm_le f h v }
end

lemma zero : (0 : normed_group_hom V₁ V₂).norm_noninc :=
λ v, by simp
Expand Down Expand Up @@ -522,9 +516,6 @@ hg.comp hf
lemma norm_noninc_of_isometry (hf : isometry f) : f.norm_noninc :=
λ v, le_of_eq $ norm_eq_of_isometry hf v

lemma bound_by_one_of_isometry (hf : isometry f) : f.bound_by 1 :=
(norm_noninc_of_isometry hf).bound_by_one

end isometry

end normed_group_hom
2 changes: 1 addition & 1 deletion src/analysis/normed_space/normed_group_quotient.lean
Expand Up @@ -434,7 +434,7 @@ def lift {N : Type*} [semi_normed_group N] (S : add_subgroup M)
normed_group_hom (quotient S) N :=
{ bound' :=
begin
obtain ⟨c : ℝ0, hcpos : (0 : ℝ) < c, hc : f.bound_by c⟩ := f.bound,
obtain ⟨c : ℝ, hcpos : (0 : ℝ) < c, hc : ∀ x, ∥f x∥ ≤ c * ∥x∥⟩ := f.bound,
refine ⟨c, λ mbar, le_of_forall_pos_le_add (λ ε hε, _)⟩,
obtain ⟨m : M, rfl : mk' S m = mbar, hmnorm : ∥m∥ < ∥mk' S m∥ + ε/c⟩ :=
norm_mk_lt mbar (div_pos hε hcpos),
Expand Down

0 comments on commit 13e02fa

Please sign in to comment.