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

Commit 21415c8

Browse files
committed
chore(topology/algebra/ordered): drop section vars, golf 2 proofs (#4706)
* Explicitly specify explicit arguments instead of using section variables; * Add `continuous_min` and `continuous_max`; * Use them for `tendsto.min` and `tendsto.max`
1 parent 0cf8a98 commit 21415c8

File tree

1 file changed

+15
-20
lines changed

1 file changed

+15
-20
lines changed

src/topology/algebra/ordered.lean

Lines changed: 15 additions & 20 deletions
Original file line numberDiff line numberDiff line change
@@ -513,49 +513,44 @@ section decidable_linear_order
513513
variables [topological_space α] [decidable_linear_order α] [order_closed_topology α] {f g : β → α}
514514

515515
section
516-
variables [topological_space β] (hf : continuous f) (hg : continuous g)
517-
include hf hg
516+
variables [topological_space β]
518517

519-
lemma frontier_le_subset_eq : frontier {b | f b ≤ g b} ⊆ {b | f b = g b} :=
518+
lemma frontier_le_subset_eq (hf : continuous f) (hg : continuous g) :
519+
frontier {b | f b ≤ g b} ⊆ {b | f b = g b} :=
520520
begin
521521
rw [frontier_eq_closure_inter_closure, closure_le_eq hf hg],
522522
rintros b ⟨hb₁, hb₂⟩,
523523
refine le_antisymm hb₁ (closure_lt_subset_le hg hf _),
524524
convert hb₂ using 2, simp only [not_le.symm], refl
525525
end
526526

527-
lemma frontier_lt_subset_eq : frontier {b | f b < g b} ⊆ {b | f b = g b} :=
527+
lemma frontier_lt_subset_eq (hf : continuous f) (hg : continuous g) :
528+
frontier {b | f b < g b} ⊆ {b | f b = g b} :=
528529
by rw ← frontier_compl;
529530
convert frontier_le_subset_eq hg hf; simp [ext_iff, eq_comm]
530531

531-
@[continuity] lemma continuous.min : continuous (λb, min (f b) (g b)) :=
532+
@[continuity] lemma continuous.min (hf : continuous f) (hg : continuous g) :
533+
continuous (λb, min (f b) (g b)) :=
532534
have ∀b∈frontier {b | f b ≤ g b}, f b = g b, from assume b hb, frontier_le_subset_eq hf hg hb,
533535
continuous_if this hf hg
534536

535-
@[continuity] lemma continuous.max : continuous (λb, max (f b) (g b)) :=
537+
@[continuity] lemma continuous.max (hf : continuous f) (hg : continuous g) :
538+
continuous (λb, max (f b) (g b)) :=
536539
@continuous.min (order_dual α) _ _ _ _ _ _ _ hf hg
537540

538541
end
539542

543+
lemma continuous_min : continuous (λ p : α × α, min p.1 p.2) := continuous_fst.min continuous_snd
544+
545+
lemma continuous_max : continuous (λ p : α × α, max p.1 p.2) := continuous_fst.max continuous_snd
546+
540547
lemma tendsto.max {b : filter β} {a₁ a₂ : α} (hf : tendsto f b (𝓝 a₁)) (hg : tendsto g b (𝓝 a₂)) :
541548
tendsto (λb, max (f b) (g b)) b (𝓝 (max a₁ a₂)) :=
542-
show tendsto ((λp:α×α, max p.1 p.2) ∘ (λb, (f b, g b))) b (𝓝 (max a₁ a₂)),
543-
from tendsto.comp
544-
begin
545-
rw [←nhds_prod_eq],
546-
from continuous_iff_continuous_at.mp (continuous_fst.max continuous_snd) _
547-
end
548-
(hf.prod_mk hg)
549+
(continuous_max.tendsto (a₁, a₂)).comp (hf.prod_mk_nhds hg)
549550

550551
lemma tendsto.min {b : filter β} {a₁ a₂ : α} (hf : tendsto f b (𝓝 a₁)) (hg : tendsto g b (𝓝 a₂)) :
551552
tendsto (λb, min (f b) (g b)) b (𝓝 (min a₁ a₂)) :=
552-
show tendsto ((λp:α×α, min p.1 p.2) ∘ (λb, (f b, g b))) b (𝓝 (min a₁ a₂)),
553-
from tendsto.comp
554-
begin
555-
rw [←nhds_prod_eq],
556-
from continuous_iff_continuous_at.mp (continuous_fst.min continuous_snd) _
557-
end
558-
(hf.prod_mk hg)
553+
(continuous_min.tendsto (a₁, a₂)).comp (hf.prod_mk_nhds hg)
559554

560555
end decidable_linear_order
561556

0 commit comments

Comments
 (0)