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

Commit 17ef379

Browse files
committed
refactor(analysis): change the symbol for norm to align with the unicode spec (#17575)
The characters in question are: * U+2016 DOUBLE VERTICAL LINE `‖`. The Unicode Character Database says "used in pairs to indicate norm of a matrix", for instance [here](https://www.fileformat.info/info/unicode/char/2016/index.htm) and [here](https://unicode-explorer.com/c/2016) * U+2225 PARALLEL TO `∥`. On some platforms this renders with a forward slant! Previously `norm` was the latter and `parallel` was the former. This change swaps them around: * `∥x∥`, `S ‖ T`: before * `‖x‖`, `S ∥ T`: after Code using U+2016 `‖` for fintype cardinality has been left unchanged.
1 parent 720aee9 commit 17ef379

File tree

211 files changed

+4978
-4978
lines changed

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

211 files changed

+4978
-4978
lines changed

archive/imo/imo1972_q5.lean

Lines changed: 30 additions & 30 deletions
Original file line numberDiff line numberDiff line change
@@ -16,49 +16,49 @@ Prove that `|g(x)| ≤ 1` for all `x`.
1616
-/
1717

1818
/--
19-
This proof begins by introducing the supremum of `f`, `k ≤ 1` as well as `k' = k / g y`. We then
19+
This proof begins by introducing the supremum of `f`, `k ≤ 1` as well as `k' = k / g y`. We then
2020
suppose that the conclusion does not hold (`hneg`) and show that `k ≤ k'` (by
21-
`2 * (f x * g y) ≤ 2 * k` obtained from the main hypothesis `hf1`) and that `k' < k` (obtained
21+
`2 * (f x * g y) ≤ 2 * k` obtained from the main hypothesis `hf1`) and that `k' < k` (obtained
2222
from `hneg` directly), finally raising a contradiction with `k' < k'`.
2323
2424
(Authored by Stanislas Polu inspired by Ruben Van de Velde).
2525
-/
2626
example (f g : ℝ → ℝ)
2727
(hf1 : ∀ x, ∀ y, (f(x+y) + f(x-y)) = 2 * f(x) * g(y))
28-
(hf2 : ∀ y, f(y)1)
28+
(hf2 : ∀ y, f(y)1)
2929
(hf3 : ∃ x, f(x) ≠ 0)
3030
(y : ℝ) :
31-
g(y)1 :=
31+
g(y)1 :=
3232
begin
33-
set S := set.range (λ x, f x),
33+
set S := set.range (λ x, f x),
3434
-- Introduce `k`, the supremum of `f`.
3535
let k : ℝ := Sup (S),
3636

37-
-- Show that `f x ≤ k`.
38-
have hk₁ : ∀ x, f x ≤ k,
37+
-- Show that `f x ≤ k`.
38+
have hk₁ : ∀ x, f x ≤ k,
3939
{ have h : bdd_above S, from1, set.forall_range_iff.mpr hf2⟩,
4040
intro x,
4141
exact le_cSup h (set.mem_range_self x), },
42-
-- Show that `2 * (f x * g y) ≤ 2 * k`.
43-
have hk₂ : ∀ x, 2 * (f x * g y) ≤ 2 * k,
42+
-- Show that `2 * (f x * g y) ≤ 2 * k`.
43+
have hk₂ : ∀ x, 2 * (f x * g y) ≤ 2 * k,
4444
{ intro x,
45-
calc 2 * (f x * g y)
46-
= 2 * f x * g y : by simp [abs_mul, mul_assoc]
47-
... = f (x + y) + f (x - y) : by rw hf1
48-
... ≤ f (x + y) + f (x - y) : norm_add_le _ _
45+
calc 2 * (f x * g y)
46+
= 2 * f x * g y : by simp [abs_mul, mul_assoc]
47+
... = f (x + y) + f (x - y) : by rw hf1
48+
... ≤ f (x + y) + f (x - y) : norm_add_le _ _
4949
... ≤ k + k : add_le_add (hk₁ _) (hk₁ _)
5050
... = 2 * k : (two_mul _).symm, },
5151

5252
-- Suppose the conclusion does not hold.
5353
by_contra' hneg,
54-
set k' := k / g y,
54+
set k' := k / g y,
5555

5656
-- Demonstrate that `k' < k` using `hneg`.
5757
have H₁ : k' < k,
5858
{ have h₁ : 0 < k,
5959
{ obtain ⟨x, hx⟩ := hf3,
6060
calc 0
61-
< f x : norm_pos_iff.mpr hx
61+
< f x : norm_pos_iff.mpr hx
6262
... ≤ k : hk₁ x },
6363
rw div_lt_iff,
6464
apply lt_mul_of_one_lt_right h₁ hneg,
@@ -67,8 +67,8 @@ begin
6767
-- Demonstrate that `k ≤ k'` using `hk₂`.
6868
have H₂ : k ≤ k',
6969
{ have h₁ : ∃ x : ℝ, x ∈ S,
70-
{ use f 0, exact set.mem_range_self 0, },
71-
have h₂ : ∀ x, f x ≤ k',
70+
{ use f 0, exact set.mem_range_self 0, },
71+
have h₂ : ∀ x, f x ≤ k',
7272
{ intros x,
7373
rw le_div_iff,
7474
{ apply (mul_le_mul_left zero_lt_two).mp (hk₂ x) },
@@ -95,28 +95,28 @@ This is a more concise version of the proof proposed by Ruben Van de Velde.
9595
-/
9696
example (f g : ℝ → ℝ)
9797
(hf1 : ∀ x, ∀ y, (f (x+y) + f(x-y)) = 2 * f(x) * g(y))
98-
(hf2 : bdd_above (set.range (λ x, f x)))
98+
(hf2 : bdd_above (set.range (λ x, f x)))
9999
(hf3 : ∃ x, f(x) ≠ 0)
100100
(y : ℝ) :
101-
g(y)1 :=
101+
g(y)1 :=
102102
begin
103103
obtain ⟨x, hx⟩ := hf3,
104-
set k := ⨆ x, f x,
105-
have h : ∀ x, f x ≤ k := le_csupr hf2,
104+
set k := ⨆ x, f x,
105+
have h : ∀ x, f x ≤ k := le_csupr hf2,
106106
by_contra' H,
107-
have hgy : 0 < g y,
107+
have hgy : 0 < g y,
108108
by linarith,
109109
have k_pos : 0 < k := lt_of_lt_of_le (norm_pos_iff.mpr hx) (h x),
110-
have : k / g y < k := (div_lt_iff hgy).mpr (lt_mul_of_one_lt_right k_pos H),
111-
have : k ≤ k / g y,
112-
{ suffices : ∀ x, f x ≤ k / g y, from csupr_le this,
110+
have : k / g y < k := (div_lt_iff hgy).mpr (lt_mul_of_one_lt_right k_pos H),
111+
have : k ≤ k / g y,
112+
{ suffices : ∀ x, f x ≤ k / g y, from csupr_le this,
113113
intro x,
114-
suffices : 2 * (f x * g y) ≤ 2 * k,
114+
suffices : 2 * (f x * g y) ≤ 2 * k,
115115
by { rwa [le_div_iff hgy, ←mul_le_mul_left (zero_lt_two : (0 : ℝ) < 2)] },
116-
calc 2 * (f x * g y)
117-
= 2 * f x * g y : by simp [abs_mul, mul_assoc]
118-
... = f (x + y) + f (x - y) : by rw hf1
119-
... ≤ f (x + y) + f (x - y) : abs_add _ _
116+
calc 2 * (f x * g y)
117+
= 2 * f x * g y : by simp [abs_mul, mul_assoc]
118+
... = f (x + y) + f (x - y) : by rw hf1
119+
... ≤ f (x + y) + f (x - y) : abs_add _ _
120120
... ≤ 2 * k : by linarith [h (x+y), h (x -y)] },
121121
linarith,
122122
end

counterexamples/phillips.lean

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -375,7 +375,7 @@ section
375375
-/
376376

377377
lemma norm_indicator_le_one (s : set α) (x : α) :
378-
(indicator s (1 : α → ℝ)) x1 :=
378+
(indicator s (1 : α → ℝ)) x1 :=
379379
by { simp only [indicator, pi.one_apply], split_ifs; norm_num }
380380

381381
/-- A functional in the dual space of bounded functions gives rise to a bounded additive measure,
@@ -392,8 +392,8 @@ def _root_.continuous_linear_map.to_bounded_additive_measure
392392
by { ext x, simp [indicator_union_of_disjoint hst], },
393393
rw [this, f.map_add],
394394
end,
395-
exists_bound := ⟨∥f∥, λ s, begin
396-
have I : of_normed_add_comm_group_discrete (indicator s 1) 1 (norm_indicator_le_one s)1,
395+
exists_bound := ⟨‖f‖, λ s, begin
396+
have I : of_normed_add_comm_group_discrete (indicator s 1) 1 (norm_indicator_le_one s)1,
397397
by apply norm_of_normed_add_comm_group_le _ zero_le_one,
398398
apply le_trans (f.le_op_norm _),
399399
simpa using mul_le_mul_of_nonneg_left I (norm_nonneg f),
@@ -492,7 +492,7 @@ lemma countable_spf_mem (Hcont : #ℝ = aleph 1) (y : ℝ) : {x | y ∈ spf Hcon
492492
493493
We construct a function `f` from `[0,1]` to a complete Banach space `B`, which is weakly measurable
494494
(i.e., for any continuous linear form `φ` on `B` the function `φ ∘ f` is measurable), bounded in
495-
norm (i.e., for all `x`, one has `f x ≤ 1`), and still `f` has no Pettis integral.
495+
norm (i.e., for all `x`, one has `f x ≤ 1`), and still `f` has no Pettis integral.
496496
497497
This construction, due to Phillips, requires the continuum hypothesis. We will take for `B` the
498498
space of all bounded functions on `ℝ`, with the supremum norm (no measure here, we are really
@@ -582,7 +582,7 @@ begin
582582
end
583583

584584
/-- The function `f Hcont : ℝ → (discrete_copy ℝ →ᵇ ℝ)` is uniformly bounded by `1` in norm. -/
585-
lemma norm_bound (Hcont : #ℝ = aleph 1) (x : ℝ) : f Hcont x1 :=
585+
lemma norm_bound (Hcont : #ℝ = aleph 1) (x : ℝ) : f Hcont x1 :=
586586
norm_of_normed_add_comm_group_le _ zero_le_one _
587587

588588
/-- The function `f Hcont : ℝ → (discrete_copy ℝ →ᵇ ℝ)` has no Pettis integral. -/

src/analysis/ODE/gronwall.lean

Lines changed: 8 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -9,8 +9,8 @@ import analysis.special_functions.exp_deriv
99
# Grönwall's inequality
1010
1111
The main technical result of this file is the Grönwall-like inequality
12-
`norm_le_gronwall_bound_of_norm_deriv_right_le`. It states that if `f : ℝ → E` satisfies `f a ≤ δ`
13-
and `∀ x ∈ [a, b), f' x ≤ K * f x + ε`, then for all `x ∈ [a, b]` we have `f x ≤ δ * exp (K *
12+
`norm_le_gronwall_bound_of_norm_deriv_right_le`. It states that if `f : ℝ → E` satisfies `f a ≤ δ`
13+
and `∀ x ∈ [a, b), f' x ≤ K * f x + ε`, then for all `x ∈ [a, b]` we have `f x ≤ δ * exp (K *
1414
x) + (ε / K) * (exp (K * x) - 1)`.
1515
1616
Then we use this inequality to prove some estimates on the possible rate of growth of the distance
@@ -22,7 +22,7 @@ Sec. 4.5][HubbardWest-ode], where `norm_le_gronwall_bound_of_norm_deriv_right_le
2222
2323
## TODO
2424
25-
- Once we have FTC, prove an inequality for a function satisfying `f' x ≤ K x * f x + ε`,
25+
- Once we have FTC, prove an inequality for a function satisfying `f' x ≤ K x * f x + ε`,
2626
or more generally `liminf_{y→x+0} (f y - f x)/(y - x) ≤ K x * f x + ε` with any sign
2727
of `K x` and `f x`.
2828
-/
@@ -101,7 +101,7 @@ the inequalities `f a ≤ δ` and
101101
`∀ x ∈ [a, b), liminf_{z→x+0} (f z - f x)/(z - x) ≤ K * (f x) + ε`, then `f x`
102102
is bounded by `gronwall_bound δ K ε (x - a)` on `[a, b]`.
103103
104-
See also `norm_le_gronwall_bound_of_norm_deriv_right_le` for a version bounding `f x`,
104+
See also `norm_le_gronwall_bound_of_norm_deriv_right_le` for a version bounding `f x`,
105105
`f : ℝ → E`. -/
106106
theorem le_gronwall_bound_of_liminf_deriv_right_le {f f' : ℝ → ℝ} {δ K ε : ℝ} {a b : ℝ}
107107
(hf : continuous_on f (Icc a b))
@@ -128,13 +128,13 @@ begin
128128
end
129129

130130
/-- A Grönwall-like inequality: if `f : ℝ → E` is continuous on `[a, b]`, has right derivative
131-
`f' x` at every point `x ∈ [a, b)`, and satisfies the inequalities `f a ≤ δ`,
132-
`∀ x ∈ [a, b), f' x ≤ K * f x + ε`, then `f x` is bounded by `gronwall_bound δ K ε (x - a)`
131+
`f' x` at every point `x ∈ [a, b)`, and satisfies the inequalities `f a ≤ δ`,
132+
`∀ x ∈ [a, b), f' x ≤ K * f x + ε`, then `f x` is bounded by `gronwall_bound δ K ε (x - a)`
133133
on `[a, b]`. -/
134134
theorem norm_le_gronwall_bound_of_norm_deriv_right_le {f f' : ℝ → E} {δ K ε : ℝ} {a b : ℝ}
135135
(hf : continuous_on f (Icc a b)) (hf' : ∀ x ∈ Ico a b, has_deriv_within_at f (f' x) (Ici x) x)
136-
(ha : f a ≤ δ) (bound : ∀ x ∈ Ico a b, f' x ≤ K * f x + ε) :
137-
∀ x ∈ Icc a b, f x ≤ gronwall_bound δ K ε (x - a) :=
136+
(ha : f a ≤ δ) (bound : ∀ x ∈ Ico a b, f' x ≤ K * f x + ε) :
137+
∀ x ∈ Icc a b, f x ≤ gronwall_bound δ K ε (x - a) :=
138138
le_gronwall_bound_of_liminf_deriv_right_le (continuous_norm.comp_continuous_on hf)
139139
(λ x hx r hr, (hf' x hx).liminf_right_slope_norm_le hr) ha bound
140140

src/analysis/ODE/picard_lindelof.lean

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -51,7 +51,7 @@ structure is_picard_lindelof
5151
(hR : 0 ≤ R)
5252
(lipschitz : ∀ t ∈ Icc t_min t_max, lipschitz_on_with L (v t) (closed_ball x₀ R))
5353
(cont : ∀ x ∈ closed_ball x₀ R, continuous_on (λ (t : ℝ), v t x) (Icc t_min t_max))
54-
(norm_le : ∀ (t ∈ Icc t_min t_max) (x ∈ closed_ball x₀ R), v t x ≤ C)
54+
(norm_le : ∀ (t ∈ Icc t_min t_max) (x ∈ closed_ball x₀ R), v t x ≤ C)
5555
(C_mul_le_R : (C : ℝ) * linear_order.max (t_max - t₀) (t₀ - t_min) ≤ R)
5656

5757
/-- This structure holds arguments of the Picard-Lipschitz (Cauchy-Lipschitz) theorem. It is part of
@@ -100,7 +100,7 @@ have continuous_on (uncurry (flip v)) (closed_ball v.x₀ v.R ×ˢ Icc v.t_min v
100100
this.comp continuous_swap.continuous_on (preimage_swap_prod _ _).symm.subset
101101

102102
lemma norm_le {t : ℝ} (ht : t ∈ Icc v.t_min v.t_max) {x : E} (hx : x ∈ closed_ball v.x₀ v.R) :
103-
v t x ≤ v.C :=
103+
v t x ≤ v.C :=
104104
v.is_pl.norm_le _ ht _ hx
105105

106106
/-- The maximum of distances from `t₀` to the endpoints of `[t_min, t_max]`. -/
@@ -192,7 +192,7 @@ begin
192192
exact ⟨(v.proj x).2, f.mem_closed_ball _⟩
193193
end
194194

195-
lemma norm_v_comp_le (t : ℝ) : f.v_comp t ≤ v.C :=
195+
lemma norm_v_comp_le (t : ℝ) : f.v_comp t ≤ v.C :=
196196
v.norm_le (v.proj t).2 $ f.mem_closed_ball _
197197

198198
lemma dist_apply_le_dist (f₁ f₂ : fun_space v) (t : Icc v.t_min v.t_max) :
@@ -262,7 +262,7 @@ begin
262262
simp only [dist_eq_norm, next_apply, add_sub_add_left_eq_sub,
263263
← interval_integral.integral_sub (interval_integrable_v_comp _ _ _)
264264
(interval_integrable_v_comp _ _ _), norm_integral_eq_norm_integral_Ioc] at *,
265-
calc ∫ τ in Ι (v.t₀ : ℝ) t, f₁.v_comp τ - f₂.v_comp τ
265+
calc ∫ τ in Ι (v.t₀ : ℝ) t, f₁.v_comp τ - f₂.v_comp τ
266266
≤ ∫ τ in Ι (v.t₀ : ℝ) t, v.L * ((v.L * |τ - v.t₀|) ^ n / n! * d) :
267267
begin
268268
refine norm_integral_le_of_norm_le (continuous.integrable_on_interval_oc _) _,
@@ -338,7 +338,7 @@ end picard_lindelof
338338

339339
lemma is_picard_lindelof.norm_le₀ {E : Type*} [normed_add_comm_group E]
340340
{v : ℝ → E → E} {t_min t₀ t_max : ℝ} {x₀ : E} {C R : ℝ} {L : ℝ≥0}
341-
(hpl : is_picard_lindelof v t_min t₀ t_max x₀ L R C) : v t₀ x₀ ≤ C :=
341+
(hpl : is_picard_lindelof v t_min t₀ t_max x₀ L R C) : v t₀ x₀ ≤ C :=
342342
hpl.norm_le t₀ hpl.ht₀ x₀ $ mem_closed_ball_self hpl.hR
343343

344344
/-- Picard-Lindelöf (Cauchy-Lipschitz) theorem. -/

0 commit comments

Comments
 (0)