diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 798070260732f..dfdf4c6a4850e 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -179,7 +179,7 @@ jobs: - name: check for noisy stdout lines run: | - ! grep "stdout:" stdout.log + ! grep --after-context=1 "stdout:" stdout.log - name: build library_search cache run: lake build -KCI MathlibExtras @@ -256,10 +256,7 @@ jobs: run: | git clone https://github.com/leanprover/lean4checker cd lean4checker - # Because `Lean4Checker/Tests/ReduceBool.lean` is non-deterministic (compiles only 1/4 of the time), - # we just keep rebuilding until it works! - # I'll try to come up with something better. - until lake build > /dev/null 2>&1; do :; done + lake build ./test.sh cd .. lake env lean4checker/build/bin/lean4checker diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 67ac5564b1657..5bcb0aef7b5f8 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -16,6 +16,7 @@ on: - 'nolints' # ignore staging branch used by bors, this is handled by bors.yml - 'staging' + merge_group: name: continuous integration @@ -185,7 +186,7 @@ jobs: - name: check for noisy stdout lines run: | - ! grep "stdout:" stdout.log + ! grep --after-context=1 "stdout:" stdout.log - name: build library_search cache run: lake build -KCI MathlibExtras @@ -262,10 +263,7 @@ jobs: run: | git clone https://github.com/leanprover/lean4checker cd lean4checker - # Because `Lean4Checker/Tests/ReduceBool.lean` is non-deterministic (compiles only 1/4 of the time), - # we just keep rebuilding until it works! - # I'll try to come up with something better. - until lake build > /dev/null 2>&1; do :; done + lake build ./test.sh cd .. lake env lean4checker/build/bin/lean4checker diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index a7908e14b7c8c..186b741d46a99 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -165,7 +165,7 @@ jobs: - name: check for noisy stdout lines run: | - ! grep "stdout:" stdout.log + ! grep --after-context=1 "stdout:" stdout.log - name: build library_search cache run: lake build -KCI MathlibExtras @@ -242,10 +242,7 @@ jobs: run: | git clone https://github.com/leanprover/lean4checker cd lean4checker - # Because `Lean4Checker/Tests/ReduceBool.lean` is non-deterministic (compiles only 1/4 of the time), - # we just keep rebuilding until it works! - # I'll try to come up with something better. - until lake build > /dev/null 2>&1; do :; done + lake build ./test.sh cd .. lake env lean4checker/build/bin/lean4checker diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index b0bbee304f7ff..07c0ecb93f7b0 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -183,7 +183,7 @@ jobs: - name: check for noisy stdout lines run: | - ! grep "stdout:" stdout.log + ! grep --after-context=1 "stdout:" stdout.log - name: build library_search cache run: lake build -KCI MathlibExtras @@ -260,10 +260,7 @@ jobs: run: | git clone https://github.com/leanprover/lean4checker cd lean4checker - # Because `Lean4Checker/Tests/ReduceBool.lean` is non-deterministic (compiles only 1/4 of the time), - # we just keep rebuilding until it works! - # I'll try to come up with something better. - until lake build > /dev/null 2>&1; do :; done + lake build ./test.sh cd .. lake env lean4checker/build/bin/lean4checker diff --git a/.github/workflows/label_new_contributor.yml b/.github/workflows/label_new_contributor.yml new file mode 100644 index 0000000000000..e0702e7f085af --- /dev/null +++ b/.github/workflows/label_new_contributor.yml @@ -0,0 +1,45 @@ +name: Label New Contributors + +on: + pull_request: + types: [opened] + +jobs: + label-new-contributor: + runs-on: ubuntu-latest + permissions: + pull-requests: write + + steps: + - name: Get PR author + id: pr-author + run: echo "::set-output name=author::$(jq -r .pull_request.user.login "$GITHUB_EVENT_PATH")" + shell: bash + + - name: Check if user is new contributor + id: check-contributor + uses: actions/github-script@v5 + with: + script: | + const { owner, repo } = context.repo + const author = "${{ steps.pr-author.outputs.author }}" + const prs = await github.rest.pulls.list({ + owner, + repo, + state: "closed", + author + }) + return (prs.data.length < 5) + + - name: Add label if new contributor + if: steps.check-contributor.outputs.result == 'true' + uses: actions/github-script@v5 + with: + script: | + const { owner, repo, number } = context.issue + await github.rest.issues.addLabels({ + owner, + repo, + issue_number: number, + labels: ['new-contributor'] + }) diff --git a/.github/workflows/mk_build_yml.sh b/.github/workflows/mk_build_yml.sh index e41f33b2205e0..96cf301b33a26 100755 --- a/.github/workflows/mk_build_yml.sh +++ b/.github/workflows/mk_build_yml.sh @@ -28,6 +28,7 @@ on: - 'nolints' # ignore staging branch used by bors, this is handled by bors.yml - 'staging' + merge_group: name: continuous integration EOF diff --git a/.github/workflows/nightly_bump_toolchain.yml b/.github/workflows/nightly_bump_toolchain.yml index 82e8c87bf7385..111fb488c04f5 100644 --- a/.github/workflows/nightly_bump_toolchain.yml +++ b/.github/workflows/nightly_bump_toolchain.yml @@ -10,7 +10,7 @@ jobs: steps: - name: Checkout code - uses: actions/checkout@v2 + uses: actions/checkout@v3 with: ref: nightly-testing # checkout nightly-testing branch token: ${{ secrets.NIGHTLY_TESTING }} diff --git a/.github/workflows/nightly_detect_failure.yml b/.github/workflows/nightly_detect_failure.yml index 1f1c9c4f13c58..50c1dbf53b50e 100644 --- a/.github/workflows/nightly_detect_failure.yml +++ b/.github/workflows/nightly_detect_failure.yml @@ -23,3 +23,25 @@ jobs: topic: 'CI failure on the nightly-testing branch' content: | The latest CI for branch#nightly-testing has [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.event.workflow_run.id }}). + + handle_success: + if: ${{ github.event.workflow_run.conclusion == 'success' && github.event.workflow_run.head_branch == 'nightly-testing' }} + runs-on: ubuntu-latest + + steps: + - name: Checkout code + uses: actions/checkout@v3 + with: + ref: nightly-testing # checkout nightly-testing branch + fetch-depth: 0 # checkout all branches so that we can push from `nightly-testing` to `nightly-testing-YYYY-MM-DD` + token: ${{ secrets.NIGHTLY_TESTING }} + - name: Update the nightly-testing-YYYY-MM-DD branch + run: | + toolchain=$( `(tactic| aesop) +macro "◾" : term => `(term| by aesop) + +namespace IfExpr + +/-! +We add some local simp lemmas so we can unfold the definitions of the normalization condition. +-/ +attribute [local simp] normalized hasNestedIf hasConstantIf hasRedundantIf disjoint vars + List.disjoint + +/-! +Adding these lemmas to the simp set allows Lean to handle the termination proof automatically. +-/ +attribute [local simp] Nat.lt_add_one_iff le_add_of_le_right + +/-! +Some further simp lemmas for handling if-then-else statements. +-/ +attribute [local simp] apply_ite ite_eq_iff' + +-- A copy of Lean's `decide_eq_true_eq` which unifies the `Decidable` instance +-- rather than finding it by typeclass search. +-- See https://github.com/leanprover/lean4/pull/2816 +@[simp] theorem decide_eq_true_eq {i : Decidable p} : (@decide p i = true) = p := + _root_.decide_eq_true_eq + + +/-! +Simp lemmas for `eval`. +We don't want a `simp` lemma for `(ite i t e).eval` in general, only once we know the shape of `i`. +-/ +@[simp] theorem eval_lit : (lit b).eval f = b := rfl +@[simp] theorem eval_var : (var i).eval f = f i := rfl +@[simp] theorem eval_ite_lit : + (ite (.lit b) t e).eval f = bif b then t.eval f else e.eval f := rfl +@[simp] theorem eval_ite_var : + (ite (.var i) t e).eval f = bif f i then t.eval f else e.eval f := rfl +@[simp] theorem eval_ite_ite : + (ite (ite a b c) d e).eval f = (ite a (ite b d e) (ite c d e)).eval f := by + cases h : eval f a <;> simp_all [eval] + +/-- Custom size function for if-expressions, used for proving termination. -/ +@[simp] def normSize : IfExpr → Nat + | lit _ => 0 + | var _ => 1 + | .ite i t e => 2 * normSize i + max (normSize t) (normSize e) + 1 + +/-- Normalizes the expression at the same time as assigning all variables in +`e` to the literal booleans given by `l` -/ +def normalize (l : AList (fun _ : ℕ => Bool)) : + (e : IfExpr) → { e' : IfExpr // + (∀ f, e'.eval f = e.eval (fun w => (l.lookup w).elim (f w) (fun b => b))) + ∧ e'.normalized + ∧ ∀ (v : ℕ), v ∈ vars e' → l.lookup v = none } + | lit b => ⟨lit b, ◾⟩ + | var v => + match h : l.lookup v with + | none => ⟨var v, ◾⟩ + | some b => ⟨lit b, ◾⟩ + | .ite (lit true) t e => have t' := normalize l t; ⟨t'.1, ◾⟩ + | .ite (lit false) t e => have e' := normalize l e; ⟨e'.1, ◾⟩ + | .ite (.ite a b c) t e => have i' := normalize l (.ite a (.ite b t e) (.ite c t e)); ⟨i'.1, ◾⟩ + | .ite (var v) t e => + match h : l.lookup v with + | none => + have ⟨t', ht₁, ht₂, ht₃⟩ := normalize (l.insert v true) t + have ⟨e', he₁, he₂, he₃⟩ := normalize (l.insert v false) e + ⟨if t' = e' then t' else .ite (var v) t' e', by + refine ⟨fun f => ?_, ?_, fun w b => ?_⟩ + · -- eval = eval + simp + cases hfv : f v + · simp_all + congr + ext w + by_cases w = v <;> ◾ + · simp [h, ht₁] + congr + ext w + by_cases w = v <;> ◾ + · -- normalized + have := ht₃ v + have := he₃ v + ◾ + · -- lookup = none + have := ht₃ w + have := he₃ w + by_cases w = v <;> ◾⟩ + | some b => + have i' := normalize l (.ite (lit b) t e); ⟨i'.1, ◾⟩ + termination_by normalize e => e.normSize + decreasing_by { decreasing_with simp (config := { arith := true }) [Zero.zero]; done } + +/- +We recall the statement of the if-normalization problem. + +We want a function from if-expressions to if-expressions, +that outputs normalized if-expressions and preserves meaning. +-/ +recall IfNormalization := + { Z : IfExpr → IfExpr // ∀ e, (Z e).normalized ∧ (Z e).eval = e.eval } + +example : IfNormalization := + ⟨_, fun e => ⟨(IfExpr.normalize ∅ e).2.2.1, by simp [(IfExpr.normalize ∅ e).2.1]⟩⟩ diff --git a/Archive/Examples/IfNormalization/Statement.lean b/Archive/Examples/IfNormalization/Statement.lean new file mode 100644 index 0000000000000..944efe6ae7c26 --- /dev/null +++ b/Archive/Examples/IfNormalization/Statement.lean @@ -0,0 +1,107 @@ +/- +Copyright (c) 2023 Lean FRO LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ + +/-! +# If normalization + +Rustan Leino, Stephan Merz, and Natarajan Shankar have recently been discussing challenge problems +to compare proof assistants. +(See https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Rustan's.20challenge) + +Their first suggestion was "if-normalization". + +This file contains a Lean formulation of the problem. See `Result.lean` for a Lean solution. +-/ + +/-- An if-expression is either boolean literal, a numbered variable, +or an if-then-else expression where each subexpression is an if-expression. -/ +inductive IfExpr + | lit : Bool → IfExpr + | var : Nat → IfExpr + | ite : IfExpr → IfExpr → IfExpr → IfExpr +deriving DecidableEq, Repr + +namespace IfExpr + +/-- +An if-expression has a "nested if" if it contains +an if-then-else where the "if" is itself an if-then-else. +-/ +def hasNestedIf : IfExpr → Bool + | lit _ => false + | var _ => false + | ite (ite _ _ _) _ _ => true + | ite _ t e => t.hasNestedIf || e.hasNestedIf + +/-- +An if-expression has a "constant if" if it contains +an if-then-else where the "if" is itself a literal. +-/ +def hasConstantIf : IfExpr → Bool + | lit _ => false + | var _ => false + | ite (lit _) _ _ => true + | ite i t e => i.hasConstantIf || t.hasConstantIf || e.hasConstantIf + +/-- +An if-expression has a "redundant if" if it contains +an if-then-else where the then and else clauses are identical. +-/ +def hasRedundantIf : IfExpr → Bool + | lit _ => false + | var _ => false + | ite i t e => t == e || i.hasRedundantIf || t.hasRedundantIf || e.hasRedundantIf + +/-- +All the variables appearing in an if-expressions, read left to right, without removing duplicates. +-/ +def vars : IfExpr → List Nat + | lit _ => [] + | var i => [i] + | ite i t e => i.vars ++ t.vars ++ e.vars + +/-- +A helper function to specify that two lists are disjoint. +-/ +def _root_.List.disjoint {α} [DecidableEq α] : List α → List α → Bool + | [], _ => true + | x::xs, ys => x ∉ ys && xs.disjoint ys + +/-- +An if expression evaluates each variable at most once if for each if-then-else +the variables in the if clause are disjoint from the variables in the then clause, and +the variables in the if clause are disjoint from the variables in the else clause. +-/ +def disjoint : IfExpr → Bool + | lit _ => true + | var _ => true + | ite i t e => + i.vars.disjoint t.vars && i.vars.disjoint e.vars && i.disjoint && t.disjoint && e.disjoint + +/-- +An if expression is "normalized" if it has not nested, constant, or redundant ifs, +and it evaluates each variable at most once. +-/ +def normalized (e : IfExpr) : Bool := + !e.hasNestedIf && !e.hasConstantIf && !e.hasRedundantIf && e.disjoint + +/-- +The evaluation of an if expresssion at some assignment of variables. +-/ +def eval (f : Nat → Bool) : IfExpr → Bool + | lit b => b + | var i => f i + | ite i t e => bif i.eval f then t.eval f else e.eval f + +end IfExpr + +/-- +This is the statement of the if normalization problem. + +We require a function from that transforms if expressions to normalized if expressions, +preserving all evaluations. +-/ +def IfNormalization : Type := { Z : IfExpr → IfExpr // ∀ e, (Z e).normalized ∧ (Z e).eval = e.eval } diff --git a/Archive/Examples/IfNormalization/WithoutAesop.lean b/Archive/Examples/IfNormalization/WithoutAesop.lean new file mode 100644 index 0000000000000..254ac3b982785 --- /dev/null +++ b/Archive/Examples/IfNormalization/WithoutAesop.lean @@ -0,0 +1,131 @@ +/- +Copyright (c) 2023 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Scott Morrison +-/ +import Archive.Examples.IfNormalization.Statement +import Mathlib.Data.List.AList + +/-! +# A variant of Chris Hughes' solution for the if normalization challenge. + +In this variant we eschew the use of `aesop`, and instead write out the proofs. + +(In order to avoid duplicated names with `Result.lean`, +we put primes on the declarations in the file.) +-/ + +set_option autoImplicit true + +namespace IfExpr + +attribute [local simp] eval normalized hasNestedIf hasConstantIf hasRedundantIf disjoint vars + List.disjoint max_add_add_right max_mul_mul_left Nat.lt_add_one_iff le_add_of_le_right + +-- A copy of Lean's `decide_eq_true_eq` which unifies the `Decidable` instance +-- rather than finding it by typeclass search. +-- See https://github.com/leanprover/lean4/pull/2816 +@[simp] theorem decide_eq_true_eq' {i : Decidable p} : (@decide p i = true) = p := + _root_.decide_eq_true_eq + +theorem eval_ite_ite' : + (ite (ite a b c) d e).eval f = (ite a (ite b d e) (ite c d e)).eval f := by + cases h : eval f a <;> simp_all + +/-- Custom size function for if-expressions, used for proving termination. -/ +@[simp] def normSize' : IfExpr → ℕ + | lit _ => 0 + | var _ => 1 + | .ite i t e => 2 * normSize' i + max (normSize' t) (normSize' e) + 1 + +/-- Normalizes the expression at the same time as assigning all variables in +`e` to the literal booleans given by `l` -/ +def normalize' (l : AList (fun _ : ℕ => Bool)) : + (e : IfExpr) → { e' : IfExpr // + (∀ f, e'.eval f = e.eval (fun w => (l.lookup w).elim (f w) (fun b => b))) + ∧ e'.normalized + ∧ ∀ (v : ℕ), v ∈ vars e' → l.lookup v = none } + | lit b => ⟨lit b, by simp⟩ + | var v => + match h : l.lookup v with + | none => ⟨var v, by simp_all⟩ + | some b => ⟨lit b, by simp_all⟩ + | .ite (lit true) t e => + have ⟨t', ht'⟩ := normalize' l t + ⟨t', by simp_all⟩ + | .ite (lit false) t e => + have ⟨e', he'⟩ := normalize' l e + ⟨e', by simp_all⟩ + | .ite (.ite a b c) d e => + have ⟨t', ht₁, ht₂⟩ := normalize' l (.ite a (.ite b d e) (.ite c d e)) + ⟨t', fun f => by rw [ht₁, eval_ite_ite'], ht₂⟩ + | .ite (var v) t e => + match h : l.lookup v with + | none => + have ⟨t', ht₁, ht₂, ht₃⟩ := normalize' (l.insert v true) t + have ⟨e', he₁, he₂, he₃⟩ := normalize' (l.insert v false) e + ⟨if t' = e' then t' else .ite (var v) t' e', by + refine ⟨fun f => ?_, ?_, fun w b => ?_⟩ + · simp only [eval, apply_ite, ite_eq_iff'] + cases hfv : f v + · simp (config := {contextual := true}) only [cond_false, h, he₁] + refine ⟨fun _ => ?_, fun _ => ?_⟩ + · congr + ext w + by_cases w = v <;> rename_i x + · substs h + simp_all + · simp_all + · congr + ext w + by_cases w = v <;> rename_i x + · substs h + simp_all + · simp_all + · simp only [cond_true, h, ht₁] + refine ⟨fun _ => ?_, fun _ => ?_⟩ + · congr + ext w + by_cases w = v <;> rename_i x + · substs h + simp_all + · simp_all + · congr + ext w + by_cases w = v <;> rename_i x + · substs h + simp_all + · simp_all + · have := ht₃ v + have := he₃ v + simp_all? says simp_all only [Option.elim, ne_eq, normalized, Bool.and_eq_true, + Bool.not_eq_true', AList.lookup_insert] + obtain ⟨⟨⟨tn, tc⟩, tr⟩, td⟩ := ht₂ + split <;> rename_i h' + · subst h' + simp_all + · simp_all? says simp_all only [ne_eq, hasNestedIf, Bool.or_self, hasConstantIf, + and_self, hasRedundantIf, Bool.or_false, beq_eq_false_iff_ne, not_false_eq_true, + disjoint, List.disjoint, Bool.and_true, Bool.and_eq_true, decide_eq_true_eq', + true_and] + constructor <;> assumption + · have := ht₃ w + have := he₃ w + by_cases w = v + · subst h; simp_all + · simp_all? says simp_all only [Option.elim, ne_eq, normalized, Bool.and_eq_true, + Bool.not_eq_true', not_false_eq_true, AList.lookup_insert_ne] + obtain ⟨⟨⟨en, ec⟩, er⟩, ed⟩ := he₂ + split at b <;> rename_i h' + · subst h'; simp_all + · simp_all only [ne_eq, vars, List.singleton_append, List.cons_append, + Bool.not_eq_true, List.mem_cons, List.mem_append, false_or] + cases b <;> simp_all⟩ + | some b => + have ⟨e', he'⟩ := normalize' l (.ite (lit b) t e) + ⟨e', by simp_all⟩ + termination_by normalize' e => e.normSize' + +example : IfNormalization := + ⟨fun e => (normalize' ∅ e).1, + fun e => ⟨(normalize' ∅ e).2.2.1, by simp [(normalize' ∅ e).2.1]⟩⟩ diff --git a/Archive/Imo/Imo1960Q1.lean b/Archive/Imo/Imo1960Q1.lean index 5b6a4a08aea0e..d4e98c61b6659 100644 --- a/Archive/Imo/Imo1960Q1.lean +++ b/Archive/Imo/Imo1960Q1.lean @@ -95,7 +95,12 @@ theorem right_direction {n : ℕ} : ProblemPredicate n → SolutionPredicate n : have := searchUpTo_start iterate 82 replace := - searchUpTo_step this (by norm_num1; rfl) (by norm_num1; rfl) (by norm_num1; rfl) (by norm_num) + searchUpTo_step this (by norm_num1; rfl) (by norm_num1; rfl) (by norm_num1; rfl) + (by -- This used to be just `norm_num`, but after leanprover/lean4#2790, + -- that triggers a max recursion depth exception. As a workaround, we + -- manually rewrite with `Nat.digits_of_two_le_of_pos` first. + rw [Nat.digits_of_two_le_of_pos (by norm_num) (by norm_num)] + norm_num <;> decide) exact searchUpTo_end this #align imo1960_q1.right_direction Imo1960Q1.right_direction diff --git a/Archive/Imo/Imo1962Q1.lean b/Archive/Imo/Imo1962Q1.lean index 951ab2d688130..97245b3333242 100644 --- a/Archive/Imo/Imo1962Q1.lean +++ b/Archive/Imo/Imo1962Q1.lean @@ -131,7 +131,18 @@ Now we combine these cases to show that 153846 is the smallest solution. theorem satisfied_by_153846 : ProblemPredicate 153846 := by - norm_num [ProblemPredicate] + -- This proof used to be the single line `norm_num [ProblemPredicate]`. + -- After leanprover/lean4#2790, that triggers a max recursion depth exception. + -- As a workaround, we manually apply `Nat.digits_of_two_le_of_pos` a few times + -- before invoking `norm_num`. + unfold ProblemPredicate + have two_le_ten : 2 ≤ 10 := by norm_num + rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)] + rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)] + rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)] + rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)] + norm_num + decide #align imo1962_q1.satisfied_by_153846 Imo1962Q1.satisfied_by_153846 theorem no_smaller_solutions (n : ℕ) (h1 : ProblemPredicate n) : n ≥ 153846 := by diff --git a/Archive/Imo/Imo1964Q1.lean b/Archive/Imo/Imo1964Q1.lean index 0811661468c43..d0376117a076c 100644 --- a/Archive/Imo/Imo1964Q1.lean +++ b/Archive/Imo/Imo1964Q1.lean @@ -34,7 +34,7 @@ theorem two_pow_mod_seven (n : ℕ) : 2 ^ n ≡ 2 ^ (n % 3) [MOD 7] := let t := n % 3 calc 2 ^ n = 2 ^ (3 * (n / 3) + t) := by rw [Nat.div_add_mod] _ = (2 ^ 3) ^ (n / 3) * 2 ^ t := by rw [pow_add, pow_mul] - _ ≡ 1 ^ (n / 3) * 2 ^ t [MOD 7] := by gcongr; norm_num + _ ≡ 1 ^ (n / 3) * 2 ^ t [MOD 7] := by gcongr; decide _ = 2 ^ t := by ring /-! @@ -68,5 +68,5 @@ theorem imo1964_q1b (n : ℕ) : ¬7 ∣ 2 ^ n + 1 := by have H : 2 ^ t + 1 ≡ 0 [MOD 7] · calc 2 ^ t + 1 ≡ 2 ^ n + 1 [MOD 7 ] := by gcongr ?_ + 1; exact (two_pow_mod_seven n).symm _ ≡ 0 [MOD 7] := h.modEq_zero_nat - interval_cases t <;> norm_num at H + interval_cases t <;> contradiction #align imo1964_q1b imo1964_q1b diff --git a/Archive/Imo/Imo1981Q3.lean b/Archive/Imo/Imo1981Q3.lean index 5b0765b4a9cf9..1a136a674c3ed 100644 --- a/Archive/Imo/Imo1981Q3.lean +++ b/Archive/Imo/Imo1981Q3.lean @@ -129,7 +129,7 @@ theorem imp_fib {n : ℕ} : ∀ m : ℕ, NatPredicate N m n → ∃ k : ℕ, m = obtain (rfl : 1 = n) | (h4 : 1 < n) := (succ_le_iff.mpr h2.n_pos).eq_or_lt · use 1 have h5 : 1 ≤ m := succ_le_iff.mpr h2.m_pos - simpa [fib_one, fib_two] using (h3.antisymm h5 : m = 1) + simpa [fib_one, fib_two, (by decide : 1 + 1 = 2)] using (h3.antisymm h5 : m = 1) · obtain (rfl : m = n) | (h6 : m < n) := h3.eq_or_lt · exact absurd h2.eq_imp_1 (Nat.ne_of_gt h4) · have h7 : NatPredicate N (n - m) m := h2.reduction h4 @@ -205,5 +205,7 @@ theorem imo1981_q3 : IsGreatest (specifiedSet 1981) 3524578 := by have := fun h => @solution_greatest 1981 16 h 3524578 norm_num at this apply this - norm_num [ProblemPredicate_iff] + · decide + · decide + · norm_num [ProblemPredicate_iff]; decide #align imo1981_q3 imo1981_q3 diff --git a/Archive/Imo/Imo2005Q4.lean b/Archive/Imo/Imo2005Q4.lean index 56be75dd551a7..7e9272d4a0556 100644 --- a/Archive/Imo/Imo2005Q4.lean +++ b/Archive/Imo/Imo2005Q4.lean @@ -29,7 +29,7 @@ def a (n : ℕ) : ℤ := theorem find_specified_factor {p : ℕ} (hp : Nat.Prime p) (hp2 : p ≠ 2) (hp3 : p ≠ 3) : ↑p ∣ a (p - 2) := by -- Since `p` is neither `2` nor `3`, it is coprime with `2`, `3`, and `6` - rw [Ne.def, ← Nat.prime_dvd_prime_iff_eq hp (by norm_num), ← Nat.Prime.coprime_iff_not_dvd hp] + rw [Ne.def, ← Nat.prime_dvd_prime_iff_eq hp (by decide), ← Nat.Prime.coprime_iff_not_dvd hp] at hp2 hp3 have : Int.gcd p 6 = 1 := Nat.coprime_mul_iff_right.2 ⟨hp2, hp3⟩ -- Nat arithmetic needed to deal with powers @@ -71,10 +71,10 @@ theorem imo2005_q4 {k : ℕ} (hk : 0 < k) : (∀ n : ℕ, 1 ≤ n → IsCoprime -- For `p = 2` and `p = 3`, take `n = 1` and `n = 2`, respectively by_cases hp2 : p = 2 · rw [hp2] at h - apply h 1 <;> norm_num + apply h 1 <;> decide by_cases hp3 : p = 3 · rw [hp3] at h - apply h 2 <;> norm_num + apply h 2 <;> decide -- Otherwise, take `n = p - 2` refine h (p - 2) ?_ (find_specified_factor hp hp2 hp3) calc diff --git a/Archive/Imo/Imo2019Q1.lean b/Archive/Imo/Imo2019Q1.lean index 6e486889066d0..2bcd5a9f405b9 100644 --- a/Archive/Imo/Imo2019Q1.lean +++ b/Archive/Imo/Imo2019Q1.lean @@ -27,7 +27,7 @@ theorem imo2019_q1 (f : ℤ → ℤ) : (∀ a b : ℤ, f (2 * a) + 2 * f b = f (f (a + b))) ↔ f = 0 ∨ ∃ c, f = fun x => 2 * x + c := by constructor; swap -- easy way: f(x)=0 and f(x)=2x+c work. - · rintro (rfl | ⟨c, rfl⟩) <;> intros <;> simp only [Pi.zero_apply]; ring + · rintro (rfl | ⟨c, rfl⟩) <;> intros <;> norm_num; ring -- hard way. intro hf -- functional equation diff --git a/Archive/Imo/Imo2019Q2.lean b/Archive/Imo/Imo2019Q2.lean index c301c615ed3ad..3c2131796191b 100644 --- a/Archive/Imo/Imo2019Q2.lean +++ b/Archive/Imo/Imo2019Q2.lean @@ -442,7 +442,7 @@ theorem not_collinear_QPA₂ : ¬Collinear ℝ ({cfg.Q, cfg.P, cfg.A₂} : Set P have h : Cospherical ({cfg.B, cfg.A, cfg.A₂} : Set Pt) := by refine' cfg.triangleABC.circumsphere.cospherical.subset _ simp only [Set.insert_subset_iff, cfg.A_mem_circumsphere, cfg.B_mem_circumsphere, - cfg.A₂_mem_circumsphere, Sphere.mem_coe, Set.singleton_subset_iff] + cfg.A₂_mem_circumsphere, Sphere.mem_coe, Set.singleton_subset_iff, and_true] exact h.affineIndependent_of_ne cfg.A_ne_B.symm cfg.A₂_ne_B.symm cfg.A₂_ne_A.symm #align imo2019_q2.imo2019q2_cfg.not_collinear_QPA₂ Imo2019Q2.Imo2019q2Cfg.not_collinear_QPA₂ diff --git a/Archive/Imo/Imo2019Q4.lean b/Archive/Imo/Imo2019Q4.lean index b52f6085c03ce..edde9fc8f8d3e 100644 --- a/Archive/Imo/Imo2019Q4.lean +++ b/Archive/Imo/Imo2019Q4.lean @@ -65,7 +65,7 @@ theorem upper_bound {k n : ℕ} (hk : k > 0) _ ≤ k ! := by gcongr clear h h2 induction' n, hn using Nat.le_induction with n' hn' IH - · norm_num + · decide let A := ∑ i in range n', i have le_sum : ∑ i in range 6, i ≤ A · apply sum_le_sum_of_subset @@ -87,8 +87,7 @@ theorem imo2019_q4 {k n : ℕ} (hk : k > 0) (hn : n > 0) : -- The implication `←` holds. constructor swap - · rintro (h | h) <;> simp [Prod.ext_iff] at h <;> rcases h with ⟨rfl, rfl⟩ <;> - norm_num [prod_range_succ, succ_mul] + · rintro (h | h) <;> simp [Prod.ext_iff] at h <;> rcases h with ⟨rfl, rfl⟩ <;> decide intro h -- We know that n < 6. have := Imo2019Q4.upper_bound hk h @@ -101,9 +100,9 @@ theorem imo2019_q4 {k n : ℕ} (hk : k > 0) (hn : n > 0) : exact h; rw [h]; norm_num all_goals exfalso; norm_num [prod_range_succ] at h; norm_cast at h -- n = 3 - · refine' monotone_factorial.ne_of_lt_of_lt_nat 5 _ _ _ h <;> norm_num + · refine' monotone_factorial.ne_of_lt_of_lt_nat 5 _ _ _ h <;> decide -- n = 4 - · refine' monotone_factorial.ne_of_lt_of_lt_nat 7 _ _ _ h <;> norm_num + · refine' monotone_factorial.ne_of_lt_of_lt_nat 7 _ _ _ h <;> decide -- n = 5 - · refine' monotone_factorial.ne_of_lt_of_lt_nat 10 _ _ _ h <;> norm_num + · refine' monotone_factorial.ne_of_lt_of_lt_nat 10 _ _ _ h <;> decide #align imo2019_q4 imo2019_q4 diff --git a/Archive/MiuLanguage/DecisionNec.lean b/Archive/MiuLanguage/DecisionNec.lean index e1c326253f2c1..ac22effdd4b50 100644 --- a/Archive/MiuLanguage/DecisionNec.lean +++ b/Archive/MiuLanguage/DecisionNec.lean @@ -61,7 +61,7 @@ theorem mod3_eq_1_or_mod3_eq_2 {a b : ℕ} (h1 : a % 3 = 1 ∨ a % 3 = 2) · rw [h2]; exact h1 · cases' h1 with h1 h1 · right; simp [h2, mul_mod, h1, Nat.succ_lt_succ] - · left; simp [h2, mul_mod, h1] + · left; simp only [h2, mul_mod, h1, mod_mod]; decide #align miu.mod3_eq_1_or_mod3_eq_2 Miu.mod3_eq_1_or_mod3_eq_2 /-- `count_equiv_one_or_two_mod3_of_derivable` shows any derivable string must have a `count I` that @@ -80,7 +80,7 @@ theorem count_equiv_one_or_two_mod3_of_derivable (en : Miustr) : simp_rw [count_cons_self, count_nil, count_cons, ite_false, add_right_comm, add_mod_right] simp · left; rw [count_append, count_append, count_append] - simp only [ne_eq, count_cons_of_ne, count_nil, add_zero] + simp only [ne_eq, not_false_eq_true, count_cons_of_ne, count_nil, add_zero] #align miu.count_equiv_one_or_two_mod3_of_derivable Miu.count_equiv_one_or_two_mod3_of_derivable /-- Using the above theorem, we solve the MU puzzle, showing that `"MU"` is not derivable. @@ -134,7 +134,7 @@ theorem goodm_of_rule1 (xs : Miustr) (h₁ : Derivable (xs ++ ↑[I])) (h₂ : G exact mhead · change ¬M ∈ tail (xs ++ ↑([I] ++ [U])) rw [← append_assoc, tail_append_singleton_of_ne_nil] - · simp_rw [mem_append, not_or, and_true]; exact nmtail + · simp_rw [mem_append, mem_singleton, or_false]; exact nmtail · exact append_ne_nil_of_ne_nil_left _ _ this #align miu.goodm_of_rule1 Miu.goodm_of_rule1 @@ -145,7 +145,7 @@ theorem goodm_of_rule2 (xs : Miustr) (_ : Derivable (M :: xs)) (h₂ : Goodm (M · cases' h₂ with mhead mtail contrapose! mtail rw [cons_append] at mtail - exact (or_self_iff _).mp (mem_append.mp mtail) + exact or_self_iff.mp (mem_append.mp mtail) #align miu.goodm_of_rule2 Miu.goodm_of_rule2 theorem goodm_of_rule3 (as bs : Miustr) (h₁ : Derivable (as ++ ↑[I, I, I] ++ bs)) diff --git a/Archive/MiuLanguage/DecisionSuf.lean b/Archive/MiuLanguage/DecisionSuf.lean index 7f9fdaba0ac4a..6c7ff233e8c95 100644 --- a/Archive/MiuLanguage/DecisionSuf.lean +++ b/Archive/MiuLanguage/DecisionSuf.lean @@ -267,7 +267,7 @@ theorem count_I_eq_length_of_count_U_zero_and_neg_mem {ys : Miustr} (hu : count · simpa only [count] · rw [mem_cons, not_or] at hm; exact hm.2 · -- case `x = U` gives a contradiction. - exfalso; simp only [count, countP_cons_of_pos] at hu + exfalso; simp only [count, countP_cons_of_pos (· == U) _ (rfl : U == U)] at hu exact succ_ne_zero _ hu set_option linter.uppercaseLean3 false in #align miu.count_I_eq_length_of_count_U_zero_and_neg_mem Miu.count_I_eq_length_of_count_U_zero_and_neg_mem @@ -277,7 +277,8 @@ set_option linter.uppercaseLean3 false in theorem base_case_suf (en : Miustr) (h : Decstr en) (hu : count U en = 0) : Derivable en := by rcases h with ⟨⟨mhead, nmtail⟩, hi⟩ have : en ≠ nil := by - intro k; simp only [k, count, countP, if_false, zero_mod, zero_ne_one, false_or_iff] at hi + intro k + simp only [k, count, countP, countP.go, if_false, zero_mod, zero_ne_one, false_or_iff] at hi rcases exists_cons_of_ne_nil this with ⟨y, ys, rfl⟩ rcases mhead rsuffices ⟨c, rfl, hc⟩ : ∃ c, replicate c I = ys ∧ (c % 3 = 1 ∨ c % 3 = 2) @@ -331,7 +332,7 @@ theorem ind_hyp_suf (k : ℕ) (ys : Miustr) (hu : count U ys = succ k) (hdec : D rw [cons_append, cons_append] dsimp [tail] at nmtail ⊢ rw [mem_append] at nmtail - simpa only [mem_append, mem_cons, false_or_iff, or_false_iff] using nmtail + simpa only [append_assoc, cons_append, nil_append, mem_append, mem_cons, false_or] using nmtail · rw [count_append, count_append]; rw [← cons_append, count_append] at hic simp only [count_cons_self, count_nil, count_cons, if_false] at hic ⊢ rw [add_right_comm, add_mod_right]; exact hic diff --git a/Archive/Wiedijk100Theorems/AbelRuffini.lean b/Archive/Wiedijk100Theorems/AbelRuffini.lean index 72eb193a353d3..33efd6ea814d1 100644 --- a/Archive/Wiedijk100Theorems/AbelRuffini.lean +++ b/Archive/Wiedijk100Theorems/AbelRuffini.lean @@ -94,10 +94,11 @@ theorem irreducible_Phi (p : ℕ) (hp : p.Prime) (hpa : p ∣ a) (hpb : p ∣ b) rw [mem_span_singleton] rw [degree_Phi] at hn; norm_cast at hn interval_cases hn : n <;> - simp only [Φ, coeff_X_pow, coeff_C, Int.coe_nat_dvd.mpr, hpb, if_true, coeff_C_mul, if_false, - coeff_X_zero, hpa, coeff_add, zero_add, mul_zero, coeff_sub, add_zero, zero_sub, dvd_neg, - neg_zero, dvd_mul_of_dvd_left] + simp (config := {decide := true}) only [Φ, coeff_X_pow, coeff_C, Int.coe_nat_dvd.mpr, hpb, + if_true, coeff_C_mul, if_false, coeff_X_zero, hpa, coeff_add, zero_add, mul_zero, coeff_sub, + add_zero, zero_sub, dvd_neg, neg_zero, dvd_mul_of_dvd_left] · simp only [degree_Phi, ← WithBot.coe_zero, WithBot.coe_lt_coe, Nat.succ_pos'] + decide · rw [coeff_zero_Phi, span_singleton_pow, mem_span_singleton] exact mt Int.coe_nat_dvd.mp hp2b all_goals exact Monic.isPrimitive (monic_Phi a b) @@ -136,7 +137,7 @@ theorem real_roots_Phi_ge_aux (hab : b < a) : have hfa := calc f (-a) = (a : ℝ) ^ 2 - (a : ℝ) ^ 5 + b := by - norm_num [hf, ← sq, sub_eq_add_neg, add_comm, Odd.neg_pow] + norm_num [hf, ← sq, sub_eq_add_neg, add_comm, Odd.neg_pow (by decide : Odd 5)] _ ≤ (a : ℝ) ^ 2 - (a : ℝ) ^ 3 + (a - 1) := by refine' add_le_add (sub_le_sub_left (pow_le_pow ha _) _) _ <;> linarith _ = -((a : ℝ) - 1) ^ 2 * (a + 1) := by ring @@ -163,7 +164,7 @@ theorem complex_roots_Phi (h : (Φ ℚ a b).Separable) : Fintype.card ((Φ ℚ a theorem gal_Phi (hab : b < a) (h_irred : Irreducible (Φ ℚ a b)) : Bijective (galActionHom (Φ ℚ a b) ℂ) := by apply galActionHom_bijective_of_prime_degree' h_irred - · norm_num [natDegree_Phi] + · simp only [natDegree_Phi]; decide · rw [complex_roots_Phi a b h_irred.separable, Nat.succ_le_succ_iff] exact (real_roots_Phi_le a b).trans (Nat.le_succ 3) · simp_rw [complex_roots_Phi a b h_irred.separable, Nat.succ_le_succ_iff] @@ -181,7 +182,7 @@ theorem not_solvable_by_rad (p : ℕ) (x : ℂ) (hx : aeval x (Φ ℚ a b) = 0) #align abel_ruffini.not_solvable_by_rad AbelRuffini.not_solvable_by_rad theorem not_solvable_by_rad' (x : ℂ) (hx : aeval x (Φ ℚ 4 2) = 0) : ¬IsSolvableByRad ℚ x := by - apply not_solvable_by_rad 4 2 2 x hx <;> norm_num + apply not_solvable_by_rad 4 2 2 x hx <;> decide #align abel_ruffini.not_solvable_by_rad' AbelRuffini.not_solvable_by_rad' /-- **Abel-Ruffini Theorem** -/ diff --git a/Archive/Wiedijk100Theorems/AreaOfACircle.lean b/Archive/Wiedijk100Theorems/AreaOfACircle.lean index 3320ee7301be6..c98ccfc1e0236 100644 --- a/Archive/Wiedijk100Theorems/AreaOfACircle.lean +++ b/Archive/Wiedijk100Theorems/AreaOfACircle.lean @@ -99,7 +99,8 @@ theorem area_disc : volume (disc r) = NNReal.pi * r ^ 2 := by _ = ENNReal.ofReal (∫ x in Ioc (-r : ℝ) r, (f - Neg.neg ∘ f) x) := (volume_regionBetween_eq_integral h.neg h measurableSet_Ioc fun x _ => neg_le_self (sqrt_nonneg _)) - _ = ENNReal.ofReal (∫ x in (-r : ℝ)..r, 2 * f x) := by simp [two_mul, integral_of_le] + _ = ENNReal.ofReal (∫ x in (-r : ℝ)..r, 2 * f x) := by + rw [integral_of_le] <;> simp [two_mul, neg_le_self] _ = NNReal.pi * r ^ 2 := by rw_mod_cast [this, ← ENNReal.coe_nnreal_eq] have hle := NNReal.coe_nonneg r obtain heq | hlt := hle.eq_or_lt; · simp [← heq] diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index 84660e6e21d11..f7a5a95024821 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -221,7 +221,9 @@ theorem first_vote_pos : simp [ENNReal.div_self _ _] | 0, q + 1, _ => by rw [counted_left_zero, condCount_singleton] - simp + simp only [List.replicate, Nat.add_eq, add_zero, mem_setOf_eq, List.headI_cons, Nat.cast_zero, + ENNReal.zero_div, ite_eq_right_iff] + decide | p + 1, q + 1, _ => by simp_rw [counted_succ_succ] rw [← condCount_disjoint_union ((countedSequence_finite _ _).image _) diff --git a/Archive/Wiedijk100Theorems/BirthdayProblem.lean b/Archive/Wiedijk100Theorems/BirthdayProblem.lean index 4c3224823ebb2..aadb154d8066e 100644 --- a/Archive/Wiedijk100Theorems/BirthdayProblem.lean +++ b/Archive/Wiedijk100Theorems/BirthdayProblem.lean @@ -29,7 +29,13 @@ local notation "‖" x "‖" => Fintype.card x /-- **Birthday Problem**: set cardinality interpretation. -/ theorem birthday : 2 * ‖Fin 23 ↪ Fin 365‖ < ‖Fin 23 → Fin 365‖ ∧ 2 * ‖Fin 22 ↪ Fin 365‖ > ‖Fin 22 → Fin 365‖ := by - simp only [Nat.descFactorial, Fintype.card_fin, Fintype.card_embedding_eq, Fintype.card_fun] + -- This used to be + -- `simp only [Nat.descFactorial, Fintype.card_fin, Fintype.card_embedding_eq, Fintype.card_fun]` + -- but after leanprover/lean4#2790 that triggers a max recursion depth exception. + -- As a workaround, we make some of the reduction steps more explicit. + rw [Fintype.card_embedding_eq, Fintype.card_fun, Fintype.card_fin, Fintype.card_fin] + rw [Fintype.card_embedding_eq, Fintype.card_fun, Fintype.card_fin, Fintype.card_fin] + decide #align theorems_100.birthday Theorems100.birthday section MeasureTheory @@ -72,12 +78,14 @@ theorem birthday_measure : generalize_proofs hfin have : |hfin.toFinset| = 42200819302092359872395663074908957253749760700776448000000 := by trans ‖Fin 23 ↪ Fin 365‖ - · simp_rw [← Fintype.card_coe, Set.Finite.coeSort_toFinset, Set.coe_setOf] - exact Fintype.card_congr (Equiv.subtypeInjectiveEquivEmbedding _ _) - · simp only [Fintype.card_embedding_eq, Fintype.card_fin, Nat.descFactorial] + · rw [← Fintype.card_coe] + apply Fintype.card_congr + rw [Set.Finite.coeSort_toFinset, Set.coe_setOf] + exact Equiv.subtypeInjectiveEquivEmbedding _ _ + · rw [Fintype.card_embedding_eq, Fintype.card_fin, Fintype.card_fin] + rfl rw [this, ENNReal.lt_div_iff_mul_lt, mul_comm, mul_div, ENNReal.div_lt_iff] - rotate_left; (iterate 2 right; norm_num); (iterate 2 left; norm_num) - norm_cast + rotate_left; (iterate 2 right; norm_num); decide; (iterate 2 left; norm_num) simp only [Fintype.card_pi] norm_num #align theorems_100.birthday_measure Theorems100.birthday_measure diff --git a/Archive/Wiedijk100Theorems/Konigsberg.lean b/Archive/Wiedijk100Theorems/Konigsberg.lean index 816ba1994e37b..a277213885050 100644 --- a/Archive/Wiedijk100Theorems/Konigsberg.lean +++ b/Archive/Wiedijk100Theorems/Konigsberg.lean @@ -75,7 +75,7 @@ lemma degree_eq_degree (v : Verts) : graph.degree v = degree v := by cases v <;> #align konigsberg.degree_eq_degree Konigsberg.degree_eq_degree lemma not_even_degree_iff (w : Verts) : ¬Even (degree w) ↔ w = V1 ∨ w = V2 ∨ w = V3 ∨ w = V4 := by - cases w <;> simp + cases w <;> decide lemma setOf_odd_degree_eq : {v | Odd (graph.degree v)} = {Verts.V1, Verts.V2, Verts.V3, Verts.V4} := by diff --git a/Archive/Wiedijk100Theorems/Partition.lean b/Archive/Wiedijk100Theorems/Partition.lean index 93bb09854ca4f..4ee4b66145fb6 100644 --- a/Archive/Wiedijk100Theorems/Partition.lean +++ b/Archive/Wiedijk100Theorems/Partition.lean @@ -123,7 +123,7 @@ theorem mem_cut {ι : Type*} (s : Finset ι) (n : ℕ) (f : ι → ℕ) : #align theorems_100.mem_cut Theorems100.mem_cut theorem cut_equiv_antidiag (n : ℕ) : - Equiv.finsetCongr (Equiv.boolArrowEquivProd _) (cut univ n) = Nat.antidiagonal n := by + Equiv.finsetCongr (Equiv.boolArrowEquivProd _) (cut univ n) = antidiagonal n := by ext ⟨x₁, x₂⟩ simp_rw [Equiv.finsetCongr_apply, mem_map, Equiv.toEmbedding, Function.Embedding.coeFn_mk, ← Equiv.eq_symm_apply] @@ -157,14 +157,14 @@ theorem cut_empty_succ {ι : Type*} (n : ℕ) : cut (∅ : Finset ι) (n + 1) = theorem cut_insert {ι : Type*} (n : ℕ) (a : ι) (s : Finset ι) (h : a ∉ s) : cut (insert a s) n = - (Nat.antidiagonal n).biUnion fun p : ℕ × ℕ => + (antidiagonal n).biUnion fun p : ℕ × ℕ => (cut s p.snd).map ⟨fun f => f + fun t => if t = a then p.fst else 0, add_left_injective _⟩ := by ext f rw [mem_cut, mem_biUnion, sum_insert h] constructor · rintro ⟨rfl, h₁⟩ - simp only [exists_prop, Function.Embedding.coeFn_mk, mem_map, Nat.mem_antidiagonal, Prod.exists] + simp only [exists_prop, Function.Embedding.coeFn_mk, mem_map, mem_antidiagonal, Prod.exists] refine' ⟨f a, s.sum f, rfl, fun i => if i = a then 0 else f i, _, _⟩ · rw [mem_cut] refine' ⟨_, _⟩ @@ -183,7 +183,7 @@ theorem cut_insert {ι : Type*} (n : ℕ) (a : ι) (s : Finset ι) (h : a ∉ s) obtain rfl | h := eq_or_ne x a · simp · simp [if_neg h] - · simp only [mem_insert, Function.Embedding.coeFn_mk, mem_map, Nat.mem_antidiagonal, Prod.exists, + · simp only [mem_insert, Function.Embedding.coeFn_mk, mem_map, mem_antidiagonal, Prod.exists, exists_prop, mem_cut, not_or] rintro ⟨p, q, rfl, g, ⟨rfl, hg₂⟩, rfl⟩ refine' ⟨_, _⟩ @@ -216,7 +216,7 @@ theorem coeff_prod_range [CommSemiring α] {ι : Type*} (s : Finset ι) (f : ι rw [if_neg, add_zero] exact ne_of_mem_of_not_mem hk hi · simp only [Set.PairwiseDisjoint, Set.Pairwise, Prod.forall, not_and, Ne.def, - Nat.mem_antidiagonal, disjoint_left, mem_map, exists_prop, Function.Embedding.coeFn_mk, + mem_antidiagonal, disjoint_left, mem_map, exists_prop, Function.Embedding.coeFn_mk, exists_imp, not_exists, Finset.mem_coe, Function.onFun, mem_cut, and_imp] rintro p₁ q₁ rfl p₂ q₂ h t x p hp _ hp3 q hq _ hq3 have z := hp3.trans hq3.symm @@ -275,14 +275,14 @@ theorem num_series' [Field α] (i : ℕ) : symm split_ifs with h · suffices - ((Nat.antidiagonal n.succ).filter fun a : ℕ × ℕ => i + 1 ∣ a.fst ∧ a.snd = i + 1).card = + ((antidiagonal n.succ).filter fun a : ℕ × ℕ => i + 1 ∣ a.fst ∧ a.snd = i + 1).card = 1 by simp only [Set.mem_setOf_eq]; convert congr_arg ((↑) : ℕ → α) this; norm_cast rw [card_eq_one] cases' h with p hp refine' ⟨((i + 1) * (p - 1), i + 1), _⟩ ext ⟨a₁, a₂⟩ - simp only [mem_filter, Prod.mk.inj_iff, Nat.mem_antidiagonal, mem_singleton] + simp only [mem_filter, Prod.mk.inj_iff, mem_antidiagonal, mem_singleton] constructor · rintro ⟨a_left, ⟨a, rfl⟩, rfl⟩ refine' ⟨_, rfl⟩ @@ -292,12 +292,12 @@ theorem num_series' [Field α] (i : ℕ) : | 0 => rw [mul_zero] at hp; cases hp | p + 1 => rw [hp]; simp [mul_add] · suffices - (filter (fun a : ℕ × ℕ => i + 1 ∣ a.fst ∧ a.snd = i + 1) (Nat.antidiagonal n.succ)).card = + (filter (fun a : ℕ × ℕ => i + 1 ∣ a.fst ∧ a.snd = i + 1) (antidiagonal n.succ)).card = 0 by simp only [Set.mem_setOf_eq]; convert congr_arg ((↑) : ℕ → α) this; norm_cast rw [card_eq_zero] apply eq_empty_of_forall_not_mem - simp only [Prod.forall, mem_filter, not_and, Nat.mem_antidiagonal] + simp only [Prod.forall, mem_filter, not_and, mem_antidiagonal] rintro _ h₁ h₂ ⟨a, rfl⟩ rfl apply h simp [← h₂] diff --git a/Archive/Wiedijk100Theorems/PerfectNumbers.lean b/Archive/Wiedijk100Theorems/PerfectNumbers.lean index ff1f483234934..c610c03973978 100644 --- a/Archive/Wiedijk100Theorems/PerfectNumbers.lean +++ b/Archive/Wiedijk100Theorems/PerfectNumbers.lean @@ -39,8 +39,8 @@ namespace Nat open Nat.ArithmeticFunction Finset theorem sigma_two_pow_eq_mersenne_succ (k : ℕ) : σ 1 (2 ^ k) = mersenne (k + 1) := by - simp [sigma_one_apply, mersenne, Nat.prime_two, show 2 = 1 + 1 from rfl, - ← geom_sum_mul_add 1 (k + 1)] + simp_rw [sigma_one_apply, mersenne, show 2 = 1 + 1 from rfl, ← geom_sum_mul_add 1 (k + 1)] + norm_num #align theorems_100.nat.sigma_two_pow_eq_mersenne_succ Theorems100.Nat.sigma_two_pow_eq_mersenne_succ /-- Euclid's theorem that Mersenne primes induce perfect numbers -/ diff --git a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean index 5713e9608429c..05b6a23992908 100644 --- a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean +++ b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean @@ -144,7 +144,7 @@ The number of `e < x` for which `e + 1` is a squarefree product of primes smalle theorem card_le_two_pow {x k : ℕ} : card (Finset.filter (fun e => Squarefree (e + 1)) (M x k)) ≤ 2 ^ k := by let M₁ := Finset.filter (fun e => Squarefree (e + 1)) (M x k) - let f s := (Finset.prod s fun a => a) - 1 + let f s := (∏ a in s, a) - 1 let K := powerset (image Nat.succ (range k)) -- Take `e` in `M x k`. If `e + 1` is squarefree, then it is the product of a subset of `[1, k]`. -- It follows that `e` is one less than such a product. diff --git a/Archive/ZagierTwoSquares.lean b/Archive/ZagierTwoSquares.lean index 686263c38918f..87dfcf5a23c7f 100644 --- a/Archive/ZagierTwoSquares.lean +++ b/Archive/ZagierTwoSquares.lean @@ -48,9 +48,10 @@ lemma zagierSet_lower_bound {x y z : ℕ} (h : (x, y, z) ∈ zagierSet k) : 0 < all_goals cases' (Nat.dvd_prime hk.out).1 (dvd_of_mul_left_eq _ h) with e e all_goals - simp only [e, self_eq_add_left, ne_eq, add_eq_zero, and_false, mul_eq_left₀] at h + simp only [e, self_eq_add_left, ne_eq, add_eq_zero, and_false, not_false_eq_true, + mul_eq_left₀] at h simp only [h, zero_add] at hk - exact hk.out + exact Nat.not_prime_one hk.out lemma zagierSet_upper_bound {x y z : ℕ} (h : (x, y, z) ∈ zagierSet k) : x ≤ k + 1 ∧ y ≤ k ∧ z ≤ k := by @@ -143,11 +144,12 @@ theorem eq_of_mem_fixedPoints {t : zagierSet k} (mem : t ∈ fixedPoints (comple rw [mem_fixedPoints_iff, complexInvo, Subtype.mk.injEq] at mem split_ifs at mem with less more <;> -- less (completely handled by the pre-applied `simp_all only`) - simp_all only [not_lt, Prod.mk.injEq, add_right_eq_self, mul_eq_zero, false_or] + simp_all only [not_lt, Prod.mk.injEq, add_right_eq_self, mul_eq_zero, false_or, + lt_self_iff_false] · -- more obtain ⟨_, _, _⟩ := mem; simp_all · -- middle (the one fixed point falls under this case) - simp only [zagierSet, Set.mem_setOf_eq] at h + simp [zagierSet, Set.mem_setOf_eq] at h replace mem := mem.1 rw [tsub_eq_iff_eq_add_of_le more, ← two_mul] at mem replace mem := (mul_left_cancel₀ two_ne_zero mem).symm @@ -194,3 +196,4 @@ theorem Nat.Prime.sq_add_sq' {p : ℕ} [h : Fact p.Prime] (hp : p % 4 = 1) : contrapose key rw [Set.not_nonempty_iff_eq_empty] at key simp_rw [key, Fintype.card_of_isEmpty, card_fixedPoints_eq_one] + decide diff --git a/Counterexamples.lean b/Counterexamples.lean index 18cd76a5ceb5f..dd67bf15a550c 100644 --- a/Counterexamples.lean +++ b/Counterexamples.lean @@ -1,5 +1,6 @@ import Counterexamples.CanonicallyOrderedCommSemiringTwoMul import Counterexamples.CharPZeroNeCharZero +import Counterexamples.CliffordAlgebra_not_injective import Counterexamples.Cyclotomic105 import Counterexamples.DirectSumIsInternal import Counterexamples.Girard diff --git a/Counterexamples/CliffordAlgebra_not_injective.lean b/Counterexamples/CliffordAlgebra_not_injective.lean new file mode 100644 index 0000000000000..27134e15dcf2e --- /dev/null +++ b/Counterexamples/CliffordAlgebra_not_injective.lean @@ -0,0 +1,280 @@ +/- +Copyright (c) 2021 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Algebra.CharP.Pi +import Mathlib.Algebra.CharP.Quotient +import Mathlib.Algebra.CharP.Two +import Mathlib.Data.MvPolynomial.CommRing +import Mathlib.Data.ZMod.Basic +import Mathlib.LinearAlgebra.CliffordAlgebra.Basic +import Mathlib.LinearAlgebra.Finsupp +import Mathlib.RingTheory.MvPolynomial.Basic +import Mathlib.RingTheory.MvPolynomial.Ideal + +/-! # `algebraMap R (CliffordAlgebra Q)` is not always injective. + +A formalization of [Darij Grinberg's answer](https://mathoverflow.net/questions/60596/clifford-pbw-theorem-for-quadratic-form/87958#87958) +to a "Clifford PBW theorem for quadratic form" post on MathOverflow, that provides a counterexample +to `Function.Injective (algebraMap R (CliffordAlgebra Q))`. + +The outline is that we define: + +* $k$ (`Q60596.K`) as the commutative ring $𝔽₂[α, β, γ] / (α², β², γ²)$ +* $L$ (`Q60596.L`) as the $k$-module $⟨x,y,z⟩ / ⟨αx + βy + γz⟩$ +* $Q$ (`Q60596.Q`) as the quadratic form sending $Q(\overline{ax + by = cz}) = a² + b² + c²$ + +and discover that $αβγ ≠ 0$ as an element of $K$, but $αβγ = 0$ as an element of $𝒞l(Q)$. + +Some Zulip discussion at https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/.F0.9D.94.BD.E2.82.82.5B.CE.B1.2C.20.CE.B2.2C.20.CE.B3.5D.20.2F.20.28.CE.B1.C2.B2.2C.20.CE.B2.C2.B2.2C.20.CE.B3.C2.B2.29/near/222716333. +-/ + +noncomputable section + +open scoped BigOperators + +namespace Q60596 + +open MvPolynomial + +/-- The monomial ideal generated by terms of the form $x_ix_i$. -/ +def kIdeal : Ideal (MvPolynomial (Fin 3) (ZMod 2)) := + Ideal.span (Set.range fun i => (X i * X i : MvPolynomial (Fin 3) (ZMod 2))) + +theorem mem_kIdeal_iff (x : MvPolynomial (Fin 3) (ZMod 2)) : + x ∈ kIdeal ↔ ∀ m : Fin 3 →₀ ℕ, m ∈ x.support → ∃ i, 2 ≤ m i := by + have : + kIdeal = Ideal.span ((monomial · (1 : ZMod 2)) '' Set.range (Finsupp.single · 2)) := by + simp_rw [kIdeal, X, monomial_mul, one_mul, ← Finsupp.single_add, ← Set.range_comp, + Function.comp] + rw [this, mem_ideal_span_monomial_image] + simp + +theorem X0_X1_X2_not_mem_kIdeal : (X 0 * X 1 * X 2 : MvPolynomial (Fin 3) (ZMod 2)) ∉ kIdeal := by + intro h + simp_rw [mem_kIdeal_iff, support_mul_X, support_X, Finset.map_singleton, addRightEmbedding_apply, + Finset.mem_singleton, forall_eq, ← Fin.sum_univ_three fun i => Finsupp.single i 1] at h + simp (config := {decide := true}) [← Finsupp.equivFunOnFinite_symm_eq_sum] at h + +theorem mul_self_mem_kIdeal_of_X0_X1_X2_mul_mem {x : MvPolynomial (Fin 3) (ZMod 2)} + (h : X 0 * X 1 * X 2 * x ∈ kIdeal) : x * x ∈ kIdeal := by + rw [mem_kIdeal_iff] at h + have : x ∈ Ideal.span ((X : Fin 3 → MvPolynomial _ (ZMod 2)) '' Set.univ) := by + rw [mem_ideal_span_X_image] + intro m hm + simp_rw [mul_assoc, support_X_mul, Finset.map_map, Finset.mem_map, + Function.Embedding.trans_apply, addLeftEmbedding_apply, forall_exists_index, + and_imp, forall_apply_eq_imp_iff₂, ← add_assoc, ← + Fin.sum_univ_three fun i => Finsupp.single i 1, ← Finsupp.equivFunOnFinite_symm_eq_sum, + Finsupp.add_apply, Finsupp.equivFunOnFinite_symm_apply_toFun] at h + refine (h _ hm).imp fun i hi => ⟨Set.mem_univ _, ?_⟩ + rintro hmi + rw [hmi] at hi + norm_num at hi + rw [as_sum x, CharTwo.sum_mul_self] + refine sum_mem fun m hm => ?_ + rw [mem_kIdeal_iff, monomial_mul] + intro m' hm' + obtain rfl := Finset.mem_singleton.1 (support_monomial_subset hm') + rw [mem_ideal_span_X_image] at this + obtain ⟨i, _, hi⟩ := this m hm + simp_rw [←one_add_one_eq_two] + refine ⟨i, Nat.add_le_add ?_ ?_⟩ <;> rwa [Nat.one_le_iff_ne_zero] + +/-- `𝔽₂[α, β, γ] / (α², β², γ²)` -/ +def K : Type _ := _ ⧸ kIdeal + +instance : CommRing K := Ideal.Quotient.commRing _ + +theorem comap_C_kIdeal : kIdeal.comap (C : ZMod 2 →+* MvPolynomial (Fin 3) (ZMod 2)) = ⊥ := by + refine bot_unique ?_ + refine (Ideal.comap_le_map_of_inverse _ _ _ (constantCoeff_C _)).trans ?_ + rw [kIdeal, Ideal.map_span] + refine (Ideal.span_le).2 ?_ + rintro x ⟨_, ⟨i, rfl⟩, rfl⟩ + rw [RingHom.map_mul, constantCoeff_X, mul_zero, Submodule.bot_coe, + Set.mem_singleton_iff] + +/-- `k` has characteristic 2. -/ +instance K.charP : CharP K 2 := by + dsimp only [K] + rw [CharP.quotient_iff_le_ker_natCast] + have : Nat.castRingHom (MvPolynomial (Fin 3) (ZMod 2)) = C.comp (Nat.castRingHom _) := by + ext1 r; rfl + rw [this, ← Ideal.comap_comap, ← RingHom.comap_ker, comap_C_kIdeal] + exact Ideal.comap_mono bot_le + +/-- The generators of `K`. -/ +def K.gen (i : Fin 3) : K := Ideal.Quotient.mk _ (MvPolynomial.X i) + +local notation "α" => K.gen 0 +local notation "β" => K.gen 1 +local notation "γ" => K.gen 2 + +/-- The elements above square to zero -/ +@[simp] +theorem X_sq (i : Fin 3) : K.gen i * K.gen i = (0 : K) := by + change Ideal.Quotient.mk _ _ = _ + rw [Ideal.Quotient.eq_zero_iff_mem] + exact Ideal.subset_span ⟨i, rfl⟩ + +/-- If an element multiplied by `αβγ` is zero then it squares to zero. -/ +theorem sq_zero_of_αβγ_mul {x : K} : α * β * γ * x = 0 → x * x = 0 := by + induction x using Quotient.inductionOn' + change Ideal.Quotient.mk _ _ = 0 → Ideal.Quotient.mk _ _ = 0 + rw [Ideal.Quotient.eq_zero_iff_mem, Ideal.Quotient.eq_zero_iff_mem] + exact mul_self_mem_kIdeal_of_X0_X1_X2_mul_mem + +/-- Though `αβγ` is not itself zero-/ +theorem αβγ_ne_zero : α * β * γ ≠ 0 := fun h => + X0_X1_X2_not_mem_kIdeal <| Ideal.Quotient.eq_zero_iff_mem.1 h + +-- A variant of lean4#2220 +local macro_rules | `($x • $y) => `(@HSMul.hSMul _ _ _ instHSMul $x $y) + +/-- The 1-form on $K^3$, the kernel of which we will take a quotient by. + +Our source uses $αx - βy - γz$, though since this is characteristic two we just use $αx + βy + γz$. + -/ +@[simps!] +def lFunc : (Fin 3 → K) →ₗ[K] K := + letI proj : Fin 3 → (Fin 3 → K) →ₗ[K] K := LinearMap.proj + α • proj 0 + β • proj 1 + γ • proj 2 + +/-- The quotient of `K^3` by the specified relation. -/ +abbrev L : Type _ := _ ⧸ LinearMap.ker lFunc + +/-- The quadratic form corresponding to squaring a single coefficient. -/ +def sq {ι R : Type*} [CommRing R] (i : ι) : QuadraticForm R (ι → R) := + QuadraticForm.sq.comp <| LinearMap.proj i + +theorem sq_map_add_char_two {ι R : Type*} [CommRing R] [CharP R 2] (i : ι) (a b : ι → R) : + sq i (a + b) = sq i a + sq i b := + CharTwo.add_mul_self _ _ + +theorem sq_map_sub_char_two {ι R : Type*} [CommRing R] [CharP R 2] (i : ι) (a b : ι → R) : + sq i (a - b) = sq i a - sq i b := by + haveI : Nonempty ι := ⟨i⟩ + rw [CharTwo.sub_eq_add, CharTwo.sub_eq_add, sq_map_add_char_two] + +open scoped BigOperators + +/-- The quadratic form (metric) is just euclidean -/ +def Q' : QuadraticForm K (Fin 3 → K) := + ∑ i, sq i + +theorem Q'_add (x y : Fin 3 → K) : Q' (x + y) = Q' x + Q' y := by + simp only [Q', QuadraticForm.sum_apply, sq_map_add_char_two, Finset.sum_add_distrib] + +theorem Q'_sub (x y : Fin 3 → K) : Q' (x - y) = Q' x - Q' y := by + simp only [Q', QuadraticForm.sum_apply, sq_map_sub_char_two, Finset.sum_sub_distrib] + +theorem Q'_apply (a : Fin 3 → K) : Q' a = a 0 * a 0 + a 1 * a 1 + a 2 * a 2 := + calc + Q' a = a 0 * a 0 + (a 1 * a 1 + (a 2 * a 2 + 0)) := rfl + _ = _ := by ring + +theorem Q'_apply_single (i : Fin 3) (x : K) : Q' (Pi.single i x) = x * x := + calc + Q' (Pi.single i x) = ∑ j : Fin 3, (Pi.single i x * Pi.single i x : Fin 3 → K) j := by + simp [Q', sq] + _ = _ := by simp_rw [← Pi.single_mul, Finset.sum_pi_single', Finset.mem_univ, if_pos] + +theorem Q'_zero_under_ideal (v : Fin 3 → K) (hv : v ∈ LinearMap.ker lFunc) : Q' v = 0 := by + rw [LinearMap.mem_ker, lFunc_apply] at hv + have h0 : α * β * γ * v 0 = 0 := by + have := congr_arg (β * γ * ·) hv + simp only [mul_zero, mul_add, ← mul_assoc] at this + rw [mul_comm (β * γ) α, ← mul_assoc, mul_right_comm β γ β, mul_assoc β γ γ, X_sq, X_sq] at this + simpa only [mul_zero, zero_mul, add_zero, zero_add] using this + have h1 : α * β * γ * v 1 = 0 := by + have := congr_arg (α * γ * ·) hv + simp only [mul_zero, mul_add, ← mul_assoc] at this + rw [mul_right_comm α γ α, mul_assoc α γ γ, mul_right_comm α γ β, X_sq, X_sq] at this + simpa only [mul_zero, zero_mul, add_zero, zero_add] using this + have h2 : α * β * γ * v 2 = 0 := by + have := congr_arg (α * β * ·) hv + simp only [mul_zero, mul_add, ← mul_assoc] at this + rw [mul_right_comm α β α, mul_assoc α β β, X_sq, X_sq] at this + simpa only [mul_zero, zero_mul, add_zero, zero_add] using this + rw [Q'_apply, sq_zero_of_αβγ_mul h0, sq_zero_of_αβγ_mul h1, sq_zero_of_αβγ_mul h2, add_zero, + add_zero] + +/-- `Q'`, lifted to operate on the quotient space `L`. -/ +@[simps!] +def Q : QuadraticForm K L := + QuadraticForm.ofPolar + (fun x => + Quotient.liftOn' x Q' fun a b h => by + rw [Submodule.quotientRel_r_def] at h + suffices Q' (a - b) = 0 by rwa [Q'_sub, sub_eq_zero] at this + apply Q'_zero_under_ideal (a - b) h) + (fun a x => by + induction x using Quotient.inductionOn + exact Q'.toFun_smul a _) + (by rintro ⟨x⟩ ⟨x'⟩ ⟨y⟩; exact Q'.polar_add_left x x' y) + (by rintro c ⟨x⟩ ⟨y⟩; exact Q'.polar_smul_left c x y) + +open CliffordAlgebra + +/-- Basis vectors in the Clifford algebra -/ +def gen (i : Fin 3) : CliffordAlgebra Q := ι Q <| Submodule.Quotient.mk (Pi.single i 1) + +local notation "x'" => gen 0 +local notation "y'" => gen 1 +local notation "z'" => gen 2 + +/-- The basis vectors square to one -/ +@[simp] +theorem gen_mul_gen (i) : gen i * gen i = 1 := by + dsimp only [gen] + simp_rw [CliffordAlgebra.ι_sq_scalar, Q_apply, ← Submodule.Quotient.mk''_eq_mk, + Quotient.liftOn'_mk'', Q'_apply_single, mul_one, map_one] + +/-- By virtue of the quotient, terms of this form are zero -/ +theorem quot_obv : α • x' - β • y' - γ • z' = 0 := by + dsimp only [gen] + simp_rw [← LinearMap.map_smul, ←LinearMap.map_sub, ← Submodule.Quotient.mk_smul _ (_ : K), + ← Submodule.Quotient.mk_sub] + convert LinearMap.map_zero _ using 2 + rw [Submodule.Quotient.mk_eq_zero] + simp (config := {decide := true}) [sub_zero, Ideal.span, Pi.single_apply] + +/-- The core of the proof - scaling `1` by `α * β * γ` gives zero -/ +theorem αβγ_smul_eq_zero : (α * β * γ) • (1 : CliffordAlgebra Q) = 0 := by + suffices α • 1 - β • (y' * x') - γ • (z' * x') = 0 by + have := congr_arg (fun x => (β * γ) • x) this + dsimp only at this + simp_rw [smul_sub, smul_smul] at this + rwa [mul_assoc β γ γ, mul_right_comm β γ β, mul_right_comm β γ α, mul_comm β α, X_sq, X_sq, + zero_mul, mul_zero, zero_smul, zero_smul, sub_zero, sub_zero, smul_zero] at this + have : (α • x' - β • y' - γ • z') * x' = α • 1 - β • (y' * x') - γ • (z' * x') := by + simp_rw [sub_mul, smul_mul_assoc, gen_mul_gen] + rw [← this] + rw [quot_obv, zero_mul] + +theorem algebraMap_αβγ_eq_zero : algebraMap K (CliffordAlgebra Q) (α * β * γ) = 0 := by + rw [Algebra.algebraMap_eq_smul_one, αβγ_smul_eq_zero] + +/-- Our final result: for the quadratic form `Q60596.Q`, the algebra map to the Clifford algebra +is not injective, as it sends the non-zero `α * β * γ` to zero. -/ +theorem algebraMap_not_injective : ¬Function.Injective (algebraMap K <| CliffordAlgebra Q) := + fun h => αβγ_ne_zero <| h <| by rw [algebraMap_αβγ_eq_zero, RingHom.map_zero] + +end Q60596 + +open Q60596 in +/-- The general statement: not every Clifford algebra over a module has an injective algebra map. -/ +theorem CliffordAlgebra.not_forall_algebraMap_injective.{v} : + -- TODO: make `R` universe polymorphic + ¬∀ (R : Type) (M : Type v) [CommRing R] [AddCommGroup M] [Module R M] (Q : QuadraticForm R M), + Function.Injective (algebraMap R <| CliffordAlgebra Q) := + fun h => algebraMap_not_injective fun x y hxy => by + let uU := ULift.moduleEquiv (R := K) (M := L) + let uQ := Q.comp uU.toLinearMap + let f : Q →qᵢ uQ := { uU.symm with map_app' := fun _ => rfl } + refine h K (ULift L) (Q.comp uU.toLinearMap) ?_ + let uC := CliffordAlgebra.map f + have := uC.congr_arg hxy + rwa [AlgHom.commutes, AlgHom.commutes] at this diff --git a/Counterexamples/Cyclotomic105.lean b/Counterexamples/Cyclotomic105.lean index f05b5d44f1eba..07a1cd6df40d5 100644 --- a/Counterexamples/Cyclotomic105.lean +++ b/Counterexamples/Cyclotomic105.lean @@ -109,7 +109,7 @@ theorem cyclotomic_105 : #align counterexample.cyclotomic_105 Counterexample.cyclotomic_105 theorem coeff_cyclotomic_105 : coeff (cyclotomic 105 ℤ) 7 = -2 := by - simp [cyclotomic_105, coeff_X_pow, coeff_one, coeff_X_of_ne_one, coeff_bit0_mul, two_mul] + simp [cyclotomic_105, coeff_one, coeff_X_of_ne_one] #align counterexample.coeff_cyclotomic_105 Counterexample.coeff_cyclotomic_105 theorem not_forall_coeff_cyclotomic_neg_one_zero_one : diff --git a/Counterexamples/HomogeneousPrimeNotPrime.lean b/Counterexamples/HomogeneousPrimeNotPrime.lean index dd776d4703975..d3d2dcdb4aec1 100644 --- a/Counterexamples/HomogeneousPrimeNotPrime.lean +++ b/Counterexamples/HomogeneousPrimeNotPrime.lean @@ -175,7 +175,7 @@ theorem homogeneous_mem_or_mem {x y : R × R} (hx : SetLike.Homogeneous (grading -- Porting note: added `h2` for later use; the proof is hideous have h2 : Prime (2:R) := by unfold Prime - simp only [true_and] + refine ⟨by decide, by decide, ?_⟩ intro a b have aux2 : (Fin.mk 2 _ : R) = 2 := rfl have aux3 : (Fin.mk 3 _ : R) = -1 := rfl diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index b16db995a3846..45437cd0ecd00 100644 --- a/Counterexamples/Phillips.lean +++ b/Counterexamples/Phillips.lean @@ -571,7 +571,7 @@ theorem countable_ne (Hcont : #ℝ = aleph 1) (φ : (DiscreteCopy ℝ →ᵇ ℝ {x | φ.toBoundedAdditiveMeasure.discreteSupport ∩ spf Hcont x ≠ ∅} := by intro x hx contrapose! hx - simp only [Classical.not_not, mem_setOf_eq] at hx + simp only [Classical.not_not, mem_setOf_eq, not_nonempty_iff_eq_empty] at hx simp [apply_f_eq_continuousPart Hcont φ x hx] have B : {x | φ.toBoundedAdditiveMeasure.discreteSupport ∩ spf Hcont x ≠ ∅} ⊆ diff --git a/Counterexamples/SorgenfreyLine.lean b/Counterexamples/SorgenfreyLine.lean index 76e3c6e772b24..5b1742059d159 100644 --- a/Counterexamples/SorgenfreyLine.lean +++ b/Counterexamples/SorgenfreyLine.lean @@ -6,7 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Topology.Instances.Irrational import Mathlib.Topology.Algebra.Order.Archimedean import Mathlib.Topology.Compactness.Paracompact -import Mathlib.Topology.MetricSpace.Metrizable +import Mathlib.Topology.Metrizable.Urysohn import Mathlib.Topology.EMetricSpace.Paracompact import Mathlib.Data.Set.Intervals.Monotone import Mathlib.Topology.Separation.NotNormal diff --git a/Mathlib.lean b/Mathlib.lean index f76102025e298..3be0ab21b3e2d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4,6 +4,7 @@ import Mathlib.Algebra.Algebra.Basic import Mathlib.Algebra.Algebra.Bilinear import Mathlib.Algebra.Algebra.Equiv import Mathlib.Algebra.Algebra.Hom +import Mathlib.Algebra.Algebra.NonUnitalHom import Mathlib.Algebra.Algebra.NonUnitalSubalgebra import Mathlib.Algebra.Algebra.Operations import Mathlib.Algebra.Algebra.Opposite @@ -159,15 +160,24 @@ import Mathlib.Algebra.GCDMonoid.Multiset import Mathlib.Algebra.GeomSum import Mathlib.Algebra.GradedMonoid import Mathlib.Algebra.GradedMulAction +import Mathlib.Algebra.Group.Aut import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.Group.Commutator import Mathlib.Algebra.Group.Commute.Basic import Mathlib.Algebra.Group.Commute.Defs +import Mathlib.Algebra.Group.Commute.Hom import Mathlib.Algebra.Group.Commute.Units import Mathlib.Algebra.Group.Conj import Mathlib.Algebra.Group.ConjFinite import Mathlib.Algebra.Group.Defs +import Mathlib.Algebra.Group.Embedding +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Algebra.Group.Equiv.TypeTags import Mathlib.Algebra.Group.Ext +import Mathlib.Algebra.Group.Freiman +import Mathlib.Algebra.Group.Hom.Basic +import Mathlib.Algebra.Group.Hom.Defs +import Mathlib.Algebra.Group.Hom.Instances import Mathlib.Algebra.Group.InjSurj import Mathlib.Algebra.Group.MinimalAxioms import Mathlib.Algebra.Group.Opposite @@ -181,11 +191,15 @@ import Mathlib.Algebra.Group.TypeTags import Mathlib.Algebra.Group.ULift import Mathlib.Algebra.Group.UniqueProds import Mathlib.Algebra.Group.Units +import Mathlib.Algebra.Group.Units.Equiv +import Mathlib.Algebra.Group.Units.Hom import Mathlib.Algebra.Group.WithOne.Basic import Mathlib.Algebra.Group.WithOne.Defs import Mathlib.Algebra.Group.WithOne.Units import Mathlib.Algebra.GroupPower.Basic +import Mathlib.Algebra.GroupPower.CovariantClass import Mathlib.Algebra.GroupPower.Identities +import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.GroupPower.Lemmas import Mathlib.Algebra.GroupPower.NegOnePow import Mathlib.Algebra.GroupPower.Order @@ -203,26 +217,9 @@ import Mathlib.Algebra.GroupWithZero.NeZero import Mathlib.Algebra.GroupWithZero.Power import Mathlib.Algebra.GroupWithZero.Semiconj import Mathlib.Algebra.GroupWithZero.Units.Basic +import Mathlib.Algebra.GroupWithZero.Units.Equiv import Mathlib.Algebra.GroupWithZero.Units.Lemmas import Mathlib.Algebra.HierarchyDesign -import Mathlib.Algebra.Hom.Aut -import Mathlib.Algebra.Hom.Centroid -import Mathlib.Algebra.Hom.Commute -import Mathlib.Algebra.Hom.Embedding -import Mathlib.Algebra.Hom.Equiv.Basic -import Mathlib.Algebra.Hom.Equiv.TypeTags -import Mathlib.Algebra.Hom.Equiv.Units.Basic -import Mathlib.Algebra.Hom.Equiv.Units.GroupWithZero -import Mathlib.Algebra.Hom.Freiman -import Mathlib.Algebra.Hom.Group.Basic -import Mathlib.Algebra.Hom.Group.Defs -import Mathlib.Algebra.Hom.GroupAction -import Mathlib.Algebra.Hom.GroupInstances -import Mathlib.Algebra.Hom.Iterate -import Mathlib.Algebra.Hom.NonUnitalAlg -import Mathlib.Algebra.Hom.Ring.Basic -import Mathlib.Algebra.Hom.Ring.Defs -import Mathlib.Algebra.Hom.Units import Mathlib.Algebra.Homology.Additive import Mathlib.Algebra.Homology.Augment import Mathlib.Algebra.Homology.ComplexShape @@ -241,8 +238,10 @@ import Mathlib.Algebra.Homology.LocalCohomology import Mathlib.Algebra.Homology.ModuleCat import Mathlib.Algebra.Homology.Opposite import Mathlib.Algebra.Homology.QuasiIso +import Mathlib.Algebra.Homology.ShortComplex.Ab import Mathlib.Algebra.Homology.ShortComplex.Abelian import Mathlib.Algebra.Homology.ShortComplex.Basic +import Mathlib.Algebra.Homology.ShortComplex.ConcreteCategory import Mathlib.Algebra.Homology.ShortComplex.Exact import Mathlib.Algebra.Homology.ShortComplex.FunctorEquivalence import Mathlib.Algebra.Homology.ShortComplex.HomologicalComplex @@ -253,6 +252,8 @@ import Mathlib.Algebra.Homology.ShortComplex.Preadditive import Mathlib.Algebra.Homology.ShortComplex.PreservesHomology import Mathlib.Algebra.Homology.ShortComplex.QuasiIso import Mathlib.Algebra.Homology.ShortComplex.RightHomology +import Mathlib.Algebra.Homology.ShortComplex.ShortExact +import Mathlib.Algebra.Homology.ShortComplex.SnakeLemma import Mathlib.Algebra.Homology.ShortExact.Abelian import Mathlib.Algebra.Homology.ShortExact.Preadditive import Mathlib.Algebra.Homology.Single @@ -311,6 +312,7 @@ import Mathlib.Algebra.Module.Projective import Mathlib.Algebra.Module.Submodule.Basic import Mathlib.Algebra.Module.Submodule.Bilinear import Mathlib.Algebra.Module.Submodule.Lattice +import Mathlib.Algebra.Module.Submodule.Map import Mathlib.Algebra.Module.Submodule.Pointwise import Mathlib.Algebra.Module.Torsion import Mathlib.Algebra.Module.ULift @@ -421,6 +423,7 @@ import Mathlib.Algebra.Ring.AddAut import Mathlib.Algebra.Ring.Aut import Mathlib.Algebra.Ring.Basic import Mathlib.Algebra.Ring.BooleanRing +import Mathlib.Algebra.Ring.CentroidHom import Mathlib.Algebra.Ring.Commute import Mathlib.Algebra.Ring.CompTypeclasses import Mathlib.Algebra.Ring.Defs @@ -428,6 +431,8 @@ import Mathlib.Algebra.Ring.Divisibility.Basic import Mathlib.Algebra.Ring.Divisibility.Lemmas import Mathlib.Algebra.Ring.Equiv import Mathlib.Algebra.Ring.Fin +import Mathlib.Algebra.Ring.Hom.Basic +import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Algebra.Ring.Idempotents import Mathlib.Algebra.Ring.InjSurj import Mathlib.Algebra.Ring.MinimalAxioms @@ -882,6 +887,7 @@ import Mathlib.CategoryTheory.Abelian.NonPreadditive import Mathlib.CategoryTheory.Abelian.Opposite import Mathlib.CategoryTheory.Abelian.Projective import Mathlib.CategoryTheory.Abelian.Pseudoelements +import Mathlib.CategoryTheory.Abelian.Refinements import Mathlib.CategoryTheory.Abelian.RightDerived import Mathlib.CategoryTheory.Abelian.Subobject import Mathlib.CategoryTheory.Abelian.Transfer @@ -942,6 +948,7 @@ import Mathlib.CategoryTheory.Closed.Zero import Mathlib.CategoryTheory.CofilteredSystem import Mathlib.CategoryTheory.CommSq import Mathlib.CategoryTheory.Comma +import Mathlib.CategoryTheory.ComposableArrows import Mathlib.CategoryTheory.ConcreteCategory.Basic import Mathlib.CategoryTheory.ConcreteCategory.Bundled import Mathlib.CategoryTheory.ConcreteCategory.BundledHom @@ -1029,6 +1036,7 @@ import Mathlib.CategoryTheory.Limits.Filtered import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesFiniteLimit import Mathlib.CategoryTheory.Limits.Final import Mathlib.CategoryTheory.Limits.FinallySmall +import Mathlib.CategoryTheory.Limits.FintypeCat import Mathlib.CategoryTheory.Limits.Fubini import Mathlib.CategoryTheory.Limits.FullSubcategory import Mathlib.CategoryTheory.Limits.FunctorCategory @@ -1088,12 +1096,14 @@ import Mathlib.CategoryTheory.Limits.Shapes.ZeroObjects import Mathlib.CategoryTheory.Limits.SmallComplete import Mathlib.CategoryTheory.Limits.Types import Mathlib.CategoryTheory.Limits.Unit +import Mathlib.CategoryTheory.Limits.VanKampen import Mathlib.CategoryTheory.Limits.Yoneda import Mathlib.CategoryTheory.Linear.Basic import Mathlib.CategoryTheory.Linear.FunctorCategory import Mathlib.CategoryTheory.Linear.LinearFunctor import Mathlib.CategoryTheory.Linear.Yoneda import Mathlib.CategoryTheory.Localization.Adjunction +import Mathlib.CategoryTheory.Localization.CalculusOfFractions import Mathlib.CategoryTheory.Localization.Composition import Mathlib.CategoryTheory.Localization.Construction import Mathlib.CategoryTheory.Localization.Equivalence @@ -1282,6 +1292,7 @@ import Mathlib.Combinatorics.Quiver.Subquiver import Mathlib.Combinatorics.Quiver.Symmetric import Mathlib.Combinatorics.SetFamily.Compression.Down import Mathlib.Combinatorics.SetFamily.Compression.UV +import Mathlib.Combinatorics.SetFamily.FourFunctions import Mathlib.Combinatorics.SetFamily.HarrisKleitman import Mathlib.Combinatorics.SetFamily.Intersecting import Mathlib.Combinatorics.SetFamily.Kleitman @@ -1427,6 +1438,7 @@ import Mathlib.Data.Finite.Card import Mathlib.Data.Finite.Defs import Mathlib.Data.Finite.Set import Mathlib.Data.Finmap +import Mathlib.Data.Finset.Antidiagonal import Mathlib.Data.Finset.Basic import Mathlib.Data.Finset.Card import Mathlib.Data.Finset.Fin @@ -1605,6 +1617,9 @@ import Mathlib.Data.Matrix.Notation import Mathlib.Data.Matrix.PEquiv import Mathlib.Data.Matrix.Rank import Mathlib.Data.Matrix.Reflection +import Mathlib.Data.Matrix.RowCol +import Mathlib.Data.Matroid.Basic +import Mathlib.Data.Matroid.Init import Mathlib.Data.Multiset.Antidiagonal import Mathlib.Data.Multiset.Basic import Mathlib.Data.Multiset.Bind @@ -1953,6 +1968,7 @@ import Mathlib.Dynamics.Minimal import Mathlib.Dynamics.OmegaLimit import Mathlib.Dynamics.PeriodicPts import Mathlib.FieldTheory.AbelRuffini +import Mathlib.FieldTheory.AbsoluteGaloisGroup import Mathlib.FieldTheory.Adjoin import Mathlib.FieldTheory.AxGrothendieck import Mathlib.FieldTheory.Cardinality @@ -2082,6 +2098,7 @@ import Mathlib.GroupTheory.GroupAction.DomAct.Basic import Mathlib.GroupTheory.GroupAction.Embedding import Mathlib.GroupTheory.GroupAction.FixingSubgroup import Mathlib.GroupTheory.GroupAction.Group +import Mathlib.GroupTheory.GroupAction.Hom import Mathlib.GroupTheory.GroupAction.Opposite import Mathlib.GroupTheory.GroupAction.Option import Mathlib.GroupTheory.GroupAction.Pi @@ -2103,6 +2120,7 @@ import Mathlib.GroupTheory.PGroup import Mathlib.GroupTheory.Perm.Basic import Mathlib.GroupTheory.Perm.Cycle.Basic import Mathlib.GroupTheory.Perm.Cycle.Concrete +import Mathlib.GroupTheory.Perm.Cycle.PossibleTypes import Mathlib.GroupTheory.Perm.Cycle.Type import Mathlib.GroupTheory.Perm.Fin import Mathlib.GroupTheory.Perm.List @@ -2112,6 +2130,7 @@ import Mathlib.GroupTheory.Perm.Subgroup import Mathlib.GroupTheory.Perm.Support import Mathlib.GroupTheory.Perm.ViaEmbedding import Mathlib.GroupTheory.PresentedGroup +import Mathlib.GroupTheory.PushoutI import Mathlib.GroupTheory.QuotientGroup import Mathlib.GroupTheory.Schreier import Mathlib.GroupTheory.SchurZassenhaus @@ -2198,6 +2217,7 @@ import Mathlib.Init.ZeroOne import Mathlib.Lean.CoreM import Mathlib.Lean.Data.NameMap import Mathlib.Lean.Elab.Tactic.Basic +import Mathlib.Lean.Elab.Term import Mathlib.Lean.EnvExtension import Mathlib.Lean.Exception import Mathlib.Lean.Expr @@ -2215,6 +2235,7 @@ import Mathlib.Lean.Meta.CongrTheorems import Mathlib.Lean.Meta.DiscrTree import Mathlib.Lean.Meta.Simp import Mathlib.Lean.Name +import Mathlib.Lean.PrettyPrinter.Delaborator import Mathlib.Lean.SMap import Mathlib.Lean.System.IO import Mathlib.Lean.Thunk @@ -2674,6 +2695,7 @@ import Mathlib.Order.Antisymmetrization import Mathlib.Order.Atoms import Mathlib.Order.Atoms.Finite import Mathlib.Order.Basic +import Mathlib.Order.Birkhoff import Mathlib.Order.BooleanAlgebra import Mathlib.Order.Booleanisation import Mathlib.Order.Bounded @@ -2796,6 +2818,7 @@ import Mathlib.Order.RelIso.Group import Mathlib.Order.RelIso.Set import Mathlib.Order.RelSeries import Mathlib.Order.SemiconjSup +import Mathlib.Order.Sublattice import Mathlib.Order.SuccPred.Basic import Mathlib.Order.SuccPred.IntervalSucc import Mathlib.Order.SuccPred.Limit @@ -2862,6 +2885,7 @@ import Mathlib.RepresentationTheory.Basic import Mathlib.RepresentationTheory.Character import Mathlib.RepresentationTheory.FdRep import Mathlib.RepresentationTheory.GroupCohomology.Basic +import Mathlib.RepresentationTheory.GroupCohomology.LowDegree import Mathlib.RepresentationTheory.GroupCohomology.Resolution import Mathlib.RepresentationTheory.Invariants import Mathlib.RepresentationTheory.Maschke @@ -3004,6 +3028,7 @@ import Mathlib.RingTheory.Polynomial.Vieta import Mathlib.RingTheory.PolynomialAlgebra import Mathlib.RingTheory.PowerBasis import Mathlib.RingTheory.PowerSeries.Basic +import Mathlib.RingTheory.PowerSeries.Derivative import Mathlib.RingTheory.PowerSeries.WellKnown import Mathlib.RingTheory.Prime import Mathlib.RingTheory.PrincipalIdealDomain @@ -3096,7 +3121,6 @@ import Mathlib.Tactic.Attr.Register import Mathlib.Tactic.Backtrack import Mathlib.Tactic.Basic import Mathlib.Tactic.ByContra -import Mathlib.Tactic.Cache import Mathlib.Tactic.CancelDenoms import Mathlib.Tactic.CancelDenoms.Core import Mathlib.Tactic.Cases @@ -3156,7 +3180,6 @@ import Mathlib.Tactic.InferParam import Mathlib.Tactic.Inhabit import Mathlib.Tactic.IntervalCases import Mathlib.Tactic.IrreducibleDef -import Mathlib.Tactic.LeftRight import Mathlib.Tactic.LibrarySearch import Mathlib.Tactic.Lift import Mathlib.Tactic.LiftLets @@ -3203,12 +3226,12 @@ import Mathlib.Tactic.NormNum.Result import Mathlib.Tactic.NthRewrite import Mathlib.Tactic.Observe import Mathlib.Tactic.PPWithUniv +import Mathlib.Tactic.Peel import Mathlib.Tactic.PermuteGoals import Mathlib.Tactic.Polyrith import Mathlib.Tactic.Positivity import Mathlib.Tactic.Positivity.Basic import Mathlib.Tactic.Positivity.Core -import Mathlib.Tactic.PrintPrefix import Mathlib.Tactic.ProjectionNotation import Mathlib.Tactic.Propose import Mathlib.Tactic.ProxyType @@ -3223,9 +3246,11 @@ import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.Rename import Mathlib.Tactic.RenameBVar import Mathlib.Tactic.Replace +import Mathlib.Tactic.RewriteSearch import Mathlib.Tactic.Rewrites import Mathlib.Tactic.Ring import Mathlib.Tactic.Ring.Basic +import Mathlib.Tactic.Ring.PNat import Mathlib.Tactic.Ring.RingNF import Mathlib.Tactic.RunCmd import Mathlib.Tactic.Sat.FromLRAT @@ -3284,6 +3309,7 @@ import Mathlib.Topology.Algebra.FilterBasis import Mathlib.Topology.Algebra.Group.Basic import Mathlib.Topology.Algebra.Group.Compact import Mathlib.Topology.Algebra.Group.OpenMapping +import Mathlib.Topology.Algebra.Group.TopologicalAbelianization import Mathlib.Topology.Algebra.GroupCompletion import Mathlib.Topology.Algebra.GroupWithZero import Mathlib.Topology.Algebra.InfiniteSum.Basic @@ -3362,6 +3388,7 @@ import Mathlib.Topology.Category.Profinite.Basic import Mathlib.Topology.Category.Profinite.CofilteredLimit import Mathlib.Topology.Category.Profinite.EffectiveEpi import Mathlib.Topology.Category.Profinite.Limits +import Mathlib.Topology.Category.Profinite.Nobeling import Mathlib.Topology.Category.Profinite.Product import Mathlib.Topology.Category.Profinite.Projective import Mathlib.Topology.Category.Stonean.Adjunctions @@ -3387,6 +3414,7 @@ import Mathlib.Topology.Compactness.Compact import Mathlib.Topology.Compactness.LocallyCompact import Mathlib.Topology.Compactness.Paracompact import Mathlib.Topology.Compactness.SigmaCompact +import Mathlib.Topology.CompletelyRegular import Mathlib.Topology.Connected.Basic import Mathlib.Topology.Connected.LocallyConnected import Mathlib.Topology.Connected.PathConnected @@ -3464,8 +3492,10 @@ import Mathlib.Topology.MetricSpace.Algebra import Mathlib.Topology.MetricSpace.Antilipschitz import Mathlib.Topology.MetricSpace.Baire import Mathlib.Topology.MetricSpace.Basic +import Mathlib.Topology.MetricSpace.Bounded import Mathlib.Topology.MetricSpace.CantorScheme import Mathlib.Topology.MetricSpace.CauSeqFilter +import Mathlib.Topology.MetricSpace.Cauchy import Mathlib.Topology.MetricSpace.Closeds import Mathlib.Topology.MetricSpace.Completion import Mathlib.Topology.MetricSpace.Contracting @@ -3484,13 +3514,15 @@ import Mathlib.Topology.MetricSpace.Isometry import Mathlib.Topology.MetricSpace.Kuratowski import Mathlib.Topology.MetricSpace.Lipschitz import Mathlib.Topology.MetricSpace.MetricSeparated -import Mathlib.Topology.MetricSpace.Metrizable -import Mathlib.Topology.MetricSpace.MetrizableUniformity import Mathlib.Topology.MetricSpace.PartitionOfUnity import Mathlib.Topology.MetricSpace.PiNat import Mathlib.Topology.MetricSpace.Polish +import Mathlib.Topology.MetricSpace.PseudoMetric import Mathlib.Topology.MetricSpace.ShrinkingLemma import Mathlib.Topology.MetricSpace.ThickenedIndicator +import Mathlib.Topology.Metrizable.Basic +import Mathlib.Topology.Metrizable.Uniformity +import Mathlib.Topology.Metrizable.Urysohn import Mathlib.Topology.NhdsSet import Mathlib.Topology.NoetherianSpace import Mathlib.Topology.OmegaCompletePartialOrder @@ -3511,6 +3543,7 @@ import Mathlib.Topology.Perfect import Mathlib.Topology.ProperMap import Mathlib.Topology.QuasiSeparated import Mathlib.Topology.Semicontinuous +import Mathlib.Topology.SeparatedMap import Mathlib.Topology.Separation import Mathlib.Topology.Separation.NotNormal import Mathlib.Topology.Sequences @@ -3576,13 +3609,13 @@ import Mathlib.Util.AssertNoSorry import Mathlib.Util.AtomM import Mathlib.Util.CompileInductive import Mathlib.Util.CountHeartbeats +import Mathlib.Util.Delaborators import Mathlib.Util.DischargerAsTactic import Mathlib.Util.Export import Mathlib.Util.Imports import Mathlib.Util.IncludeStr import Mathlib.Util.LongNames import Mathlib.Util.MemoFix -import Mathlib.Util.PiNotation import Mathlib.Util.Qq import Mathlib.Util.SleepHeartbeats import Mathlib.Util.Superscript diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 60991a97600f3..1bbd5826b940b 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -244,10 +244,8 @@ section FieldDivisionRing variable (R A : Type*) [Field R] [DivisionRing A] [Algebra R A] --- porting note: todo: drop implicit args @[norm_cast] -theorem coe_ratCast (q : ℚ) : ↑(q : R) = (q : A) := - @map_ratCast (R →+* A) R A _ _ _ (algebraMap R A) q +theorem coe_ratCast (q : ℚ) : ↑(q : R) = (q : A) := map_ratCast (algebraMap R A) q #align algebra_map.coe_rat_cast algebraMap.coe_ratCast end FieldDivisionRing @@ -363,6 +361,12 @@ theorem commutes (r : R) (x : A) : algebraMap R A r * x = x * algebraMap R A r : Algebra.commutes' r x #align algebra.commutes Algebra.commutes +lemma commute_algebraMap_left (r : R) (x : A) : Commute (algebraMap R A r) x := + Algebra.commutes r x + +lemma commute_algebraMap_right (r : R) (x : A) : Commute x (algebraMap R A r) := + (Algebra.commutes r x).symm + /-- `mul_left_comm` for `Algebra`s when one element is from the base ring. -/ theorem left_comm (x : A) (r : R) (y : A) : x * (algebraMap R A r * y) = algebraMap R A r * (x * y) := by diff --git a/Mathlib/Algebra/Algebra/Bilinear.lean b/Mathlib/Algebra/Algebra/Bilinear.lean index 9cd9195bc499a..ae1d76a38ef79 100644 --- a/Mathlib/Algebra/Algebra/Bilinear.lean +++ b/Mathlib/Algebra/Algebra/Bilinear.lean @@ -5,8 +5,8 @@ Authors: Kenny Lau, Yury Kudryashov -/ import Mathlib.Algebra.Algebra.Basic import Mathlib.Algebra.Algebra.Equiv -import Mathlib.Algebra.Hom.Iterate -import Mathlib.Algebra.Hom.NonUnitalAlg +import Mathlib.Algebra.Algebra.NonUnitalHom +import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.LinearAlgebra.TensorProduct #align_import algebra.algebra.bilinear from "leanprover-community/mathlib"@"657df4339ae6ceada048c8a2980fb10e393143ec" diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index f9ac03944b436..2e2cf4b586fea 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -86,7 +86,7 @@ end AlgEquivClass namespace AlgEquiv -universe uR uA₁ uA₂ uA₃ uA₁' uA₂' uA₃' +universe uR uA₁ uA₂ uA₃ uA₁' uA₂' uA₃' variable {R : Type uR} variable {A₁ : Type uA₁} {A₂ : Type uA₂} {A₃ : Type uA₃} variable {A₁' : Type uA₁'} {A₂' : Type uA₂'} {A₃' : Type uA₃'} diff --git a/Mathlib/Algebra/Hom/NonUnitalAlg.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean similarity index 100% rename from Mathlib/Algebra/Hom/NonUnitalAlg.lean rename to Mathlib/Algebra/Algebra/NonUnitalHom.lean diff --git a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean index 40487af893ac6..963d3e7452734 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean @@ -3,10 +3,10 @@ Copyright (c) 2023 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.RingTheory.NonUnitalSubring.Basic -import Mathlib.Algebra.Hom.NonUnitalAlg +import Mathlib.Algebra.Algebra.NonUnitalHom import Mathlib.Data.Set.UnionLift import Mathlib.LinearAlgebra.Finsupp +import Mathlib.RingTheory.NonUnitalSubring.Basic /-! # Non-unital Subalgebras over Commutative Semirings @@ -56,9 +56,11 @@ add_decl_doc NonUnitalSubalgebra.toSubmodule namespace NonUnitalSubalgebra variable {F : Type v'} {R' : Type u'} {R : Type u} {A : Type v} {B : Type w} {C : Type w'} + +section NonUnitalNonAssocSemiring variable [CommSemiring R] -variable [NonUnitalNonAssocSemiring A] [Module R A] [NonUnitalNonAssocSemiring B] [Module R B] -variable [NonUnitalNonAssocSemiring C] [Module R C] [NonUnitalAlgHomClass F R A B] +variable [NonUnitalNonAssocSemiring A] [NonUnitalNonAssocSemiring B] [NonUnitalNonAssocSemiring C] +variable [Module R A] [Module R B] [Module R C] instance : SetLike (NonUnitalSubalgebra R A) A where @@ -74,11 +76,6 @@ instance instNonUnitalSubsemiringClass : NonUnitalSubsemiringClass (NonUnitalSub instance instSMulMemClass : SMulMemClass (NonUnitalSubalgebra R A) R A where smul_mem := @fun s => s.smul_mem' -instance instNonUnitalSubringClass {R : Type u} {A : Type v} [CommRing R] [NonUnitalRing A] - [Module R A] : NonUnitalSubringClass (NonUnitalSubalgebra R A) A := - { NonUnitalSubalgebra.instNonUnitalSubsemiringClass with - neg_mem := @fun _ x hx => neg_one_smul R x ▸ SMulMemClass.smul_mem _ hx } - theorem mem_carrier {s : NonUnitalSubalgebra R A} {x : A} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl @@ -138,35 +135,44 @@ theorem coe_copy (S : NonUnitalSubalgebra R A) (s : Set A) (hs : s = ↑S) : theorem copy_eq (S : NonUnitalSubalgebra R A) (s : Set A) (hs : s = ↑S) : S.copy s hs = S := SetLike.coe_injective hs -variable (S : NonUnitalSubalgebra R A) +instance (S : NonUnitalSubalgebra R A) : Inhabited S := + ⟨(0 : S.toNonUnitalSubsemiring)⟩ + +end NonUnitalNonAssocSemiring + +section NonUnitalNonAssocRing +variable [CommRing R] +variable [NonUnitalNonAssocRing A] [NonUnitalNonAssocRing B] [NonUnitalNonAssocRing C] +variable [Module R A] [Module R B] [Module R C] + +instance instNonUnitalSubringClass : NonUnitalSubringClass (NonUnitalSubalgebra R A) A := + { NonUnitalSubalgebra.instNonUnitalSubsemiringClass with + neg_mem := @fun _ x hx => neg_one_smul R x ▸ SMulMemClass.smul_mem _ hx } /-- A non-unital subalgebra over a ring is also a `Subring`. -/ -def toNonUnitalSubring {R : Type u} {A : Type v} [CommRing R] [NonUnitalRing A] [Module R A] - (S : NonUnitalSubalgebra R A) : NonUnitalSubring A where +def toNonUnitalSubring (S : NonUnitalSubalgebra R A) : NonUnitalSubring A where toNonUnitalSubsemiring := S.toNonUnitalSubsemiring neg_mem' := neg_mem (s := S) @[simp] -theorem mem_toNonUnitalSubring {R : Type u} {A : Type v} [CommRing R] [NonUnitalRing A] [Module R A] - {S : NonUnitalSubalgebra R A} {x} : x ∈ S.toNonUnitalSubring ↔ x ∈ S := +theorem mem_toNonUnitalSubring {S : NonUnitalSubalgebra R A} {x} : + x ∈ S.toNonUnitalSubring ↔ x ∈ S := Iff.rfl @[simp] -theorem coe_toNonUnitalSubring {R : Type u} {A : Type v} [CommRing R] [NonUnitalRing A] [Module R A] - (S : NonUnitalSubalgebra R A) : (↑S.toNonUnitalSubring : Set A) = S := +theorem coe_toNonUnitalSubring (S : NonUnitalSubalgebra R A) : + (↑S.toNonUnitalSubring : Set A) = S := rfl -theorem toNonUnitalSubring_injective {R : Type u} {A : Type v} [CommRing R] [NonUnitalRing A] - [Module R A] : +theorem toNonUnitalSubring_injective : Function.Injective (toNonUnitalSubring : NonUnitalSubalgebra R A → NonUnitalSubring A) := fun S T h => ext fun x => by rw [← mem_toNonUnitalSubring, ← mem_toNonUnitalSubring, h] -theorem toNonUnitalSubring_inj {R : Type u} {A : Type v} [CommRing R] [NonUnitalRing A] [Module R A] - {S U : NonUnitalSubalgebra R A} : S.toNonUnitalSubring = U.toNonUnitalSubring ↔ S = U := +theorem toNonUnitalSubring_inj {S U : NonUnitalSubalgebra R A} : + S.toNonUnitalSubring = U.toNonUnitalSubring ↔ S = U := toNonUnitalSubring_injective.eq_iff -instance : Inhabited S := - ⟨(0 : S.toNonUnitalSubsemiring)⟩ +end NonUnitalNonAssocRing section @@ -174,27 +180,35 @@ section coercions. -/ -instance toNonUnitalSemiring {R A} [CommSemiring R] [NonUnitalSemiring A] [Module R A] +instance toNonUnitalNonAssocSemiring [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] + (S : NonUnitalSubalgebra R A) : NonUnitalNonAssocSemiring S := + inferInstance + +instance toNonUnitalSemiring [CommSemiring R] [NonUnitalSemiring A] [Module R A] (S : NonUnitalSubalgebra R A) : NonUnitalSemiring S := inferInstance -instance toNonUnitalCommSemiring {R A} [CommSemiring R] [NonUnitalCommSemiring A] [Module R A] +instance toNonUnitalCommSemiring [CommSemiring R] [NonUnitalCommSemiring A] [Module R A] (S : NonUnitalSubalgebra R A) : NonUnitalCommSemiring S := inferInstance -instance toNonUnitalRing {R A} [CommRing R] [NonUnitalRing A] [Module R A] +instance toNonUnitalNonAssocRing [CommRing R] [NonUnitalNonAssocRing A] [Module R A] + (S : NonUnitalSubalgebra R A) : NonUnitalNonAssocRing S := + inferInstance + +instance toNonUnitalRing [CommRing R] [NonUnitalRing A] [Module R A] (S : NonUnitalSubalgebra R A) : NonUnitalRing S := inferInstance -instance toNonUnitalCommRing {R A} [CommRing R] [NonUnitalCommRing A] [Module R A] +instance toNonUnitalCommRing [CommRing R] [NonUnitalCommRing A] [Module R A] (S : NonUnitalSubalgebra R A) : NonUnitalCommRing S := inferInstance end /-- The forgetful map from `NonUnitalSubalgebra` to `Submodule` as an `OrderEmbedding` -/ -def toSubmodule' : NonUnitalSubalgebra R A ↪o Submodule R A - where +def toSubmodule' [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] : + NonUnitalSubalgebra R A ↪o Submodule R A where toEmbedding := { toFun := fun S => S.toSubmodule inj' := fun S T h => ext <| by apply SetLike.ext_iff.1 h } @@ -202,8 +216,8 @@ def toSubmodule' : NonUnitalSubalgebra R A ↪o Submodule R A /-- The forgetful map from `NonUnitalSubalgebra` to `NonUnitalSubsemiring` as an `OrderEmbedding` -/ -def toNonUnitalSubsemiring' : NonUnitalSubalgebra R A ↪o NonUnitalSubsemiring A - where +def toNonUnitalSubsemiring' [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] : + NonUnitalSubalgebra R A ↪o NonUnitalSubsemiring A where toEmbedding := { toFun := fun S => S.toNonUnitalSubsemiring inj' := fun S T h => ext <| by apply SetLike.ext_iff.1 h } @@ -211,14 +225,18 @@ def toNonUnitalSubsemiring' : NonUnitalSubalgebra R A ↪o NonUnitalSubsemiring /-- The forgetful map from `NonUnitalSubalgebra` to `NonUnitalSubsemiring` as an `OrderEmbedding` -/ -def toNonUnitalSubring' {R : Type u} {A : Type v} [CommRing R] [NonUnitalRing A] [Module R A] : - NonUnitalSubalgebra R A ↪o NonUnitalSubring A - where +def toNonUnitalSubring' [CommRing R] [NonUnitalNonAssocRing A] [Module R A] : + NonUnitalSubalgebra R A ↪o NonUnitalSubring A where toEmbedding := { toFun := fun S => S.toNonUnitalSubring inj' := fun S T h => ext <| by apply SetLike.ext_iff.1 h } map_rel_iff' := SetLike.coe_subset_coe.symm.trans SetLike.coe_subset_coe +variable [CommSemiring R] +variable [NonUnitalNonAssocSemiring A] [NonUnitalNonAssocSemiring B] [NonUnitalNonAssocSemiring C] +variable [Module R A] [Module R B] [Module R C] +variable {S : NonUnitalSubalgebra R A} + section /-! ### `NonUnitalSubalgebra`s inherit structure from their `Submodule` coercions. -/ @@ -243,13 +261,13 @@ instance instSMulCommClass' [Semiring R'] [SMul R' R] [Module R' A] [IsScalarTow instance instSMulCommClass [SMulCommClass R A A] : SMulCommClass R S S where smul_comm r x y := Subtype.ext <| smul_comm r (x : A) (y : A) -end - instance noZeroSMulDivisors_bot [NoZeroSMulDivisors R A] : NoZeroSMulDivisors R S := ⟨fun {c x} h => have : c = 0 ∨ (x : A) = 0 := eq_zero_or_eq_zero_of_smul_eq_zero (congr_arg ((↑) : S → A) h) this.imp_right (@Subtype.ext_iff _ _ x 0).mpr⟩ +end + protected theorem coe_add (x y : S) : (↑(x + y) : A) = ↑x + ↑y := rfl @@ -291,6 +309,8 @@ we define it as a `LinearEquiv` to avoid type equalities. -/ def toSubmoduleEquiv (S : NonUnitalSubalgebra R A) : S.toSubmodule ≃ₗ[R] S := LinearEquiv.ofEq _ _ rfl +variable [NonUnitalAlgHomClass F R A B] + /-- Transport a non-unital subalgebra via an algebra homomorphism. -/ def map (f : F) (S : NonUnitalSubalgebra R A) : NonUnitalSubalgebra R B := { S.toNonUnitalSubsemiring.map (f : A →ₙ+* B) with @@ -617,7 +637,7 @@ theorem top_toNonUnitalSubsemiring : (⊤ : NonUnitalSubalgebra R A).toNonUnital rfl @[simp] -theorem top_toSubring {R A : Type*} [CommRing R] [NonUnitalRing A] [Module R A] +theorem top_toSubring {R A : Type*} [CommRing R] [NonUnitalNonAssocRing A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] : (⊤ : NonUnitalSubalgebra R A).toNonUnitalSubring = ⊤ := rfl @@ -1027,7 +1047,7 @@ end NonUnitalSubalgebra section Nat -variable {R : Type*} [NonUnitalSemiring R] +variable {R : Type*} [NonUnitalNonAssocSemiring R] /-- A non-unital subsemiring is a non-unital `ℕ`-subalgebra. -/ def nonUnitalSubalgebraOfNonUnitalSubsemiring (S : NonUnitalSubsemiring R) : @@ -1044,9 +1064,9 @@ end Nat section Int -variable {R : Type*} [NonUnitalRing R] +variable {R : Type*} [NonUnitalNonAssocRing R] -/-- A non-unital subring is a non-unitalsubalgebra. -/ +/-- A non-unital subring is a non-unital `ℤ`-subalgebra. -/ def nonUnitalSubalgebraOfNonUnitalSubring (S : NonUnitalSubring R) : NonUnitalSubalgebra ℤ R where toNonUnitalSubsemiring := S.toNonUnitalSubsemiring smul_mem' n _x hx := zsmul_mem (K := S) hx n diff --git a/Mathlib/Algebra/Algebra/Opposite.lean b/Mathlib/Algebra/Algebra/Opposite.lean index 0f649d29c886c..ac3f660d71bbb 100644 --- a/Mathlib/Algebra/Algebra/Opposite.lean +++ b/Mathlib/Algebra/Algebra/Opposite.lean @@ -153,7 +153,7 @@ theorem toRingEquiv_op (f : A ≃ₐ[R] B) : (AlgEquiv.op f).toRingEquiv = RingEquiv.op f.toRingEquiv := rfl -/-- The 'unopposite' of an algebra iso `Aᵐᵒᵖ ≃ₐ[R] Bᵐᵒᵖ`. Inverse to `AlgEquiv.op`. -/ +/-- The 'unopposite' of an algebra iso `Aᵐᵒᵖ ≃ₐ[R] Bᵐᵒᵖ`. Inverse to `AlgEquiv.op`. -/ abbrev unop : (Aᵐᵒᵖ ≃ₐ[R] Bᵐᵒᵖ) ≃ A ≃ₐ[R] B := AlgEquiv.op.symm theorem toAlgHom_unop (f : Aᵐᵒᵖ ≃ₐ[R] Bᵐᵒᵖ) : f.unop.toAlgHom = AlgHom.unop f.toAlgHom := diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index 369fe38f150e0..ac628ca71d8cc 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -352,6 +352,9 @@ theorem mem_toSubmodule {x} : x ∈ (toSubmodule S) ↔ x ∈ S := Iff.rfl theorem coe_toSubmodule (S : Subalgebra R A) : (toSubmodule S : Set A) = S := rfl #align subalgebra.coe_to_submodule Subalgebra.coe_toSubmodule +theorem toSubmodule_injective : Function.Injective (toSubmodule : Subalgebra R A → Submodule R A) := + fun _S₁ _S₂ h => SetLike.ext (SetLike.ext_iff.mp h :) + section /-! `Subalgebra`s inherit structure from their `Submodule` coercions. -/ @@ -743,16 +746,13 @@ noncomputable def ofInjectiveField {E F : Type*} [DivisionRing E] [Semiring F] [ /-- Given an equivalence `e : A ≃ₐ[R] B` of `R`-algebras and a subalgebra `S` of `A`, `subalgebra_map` is the induced equivalence between `S` and `S.map e` -/ @[simps!] -def subalgebraMap (e : A ≃ₐ[R] B) (S : Subalgebra R A) : S ≃ₐ[R] S.map e.toAlgHom := +def subalgebraMap (e : A ≃ₐ[R] B) (S : Subalgebra R A) : S ≃ₐ[R] S.map (e : A →ₐ[R] B) := { e.toRingEquiv.subsemiringMap S.toSubsemiring with commutes' := fun r => by ext; dsimp only; erw [RingEquiv.subsemiringMap_apply_coe] exact e.commutes _ } #align alg_equiv.subalgebra_map AlgEquiv.subalgebraMap --- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing -attribute [nolint simpNF] AlgEquiv.subalgebraMap_apply_coe AlgEquiv.subalgebraMap_symm_apply_coe - end AlgEquiv namespace Algebra @@ -783,8 +783,10 @@ protected def gi : GaloisInsertion (adjoin R : Set A → Subalgebra R A) (↑) w choice_eq _ _ := Subalgebra.copy_eq _ _ _ #align algebra.gi Algebra.gi -instance : CompleteLattice (Subalgebra R A) := - GaloisInsertion.liftCompleteLattice Algebra.gi +instance : CompleteLattice (Subalgebra R A) where + __ := GaloisInsertion.liftCompleteLattice Algebra.gi + bot := (Algebra.ofId R A).range + bot_le _S := fun _a ⟨_r, hr⟩ => hr ▸ algebraMap_mem _ _ @[simp] theorem coe_top : (↑(⊤ : Subalgebra R A) : Set A) = Set.univ := rfl @@ -898,21 +900,14 @@ theorem iInf_toSubmodule {ι : Sort*} (S : ι → Subalgebra R A) : instance : Inhabited (Subalgebra R A) := ⟨⊥⟩ -theorem mem_bot {x : A} : x ∈ (⊥ : Subalgebra R A) ↔ x ∈ Set.range (algebraMap R A) := - suffices (ofId R A).range = (⊥ : Subalgebra R A) by - rw [← this, ← SetLike.mem_coe, AlgHom.coe_range] - rfl - le_bot_iff.mp fun x hx => Subalgebra.range_le _ ((ofId R A).coe_range ▸ hx) +theorem mem_bot {x : A} : x ∈ (⊥ : Subalgebra R A) ↔ x ∈ Set.range (algebraMap R A) := Iff.rfl #align algebra.mem_bot Algebra.mem_bot -theorem toSubmodule_bot : Subalgebra.toSubmodule (⊥ : Subalgebra R A) = R ∙ 1 := by - ext x - simp [mem_bot, Submodule.mem_span_singleton, Algebra.smul_def] +theorem toSubmodule_bot : Subalgebra.toSubmodule (⊥ : Subalgebra R A) = 1 := rfl #align algebra.to_submodule_bot Algebra.toSubmodule_bot @[simp] -theorem coe_bot : ((⊥ : Subalgebra R A) : Set A) = Set.range (algebraMap R A) := by - simp [Set.ext_iff, Algebra.mem_bot] +theorem coe_bot : ((⊥ : Subalgebra R A) : Set A) = Set.range (algebraMap R A) := rfl #align algebra.coe_bot Algebra.coe_bot theorem eq_top_iff {S : Subalgebra R A} : S = ⊤ ↔ ∀ x : A, x ∈ S := @@ -937,8 +932,7 @@ theorem map_top (f : A →ₐ[R] B) : (⊤ : Subalgebra R A).map f = f.range := @[simp] theorem map_bot (f : A →ₐ[R] B) : (⊥ : Subalgebra R A).map f = ⊥ := - SetLike.coe_injective <| by - simp only [← Set.range_comp, (· ∘ ·), Algebra.coe_bot, Subalgebra.coe_map, f.commutes] + Subalgebra.toSubmodule_injective <| Submodule.map_one _ #align algebra.map_bot Algebra.map_bot @[simp] @@ -971,9 +965,7 @@ noncomputable def botEquivOfInjective (h : Function.Injective (algebraMap R A)) (⊥ : Subalgebra R A) ≃ₐ[R] R := AlgEquiv.symm <| AlgEquiv.ofBijective (Algebra.ofId R _) - ⟨fun _x _y hxy => h (congr_arg Subtype.val hxy : _), fun ⟨_y, hy⟩ => - let ⟨x, hx⟩ := Algebra.mem_bot.1 hy - ⟨x, Subtype.eq hx⟩⟩ + ⟨fun _x _y hxy => h (congr_arg Subtype.val hxy : _), fun ⟨_y, x, hx⟩ => ⟨x, Subtype.eq hx⟩⟩ #align algebra.bot_equiv_of_injective Algebra.botEquivOfInjective /-- The bottom subalgebra is isomorphic to the field. -/ diff --git a/Mathlib/Algebra/Algebra/Unitization.lean b/Mathlib/Algebra/Algebra/Unitization.lean index 7fe7784803bcd..f36570339b4ac 100644 --- a/Mathlib/Algebra/Algebra/Unitization.lean +++ b/Mathlib/Algebra/Algebra/Unitization.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Algebra.Algebra.Basic -import Mathlib.LinearAlgebra.Prod -import Mathlib.Algebra.Hom.NonUnitalAlg -import Mathlib.Algebra.Star.StarAlgHom +import Mathlib.Algebra.Algebra.NonUnitalHom import Mathlib.Algebra.Star.Module +import Mathlib.Algebra.Star.StarAlgHom +import Mathlib.LinearAlgebra.Prod #align_import algebra.algebra.unitization from "leanprover-community/mathlib"@"8f66240cab125b938b327d3850169d490cfbcdd8" diff --git a/Mathlib/Algebra/BigOperators/Basic.lean b/Mathlib/Algebra/BigOperators/Basic.lean index 2632d094c0976..753ff78903ad4 100644 --- a/Mathlib/Algebra/BigOperators/Basic.lean +++ b/Mathlib/Algebra/BigOperators/Basic.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Algebra.BigOperators.Multiset.Lemmas +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.Pi import Mathlib.Algebra.GroupPower.Lemmas -import Mathlib.Algebra.Hom.Equiv.Basic import Mathlib.Algebra.Ring.Opposite import Mathlib.Data.Finset.Powerset import Mathlib.Data.Finset.Sigma @@ -278,9 +278,11 @@ theorem MonoidHom.coe_finset_prod [MulOneClass β] [CommMonoid γ] (f : α → #align monoid_hom.coe_finset_prod MonoidHom.coe_finset_prod #align add_monoid_hom.coe_finset_sum AddMonoidHom.coe_finset_sum --- See also `Finset.prod_apply`, with the same conclusion --- but with the weaker hypothesis `f : α → β → γ`. -@[to_additive (attr := simp)] +/-- See also `Finset.prod_apply`, with the same conclusion but with the weaker hypothesis +`f : α → β → γ` -/ +@[to_additive (attr := simp) + "See also `Finset.sum_apply`, with the same conclusion but with the weaker hypothesis + `f : α → β → γ`"] theorem MonoidHom.finset_prod_apply [MulOneClass β] [CommMonoid γ] (f : α → β →* γ) (s : Finset α) (b : β) : (∏ x in s, f x) b = ∏ x in s, f x b := (MonoidHom.eval b).map_prod _ _ @@ -340,6 +342,10 @@ theorem prod_insert_one [DecidableEq α] (h : f a = 1) : ∏ x in insert a s, f #align finset.prod_insert_one Finset.prod_insert_one #align finset.sum_insert_zero Finset.sum_insert_zero +@[to_additive] +theorem prod_insert_div {M : Type*} [CommGroup M] [DecidableEq α] (ha : a ∉ s) {f : α → M} : + (∏ x in insert a s, f x) / f a = ∏ x in s, f x := by simp [ha] + @[to_additive (attr := simp)] theorem prod_singleton (f : α → β) (a : α) : ∏ x in singleton a, f x = f a := Eq.trans fold_singleton <| mul_one _ @@ -683,7 +689,7 @@ theorem prod_product_right {s : Finset γ} {t : Finset α} {f : γ × α → β} #align finset.sum_product_right Finset.sum_product_right /-- An uncurried version of `Finset.prod_product_right`. -/ -@[to_additive "An uncurried version of `Finset.prod_product_right`"] +@[to_additive "An uncurried version of `Finset.sum_product_right`"] theorem prod_product_right' {s : Finset γ} {t : Finset α} {f : γ → α → β} : ∏ x in s ×ˢ t, f x.1 x.2 = ∏ y in t, ∏ x in s, f x y := prod_product_right @@ -2123,7 +2129,7 @@ theorem finset_sum_eq_sup_iff_disjoint {β : Type*} {i : Finset β} {f : β → · simp only [Finset.not_mem_empty, IsEmpty.forall_iff, imp_true_iff, Finset.sum_empty, Finset.sup_empty, bot_eq_zero, eq_self_iff_true] · simp_rw [Finset.sum_cons hz, Finset.sup_cons, Finset.mem_cons, Multiset.sup_eq_union, - forall_eq_or_imp, Ne.def, IsEmpty.forall_iff, true_and_iff, + forall_eq_or_imp, Ne.def, not_true_eq_false, IsEmpty.forall_iff, true_and_iff, imp_and, forall_and, ← hr, @eq_comm _ z] have := fun x (H : x ∈ i) => ne_of_mem_of_not_mem H hz simp (config := { contextual := true }) only [this, not_false_iff, true_imp_iff] @@ -2282,7 +2288,7 @@ theorem nat_abs_sum_le {ι : Type*} (s : Finset ι) (f : ι → ℤ) : (∑ i in s, f i).natAbs ≤ ∑ i in s, (f i).natAbs := by classical induction' s using Finset.induction_on with i s his IH - · simp only [Finset.sum_empty, Int.natAbs_zero] + · simp only [Finset.sum_empty, Int.natAbs_zero, le_refl] · simp only [his, Finset.sum_insert, not_false_iff] exact (Int.natAbs_add_le _ _).trans (add_le_add le_rfl IH) #align nat_abs_sum_le nat_abs_sum_le diff --git a/Mathlib/Algebra/BigOperators/Finprod.lean b/Mathlib/Algebra/BigOperators/Finprod.lean index 66cd2ae32d77b..942500fb3026a 100644 --- a/Mathlib/Algebra/BigOperators/Finprod.lean +++ b/Mathlib/Algebra/BigOperators/Finprod.lean @@ -108,19 +108,15 @@ end open Std.ExtendedBinder --- Porting note: removed scoped[BigOperators], `notation3` doesn't mesh with `scoped[Foo]` - /-- `∑ᶠ x, f x` is notation for `finsum f`. It is the sum of `f x`, where `x` ranges over the support of `f`, if it's finite, zero otherwise. Taking the sum over multiple arguments or conditions is possible, e.g. `∏ᶠ (x) (y), f x y` and `∏ᶠ (x) (h: x ∈ s), f x`-/ -notation3"∑ᶠ "(...)", "r:67:(scoped f => finsum f) => r - --- Porting note: removed scoped[BigOperators], `notation3` doesn't mesh with `scoped[Foo]` +scoped[BigOperators] notation3"∑ᶠ "(...)", "r:67:(scoped f => finsum f) => r /-- `∏ᶠ x, f x` is notation for `finprod f`. It is the sum of `f x`, where `x` ranges over the multiplicative support of `f`, if it's finite, one otherwise. Taking the product over multiple arguments or conditions is possible, e.g. `∏ᶠ (x) (y), f x y` and `∏ᶠ (x) (h: x ∈ s), f x`-/ -notation3"∏ᶠ "(...)", "r:67:(scoped f => finprod f) => r +scoped[BigOperators] notation3"∏ᶠ "(...)", "r:67:(scoped f => finprod f) => r -- Porting note: The following ports the lean3 notation for this file, but is currently very fickle. @@ -651,7 +647,7 @@ theorem finprod_mem_mul_distrib' (hf : (s ∩ mulSupport f).Finite) (hg : (s ∩ #align finsum_mem_add_distrib' finsum_mem_add_distrib' /-- The product of the constant function `1` over any set equals `1`. -/ -@[to_additive "The product of the constant function `0` over any set equals `0`."] +@[to_additive "The sum of the constant function `0` over any set equals `0`."] theorem finprod_mem_one (s : Set α) : (∏ᶠ i ∈ s, (1 : M)) = 1 := by simp #align finprod_mem_one finprod_mem_one #align finsum_mem_zero finsum_mem_zero diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index 27db25fb7a2ca..c089facb26a1d 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -220,7 +220,7 @@ theorem prod_eq_single {f : α →₀ M} (a : α) {g : α → M → N} · exact h₀ b (mem_support_iff.mp hb₁) hb₂ · simp only [not_mem_support_iff] at h rw [h] - refine h₁ h + exact h₁ h end SumProd diff --git a/Mathlib/Algebra/BigOperators/Intervals.lean b/Mathlib/Algebra/BigOperators/Intervals.lean index 7cfb960455fe1..9abf4181886ad 100644 --- a/Mathlib/Algebra/BigOperators/Intervals.lean +++ b/Mathlib/Algebra/BigOperators/Intervals.lean @@ -203,7 +203,7 @@ theorem sum_range_id_mul_two (n : ℕ) : (∑ i in range n, i) * 2 = n * (n - 1) by rw [sum_range_reflect (fun i => i) n, mul_two] _ = ∑ i in range n, (i + (n - 1 - i)) := sum_add_distrib.symm _ = ∑ i in range n, (n - 1) := - sum_congr rfl fun i hi => add_tsub_cancel_of_le <| Nat.le_pred_of_lt <| mem_range.1 hi + sum_congr rfl fun i hi => add_tsub_cancel_of_le <| Nat.le_sub_one_of_lt <| mem_range.1 hi _ = n * (n - 1) := by rw [sum_const, card_range, Nat.nsmul_eq_mul] #align finset.sum_range_id_mul_two Finset.sum_range_id_mul_two @@ -275,7 +275,7 @@ theorem sum_Ico_by_parts (hmn : m < n) : have h₂ : (∑ i in Ico (m + 1) n, f i • G (i + 1)) = (∑ i in Ico m (n - 1), f i • G (i + 1)) + f (n - 1) • G n - f m • G (m + 1) := by - rw [← sum_Ico_sub_bot _ hmn, ← sum_Ico_succ_sub_top _ (Nat.le_pred_of_lt hmn), + rw [← sum_Ico_sub_bot _ hmn, ← sum_Ico_succ_sub_top _ (Nat.le_sub_one_of_lt hmn), Nat.sub_add_cancel (pos_of_gt hmn), sub_add_cancel] rw [sum_eq_sum_Ico_succ_bot hmn] -- porting note: the following used to be done with `conv` diff --git a/Mathlib/Algebra/BigOperators/NatAntidiagonal.lean b/Mathlib/Algebra/BigOperators/NatAntidiagonal.lean index 989f0df8df4ce..d72a61d044f54 100644 --- a/Mathlib/Algebra/BigOperators/NatAntidiagonal.lean +++ b/Mathlib/Algebra/BigOperators/NatAntidiagonal.lean @@ -53,13 +53,13 @@ theorem sum_antidiagonal_succ' {n : ℕ} {f : ℕ × ℕ → N} : @[to_additive] theorem prod_antidiagonal_subst {n : ℕ} {f : ℕ × ℕ → ℕ → M} : ∏ p in antidiagonal n, f p n = ∏ p in antidiagonal n, f p (p.1 + p.2) := - prod_congr rfl fun p hp ↦ by rw [Nat.mem_antidiagonal.1 hp] + prod_congr rfl fun p hp ↦ by rw [mem_antidiagonal.mp hp] #align finset.nat.prod_antidiagonal_subst Finset.Nat.prod_antidiagonal_subst #align finset.nat.sum_antidiagonal_subst Finset.Nat.sum_antidiagonal_subst @[to_additive] theorem prod_antidiagonal_eq_prod_range_succ_mk {M : Type*} [CommMonoid M] (f : ℕ × ℕ → M) - (n : ℕ) : ∏ ij in Finset.Nat.antidiagonal n, f ij = ∏ k in range n.succ, f (k, n - k) := + (n : ℕ) : ∏ ij in antidiagonal n, f ij = ∏ k in range n.succ, f (k, n - k) := Finset.prod_map (range n.succ) ⟨fun i ↦ (i, n - i), fun _ _ h ↦ (Prod.mk.inj h).1⟩ f #align finset.nat.prod_antidiagonal_eq_prod_range_succ_mk Finset.Nat.prod_antidiagonal_eq_prod_range_succ_mk #align finset.nat.sum_antidiagonal_eq_sum_range_succ_mk Finset.Nat.sum_antidiagonal_eq_sum_range_succ_mk @@ -69,7 +69,7 @@ using `rw ←`. -/ @[to_additive "This lemma matches more generally than `Finset.Nat.sum_antidiagonal_eq_sum_range_succ_mk` when using `rw ←`."] theorem prod_antidiagonal_eq_prod_range_succ {M : Type*} [CommMonoid M] (f : ℕ → ℕ → M) (n : ℕ) : - ∏ ij in Finset.Nat.antidiagonal n, f ij.1 ij.2 = ∏ k in range n.succ, f k (n - k) := + ∏ ij in antidiagonal n, f ij.1 ij.2 = ∏ k in range n.succ, f k (n - k) := prod_antidiagonal_eq_prod_range_succ_mk _ _ #align finset.nat.prod_antidiagonal_eq_prod_range_succ Finset.Nat.prod_antidiagonal_eq_prod_range_succ #align finset.nat.sum_antidiagonal_eq_sum_range_succ Finset.Nat.sum_antidiagonal_eq_sum_range_succ diff --git a/Mathlib/Algebra/Bounds.lean b/Mathlib/Algebra/Bounds.lean index 67249c993116e..a547945ca5be9 100644 --- a/Mathlib/Algebra/Bounds.lean +++ b/Mathlib/Algebra/Bounds.lean @@ -156,6 +156,14 @@ theorem ciSup_div (hf : BddAbove (Set.range f)) (a : G) : (⨆ i, f i) / a = ⨆ #align csupr_div ciSup_div #align csupr_sub ciSup_sub +@[to_additive] +theorem ciInf_mul (hf : BddBelow (Set.range f)) (a : G) : (⨅ i, f i) * a = ⨅ i, f i * a := + (OrderIso.mulRight a).map_ciInf hf + +@[to_additive] +theorem ciInf_div (hf : BddBelow (Set.range f)) (a : G) : (⨅ i, f i) / a = ⨅ i, f i / a := by + simp only [div_eq_mul_inv, ciInf_mul hf] + end Right section Left @@ -169,6 +177,10 @@ theorem mul_ciSup (hf : BddAbove (Set.range f)) (a : G) : (a * ⨆ i, f i) = ⨆ #align mul_csupr mul_ciSup #align add_csupr add_ciSup +@[to_additive] +theorem mul_ciInf (hf : BddBelow (Set.range f)) (a : G) : (a * ⨅ i, f i) = ⨅ i, a * f i := + (OrderIso.mulLeft a).map_ciInf hf + end Left end ConditionallyCompleteLattice diff --git a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean index 9ec6f0287943d..a274031e4c7cf 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean @@ -76,13 +76,13 @@ set_option maxHeartbeats 800000 in noncomputable instance instMonoidalCategory : MonoidalCategory (AlgebraCat.{u} R) := Monoidal.induced (forget₂ (AlgebraCat R) (ModuleCat R)) - { μIsoSymm := fun X Y => Iso.refl _ - εIsoSymm := Iso.refl _ + { μIso := fun X Y => Iso.refl _ + εIso := Iso.refl _ associator_eq := fun X Y Z => by dsimp only [forget₂_module_obj, forget₂_map_associator_hom] simp only [eqToIso_refl, Iso.refl_trans, Iso.refl_symm, Iso.trans_hom, tensorIso_hom, Iso.refl_hom, MonoidalCategory.tensor_id] - erw [Category.id_comp, Category.comp_id, MonoidalCategory.tensor_id, Category.comp_id] + erw [Category.id_comp, Category.comp_id, MonoidalCategory.tensor_id, Category.id_comp] rightUnitor_eq := fun X => by dsimp erw [Category.id_comp, MonoidalCategory.tensor_id, Category.id_comp] diff --git a/Mathlib/Algebra/Category/GroupCat/Basic.lean b/Mathlib/Algebra/Category/GroupCat/Basic.lean index 1359ca22fac04..5690b845346c6 100644 --- a/Mathlib/Algebra/Category/GroupCat/Basic.lean +++ b/Mathlib/Algebra/Category/GroupCat/Basic.lean @@ -145,7 +145,7 @@ set_option linter.uppercaseLean3 false in /-- Typecheck an `AddMonoidHom` as a morphism in `AddGroup`. -/ add_decl_doc AddGroupCat.ofHom -@[to_additive (attr := simp)] +@[to_additive] theorem ofHom_apply {X Y : Type _} [Group X] [Group Y] (f : X →* Y) (x : X) : (ofHom f) x = f x := rfl @@ -154,9 +154,6 @@ set_option linter.uppercaseLean3 false in set_option linter.uppercaseLean3 false in #align AddGroup.of_hom_apply AddGroupCat.ofHom_apply --- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing -attribute [nolint simpNF] AddGroupCat.ofHom_apply GroupCat.ofHom_apply - @[to_additive] instance ofUnique (G : Type*) [Group G] [i : Unique G] : Unique (GroupCat.of G) := i set_option linter.uppercaseLean3 false in diff --git a/Mathlib/Algebra/Category/GroupCat/EquivalenceGroupAddGroup.lean b/Mathlib/Algebra/Category/GroupCat/EquivalenceGroupAddGroup.lean index 23eeb82106492..edf394914e69a 100644 --- a/Mathlib/Algebra/Category/GroupCat/EquivalenceGroupAddGroup.lean +++ b/Mathlib/Algebra/Category/GroupCat/EquivalenceGroupAddGroup.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import Mathlib.Algebra.Category.GroupCat.Basic -import Mathlib.Algebra.Hom.Equiv.TypeTags +import Mathlib.Algebra.Group.Equiv.TypeTags #align_import algebra.category.Group.equivalence_Group_AddGroup from "leanprover-community/mathlib"@"47b51515e69f59bca5cf34ef456e6000fe205a69" diff --git a/Mathlib/Algebra/Category/ModuleCat/Free.lean b/Mathlib/Algebra/Category/ModuleCat/Free.lean index f88edbf45b70c..87f060179cdf0 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Free.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Free.lean @@ -176,7 +176,7 @@ theorem free_shortExact_finrank_add {M : ModuleCat R} {f : N ⟶ M} [Module.Free R P] [Module.Finite R P] (hN : FiniteDimensional.finrank R N = n) (hP : FiniteDimensional.finrank R P = p) - [StrongRankCondition R]: + [StrongRankCondition R] : FiniteDimensional.finrank R M = n + p := by apply FiniteDimensional.finrank_eq_of_rank_eq rw [free_shortExact_rank_add h, ← hN, ← hP] diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean index a39c08fe1dd57..b58421dd4d2b1 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean @@ -66,8 +66,6 @@ open MonoidalCategory -- porting note: `CoeFun` was replaced by `FunLike` -- I can't seem to express the function coercion here without writing `@FunLike.coe`. --- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing -@[simp, nolint simpNF] theorem monoidalClosed_curry {M N P : ModuleCat.{u} R} (f : M ⊗ N ⟶ P) (x : M) (y : N) : @FunLike.coe _ _ _ LinearMap.instFunLike ((MonoidalClosed.curry f : N →ₗ[R] M →ₗ[R] P) y) x = f (x ⊗ₜ[R] y) := diff --git a/Mathlib/Algebra/Category/SemigroupCat/Basic.lean b/Mathlib/Algebra/Category/SemigroupCat/Basic.lean index f4f648fd071ff..653066329e4b7 100644 --- a/Mathlib/Algebra/Category/SemigroupCat/Basic.lean +++ b/Mathlib/Algebra/Category/SemigroupCat/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Julian Kuelshammer -/ import Mathlib.Algebra.PEmptyInstances -import Mathlib.Algebra.Hom.Equiv.Basic +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.CategoryTheory.ConcreteCategory.BundledHom import Mathlib.CategoryTheory.Functor.ReflectsIso import Mathlib.CategoryTheory.Elementwise diff --git a/Mathlib/Algebra/CharP/Basic.lean b/Mathlib/Algebra/CharP/Basic.lean index 1189f3e12fb17..2dce592d4520c 100644 --- a/Mathlib/Algebra/CharP/Basic.lean +++ b/Mathlib/Algebra/CharP/Basic.lean @@ -511,8 +511,6 @@ end CommRing section Semiring -open Nat - variable [NonAssocSemiring R] theorem char_ne_one [Nontrivial R] (p : ℕ) [hc : CharP R p] : p ≠ 1 := fun hp : p = 1 => @@ -529,7 +527,7 @@ theorem char_is_prime_of_two_le (p : ℕ) [hc : CharP R p] (hp : 2 ≤ p) : Nat. fun (d : ℕ) (hdvd : ∃ e, p = d * e) => let ⟨e, hmul⟩ := hdvd have : (p : R) = 0 := (cast_eq_zero_iff R p p).mpr (dvd_refl p) - have : (d : R) * e = 0 := @cast_mul R _ d e ▸ hmul ▸ this + have : (d : R) * e = 0 := @Nat.cast_mul R _ d e ▸ hmul ▸ this Or.elim (eq_zero_or_eq_zero_of_mul_eq_zero this) (fun hd : (d : R) = 0 => have : p ∣ d := (cast_eq_zero_iff R p d).mp hd @@ -718,8 +716,8 @@ theorem Int.cast_injOn_of_ringChar_ne_two {R : Type*} [NonAssocRing R] [Nontrivi rintro _ (rfl | rfl | rfl) _ (rfl | rfl | rfl) h <;> simp only [cast_neg, cast_one, cast_zero, neg_eq_zero, one_ne_zero, zero_ne_one, zero_eq_neg] at h ⊢ - · exact (Ring.neg_one_ne_one_of_char_ne_two hR).symm h - · exact (Ring.neg_one_ne_one_of_char_ne_two hR) h + · exact ((Ring.neg_one_ne_one_of_char_ne_two hR).symm h).elim + · exact ((Ring.neg_one_ne_one_of_char_ne_two hR) h).elim #align int.cast_inj_on_of_ring_char_ne_two Int.cast_injOn_of_ringChar_ne_two end diff --git a/Mathlib/Algebra/CharP/ExpChar.lean b/Mathlib/Algebra/CharP/ExpChar.lean index 55ca94dbddbb2..0374f7e8d4f29 100644 --- a/Mathlib/Algebra/CharP/ExpChar.lean +++ b/Mathlib/Algebra/CharP/ExpChar.lean @@ -71,9 +71,9 @@ theorem char_zero_of_expChar_one (p : ℕ) [hp : CharP R p] [hq : ExpChar R 1] : · exact False.elim (CharP.char_ne_one R 1 rfl) #align char_zero_of_exp_char_one char_zero_of_expChar_one --- see Note [lower instance priority] +-- This could be an instance, but there are no `ExpChar R 1` instances in mathlib. /-- The characteristic is zero if the exponential characteristic is one. -/ -instance (priority := 100) charZero_of_expChar_one' [hq : ExpChar R 1] : CharZero R := by +theorem charZero_of_expChar_one' [hq : ExpChar R 1] : CharZero R := by cases hq · assumption · exact False.elim (CharP.char_ne_one R 1 rfl) diff --git a/Mathlib/Algebra/CharP/Quotient.lean b/Mathlib/Algebra/CharP/Quotient.lean index 53563933d2eb8..6de49e382f07b 100644 --- a/Mathlib/Algebra/CharP/Quotient.lean +++ b/Mathlib/Algebra/CharP/Quotient.lean @@ -60,9 +60,6 @@ theorem Ideal.Quotient.index_eq_zero {R : Type*} [CommRing R] (I : Ideal R) : (↑I.toAddSubgroup.index : R ⧸ I) = 0 := by rw [AddSubgroup.index, Nat.card_eq] split_ifs with hq; swap; simp - by_contra h - -- TODO: can we avoid rewriting the `I.to_add_subgroup` here? letI : Fintype (R ⧸ I) := @Fintype.ofFinite _ hq - have h : (Fintype.card (R ⧸ I) : R ⧸ I) ≠ 0 := h - simp at h + exact CharP.cast_card_eq_zero (R ⧸ I) #align ideal.quotient.index_eq_zero Ideal.Quotient.index_eq_zero diff --git a/Mathlib/Algebra/ContinuedFractions/Translations.lean b/Mathlib/Algebra/ContinuedFractions/Translations.lean index d939655bfb24f..e1d4532344f62 100644 --- a/Mathlib/Algebra/ContinuedFractions/Translations.lean +++ b/Mathlib/Algebra/ContinuedFractions/Translations.lean @@ -80,7 +80,7 @@ section WithDivisionRing /-! ### Translations Between Computational Functions -Here we give some basic translations that hold by definition for the computational methods of a +Here we give some basic translations that hold by definition for the computational methods of a continued fraction. -/ diff --git a/Mathlib/Algebra/DirectLimit.lean b/Mathlib/Algebra/DirectLimit.lean index afd41080e5552..6e62ea3518c24 100644 --- a/Mathlib/Algebra/DirectLimit.lean +++ b/Mathlib/Algebra/DirectLimit.lean @@ -128,8 +128,7 @@ theorem exists_of [Nonempty ι] [IsDirected ι (· ≤ ·)] (z : DirectLimit G f let ⟨k, hik, hjk⟩ := exists_ge_ge i j ⟨k, f i k hik x + f j k hjk y, by rw [LinearMap.map_add, of_f, of_f, ihx, ihy] - -- porting note: was `rfl` - simp only [Submodule.Quotient.mk''_eq_mk, Quotient.mk_add]⟩ + rfl ⟩ #align module.direct_limit.exists_of Module.DirectLimit.exists_of @[elab_as_elim] diff --git a/Mathlib/Algebra/DirectSum/Basic.lean b/Mathlib/Algebra/DirectSum/Basic.lean index cb238664a53b0..bdddfdbc4bc7c 100644 --- a/Mathlib/Algebra/DirectSum/Basic.lean +++ b/Mathlib/Algebra/DirectSum/Basic.lean @@ -53,11 +53,9 @@ instance [∀ i, AddCommMonoid (β i)] : FunLike (DirectSum ι β) _ fun i : ι instance [∀ i, AddCommMonoid (β i)] : CoeFun (DirectSum ι β) fun _ => ∀ i : ι, β i := inferInstanceAs (CoeFun (Π₀ i, β i) fun _ => ∀ i : ι, β i) --- Porting note: scoped does not work with notation3; TODO rewrite as lean4 notation? --- scoped[DirectSum] /-- `⨁ i, f i` is notation for `DirectSum _ f` and equals the direct sum of `fun i ↦ f i`. Taking the direct sum over multiple arguments is possible, e.g. `⨁ (i) (j), f i j`. -/ -notation3 "⨁ "(...)", "r:(scoped f => DirectSum _ f) => r +scoped[DirectSum] notation3 "⨁ "(...)", "r:(scoped f => DirectSum _ f) => r -- Porting note: The below recreates some of the lean3 notation, not fully yet -- section diff --git a/Mathlib/Algebra/Divisibility/Basic.lean b/Mathlib/Algebra/Divisibility/Basic.lean index e68ae2796ab6e..d864745f10791 100644 --- a/Mathlib/Algebra/Divisibility/Basic.lean +++ b/Mathlib/Algebra/Divisibility/Basic.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Amelia Livingston, Yury Kudryashov, Neil Strickland, Aaron Anderson -/ -import Mathlib.Algebra.Hom.Group.Defs import Mathlib.Algebra.Group.Basic +import Mathlib.Algebra.Group.Hom.Defs #align_import algebra.divisibility.basic from "leanprover-community/mathlib"@"e8638a0fcaf73e4500469f368ef9494e495099b3" diff --git a/Mathlib/Algebra/DualNumber.lean b/Mathlib/Algebra/DualNumber.lean index 2e378003953c6..2a91e789794e0 100644 --- a/Mathlib/Algebra/DualNumber.lean +++ b/Mathlib/Algebra/DualNumber.lean @@ -39,20 +39,21 @@ Rather than duplicating the API of `TrivSqZeroExt`, this file reuses the functio variable {R : Type*} -/-- The type of dual numbers, numbers of the form $a + bε$ where $ε^2 = 0$.-/ +/-- The type of dual numbers, numbers of the form $a + bε$ where $ε^2 = 0$. +`R[ε]` is notation for `DualNumber R`. -/ abbrev DualNumber (R : Type*) : Type _ := TrivSqZeroExt R R #align dual_number DualNumber -/-- The unit element $ε$ that squares to zero. -/ +/-- The unit element $ε$ that squares to zero, with notation `ε`. -/ def DualNumber.eps [Zero R] [One R] : DualNumber R := TrivSqZeroExt.inr 1 #align dual_number.eps DualNumber.eps --- mathport name: dual_number.eps +@[inherit_doc] scoped[DualNumber] notation "ε" => DualNumber.eps --- mathport name: dual_number +@[inherit_doc] scoped[DualNumber] postfix:1024 "[ε]" => DualNumber open DualNumber @@ -87,6 +88,13 @@ theorem inr_eq_smul_eps [MulZeroOneClass R] (r : R) : inr r = (r • ε : R[ε]) ext (mul_zero r).symm (mul_one r).symm #align dual_number.inr_eq_smul_eps DualNumber.inr_eq_smul_eps +/-- `ε` commutes with every element of the algebra. -/ +theorem commute_eps_left [Semiring R] (x : DualNumber R) : Commute ε x := by + ext <;> simp + +/-- `ε` commutes with every element of the algebra. -/ +theorem commute_eps_right [Semiring R] (x : DualNumber R) : Commute x ε := (commute_eps_left x).symm + /-- For two algebra morphisms out of `R[ε]` to agree, it suffices for them to agree on `ε`. -/ @[ext] theorem algHom_ext {A} [CommSemiring R] [Semiring A] [Algebra R A] ⦃f g : R[ε] →ₐ[R] A⦄ @@ -111,13 +119,13 @@ def lift : { e : A // e * e = 0 } ≃ (R[ε] →ₐ[R] A) := TrivSqZeroExt.lift #align dual_number.lift DualNumber.lift --- When applied to `ε`, `DualNumber.lift` produces the element of `A` that squares to 0. +/-- When applied to `ε`, `DualNumber.lift` produces the element of `A` that squares to 0. -/ -- @[simp] -- Porting note: simp can prove this theorem lift_apply_eps (e : { e : A // e * e = 0 }) : @lift R _ _ _ _ e (ε : R[ε]) = e := by simp only [lift_apply_apply, fst_eps, map_zero, snd_eps, one_smul, zero_add] #align dual_number.lift_apply_eps DualNumber.lift_apply_eps --- Lifting `DualNumber.eps` itself gives the identity. +/-- Lifting `DualNumber.eps` itself gives the identity. -/ @[simp] theorem lift_eps : lift ⟨ε, eps_mul_eps⟩ = AlgHom.id R R[ε] := algHom_ext <| lift_apply_eps _ diff --git a/Mathlib/Algebra/Field/Basic.lean b/Mathlib/Algebra/Field/Basic.lean index 0496e42650871..720e6cc73e8ce 100644 --- a/Mathlib/Algebra/Field/Basic.lean +++ b/Mathlib/Algebra/Field/Basic.lean @@ -5,8 +5,8 @@ Authors: Robert Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Field.Defs import Mathlib.Algebra.GroupWithZero.Units.Lemmas -import Mathlib.Algebra.Hom.Ring.Defs import Mathlib.Algebra.Ring.Commute +import Mathlib.Algebra.Ring.Hom.Defs #align_import algebra.field.basic from "leanprover-community/mathlib"@"05101c3df9d9cfe9430edc205860c79b6d660102" diff --git a/Mathlib/Algebra/Free.lean b/Mathlib/Algebra/Free.lean index 47c09c91c7196..c9a05e667bc8b 100644 --- a/Mathlib/Algebra/Free.lean +++ b/Mathlib/Algebra/Free.lean @@ -3,12 +3,12 @@ Copyright (c) 2019 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.Hom.Group.Defs -import Mathlib.Algebra.Hom.Equiv.Basic +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Control.Applicative import Mathlib.Control.Traversable.Basic -import Mathlib.Logic.Equiv.Defs import Mathlib.Data.List.Basic +import Mathlib.Logic.Equiv.Defs #align_import algebra.free from "leanprover-community/mathlib"@"6d0adfa76594f304b4650d098273d4366edeb61b" diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index e5aec67a3a769..c456e7dbe6011 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -147,7 +147,7 @@ theorem geom_sum₂_self {α : Type*} [CommRing α] (x : α) (n : ℕ) : by simp_rw [← pow_add] _ = ∑ i in Finset.range n, x ^ (n - 1) := Finset.sum_congr rfl fun i hi => - congr_arg _ <| add_tsub_cancel_of_le <| Nat.le_pred_of_lt <| Finset.mem_range.1 hi + congr_arg _ <| add_tsub_cancel_of_le <| Nat.le_sub_one_of_lt <| Finset.mem_range.1 hi _ = (Finset.range n).card • x ^ (n - 1) := Finset.sum_const _ _ = n * x ^ (n - 1) := by rw [Finset.card_range, nsmul_eq_mul] #align geom_sum₂_self geom_sum₂_self @@ -309,7 +309,7 @@ protected theorem Commute.geom_sum₂_succ_eq {α : Type u} [Ring α] {x y : α} suffices n - 1 - i + 1 = n - i by rw [this] cases' n with n · exact absurd (List.mem_range.mp hi) i.not_lt_zero - · rw [tsub_add_eq_add_tsub (Nat.le_pred_of_lt (List.mem_range.mp hi)), + · rw [tsub_add_eq_add_tsub (Nat.le_sub_one_of_lt (List.mem_range.mp hi)), tsub_add_cancel_of_le (Nat.succ_le_iff.mpr n.succ_pos)] #align commute.geom_sum₂_succ_eq Commute.geom_sum₂_succ_eq @@ -483,7 +483,7 @@ theorem geom_sum_alternating_of_le_neg_one [StrictOrderedRing α] (hx : x + 1 if Even n then (∑ i in range n, x ^ i) ≤ 0 else 1 ≤ ∑ i in range n, x ^ i := by have hx0 : x ≤ 0 := (le_add_of_nonneg_right zero_le_one).trans hx induction' n with n ih - · simp only [Nat.zero_eq, range_zero, sum_empty, le_refl, ite_true] + · simp only [Nat.zero_eq, range_zero, sum_empty, le_refl, ite_true, even_zero] simp only [Nat.even_add_one, geom_sum_succ] split_ifs at ih with h · rw [if_neg (not_not_intro h), le_add_iff_nonneg_left] @@ -497,7 +497,7 @@ theorem geom_sum_alternating_of_lt_neg_one [StrictOrderedRing α] (hx : x + 1 < if Even n then (∑ i in range n, x ^ i) < 0 else 1 < ∑ i in range n, x ^ i := by have hx0 : x < 0 := ((le_add_iff_nonneg_right _).2 zero_le_one).trans_lt hx refine' Nat.le_induction _ _ n (show 2 ≤ n from hn) - · simp only [geom_sum_two, lt_add_iff_pos_left, ite_true, gt_iff_lt, hx] + · simp only [geom_sum_two, lt_add_iff_pos_left, ite_true, gt_iff_lt, hx, even_two] clear hn intro n _ ihn simp only [Nat.even_add_one, geom_sum_succ] @@ -571,7 +571,7 @@ theorem geom_sum_eq_zero_iff_neg_one [LinearOrderedRing α] (hn : n ≠ 0) : have hx := eq_or_ne x (-1) cases' hx with hx hx · rw [hx, neg_one_geom_sum] - simp only [h hx, ne_eq, ite_eq_left_iff, one_ne_zero, not_forall, exists_prop, and_true] + simp only [h hx, ite_false, ne_eq, one_ne_zero, not_false_eq_true] · exact geom_sum_ne_zero hx hn #align geom_sum_eq_zero_iff_neg_one geom_sum_eq_zero_iff_neg_one diff --git a/Mathlib/Algebra/GradedMulAction.lean b/Mathlib/Algebra/GradedMulAction.lean index 987d78c0a0240..b6e6210ddcc12 100644 --- a/Mathlib/Algebra/GradedMulAction.lean +++ b/Mathlib/Algebra/GradedMulAction.lean @@ -47,7 +47,7 @@ graded action -/ -variable {ιA ιB ιM : Type*} +variable {ιA ιB ιM : Type*} namespace GradedMonoid diff --git a/Mathlib/Algebra/Hom/Aut.lean b/Mathlib/Algebra/Group/Aut.lean similarity index 100% rename from Mathlib/Algebra/Hom/Aut.lean rename to Mathlib/Algebra/Group/Aut.lean diff --git a/Mathlib/Algebra/Hom/Commute.lean b/Mathlib/Algebra/Group/Commute/Hom.lean similarity index 97% rename from Mathlib/Algebra/Hom/Commute.lean rename to Mathlib/Algebra/Group/Commute/Hom.lean index 906983a062f1d..21db9705aa562 100644 --- a/Mathlib/Algebra/Hom/Commute.lean +++ b/Mathlib/Algebra/Group/Commute/Hom.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Kevin Buzzard, Scott Morrison, Johan Commelin, Chris Hughes, Johannes Hölzl, Yury Kudryashov -/ -import Mathlib.Algebra.Hom.Group.Defs import Mathlib.Algebra.Group.Commute.Defs +import Mathlib.Algebra.Group.Hom.Defs #align_import algebra.hom.commute from "leanprover-community/mathlib"@"6eb334bd8f3433d5b08ba156b8ec3e6af47e1904" diff --git a/Mathlib/Algebra/Group/Commute/Units.lean b/Mathlib/Algebra/Group/Commute/Units.lean index 98b147ebaa223..ab88e7de6d631 100644 --- a/Mathlib/Algebra/Group/Commute/Units.lean +++ b/Mathlib/Algebra/Group/Commute/Units.lean @@ -94,8 +94,7 @@ theorem isUnit_mul_iff (h : Commute a b) : IsUnit (a * b) ↔ IsUnit a ∧ IsUni @[to_additive (attr := simp)] theorem _root_.isUnit_mul_self_iff : IsUnit (a * a) ↔ IsUnit a := - (Commute.refl a).isUnit_mul_iff.trans (and_self_iff _) - -- porting note: `and_self_iff` now has an implicit argument instead of an explicit one. + (Commute.refl a).isUnit_mul_iff.trans and_self_iff #align is_unit_mul_self_iff isUnit_mul_self_iff #align is_add_unit_add_self_iff isAddUnit_add_self_iff diff --git a/Mathlib/Algebra/Group/Conj.lean b/Mathlib/Algebra/Group/Conj.lean index e3d9cfcea2b27..2f7ca4622f96d 100644 --- a/Mathlib/Algebra/Group/Conj.lean +++ b/Mathlib/Algebra/Group/Conj.lean @@ -3,10 +3,10 @@ Copyright (c) 2018 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Chris Hughes, Michael Howes -/ +import Mathlib.Algebra.Group.Aut +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Algebra.Group.Semiconj.Defs import Mathlib.Algebra.GroupWithZero.Basic -import Mathlib.Algebra.Hom.Aut -import Mathlib.Algebra.Hom.Group.Defs #align_import algebra.group.conj from "leanprover-community/mathlib"@"0743cc5d9d86bcd1bba10f480e948a257d65056f" diff --git a/Mathlib/Algebra/Hom/Embedding.lean b/Mathlib/Algebra/Group/Embedding.lean similarity index 100% rename from Mathlib/Algebra/Hom/Embedding.lean rename to Mathlib/Algebra/Group/Embedding.lean diff --git a/Mathlib/Algebra/Hom/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean similarity index 99% rename from Mathlib/Algebra/Hom/Equiv/Basic.lean rename to Mathlib/Algebra/Group/Equiv/Basic.lean index 5249423ddf6f7..2518b66304d03 100644 --- a/Mathlib/Algebra/Hom/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -3,10 +3,10 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ -import Mathlib.Algebra.Hom.Group.Basic +import Mathlib.Algebra.Group.Hom.Basic import Mathlib.Data.FunLike.Equiv -import Mathlib.Logic.Equiv.Basic import Mathlib.Data.Pi.Algebra +import Mathlib.Logic.Equiv.Basic #align_import algebra.hom.equiv.basic from "leanprover-community/mathlib"@"1ac8d4304efba9d03fa720d06516fac845aa5353" diff --git a/Mathlib/Algebra/Hom/Equiv/TypeTags.lean b/Mathlib/Algebra/Group/Equiv/TypeTags.lean similarity index 99% rename from Mathlib/Algebra/Hom/Equiv/TypeTags.lean rename to Mathlib/Algebra/Group/Equiv/TypeTags.lean index a154b6573fbec..73f028cf7a279 100644 --- a/Mathlib/Algebra/Hom/Equiv/TypeTags.lean +++ b/Mathlib/Algebra/Group/Equiv/TypeTags.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ -import Mathlib.Algebra.Hom.Equiv.Basic +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.TypeTags #align_import algebra.hom.equiv.type_tags from "leanprover-community/mathlib"@"3342d1b2178381196f818146ff79bc0e7ccd9e2d" diff --git a/Mathlib/Algebra/Group/Ext.lean b/Mathlib/Algebra/Group/Ext.lean index f433b7c9a7592..379d7ad46e9dd 100644 --- a/Mathlib/Algebra/Group/Ext.lean +++ b/Mathlib/Algebra/Group/Ext.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Bryan Gin-ge Chen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bryan Gin-ge Chen, Yury Kudryashov -/ -import Mathlib.Algebra.Hom.Group.Defs +import Mathlib.Algebra.Group.Hom.Defs #align_import algebra.group.ext from "leanprover-community/mathlib"@"e574b1a4e891376b0ef974b926da39e05da12a06" diff --git a/Mathlib/Algebra/Hom/Freiman.lean b/Mathlib/Algebra/Group/Freiman.lean similarity index 99% rename from Mathlib/Algebra/Hom/Freiman.lean rename to Mathlib/Algebra/Group/Freiman.lean index ccf0501eb826c..9f5e964da35e1 100644 --- a/Mathlib/Algebra/Hom/Freiman.lean +++ b/Mathlib/Algebra/Group/Freiman.lean @@ -422,7 +422,7 @@ instance commMonoid : CommMonoid (A →*[n] β) where @[to_additive "If `β` is an additive commutative group, then `A →*[n] β` is an additive commutative group too."] -instance commGroup {β} [CommGroup β]: CommGroup (A →*[n] β) := +instance commGroup {β} [CommGroup β] : CommGroup (A →*[n] β) := { FreimanHom.commMonoid with div_eq_mul_inv := by intros diff --git a/Mathlib/Algebra/Hom/Group/Basic.lean b/Mathlib/Algebra/Group/Hom/Basic.lean similarity index 99% rename from Mathlib/Algebra/Hom/Group/Basic.lean rename to Mathlib/Algebra/Group/Hom/Basic.lean index 95d5dd2af20a6..5278a7d454605 100644 --- a/Mathlib/Algebra/Hom/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Hom/Basic.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Kevin Buzzard, Scott Morrison, Johan Commelin, Chris Hughes, Johannes Hölzl, Yury Kudryashov -/ -import Mathlib.Algebra.Hom.Group.Defs -import Mathlib.Algebra.NeZero import Mathlib.Algebra.Group.Basic +import Mathlib.Algebra.Group.Hom.Defs +import Mathlib.Algebra.NeZero #align_import algebra.hom.group from "leanprover-community/mathlib"@"a148d797a1094ab554ad4183a4ad6f130358ef64" diff --git a/Mathlib/Algebra/Hom/Group/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean similarity index 100% rename from Mathlib/Algebra/Hom/Group/Defs.lean rename to Mathlib/Algebra/Group/Hom/Defs.lean diff --git a/Mathlib/Algebra/Hom/GroupInstances.lean b/Mathlib/Algebra/Group/Hom/Instances.lean similarity index 99% rename from Mathlib/Algebra/Hom/GroupInstances.lean rename to Mathlib/Algebra/Group/Hom/Instances.lean index 06a61f75b2e63..b97ba32e2c755 100644 --- a/Mathlib/Algebra/Hom/GroupInstances.lean +++ b/Mathlib/Algebra/Group/Hom/Instances.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Kevin Buzzard, Scott Morrison, Johan Commelin, Chris Hughes, Johannes Hölzl, Yury Kudryashov -/ +import Mathlib.Algebra.Group.Hom.Basic import Mathlib.Algebra.GroupPower.Basic import Mathlib.Algebra.Ring.Basic -import Mathlib.Algebra.Hom.Group.Basic #align_import algebra.hom.group_instances from "leanprover-community/mathlib"@"2ed7e4aec72395b6a7c3ac4ac7873a7a43ead17c" diff --git a/Mathlib/Algebra/Group/Opposite.lean b/Mathlib/Algebra/Group/Opposite.lean index 2381b6bdccd91..c714f35061716 100644 --- a/Mathlib/Algebra/Group/Opposite.lean +++ b/Mathlib/Algebra/Group/Opposite.lean @@ -3,10 +3,10 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ +import Mathlib.Algebra.Group.Commute.Defs +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.InjSurj import Mathlib.Algebra.Group.Units -import Mathlib.Algebra.Group.Commute.Defs -import Mathlib.Algebra.Hom.Equiv.Basic import Mathlib.Algebra.Opposites import Mathlib.Data.Int.Cast.Defs import Mathlib.Tactic.Spread diff --git a/Mathlib/Algebra/Group/Pi.lean b/Mathlib/Algebra/Group/Pi.lean index cc2156684d33b..b516b17cde874 100644 --- a/Mathlib/Algebra/Group/Pi.lean +++ b/Mathlib/Algebra/Group/Pi.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot -/ import Mathlib.Init.CCLemmas -import Mathlib.Logic.Pairwise -import Mathlib.Algebra.Hom.GroupInstances +import Mathlib.Algebra.Group.Hom.Instances import Mathlib.Data.Pi.Algebra import Mathlib.Data.Set.Function +import Mathlib.Logic.Pairwise #align_import algebra.group.pi from "leanprover-community/mathlib"@"e4bc74cbaf429d706cb9140902f7ca6c431e75a4" diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index 1003edceeac13..ab9efddf02405 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot, Yury Kudryashov -/ import Mathlib.Algebra.Group.Opposite +import Mathlib.Algebra.Group.Units.Hom import Mathlib.Algebra.GroupWithZero.Units.Basic -import Mathlib.Algebra.Hom.Units #align_import algebra.group.prod from "leanprover-community/mathlib"@"cd391184c85986113f8c00844cfe6dda1d34be3d" @@ -378,7 +378,7 @@ variable {M' : Type*} {N' : Type*} [Mul M] [Mul N] [Mul M'] [Mul N'] [Mul P] (f (g : N →ₙ* N') /-- `Prod.map` as a `MonoidHom`. -/ -@[to_additive prodMap "`prod.map` as an `AddMonoidHom`"] +@[to_additive prodMap "`Prod.map` as an `AddMonoidHom`"] def prodMap : M × N →ₙ* M' × N' := (f.comp (fst M N)).prod (g.comp (snd M N)) #align mul_hom.prod_map MulHom.prodMap diff --git a/Mathlib/Algebra/Group/TypeTags.lean b/Mathlib/Algebra/Group/TypeTags.lean index cf5ed1aa6b851..d3a459693eb92 100644 --- a/Mathlib/Algebra/Group/TypeTags.lean +++ b/Mathlib/Algebra/Group/TypeTags.lean @@ -3,10 +3,10 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Hom.Group.Defs -import Mathlib.Logic.Nontrivial.Basic -import Mathlib.Logic.Equiv.Defs +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Data.Finite.Defs +import Mathlib.Logic.Equiv.Defs +import Mathlib.Logic.Nontrivial.Basic #align_import algebra.group.type_tags from "leanprover-community/mathlib"@"2e0975f6a25dd3fbfb9e41556a77f075f6269748" @@ -132,6 +132,10 @@ instance [h: Infinite α] : Infinite (Additive α) := h instance [h: Infinite α] : Infinite (Multiplicative α) := h +instance [h : DecidableEq α] : DecidableEq (Multiplicative α) := h + +instance [h : DecidableEq α] : DecidableEq (Additive α) := h + instance instNontrivialAdditive [Nontrivial α] : Nontrivial (Additive α) := ofMul.injective.nontrivial #align additive.nontrivial instNontrivialAdditive diff --git a/Mathlib/Algebra/Group/ULift.lean b/Mathlib/Algebra/Group/ULift.lean index 7f5f473305ef3..cf47cfe2ac986 100644 --- a/Mathlib/Algebra/Group/ULift.lean +++ b/Mathlib/Algebra/Group/ULift.lean @@ -3,10 +3,10 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Mathlib.Logic.Nontrivial.Basic -import Mathlib.Data.Int.Cast.Defs -import Mathlib.Algebra.Hom.Equiv.Basic +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.GroupWithZero.InjSurj +import Mathlib.Data.Int.Cast.Defs +import Mathlib.Logic.Nontrivial.Basic #align_import algebra.group.ulift from "leanprover-community/mathlib"@"564bcc44d2b394a50c0cd6340c14a6b02a50a99a" diff --git a/Mathlib/Algebra/Hom/Equiv/Units/Basic.lean b/Mathlib/Algebra/Group/Units/Equiv.lean similarity index 99% rename from Mathlib/Algebra/Hom/Equiv/Units/Basic.lean rename to Mathlib/Algebra/Group/Units/Equiv.lean index 3ae03a72e7a49..111ffd553fc46 100644 --- a/Mathlib/Algebra/Hom/Equiv/Units/Basic.lean +++ b/Mathlib/Algebra/Group/Units/Equiv.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ -import Mathlib.Algebra.Hom.Equiv.Basic -import Mathlib.Algebra.Hom.Units +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Algebra.Group.Units.Hom #align_import algebra.hom.equiv.units.basic from "leanprover-community/mathlib"@"a95b16cbade0f938fc24abd05412bde1e84bab9b" diff --git a/Mathlib/Algebra/Hom/Units.lean b/Mathlib/Algebra/Group/Units/Hom.lean similarity index 99% rename from Mathlib/Algebra/Hom/Units.lean rename to Mathlib/Algebra/Group/Units/Hom.lean index 1518f577ec05d..02150f895524f 100644 --- a/Mathlib/Algebra/Hom/Units.lean +++ b/Mathlib/Algebra/Group/Units/Hom.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Johan Commelin All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Chris Hughes, Kevin Buzzard -/ -import Mathlib.Algebra.Hom.Group.Defs +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Algebra.Group.Units #align_import algebra.hom.units from "leanprover-community/mathlib"@"a07d750983b94c530ab69a726862c2ab6802b38c" diff --git a/Mathlib/Algebra/Group/WithOne/Basic.lean b/Mathlib/Algebra/Group/WithOne/Basic.lean index a0f775bbcd762..3f7ae36f5602a 100644 --- a/Mathlib/Algebra/Group/WithOne/Basic.lean +++ b/Mathlib/Algebra/Group/WithOne/Basic.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Johan Commelin -/ +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.WithOne.Defs -import Mathlib.Algebra.Hom.Equiv.Basic #align_import algebra.group.with_one.basic from "leanprover-community/mathlib"@"4dc134b97a3de65ef2ed881f3513d56260971562" diff --git a/Mathlib/Algebra/GroupPower/CovariantClass.lean b/Mathlib/Algebra/GroupPower/CovariantClass.lean new file mode 100644 index 0000000000000..980b1efe151d9 --- /dev/null +++ b/Mathlib/Algebra/GroupPower/CovariantClass.lean @@ -0,0 +1,361 @@ +/- +Copyright (c) 2015 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Robert Y. Lewis, Yury G. Kudryashov +-/ +import Mathlib.Algebra.GroupPower.Basic +import Mathlib.Algebra.Order.Monoid.OrderDual +import Mathlib.Data.Nat.Basic +import Mathlib.Tactic.Monotonicity.Attr + +/-! +# Lemmas about the interaction of power operations with order in terms of CovariantClass + +More lemmas with bundled algebra+order typeclasses are in `Algebra/GroupPower/Order.lean` +and `Algebra/GroupPower/Lemmas.lean`. +-/ + +open Function + +variable {β G M : Type*} + +section Monoid + +variable [Monoid M] + +section Preorder + +variable [Preorder M] + +section Left + +variable [CovariantClass M M (· * ·) (· ≤ ·)] {x : M} + +@[to_additive (attr := mono) nsmul_le_nsmul_of_le_right] +theorem pow_le_pow_of_le_left' [CovariantClass M M (swap (· * ·)) (· ≤ ·)] {a b : M} (hab : a ≤ b) : + ∀ i : ℕ, a ^ i ≤ b ^ i + | 0 => by simp + | k + 1 => by + rw [pow_succ, pow_succ] + exact mul_le_mul' hab (pow_le_pow_of_le_left' hab k) +#align pow_le_pow_of_le_left' pow_le_pow_of_le_left' +#align nsmul_le_nsmul_of_le_right nsmul_le_nsmul_of_le_right + +@[to_additive nsmul_nonneg] +theorem one_le_pow_of_one_le' {a : M} (H : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n + | 0 => by simp + | k + 1 => by + rw [pow_succ] + exact one_le_mul H (one_le_pow_of_one_le' H k) +#align one_le_pow_of_one_le' one_le_pow_of_one_le' +#align nsmul_nonneg nsmul_nonneg + +@[to_additive nsmul_nonpos] +theorem pow_le_one' {a : M} (H : a ≤ 1) (n : ℕ) : a ^ n ≤ 1 := + @one_le_pow_of_one_le' Mᵒᵈ _ _ _ _ H n +#align pow_le_one' pow_le_one' +#align nsmul_nonpos nsmul_nonpos + +@[to_additive nsmul_le_nsmul] +theorem pow_le_pow' {a : M} {n m : ℕ} (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := + let ⟨k, hk⟩ := Nat.le.dest h + calc + a ^ n ≤ a ^ n * a ^ k := le_mul_of_one_le_right' (one_le_pow_of_one_le' ha _) + _ = a ^ m := by rw [← hk, pow_add] +#align pow_le_pow' pow_le_pow' +#align nsmul_le_nsmul nsmul_le_nsmul + +@[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' Mᵒᵈ _ _ _ _ _ _ ha h +#align pow_le_pow_of_le_one' pow_le_pow_of_le_one' +#align nsmul_le_nsmul_of_nonpos nsmul_le_nsmul_of_nonpos + +@[to_additive nsmul_pos] +theorem one_lt_pow' {a : M} (ha : 1 < a) {k : ℕ} (hk : k ≠ 0) : 1 < a ^ k := by + rcases Nat.exists_eq_succ_of_ne_zero hk with ⟨l, rfl⟩ + clear hk + induction' l with l IH + · rw [pow_succ]; simpa using ha + · rw [pow_succ] + exact one_lt_mul'' ha IH +#align one_lt_pow' one_lt_pow' +#align nsmul_pos nsmul_pos + +@[to_additive nsmul_neg] +theorem pow_lt_one' {a : M} (ha : a < 1) {k : ℕ} (hk : k ≠ 0) : a ^ k < 1 := + @one_lt_pow' Mᵒᵈ _ _ _ _ ha k hk +#align pow_lt_one' pow_lt_one' +#align nsmul_neg nsmul_neg + +@[to_additive nsmul_lt_nsmul] +theorem pow_lt_pow' [CovariantClass M M (· * ·) (· < ·)] {a : M} {n m : ℕ} (ha : 1 < a) + (h : n < m) : a ^ n < a ^ m := by + rcases Nat.le.dest h with ⟨k, rfl⟩; clear h + rw [pow_add, pow_succ', mul_assoc, ← pow_succ] + exact lt_mul_of_one_lt_right' _ (one_lt_pow' ha k.succ_ne_zero) +#align pow_lt_pow' pow_lt_pow' +#align nsmul_lt_nsmul nsmul_lt_nsmul + +@[to_additive nsmul_strictMono_right] +theorem pow_strictMono_left [CovariantClass M M (· * ·) (· < ·)] {a : M} (ha : 1 < a) : + StrictMono ((· ^ ·) a : ℕ → M) := fun _ _ => pow_lt_pow' ha +#align pow_strict_mono_left pow_strictMono_left +#align nsmul_strict_mono_right nsmul_strictMono_right + +@[to_additive Left.pow_nonneg] +theorem Left.one_le_pow_of_le (hx : 1 ≤ x) : ∀ {n : ℕ}, 1 ≤ x ^ n + | 0 => (pow_zero x).ge + | n + 1 => by + rw [pow_succ] + exact Left.one_le_mul hx <| Left.one_le_pow_of_le hx +#align left.one_le_pow_of_le Left.one_le_pow_of_le +#align left.pow_nonneg Left.pow_nonneg + +@[to_additive Left.pow_nonpos] +theorem Left.pow_le_one_of_le (hx : x ≤ 1) : ∀ {n : ℕ}, x ^ n ≤ 1 + | 0 => (pow_zero _).le + | n + 1 => by + rw [pow_succ] + exact Left.mul_le_one hx <| Left.pow_le_one_of_le hx +#align left.pow_le_one_of_le Left.pow_le_one_of_le +#align left.pow_nonpos Left.pow_nonpos + +end Left + +section Right + +variable [CovariantClass M M (swap (· * ·)) (· ≤ ·)] {x : M} + +@[to_additive Right.pow_nonneg] +theorem Right.one_le_pow_of_le (hx : 1 ≤ x) : ∀ {n : ℕ}, 1 ≤ x ^ n + | 0 => (pow_zero _).ge + | n + 1 => by + rw [pow_succ] + exact Right.one_le_mul hx <| Right.one_le_pow_of_le hx +#align right.one_le_pow_of_le Right.one_le_pow_of_le +#align right.pow_nonneg Right.pow_nonneg + +@[to_additive Right.pow_nonpos] +theorem Right.pow_le_one_of_le (hx : x ≤ 1) : ∀ {n : ℕ}, x ^ n ≤ 1 + | 0 => (pow_zero _).le + | n + 1 => by + rw [pow_succ] + exact Right.mul_le_one hx <| Right.pow_le_one_of_le hx +#align right.pow_le_one_of_le Right.pow_le_one_of_le +#align right.pow_nonpos Right.pow_nonpos + +end Right + +section CovariantLTSwap + +variable [Preorder β] [CovariantClass M M (· * ·) (· < ·)] + [CovariantClass M M (swap (· * ·)) (· < ·)] {f : β → M} + +@[to_additive StrictMono.nsmul_left] +theorem StrictMono.pow_right' (hf : StrictMono f) : ∀ {n : ℕ}, n ≠ 0 → StrictMono fun a => f a ^ n + | 0, hn => (hn rfl).elim + | 1, _ => by simpa + | Nat.succ <| Nat.succ n, _ => by + simp_rw [pow_succ _ (n + 1)] + exact hf.mul' (StrictMono.pow_right' hf n.succ_ne_zero) +#align strict_mono.pow_right' StrictMono.pow_right' +#align strict_mono.nsmul_left StrictMono.nsmul_left + +/-- See also `pow_strictMono_right` -/ +@[to_additive nsmul_strictMono_left] -- Porting note: nolint to_additive_doc +theorem pow_strictMono_right' {n : ℕ} (hn : n ≠ 0) : StrictMono fun a : M => a ^ n := + strictMono_id.pow_right' hn +#align pow_strict_mono_right' pow_strictMono_right' +#align nsmul_strict_mono_left nsmul_strictMono_left + +end CovariantLTSwap + +section CovariantLESwap + +variable [Preorder β] [CovariantClass M M (· * ·) (· ≤ ·)] + [CovariantClass M M (swap (· * ·)) (· ≤ ·)] + +@[to_additive Monotone.nsmul_left] +theorem Monotone.pow_right {f : β → M} (hf : Monotone f) : ∀ n : ℕ, Monotone fun a => f a ^ n + | 0 => by simpa using monotone_const + | n + 1 => by + simp_rw [pow_succ] + exact hf.mul' (Monotone.pow_right hf _) +#align monotone.pow_right Monotone.pow_right +#align monotone.nsmul_left Monotone.nsmul_left + +@[to_additive nsmul_mono_left] +theorem pow_mono_right (n : ℕ) : Monotone fun a : M => a ^ n := + monotone_id.pow_right _ +#align pow_mono_right pow_mono_right +#align nsmul_mono_left nsmul_mono_left + +end CovariantLESwap + +@[to_additive Left.pow_neg] +theorem Left.pow_lt_one_of_lt [CovariantClass M M (· * ·) (· < ·)] {n : ℕ} {x : M} (hn : 0 < n) + (h : x < 1) : x ^ n < 1 := + Nat.le_induction ((pow_one _).trans_lt h) + (fun n _ ih => by + rw [pow_succ] + exact mul_lt_one h ih) + _ (Nat.succ_le_iff.2 hn) +#align left.pow_lt_one_of_lt Left.pow_lt_one_of_lt +#align left.pow_neg Left.pow_neg + +@[to_additive Right.pow_neg] +theorem Right.pow_lt_one_of_lt [CovariantClass M M (swap (· * ·)) (· < ·)] {n : ℕ} {x : M} + (hn : 0 < n) (h : x < 1) : x ^ n < 1 := + Nat.le_induction ((pow_one _).trans_lt h) + (fun n _ ih => by + rw [pow_succ] + exact Right.mul_lt_one h ih) + _ (Nat.succ_le_iff.2 hn) +#align right.pow_lt_one_of_lt Right.pow_lt_one_of_lt +#align right.pow_neg Right.pow_neg + +end Preorder + +section LinearOrder + +variable [LinearOrder M] + +section CovariantLE + +variable [CovariantClass M M (· * ·) (· ≤ ·)] + +@[to_additive nsmul_nonneg_iff] +theorem one_le_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 ≤ x ^ n ↔ 1 ≤ x := + ⟨le_imp_le_of_lt_imp_lt fun h => pow_lt_one' h hn, fun h => one_le_pow_of_one_le' h n⟩ +#align one_le_pow_iff one_le_pow_iff +#align nsmul_nonneg_iff nsmul_nonneg_iff + +@[to_additive] +theorem pow_le_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n ≤ 1 ↔ x ≤ 1 := + @one_le_pow_iff Mᵒᵈ _ _ _ _ _ hn +#align pow_le_one_iff pow_le_one_iff +#align nsmul_nonpos_iff nsmul_nonpos_iff + +@[to_additive nsmul_pos_iff] +theorem one_lt_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 < x ^ n ↔ 1 < x := + lt_iff_lt_of_le_iff_le (pow_le_one_iff hn) +#align one_lt_pow_iff one_lt_pow_iff +#align nsmul_pos_iff nsmul_pos_iff + +@[to_additive] +theorem pow_lt_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n < 1 ↔ x < 1 := + lt_iff_lt_of_le_iff_le (one_le_pow_iff hn) +#align pow_lt_one_iff pow_lt_one_iff +#align nsmul_neg_iff nsmul_neg_iff + +@[to_additive] +theorem pow_eq_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n = 1 ↔ x = 1 := by + simp only [le_antisymm_iff] + rw [pow_le_one_iff hn, one_le_pow_iff hn] +#align pow_eq_one_iff pow_eq_one_iff +#align nsmul_eq_zero_iff nsmul_eq_zero_iff + +variable [CovariantClass M M (· * ·) (· < ·)] {a : M} {m n : ℕ} + +@[to_additive nsmul_le_nsmul_iff] +theorem pow_le_pow_iff' (ha : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := + (pow_strictMono_left ha).le_iff_le +#align pow_le_pow_iff' pow_le_pow_iff' +#align nsmul_le_nsmul_iff nsmul_le_nsmul_iff + +@[to_additive nsmul_lt_nsmul_iff] +theorem pow_lt_pow_iff' (ha : 1 < a) : a ^ m < a ^ n ↔ m < n := + (pow_strictMono_left ha).lt_iff_lt +#align pow_lt_pow_iff' pow_lt_pow_iff' +#align nsmul_lt_nsmul_iff nsmul_lt_nsmul_iff + +end CovariantLE + +section CovariantLESwap + +variable [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (swap (· * ·)) (· ≤ ·)] + +@[to_additive lt_of_nsmul_lt_nsmul] +theorem lt_of_pow_lt_pow' {a b : M} (n : ℕ) : a ^ n < b ^ n → a < b := + (pow_mono_right _).reflect_lt +#align lt_of_pow_lt_pow' lt_of_pow_lt_pow' +#align lt_of_nsmul_lt_nsmul lt_of_nsmul_lt_nsmul + +@[to_additive min_lt_of_add_lt_two_nsmul] +theorem min_lt_of_mul_lt_sq {a b c : M} (h : a * b < c ^ 2) : min a b < c := by + simpa using min_lt_max_of_mul_lt_mul (h.trans_eq <| pow_two _) +#align min_lt_of_mul_lt_sq min_lt_of_mul_lt_sq +#align min_lt_of_add_lt_two_nsmul min_lt_of_add_lt_two_nsmul + +@[to_additive lt_max_of_two_nsmul_lt_add] +theorem lt_max_of_sq_lt_mul {a b c : M} (h : a ^ 2 < b * c) : a < max b c := by + simpa using min_lt_max_of_mul_lt_mul ((pow_two _).symm.trans_lt h) +#align lt_max_of_sq_lt_mul lt_max_of_sq_lt_mul +#align lt_max_of_two_nsmul_lt_add lt_max_of_two_nsmul_lt_add + +end CovariantLESwap + +section CovariantLTSwap + +variable [CovariantClass M M (· * ·) (· < ·)] [CovariantClass M M (swap (· * ·)) (· < ·)] + +@[to_additive le_of_nsmul_le_nsmul] +theorem le_of_pow_le_pow' {a b : M} {n : ℕ} (hn : n ≠ 0) : a ^ n ≤ b ^ n → a ≤ b := + (pow_strictMono_right' hn).le_iff_le.1 +#align le_of_pow_le_pow' le_of_pow_le_pow' +#align le_of_nsmul_le_nsmul le_of_nsmul_le_nsmul + +@[to_additive min_le_of_add_le_two_nsmul] +theorem min_le_of_mul_le_sq {a b c : M} (h : a * b ≤ c ^ 2) : min a b ≤ c := by + simpa using min_le_max_of_mul_le_mul (h.trans_eq <| pow_two _) +#align min_le_of_mul_le_sq min_le_of_mul_le_sq +#align min_le_of_add_le_two_nsmul min_le_of_add_le_two_nsmul + +@[to_additive le_max_of_two_nsmul_le_add] +theorem le_max_of_sq_le_mul {a b c : M} (h : a ^ 2 ≤ b * c) : a ≤ max b c := by + simpa using min_le_max_of_mul_le_mul ((pow_two _).symm.trans_le h) +#align le_max_of_sq_le_mul le_max_of_sq_le_mul +#align le_max_of_two_nsmul_le_add le_max_of_two_nsmul_le_add + +end CovariantLTSwap + +@[to_additive Left.nsmul_neg_iff] +theorem Left.pow_lt_one_iff' [CovariantClass M M (· * ·) (· < ·)] {n : ℕ} {x : M} (hn : 0 < n) : + x ^ n < 1 ↔ x < 1 := + haveI := covariantClass_le_of_lt M M (· * ·) + pow_lt_one_iff hn.ne' +#align left.nsmul_neg_iff Left.nsmul_neg_iff + +theorem Left.pow_lt_one_iff [CovariantClass M M (· * ·) (· < ·)] {n : ℕ} {x : M} (hn : 0 < n) : + x ^ n < 1 ↔ x < 1 := Left.pow_lt_one_iff' hn +#align left.pow_lt_one_iff Left.pow_lt_one_iff + +@[to_additive] +theorem Right.pow_lt_one_iff [CovariantClass M M (swap (· * ·)) (· < ·)] {n : ℕ} {x : M} + (hn : 0 < n) : x ^ n < 1 ↔ x < 1 := + ⟨fun H => + not_le.mp fun k => + haveI := covariantClass_le_of_lt M M (swap (· * ·)) + H.not_le <| Right.one_le_pow_of_le k, + Right.pow_lt_one_of_lt hn⟩ +#align right.pow_lt_one_iff Right.pow_lt_one_iff +#align right.nsmul_neg_iff Right.nsmul_neg_iff + +end LinearOrder + +end Monoid + +section DivInvMonoid + +variable [DivInvMonoid G] [Preorder G] [CovariantClass G G (· * ·) (· ≤ ·)] + +@[to_additive zsmul_nonneg] +theorem one_le_zpow {x : G} (H : 1 ≤ x) {n : ℤ} (hn : 0 ≤ n) : 1 ≤ x ^ n := by + lift n to ℕ using hn + rw [zpow_ofNat] + apply one_le_pow_of_one_le' H +#align one_le_zpow one_le_zpow +#align zsmul_nonneg zsmul_nonneg + +end DivInvMonoid diff --git a/Mathlib/Algebra/Hom/Iterate.lean b/Mathlib/Algebra/GroupPower/IterateHom.lean similarity index 100% rename from Mathlib/Algebra/Hom/Iterate.lean rename to Mathlib/Algebra/GroupPower/IterateHom.lean diff --git a/Mathlib/Algebra/GroupPower/NegOnePow.lean b/Mathlib/Algebra/GroupPower/NegOnePow.lean index 06b6b2f499283..d53c1e0bf8040 100644 --- a/Mathlib/Algebra/GroupPower/NegOnePow.lean +++ b/Mathlib/Algebra/GroupPower/NegOnePow.lean @@ -56,6 +56,7 @@ lemma negOnePow_eq_one_iff (n : ℤ) : n.negOnePow = 1 ↔ Even n := by rw [Int.even_iff_not_odd] intro h' simp only [negOnePow_odd _ h'] at h + contradiction · exact negOnePow_even n lemma negOnePow_eq_neg_one_iff (n : ℤ) : n.negOnePow = -1 ↔ Odd n := by @@ -64,7 +65,7 @@ lemma negOnePow_eq_neg_one_iff (n : ℤ) : n.negOnePow = -1 ↔ Odd n := by rw [Int.odd_iff_not_even] intro h' rw [negOnePow_even _ h'] at h - simp only at h + contradiction · exact negOnePow_odd n @[simp] diff --git a/Mathlib/Algebra/GroupPower/Order.lean b/Mathlib/Algebra/GroupPower/Order.lean index 8d25463a69c78..fce3a5168d8e5 100644 --- a/Mathlib/Algebra/GroupPower/Order.lean +++ b/Mathlib/Algebra/GroupPower/Order.lean @@ -5,10 +5,9 @@ Authors: Jeremy Avigad, Robert Y. Lewis -/ import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Order.WithZero +import Mathlib.Algebra.GroupPower.CovariantClass import Mathlib.Algebra.GroupPower.Ring import Mathlib.Data.Set.Intervals.Basic -import Mathlib.Data.Nat.Basic -import Mathlib.Init.Data.Nat.Basic import Mathlib.Data.Nat.Order.Basic #align_import algebra.group_power.order from "leanprover-community/mathlib"@"00f91228655eecdcd3ac97a7fd8dbcb139fe990a" @@ -20,351 +19,7 @@ Note that some lemmas are in `Algebra/GroupPower/Lemmas.lean` as they import fil depend on this file. -/ - -open Function - -variable {β A G M R : Type*} - -section Monoid - -variable [Monoid M] - -section Preorder - -variable [Preorder M] - -section Left - -variable [CovariantClass M M (· * ·) (· ≤ ·)] {x : M} - -@[to_additive (attr := mono) nsmul_le_nsmul_of_le_right] -theorem pow_le_pow_of_le_left' [CovariantClass M M (swap (· * ·)) (· ≤ ·)] {a b : M} (hab : a ≤ b) : - ∀ i : ℕ, a ^ i ≤ b ^ i - | 0 => by simp - | k + 1 => by - rw [pow_succ, pow_succ] - exact mul_le_mul' hab (pow_le_pow_of_le_left' hab k) -#align pow_le_pow_of_le_left' pow_le_pow_of_le_left' -#align nsmul_le_nsmul_of_le_right nsmul_le_nsmul_of_le_right - -@[to_additive nsmul_nonneg] -theorem one_le_pow_of_one_le' {a : M} (H : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n - | 0 => by simp - | k + 1 => by - rw [pow_succ] - exact one_le_mul H (one_le_pow_of_one_le' H k) -#align one_le_pow_of_one_le' one_le_pow_of_one_le' -#align nsmul_nonneg nsmul_nonneg - -@[to_additive nsmul_nonpos] -theorem pow_le_one' {a : M} (H : a ≤ 1) (n : ℕ) : a ^ n ≤ 1 := - @one_le_pow_of_one_le' Mᵒᵈ _ _ _ _ H n -#align pow_le_one' pow_le_one' -#align nsmul_nonpos nsmul_nonpos - -@[to_additive nsmul_le_nsmul] -theorem pow_le_pow' {a : M} {n m : ℕ} (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := - let ⟨k, hk⟩ := Nat.le.dest h - calc - a ^ n ≤ a ^ n * a ^ k := le_mul_of_one_le_right' (one_le_pow_of_one_le' ha _) - _ = a ^ m := by rw [← hk, pow_add] -#align pow_le_pow' pow_le_pow' -#align nsmul_le_nsmul nsmul_le_nsmul - -@[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' Mᵒᵈ _ _ _ _ _ _ ha h -#align pow_le_pow_of_le_one' pow_le_pow_of_le_one' -#align nsmul_le_nsmul_of_nonpos nsmul_le_nsmul_of_nonpos - -@[to_additive nsmul_pos] -theorem one_lt_pow' {a : M} (ha : 1 < a) {k : ℕ} (hk : k ≠ 0) : 1 < a ^ k := by - rcases Nat.exists_eq_succ_of_ne_zero hk with ⟨l, rfl⟩ - clear hk - induction' l with l IH - · rw [pow_succ]; simpa using ha - · rw [pow_succ] - exact one_lt_mul'' ha IH -#align one_lt_pow' one_lt_pow' -#align nsmul_pos nsmul_pos - -@[to_additive nsmul_neg] -theorem pow_lt_one' {a : M} (ha : a < 1) {k : ℕ} (hk : k ≠ 0) : a ^ k < 1 := - @one_lt_pow' Mᵒᵈ _ _ _ _ ha k hk -#align pow_lt_one' pow_lt_one' -#align nsmul_neg nsmul_neg - -@[to_additive nsmul_lt_nsmul] -theorem pow_lt_pow' [CovariantClass M M (· * ·) (· < ·)] {a : M} {n m : ℕ} (ha : 1 < a) - (h : n < m) : a ^ n < a ^ m := by - rcases Nat.le.dest h with ⟨k, rfl⟩; clear h - rw [pow_add, pow_succ', mul_assoc, ← pow_succ] - exact lt_mul_of_one_lt_right' _ (one_lt_pow' ha k.succ_ne_zero) -#align pow_lt_pow' pow_lt_pow' -#align nsmul_lt_nsmul nsmul_lt_nsmul - -@[to_additive nsmul_strictMono_right] -theorem pow_strictMono_left [CovariantClass M M (· * ·) (· < ·)] {a : M} (ha : 1 < a) : - StrictMono ((· ^ ·) a : ℕ → M) := fun _ _ => pow_lt_pow' ha -#align pow_strict_mono_left pow_strictMono_left -#align nsmul_strict_mono_right nsmul_strictMono_right - -@[to_additive Left.pow_nonneg] -theorem Left.one_le_pow_of_le (hx : 1 ≤ x) : ∀ {n : ℕ}, 1 ≤ x ^ n - | 0 => (pow_zero x).ge - | n + 1 => by - rw [pow_succ] - exact Left.one_le_mul hx <| Left.one_le_pow_of_le hx -#align left.one_le_pow_of_le Left.one_le_pow_of_le -#align left.pow_nonneg Left.pow_nonneg - -@[to_additive Left.pow_nonpos] -theorem Left.pow_le_one_of_le (hx : x ≤ 1) : ∀ {n : ℕ}, x ^ n ≤ 1 - | 0 => (pow_zero _).le - | n + 1 => by - rw [pow_succ] - exact Left.mul_le_one hx <| Left.pow_le_one_of_le hx -#align left.pow_le_one_of_le Left.pow_le_one_of_le -#align left.pow_nonpos Left.pow_nonpos - -end Left - -section Right - -variable [CovariantClass M M (swap (· * ·)) (· ≤ ·)] {x : M} - -@[to_additive Right.pow_nonneg] -theorem Right.one_le_pow_of_le (hx : 1 ≤ x) : ∀ {n : ℕ}, 1 ≤ x ^ n - | 0 => (pow_zero _).ge - | n + 1 => by - rw [pow_succ] - exact Right.one_le_mul hx <| Right.one_le_pow_of_le hx -#align right.one_le_pow_of_le Right.one_le_pow_of_le -#align right.pow_nonneg Right.pow_nonneg - -@[to_additive Right.pow_nonpos] -theorem Right.pow_le_one_of_le (hx : x ≤ 1) : ∀ {n : ℕ}, x ^ n ≤ 1 - | 0 => (pow_zero _).le - | n + 1 => by - rw [pow_succ] - exact Right.mul_le_one hx <| Right.pow_le_one_of_le hx -#align right.pow_le_one_of_le Right.pow_le_one_of_le -#align right.pow_nonpos Right.pow_nonpos - -end Right - -section CovariantLTSwap - -variable [Preorder β] [CovariantClass M M (· * ·) (· < ·)] - [CovariantClass M M (swap (· * ·)) (· < ·)] {f : β → M} - -@[to_additive StrictMono.nsmul_left] -theorem StrictMono.pow_right' (hf : StrictMono f) : ∀ {n : ℕ}, n ≠ 0 → StrictMono fun a => f a ^ n - | 0, hn => (hn rfl).elim - | 1, _ => by simpa - | Nat.succ <| Nat.succ n, _ => by - simp_rw [pow_succ _ (n + 1)] - exact hf.mul' (StrictMono.pow_right' hf n.succ_ne_zero) -#align strict_mono.pow_right' StrictMono.pow_right' -#align strict_mono.nsmul_left StrictMono.nsmul_left - -/-- See also `pow_strictMono_right` -/ -@[to_additive nsmul_strictMono_left] -- Porting note: nolint to_additive_doc -theorem pow_strictMono_right' {n : ℕ} (hn : n ≠ 0) : StrictMono fun a : M => a ^ n := - strictMono_id.pow_right' hn -#align pow_strict_mono_right' pow_strictMono_right' -#align nsmul_strict_mono_left nsmul_strictMono_left - -end CovariantLTSwap - -section CovariantLESwap - -variable [Preorder β] [CovariantClass M M (· * ·) (· ≤ ·)] - [CovariantClass M M (swap (· * ·)) (· ≤ ·)] - -@[to_additive Monotone.nsmul_left] -theorem Monotone.pow_right {f : β → M} (hf : Monotone f) : ∀ n : ℕ, Monotone fun a => f a ^ n - | 0 => by simpa using monotone_const - | n + 1 => by - simp_rw [pow_succ] - exact hf.mul' (Monotone.pow_right hf _) -#align monotone.pow_right Monotone.pow_right -#align monotone.nsmul_left Monotone.nsmul_left - -@[to_additive nsmul_mono_left] -theorem pow_mono_right (n : ℕ) : Monotone fun a : M => a ^ n := - monotone_id.pow_right _ -#align pow_mono_right pow_mono_right -#align nsmul_mono_left nsmul_mono_left - -end CovariantLESwap - -@[to_additive Left.pow_neg] -theorem Left.pow_lt_one_of_lt [CovariantClass M M (· * ·) (· < ·)] {n : ℕ} {x : M} (hn : 0 < n) - (h : x < 1) : x ^ n < 1 := - Nat.le_induction ((pow_one _).trans_lt h) - (fun n _ ih => by - rw [pow_succ] - exact mul_lt_one h ih) - _ (Nat.succ_le_iff.2 hn) -#align left.pow_lt_one_of_lt Left.pow_lt_one_of_lt -#align left.pow_neg Left.pow_neg - -@[to_additive Right.pow_neg] -theorem Right.pow_lt_one_of_lt [CovariantClass M M (swap (· * ·)) (· < ·)] {n : ℕ} {x : M} - (hn : 0 < n) (h : x < 1) : x ^ n < 1 := - Nat.le_induction ((pow_one _).trans_lt h) - (fun n _ ih => by - rw [pow_succ] - exact Right.mul_lt_one h ih) - _ (Nat.succ_le_iff.2 hn) -#align right.pow_lt_one_of_lt Right.pow_lt_one_of_lt -#align right.pow_neg Right.pow_neg - -end Preorder - -section LinearOrder - -variable [LinearOrder M] - -section CovariantLE - -variable [CovariantClass M M (· * ·) (· ≤ ·)] - -@[to_additive nsmul_nonneg_iff] -theorem one_le_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 ≤ x ^ n ↔ 1 ≤ x := - ⟨le_imp_le_of_lt_imp_lt fun h => pow_lt_one' h hn, fun h => one_le_pow_of_one_le' h n⟩ -#align one_le_pow_iff one_le_pow_iff -#align nsmul_nonneg_iff nsmul_nonneg_iff - -@[to_additive] -theorem pow_le_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n ≤ 1 ↔ x ≤ 1 := - @one_le_pow_iff Mᵒᵈ _ _ _ _ _ hn -#align pow_le_one_iff pow_le_one_iff -#align nsmul_nonpos_iff nsmul_nonpos_iff - -@[to_additive nsmul_pos_iff] -theorem one_lt_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 < x ^ n ↔ 1 < x := - lt_iff_lt_of_le_iff_le (pow_le_one_iff hn) -#align one_lt_pow_iff one_lt_pow_iff -#align nsmul_pos_iff nsmul_pos_iff - -@[to_additive] -theorem pow_lt_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n < 1 ↔ x < 1 := - lt_iff_lt_of_le_iff_le (one_le_pow_iff hn) -#align pow_lt_one_iff pow_lt_one_iff -#align nsmul_neg_iff nsmul_neg_iff - -@[to_additive] -theorem pow_eq_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n = 1 ↔ x = 1 := by - simp only [le_antisymm_iff] - rw [pow_le_one_iff hn, one_le_pow_iff hn] -#align pow_eq_one_iff pow_eq_one_iff -#align nsmul_eq_zero_iff nsmul_eq_zero_iff - -variable [CovariantClass M M (· * ·) (· < ·)] {a : M} {m n : ℕ} - -@[to_additive nsmul_le_nsmul_iff] -theorem pow_le_pow_iff' (ha : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := - (pow_strictMono_left ha).le_iff_le -#align pow_le_pow_iff' pow_le_pow_iff' -#align nsmul_le_nsmul_iff nsmul_le_nsmul_iff - -@[to_additive nsmul_lt_nsmul_iff] -theorem pow_lt_pow_iff' (ha : 1 < a) : a ^ m < a ^ n ↔ m < n := - (pow_strictMono_left ha).lt_iff_lt -#align pow_lt_pow_iff' pow_lt_pow_iff' -#align nsmul_lt_nsmul_iff nsmul_lt_nsmul_iff - -end CovariantLE - -section CovariantLESwap - -variable [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (swap (· * ·)) (· ≤ ·)] - -@[to_additive lt_of_nsmul_lt_nsmul] -theorem lt_of_pow_lt_pow' {a b : M} (n : ℕ) : a ^ n < b ^ n → a < b := - (pow_mono_right _).reflect_lt -#align lt_of_pow_lt_pow' lt_of_pow_lt_pow' -#align lt_of_nsmul_lt_nsmul lt_of_nsmul_lt_nsmul - -@[to_additive min_lt_of_add_lt_two_nsmul] -theorem min_lt_of_mul_lt_sq {a b c : M} (h : a * b < c ^ 2) : min a b < c := by - simpa using min_lt_max_of_mul_lt_mul (h.trans_eq <| pow_two _) -#align min_lt_of_mul_lt_sq min_lt_of_mul_lt_sq -#align min_lt_of_add_lt_two_nsmul min_lt_of_add_lt_two_nsmul - -@[to_additive lt_max_of_two_nsmul_lt_add] -theorem lt_max_of_sq_lt_mul {a b c : M} (h : a ^ 2 < b * c) : a < max b c := by - simpa using min_lt_max_of_mul_lt_mul ((pow_two _).symm.trans_lt h) -#align lt_max_of_sq_lt_mul lt_max_of_sq_lt_mul -#align lt_max_of_two_nsmul_lt_add lt_max_of_two_nsmul_lt_add - -end CovariantLESwap - -section CovariantLTSwap - -variable [CovariantClass M M (· * ·) (· < ·)] [CovariantClass M M (swap (· * ·)) (· < ·)] - -@[to_additive le_of_nsmul_le_nsmul] -theorem le_of_pow_le_pow' {a b : M} {n : ℕ} (hn : n ≠ 0) : a ^ n ≤ b ^ n → a ≤ b := - (pow_strictMono_right' hn).le_iff_le.1 -#align le_of_pow_le_pow' le_of_pow_le_pow' -#align le_of_nsmul_le_nsmul le_of_nsmul_le_nsmul - -@[to_additive min_le_of_add_le_two_nsmul] -theorem min_le_of_mul_le_sq {a b c : M} (h : a * b ≤ c ^ 2) : min a b ≤ c := by - simpa using min_le_max_of_mul_le_mul (h.trans_eq <| pow_two _) -#align min_le_of_mul_le_sq min_le_of_mul_le_sq -#align min_le_of_add_le_two_nsmul min_le_of_add_le_two_nsmul - -@[to_additive le_max_of_two_nsmul_le_add] -theorem le_max_of_sq_le_mul {a b c : M} (h : a ^ 2 ≤ b * c) : a ≤ max b c := by - simpa using min_le_max_of_mul_le_mul ((pow_two _).symm.trans_le h) -#align le_max_of_sq_le_mul le_max_of_sq_le_mul -#align le_max_of_two_nsmul_le_add le_max_of_two_nsmul_le_add - -end CovariantLTSwap - -@[to_additive Left.nsmul_neg_iff] -theorem Left.pow_lt_one_iff' [CovariantClass M M (· * ·) (· < ·)] {n : ℕ} {x : M} (hn : 0 < n) : - x ^ n < 1 ↔ x < 1 := - haveI := covariantClass_le_of_lt M M (· * ·) - pow_lt_one_iff hn.ne' -#align left.nsmul_neg_iff Left.nsmul_neg_iff - -theorem Left.pow_lt_one_iff [CovariantClass M M (· * ·) (· < ·)] {n : ℕ} {x : M} (hn : 0 < n) : - x ^ n < 1 ↔ x < 1 := Left.pow_lt_one_iff' hn -#align left.pow_lt_one_iff Left.pow_lt_one_iff - -@[to_additive] -theorem Right.pow_lt_one_iff [CovariantClass M M (swap (· * ·)) (· < ·)] {n : ℕ} {x : M} - (hn : 0 < n) : x ^ n < 1 ↔ x < 1 := - ⟨fun H => - not_le.mp fun k => - haveI := covariantClass_le_of_lt M M (swap (· * ·)) - H.not_le <| Right.one_le_pow_of_le k, - Right.pow_lt_one_of_lt hn⟩ -#align right.pow_lt_one_iff Right.pow_lt_one_iff -#align right.nsmul_neg_iff Right.nsmul_neg_iff - -end LinearOrder - -end Monoid - -section DivInvMonoid - -variable [DivInvMonoid G] [Preorder G] [CovariantClass G G (· * ·) (· ≤ ·)] - -@[to_additive zsmul_nonneg] -theorem one_le_zpow {x : G} (H : 1 ≤ x) {n : ℤ} (hn : 0 ≤ n) : 1 ≤ x ^ n := by - lift n to ℕ using hn - rw [zpow_ofNat] - apply one_le_pow_of_one_le' H -#align one_le_zpow one_le_zpow -#align zsmul_nonneg zsmul_nonneg - -end DivInvMonoid +variable {M R : Type*} namespace CanonicallyOrderedCommSemiring diff --git a/Mathlib/Algebra/GroupPower/Ring.lean b/Mathlib/Algebra/GroupPower/Ring.lean index 11a9e9b89ee44..a8ad33a2d07a7 100644 --- a/Mathlib/Algebra/GroupPower/Ring.lean +++ b/Mathlib/Algebra/GroupPower/Ring.lean @@ -4,13 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis -/ +import Mathlib.Algebra.Group.Units.Hom import Mathlib.Algebra.GroupPower.Basic import Mathlib.Algebra.GroupWithZero.Commute -import Mathlib.Algebra.Hom.Ring.Defs -import Mathlib.Algebra.Hom.Units -import Mathlib.Algebra.Ring.Commute import Mathlib.Algebra.GroupWithZero.Divisibility +import Mathlib.Algebra.Ring.Commute import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Data.Nat.Order.Basic #align_import algebra.group_power.ring from "leanprover-community/mathlib"@"fc2ed6f838ce7c9b7c7171e58d78eaf7b438fb0e" diff --git a/Mathlib/Algebra/GroupRingAction/Invariant.lean b/Mathlib/Algebra/GroupRingAction/Invariant.lean index 45d0b1e884e94..5345722af0806 100644 --- a/Mathlib/Algebra/GroupRingAction/Invariant.lean +++ b/Mathlib/Algebra/GroupRingAction/Invariant.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Algebra.Hom.GroupAction +import Mathlib.GroupTheory.GroupAction.Hom import Mathlib.RingTheory.Subring.Pointwise #align_import algebra.group_ring_action.invariant from "leanprover-community/mathlib"@"e7bab9a85e92cf46c02cb4725a7be2f04691e3a7" diff --git a/Mathlib/Algebra/Hom/Equiv/Units/GroupWithZero.lean b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean similarity index 97% rename from Mathlib/Algebra/Hom/Equiv/Units/GroupWithZero.lean rename to Mathlib/Algebra/GroupWithZero/Units/Equiv.lean index 929c02db5fa6d..312f7e28587fe 100644 --- a/Mathlib/Algebra/Hom/Equiv/Units/GroupWithZero.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ -import Mathlib.Algebra.Hom.Equiv.Units.Basic +import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Algebra.GroupWithZero.Units.Basic #align_import algebra.hom.equiv.units.group_with_zero from "leanprover-community/mathlib"@"655994e298904d7e5bbd1e18c95defd7b543eb94" diff --git a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean index da7cf0278b389..8d55661eeadcb 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean @@ -3,11 +3,11 @@ Copyright (c) 2020 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ +import Mathlib.Algebra.Group.Hom.Basic +import Mathlib.Algebra.Group.Units.Hom import Mathlib.Algebra.GroupWithZero.Commute -import Mathlib.Algebra.Hom.Units -import Mathlib.Algebra.Hom.Group.Basic -import Mathlib.GroupTheory.GroupAction.Units import Mathlib.Algebra.GroupWithZero.Units.Basic +import Mathlib.GroupTheory.GroupAction.Units #align_import algebra.group_with_zero.units.lemmas from "leanprover-community/mathlib"@"dc6c365e751e34d100e80fe6e314c3c3e0fd2988" diff --git a/Mathlib/Algebra/Homology/Additive.lean b/Mathlib/Algebra/Homology/Additive.lean index 9f83edd0b8509..f5027b147ad29 100644 --- a/Mathlib/Algebra/Homology/Additive.lean +++ b/Mathlib/Algebra/Homology/Additive.lean @@ -120,8 +120,8 @@ namespace HomologicalComplex instance eval_additive (i : ι) : (eval V c i).Additive where #align homological_complex.eval_additive HomologicalComplex.eval_additive -instance cycles_additive [HasEqualizers V] : (cyclesFunctor V c i).Additive where -#align homological_complex.cycles_additive HomologicalComplex.cycles_additive +instance cycles'_additive [HasEqualizers V] : (cycles'Functor V c i).Additive where +#align homological_complex.cycles_additive HomologicalComplex.cycles'_additive variable [HasImages V] [HasImageMaps V] @@ -130,11 +130,11 @@ instance boundaries_additive : (boundariesFunctor V c i).Additive where variable [HasEqualizers V] [HasCokernels V] -instance homology_additive : (homologyFunctor V c i).Additive where +instance homology_additive : (homology'Functor V c i).Additive where map_add {_ _ f g} := by - dsimp [homologyFunctor] + dsimp [homology'Functor] ext - simp only [homology.π_map, Preadditive.comp_add, ← Preadditive.add_comp] + simp only [homology'.π_map, Preadditive.comp_add, ← Preadditive.add_comp] congr ext simp diff --git a/Mathlib/Algebra/Homology/Augment.lean b/Mathlib/Algebra/Homology/Augment.lean index fe2d9689017c8..b19030893c688 100644 --- a/Mathlib/Algebra/Homology/Augment.lean +++ b/Mathlib/Algebra/Homology/Augment.lean @@ -239,7 +239,7 @@ def augment (C : CochainComplex V ℕ) {X : V} (f : X ⟶ C.X 0) (w : f ≫ C.d shape i j s := by simp at s rcases j with (_ | _ | j) <;> cases i <;> try simp - · simp at s + · contradiction · rw [C.shape] simp only [ComplexShape.up_Rel] contrapose! s diff --git a/Mathlib/Algebra/Homology/Exact.lean b/Mathlib/Algebra/Homology/Exact.lean index 1f56be7810bc1..3167d21fea54b 100644 --- a/Mathlib/Algebra/Homology/Exact.lean +++ b/Mathlib/Algebra/Homology/Exact.lean @@ -37,6 +37,9 @@ these results are found in `CategoryTheory/Abelian/Exact.lean`. categories?) * Two adjacent maps in a chain complex are exact iff the homology vanishes +Note: It is planned that the definition in this file will be replaced by the new +homology API, in particular by the content of `Algebra.Homology.ShortComplex.Exact`. + -/ @@ -84,22 +87,22 @@ open ZeroObject /-- In any preadditive category, composable morphisms `f g` are exact iff they compose to zero and the homology vanishes. -/ -theorem Preadditive.exact_iff_homology_zero {A B C : V} (f : A ⟶ B) (g : B ⟶ C) : - Exact f g ↔ ∃ w : f ≫ g = 0, Nonempty (homology f g w ≅ 0) := +theorem Preadditive.exact_iff_homology'_zero {A B C : V} (f : A ⟶ B) (g : B ⟶ C) : + Exact f g ↔ ∃ w : f ≫ g = 0, Nonempty (homology' f g w ≅ 0) := ⟨fun h => ⟨h.w, ⟨by haveI := h.epi exact cokernel.ofEpi _⟩⟩, fun h => by obtain ⟨w, ⟨i⟩⟩ := h exact ⟨w, Preadditive.epi_of_cokernel_zero ((cancel_mono i.hom).mp (by ext))⟩⟩ -#align category_theory.preadditive.exact_iff_homology_zero CategoryTheory.Preadditive.exact_iff_homology_zero +#align category_theory.preadditive.exact_iff_homology_zero CategoryTheory.Preadditive.exact_iff_homology'_zero theorem Preadditive.exact_of_iso_of_exact {A₁ B₁ C₁ A₂ B₂ C₂ : V} (f₁ : A₁ ⟶ B₁) (g₁ : B₁ ⟶ C₁) (f₂ : A₂ ⟶ B₂) (g₂ : B₂ ⟶ C₂) (α : Arrow.mk f₁ ≅ Arrow.mk f₂) (β : Arrow.mk g₁ ≅ Arrow.mk g₂) (p : α.hom.right = β.hom.left) (h : Exact f₁ g₁) : Exact f₂ g₂ := by - rw [Preadditive.exact_iff_homology_zero] at h ⊢ + rw [Preadditive.exact_iff_homology'_zero] at h ⊢ rcases h with ⟨w₁, ⟨i⟩⟩ - suffices w₂ : f₂ ≫ g₂ = 0; exact ⟨w₂, ⟨(homology.mapIso w₁ w₂ α β p).symm.trans i⟩⟩ + suffices w₂ : f₂ ≫ g₂ = 0; exact ⟨w₂, ⟨(homology'.mapIso w₁ w₂ α β p).symm.trans i⟩⟩ rw [← cancel_epi α.hom.left, ← cancel_mono β.inv.right, comp_zero, zero_comp, ← w₁] have eq₁ := β.inv.w have eq₂ := α.hom.w @@ -147,7 +150,7 @@ theorem imageToKernel_isIso_of_image_eq_kernel {A B C : V} (f : A ⟶ B) (g : B IsIso (imageToKernel f g (comp_eq_zero_of_image_eq_kernel f g p)) := by refine' ⟨⟨Subobject.ofLE _ _ p.ge, _⟩⟩ dsimp [imageToKernel] - simp only [Subobject.ofLE_comp_ofLE, Subobject.ofLE_refl] + simp only [Subobject.ofLE_comp_ofLE, Subobject.ofLE_refl, and_self] #align category_theory.image_to_kernel_is_iso_of_image_eq_kernel CategoryTheory.imageToKernel_isIso_of_image_eq_kernel -- We'll prove the converse later, when `V` is abelian. diff --git a/Mathlib/Algebra/Homology/Flip.lean b/Mathlib/Algebra/Homology/Flip.lean index 678caa3a87780..8e4ad84660353 100644 --- a/Mathlib/Algebra/Homology/Flip.lean +++ b/Mathlib/Algebra/Homology/Flip.lean @@ -39,7 +39,7 @@ def flipObj (C : HomologicalComplex (HomologicalComplex V c) c') : { X := fun j => (C.X j).X i d := fun j j' => (C.d j j').f i shape := fun j j' w => by - simp_all only [shape, zero_f] + simp_all only [not_false_eq_true, shape, zero_f] d_comp_d' := fun j₁ j₂ j₃ _ _ => congr_hom (C.d_comp_d j₁ j₂ j₃) i } d i i' := { f := fun j => (C.X j).d i i' diff --git a/Mathlib/Algebra/Homology/HomologicalComplex.lean b/Mathlib/Algebra/Homology/HomologicalComplex.lean index 947cf4ac17361..9a4e4617b6176 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplex.lean @@ -803,10 +803,10 @@ theorem mk_d_1_0 : (mk X₀ X₁ X₂ d₀ d₁ s succ).d 1 0 = d₀ := by #align chain_complex.mk_d_1_0 ChainComplex.mk_d_1_0 @[simp] -theorem mk_d_2_0 : (mk X₀ X₁ X₂ d₀ d₁ s succ).d 2 1 = d₁ := by +theorem mk_d_2_1 : (mk X₀ X₁ X₂ d₀ d₁ s succ).d 2 1 = d₁ := by change ite (2 = 1 + 1) (𝟙 X₂ ≫ d₁) 0 = d₁ rw [if_pos rfl, Category.id_comp] -#align chain_complex.mk_d_2_0 ChainComplex.mk_d_2_0 +#align chain_complex.mk_d_2_0 ChainComplex.mk_d_2_1 -- TODO simp lemmas for the inductive steps? It's not entirely clear that they are needed. /-- A simpler inductive constructor for `ℕ`-indexed chain complexes. diff --git a/Mathlib/Algebra/Homology/Homology.lean b/Mathlib/Algebra/Homology/Homology.lean index b22e4e18ade5d..5bd6b088a8ae0 100644 --- a/Mathlib/Algebra/Homology/Homology.lean +++ b/Mathlib/Algebra/Homology/Homology.lean @@ -12,14 +12,21 @@ import Mathlib.CategoryTheory.GradedObject /-! # The homology of a complex -Given `C : HomologicalComplex V c`, we have `C.cycles i` and `C.boundaries i`, +Given `C : HomologicalComplex V c`, we have `C.cycles' i` and `C.boundaries i`, both defined as subobjects of `C.X i`. We show these are functorial with respect to chain maps, -as `C.cyclesMap f i` and `C.boundariesMap f i`. +as `cyclesMap' f i` and `boundariesMap f i`. -As a consequence we construct `homologyFunctor i : HomologicalComplex V c ⥤ V`, +As a consequence we construct `homologyFunctor' i : HomologicalComplex V c ⥤ V`, computing the `i`-th homology. + +Note: Some definitions (specifically, names containing components `homology`, `cycles`) +in this file have the suffix `'` so as to allow the development of the +new homology API of homological complex (starting from +`Algebra.Homology.ShortComplex.HomologicalComplex`). It is planned that these definitions +shall be removed and replaced by the new API. + -/ @@ -44,22 +51,22 @@ section Cycles variable [HasKernels V] /-- The cycles at index `i`, as a subobject. -/ -abbrev cycles (i : ι) : Subobject (C.X i) := +abbrev cycles' (i : ι) : Subobject (C.X i) := kernelSubobject (C.dFrom i) -#align homological_complex.cycles HomologicalComplex.cycles +#align homological_complex.cycles HomologicalComplex.cycles' -theorem cycles_eq_kernelSubobject {i j : ι} (r : c.Rel i j) : - C.cycles i = kernelSubobject (C.d i j) := +theorem cycles'_eq_kernelSubobject {i j : ι} (r : c.Rel i j) : + C.cycles' i = kernelSubobject (C.d i j) := C.kernel_from_eq_kernel r -#align homological_complex.cycles_eq_kernel_subobject HomologicalComplex.cycles_eq_kernelSubobject +#align homological_complex.cycles_eq_kernel_subobject HomologicalComplex.cycles'_eq_kernelSubobject -/-- The underlying object of `C.cycles i` is isomorphic to `kernel (C.d i j)`, +/-- The underlying object of `C.cycles' i` is isomorphic to `kernel (C.d i j)`, for any `j` such that `Rel i j`. -/ -def cyclesIsoKernel {i j : ι} (r : c.Rel i j) : (C.cycles i : V) ≅ kernel (C.d i j) := - Subobject.isoOfEq _ _ (C.cycles_eq_kernelSubobject r) ≪≫ kernelSubobjectIso (C.d i j) -#align homological_complex.cycles_iso_kernel HomologicalComplex.cyclesIsoKernel +def cycles'IsoKernel {i j : ι} (r : c.Rel i j) : (C.cycles' i : V) ≅ kernel (C.d i j) := + Subobject.isoOfEq _ _ (C.cycles'_eq_kernelSubobject r) ≪≫ kernelSubobjectIso (C.d i j) +#align homological_complex.cycles_iso_kernel HomologicalComplex.cycles'IsoKernel -theorem cycles_eq_top {i} (h : ¬c.Rel i (c.next i)) : C.cycles i = ⊤ := by +theorem cycles_eq_top {i} (h : ¬c.Rel i (c.next i)) : C.cycles' i = ⊤ := by rw [eq_top_iff] apply le_kernelSubobject rw [C.dFrom_eq_zero h, comp_zero] @@ -100,50 +107,52 @@ section variable [HasKernels V] [HasImages V] -theorem boundaries_le_cycles (C : HomologicalComplex V c) (i : ι) : C.boundaries i ≤ C.cycles i := +theorem boundaries_le_cycles' (C : HomologicalComplex V c) (i : ι) : + C.boundaries i ≤ C.cycles' i := image_le_kernel _ _ (C.dTo_comp_dFrom i) -#align homological_complex.boundaries_le_cycles HomologicalComplex.boundaries_le_cycles +#align homological_complex.boundaries_le_cycles HomologicalComplex.boundaries_le_cycles' -/-- The canonical map from `boundaries i` to `cycles i`. -/ -abbrev boundariesToCycles (C : HomologicalComplex V c) (i : ι) : - (C.boundaries i : V) ⟶ (C.cycles i : V) := +/-- The canonical map from `boundaries i` to `cycles' i`. -/ +abbrev boundariesToCycles' (C : HomologicalComplex V c) (i : ι) : + (C.boundaries i : V) ⟶ (C.cycles' i : V) := imageToKernel _ _ (C.dTo_comp_dFrom i) -#align homological_complex.boundaries_to_cycles HomologicalComplex.boundariesToCycles +#align homological_complex.boundaries_to_cycles HomologicalComplex.boundariesToCycles' -/-- Prefer `boundariesToCycles`. -/ +/-- Prefer `boundariesToCycles'`. -/ @[simp 1100] -theorem imageToKernel_as_boundariesToCycles (C : HomologicalComplex V c) (i : ι) (h) : - (C.boundaries i).ofLE (C.cycles i) h = C.boundariesToCycles i := rfl -#align homological_complex.image_to_kernel_as_boundaries_to_cycles HomologicalComplex.imageToKernel_as_boundariesToCycles +theorem imageToKernel_as_boundariesToCycles' (C : HomologicalComplex V c) (i : ι) (h) : + (C.boundaries i).ofLE (C.cycles' i) h = C.boundariesToCycles' i := rfl +#align homological_complex.image_to_kernel_as_boundaries_to_cycles HomologicalComplex.imageToKernel_as_boundariesToCycles' variable [HasCokernels V] /-- The homology of a complex at index `i`. -/ -abbrev homology (C : HomologicalComplex V c) (i : ι) : V := - _root_.homology (C.dTo i) (C.dFrom i) (C.dTo_comp_dFrom i) -#align homological_complex.homology HomologicalComplex.homology +abbrev homology' (C : HomologicalComplex V c) (i : ι) : V := + _root_.homology' (C.dTo i) (C.dFrom i) (C.dTo_comp_dFrom i) +#align homological_complex.homology HomologicalComplex.homology' /-- The `j`th homology of a homological complex (as kernel of 'the differential from `Cⱼ`' modulo the image of 'the differential to `Cⱼ`') is isomorphic to the kernel of `d : Cⱼ → Cₖ` modulo the image of `d : Cᵢ → Cⱼ` when `Rel i j` and `Rel j k`. -/ -def homologyIso (C : HomologicalComplex V c) {i j k : ι} (hij : c.Rel i j) (hjk : c.Rel j k) : - C.homology j ≅ _root_.homology (C.d i j) (C.d j k) (C.d_comp_d i j k) := - homology.mapIso _ _ +def homology'Iso (C : HomologicalComplex V c) {i j k : ι} (hij : c.Rel i j) (hjk : c.Rel j k) : + C.homology' j ≅ _root_.homology' (C.d i j) (C.d j k) (C.d_comp_d i j k) := + homology'.mapIso _ _ (Arrow.isoMk (C.xPrevIso hij) (Iso.refl _) <| by dsimp; rw [C.dTo_eq hij, Category.comp_id]) (Arrow.isoMk (Iso.refl _) (C.xNextIso hjk) <| by dsimp rw [C.dFrom_comp_xNextIso hjk, Category.id_comp]) rfl -#align homological_complex.homology_iso HomologicalComplex.homologyIso +#align homological_complex.homology_iso HomologicalComplex.homology'Iso end end HomologicalComplex /-- The 0th homology of a chain complex is isomorphic to the cokernel of `d : C₁ ⟶ C₀`. -/ -def ChainComplex.homologyZeroIso [HasKernels V] [HasImages V] [HasCokernels V] - (C : ChainComplex V ℕ) [Epi (factorThruImage (C.d 1 0))] : C.homology 0 ≅ cokernel (C.d 1 0) := - (homology.mapIso _ _ +def ChainComplex.homology'ZeroIso [HasKernels V] [HasImages V] [HasCokernels V] + (C : ChainComplex V ℕ) [Epi (factorThruImage (C.d 1 0))] : + C.homology' 0 ≅ cokernel (C.d 1 0) := + (homology'.mapIso _ _ (Arrow.isoMk (C.xPrevIso rfl) (Iso.refl _) <| by rw [C.dTo_eq rfl] exact (Category.comp_id _).symm : Arrow.mk (C.dTo 0) ≅ Arrow.mk (C.d 1 0)) @@ -152,38 +161,38 @@ def ChainComplex.homologyZeroIso [HasKernels V] [HasImages V] [HasCokernels V] one_ne_zero <| by rwa [ChainComplex.next_nat_zero, Nat.zero_add] at h] : Arrow.mk (C.dFrom 0) ≅ Arrow.mk 0) rfl).trans <| - homologyOfZeroRight _ -#align chain_complex.homology_zero_iso ChainComplex.homologyZeroIso + homology'OfZeroRight _ +#align chain_complex.homology_zero_iso ChainComplex.homology'ZeroIso /-- The 0th cohomology of a cochain complex is isomorphic to the kernel of `d : C₀ → C₁`. -/ -def CochainComplex.homologyZeroIso [HasZeroObject V] [HasKernels V] [HasImages V] [HasCokernels V] - (C : CochainComplex V ℕ) : C.homology 0 ≅ kernel (C.d 0 1) := - (homology.mapIso _ _ +def CochainComplex.homology'ZeroIso [HasZeroObject V] [HasKernels V] [HasImages V] [HasCokernels V] + (C : CochainComplex V ℕ) : C.homology' 0 ≅ kernel (C.d 0 1) := + (homology'.mapIso _ _ (Arrow.isoMk (C.xPrevIsoSelf (by rw [CochainComplex.prev_nat_zero]; exact one_ne_zero)) (Iso.refl _) (by simp) : Arrow.mk (C.dTo 0) ≅ Arrow.mk 0) (Arrow.isoMk (Iso.refl _) (C.xNextIso rfl) (by simp) : Arrow.mk (C.dFrom 0) ≅ Arrow.mk (C.d 0 1)) <| by simp).trans <| - homologyOfZeroLeft _ -#align cochain_complex.homology_zero_iso CochainComplex.homologyZeroIso + homology'OfZeroLeft _ +#align cochain_complex.homology_zero_iso CochainComplex.homology'ZeroIso /-- The `n + 1`th homology of a chain complex (as kernel of 'the differential from `Cₙ₊₁`' modulo the image of 'the differential to `Cₙ₊₁`') is isomorphic to the kernel of `d : Cₙ₊₁ → Cₙ` modulo the image of `d : Cₙ₊₂ → Cₙ₊₁`. -/ -def ChainComplex.homologySuccIso [HasKernels V] [HasImages V] [HasCokernels V] +def ChainComplex.homology'SuccIso [HasKernels V] [HasImages V] [HasCokernels V] (C : ChainComplex V ℕ) (n : ℕ) : - C.homology (n + 1) ≅ homology (C.d (n + 2) (n + 1)) (C.d (n + 1) n) (C.d_comp_d _ _ _) := - C.homologyIso rfl rfl -#align chain_complex.homology_succ_iso ChainComplex.homologySuccIso + C.homology' (n + 1) ≅ homology' (C.d (n + 2) (n + 1)) (C.d (n + 1) n) (C.d_comp_d _ _ _) := + C.homology'Iso rfl rfl +#align chain_complex.homology_succ_iso ChainComplex.homology'SuccIso /-- The `n + 1`th cohomology of a cochain complex (as kernel of 'the differential from `Cₙ₊₁`' modulo the image of 'the differential to `Cₙ₊₁`') is isomorphic to the kernel of `d : Cₙ₊₁ → Cₙ₊₂` modulo the image of `d : Cₙ → Cₙ₊₁`. -/ -def CochainComplex.homologySuccIso [HasKernels V] [HasImages V] [HasCokernels V] +def CochainComplex.homology'SuccIso [HasKernels V] [HasImages V] [HasCokernels V] (C : CochainComplex V ℕ) (n : ℕ) : - C.homology (n + 1) ≅ homology (C.d n (n + 1)) (C.d (n + 1) (n + 2)) (C.d_comp_d _ _ _) := - C.homologyIso rfl rfl -#align cochain_complex.homology_succ_iso CochainComplex.homologySuccIso + C.homology' (n + 1) ≅ homology' (C.d n (n + 1)) (C.d (n + 1) (n + 2)) (C.d_comp_d _ _ _) := + C.homology'Iso rfl rfl +#align cochain_complex.homology_succ_iso CochainComplex.homology'SuccIso open HomologicalComplex @@ -197,40 +206,40 @@ variable [HasKernels V] variable {C₁ C₂ C₃ : HomologicalComplex V c} (f : C₁ ⟶ C₂) /-- The morphism between cycles induced by a chain map. -/ -abbrev cyclesMap (f : C₁ ⟶ C₂) (i : ι) : (C₁.cycles i : V) ⟶ (C₂.cycles i : V) := - Subobject.factorThru _ ((C₁.cycles i).arrow ≫ f.f i) (kernelSubobject_factors _ _ (by simp)) -#align cycles_map cyclesMap +abbrev cycles'Map (f : C₁ ⟶ C₂) (i : ι) : (C₁.cycles' i : V) ⟶ (C₂.cycles' i : V) := + Subobject.factorThru _ ((C₁.cycles' i).arrow ≫ f.f i) (kernelSubobject_factors _ _ (by simp)) +#align cycles_map cycles'Map -- Porting note: Originally `@[simp, reassoc.1, elementwise]` @[reassoc, elementwise] -- @[simp] -- Porting note: simp can prove this -theorem cyclesMap_arrow (f : C₁ ⟶ C₂) (i : ι) : - cyclesMap f i ≫ (C₂.cycles i).arrow = (C₁.cycles i).arrow ≫ f.f i := by simp -#align cycles_map_arrow cyclesMap_arrow +theorem cycles'Map_arrow (f : C₁ ⟶ C₂) (i : ι) : + cycles'Map f i ≫ (C₂.cycles' i).arrow = (C₁.cycles' i).arrow ≫ f.f i := by simp +#align cycles_map_arrow cycles'Map_arrow -attribute [simp 1100] cyclesMap_arrow_assoc -attribute [simp] cyclesMap_arrow_apply +attribute [simp 1100] cycles'Map_arrow_assoc +attribute [simp] cycles'Map_arrow_apply @[simp] -theorem cyclesMap_id (i : ι) : cyclesMap (𝟙 C₁) i = 𝟙 _ := by - dsimp only [cyclesMap] +theorem cycles'Map_id (i : ι) : cycles'Map (𝟙 C₁) i = 𝟙 _ := by + dsimp only [cycles'Map] simp -#align cycles_map_id cyclesMap_id +#align cycles_map_id cycles'Map_id @[simp] -theorem cyclesMap_comp (f : C₁ ⟶ C₂) (g : C₂ ⟶ C₃) (i : ι) : - cyclesMap (f ≫ g) i = cyclesMap f i ≫ cyclesMap g i := by - dsimp only [cyclesMap] +theorem cycles'Map_comp (f : C₁ ⟶ C₂) (g : C₂ ⟶ C₃) (i : ι) : + cycles'Map (f ≫ g) i = cycles'Map f i ≫ cycles'Map g i := by + dsimp only [cycles'Map] simp [Subobject.factorThru_right] -#align cycles_map_comp cyclesMap_comp +#align cycles_map_comp cycles'Map_comp variable (V c) /-- Cycles as a functor. -/ @[simps] -def cyclesFunctor (i : ι) : HomologicalComplex V c ⥤ V where - obj C := C.cycles i - map {C₁ C₂} f := cyclesMap f i -#align cycles_functor cyclesFunctor +def cycles'Functor (i : ι) : HomologicalComplex V c ⥤ V where + obj C := C.cycles' i + map {C₁ C₂} f := cycles'Map f i +#align cycles_functor cycles'Functor end @@ -270,58 +279,37 @@ variable {C₁ C₂ : HomologicalComplex V c} (f : C₁ ⟶ C₂) -- Porting note: Originally `@[simp, reassoc.1]` @[reassoc (attr := simp)] -theorem boundariesToCycles_naturality (i : ι) : - boundariesMap f i ≫ C₂.boundariesToCycles i = C₁.boundariesToCycles i ≫ cyclesMap f i := by +theorem boundariesToCycles'_naturality (i : ι) : + boundariesMap f i ≫ C₂.boundariesToCycles' i = + C₁.boundariesToCycles' i ≫ cycles'Map f i := by ext simp -#align boundaries_to_cycles_naturality boundariesToCycles_naturality +#align boundaries_to_cycles_naturality boundariesToCycles'_naturality variable (V c) /-- The natural transformation from the boundaries functor to the cycles functor. -/ @[simps] -def boundariesToCyclesNatTrans (i : ι) : boundariesFunctor V c i ⟶ cyclesFunctor V c i where - app C := C.boundariesToCycles i - naturality _ _ f := boundariesToCycles_naturality f i -#align boundaries_to_cycles_nat_trans boundariesToCyclesNatTrans +def boundariesToCycles'NatTrans (i : ι) : boundariesFunctor V c i ⟶ cycles'Functor V c i where + app C := C.boundariesToCycles' i + naturality _ _ f := boundariesToCycles'_naturality f i +#align boundaries_to_cycles_nat_trans boundariesToCycles'NatTrans /-- The `i`-th homology, as a functor to `V`. -/ @[simps] -def homologyFunctor [HasCokernels V] (i : ι) : HomologicalComplex V c ⥤ V where +def homology'Functor [HasCokernels V] (i : ι) : HomologicalComplex V c ⥤ V where -- It would be nice if we could just write -- `cokernel (boundariesToCyclesNatTrans V c i)` -- here, but universe implementation details get in the way... - obj C := C.homology i - map {C₁ C₂} f := homology.map _ _ (f.sqTo i) (f.sqFrom i) rfl - map_id _ := by - simp only - ext1 - simp only [homology.π_map, kernelSubobjectMap_id, Hom.sqFrom_id, Category.id_comp, - Category.comp_id] - map_comp _ _ := by - simp only - ext1 - simp only [Hom.sqFrom_comp, kernelSubobjectMap_comp, homology.π_map_assoc, homology.π_map, - Category.assoc] -#align homology_functor homologyFunctor + obj C := C.homology' i + map {C₁ C₂} f := homology'.map _ _ (f.sqTo i) (f.sqFrom i) rfl +#align homology_functor homology'Functor /-- The homology functor from `ι`-indexed complexes to `ι`-graded objects in `V`. -/ @[simps] -def gradedHomologyFunctor [HasCokernels V] : HomologicalComplex V c ⥤ GradedObject ι V where - obj C i := C.homology i - map {C C'} f i := (homologyFunctor V c i).map f - map_id _ := by - ext - simp only [GradedObject.categoryOfGradedObjects_id] - ext - simp only [homology.π_map, homologyFunctor_map, kernelSubobjectMap_id, Hom.sqFrom_id, - Category.id_comp, Category.comp_id] - map_comp _ _ := by - ext - simp only [GradedObject.categoryOfGradedObjects_comp] - ext - simp only [Hom.sqFrom_comp, kernelSubobjectMap_comp, homology.π_map_assoc, homology.π_map, - homologyFunctor_map, Category.assoc] -#align graded_homology_functor gradedHomologyFunctor +def gradedHomology'Functor [HasCokernels V] : HomologicalComplex V c ⥤ GradedObject ι V where + obj C i := C.homology' i + map {C C'} f i := (homology'Functor V c i).map f +#align graded_homology_functor gradedHomology'Functor end diff --git a/Mathlib/Algebra/Homology/Homotopy.lean b/Mathlib/Algebra/Homology/Homotopy.lean index ac54a9a3487e8..a0a7d0faf0a25 100644 --- a/Mathlib/Algebra/Homology/Homotopy.lean +++ b/Mathlib/Algebra/Homology/Homotopy.lean @@ -565,7 +565,7 @@ def mkInductive : Homotopy e 0 where · cases i · dsimp [fromNext, mkInductiveAux₂] rw [dif_neg] - simp only + decide · dsimp [fromNext] simp only [ChainComplex.next_nat_succ, dite_true] rw [mkInductiveAux₃ e zero comm_zero one comm_one succ] @@ -704,7 +704,7 @@ def mkCoinductive : Homotopy e 0 where · cases i · dsimp [toPrev, mkCoinductiveAux₂] rw [dif_neg] - simp only + decide · dsimp [toPrev] simp only [CochainComplex.prev_nat_succ, dite_true] rw [mkCoinductiveAux₃ e zero comm_zero one comm_one succ] @@ -781,34 +781,34 @@ variable [HasEqualizers V] [HasCokernels V] [HasImages V] [HasImageMaps V] /-- Homotopic maps induce the same map on homology. -/ -theorem homology_map_eq_of_homotopy (h : Homotopy f g) (i : ι) : - (homologyFunctor V c i).map f = (homologyFunctor V c i).map g := by - dsimp [homologyFunctor] +theorem homology'_map_eq_of_homotopy (h : Homotopy f g) (i : ι) : + (homology'Functor V c i).map f = (homology'Functor V c i).map g := by + dsimp [homology'Functor] apply eq_of_sub_eq_zero ext - simp only [homology.π_map, comp_zero, Preadditive.comp_sub] + simp only [homology'.π_map, comp_zero, Preadditive.comp_sub] dsimp [kernelSubobjectMap] simp_rw [h.comm i] simp only [zero_add, zero_comp, dNext_eq_dFrom_fromNext, kernelSubobject_arrow_comp_assoc, Preadditive.comp_add] rw [← Preadditive.sub_comp] simp only [CategoryTheory.Subobject.factorThru_add_sub_factorThru_right] - erw [Subobject.factorThru_ofLE (D.boundaries_le_cycles i)] + erw [Subobject.factorThru_ofLE (D.boundaries_le_cycles' i)] · simp · rw [prevD_eq_toPrev_dTo, ← Category.assoc] apply imageSubobject_factors_comp_self -#align homology_map_eq_of_homotopy homology_map_eq_of_homotopy +#align homology_map_eq_of_homotopy homology'_map_eq_of_homotopy /-- Homotopy equivalent complexes have isomorphic homologies. -/ def homologyObjIsoOfHomotopyEquiv (f : HomotopyEquiv C D) (i : ι) : - (homologyFunctor V c i).obj C ≅ (homologyFunctor V c i).obj D where - hom := (homologyFunctor V c i).map f.hom - inv := (homologyFunctor V c i).map f.inv + (homology'Functor V c i).obj C ≅ (homology'Functor V c i).obj D where + hom := (homology'Functor V c i).map f.hom + inv := (homology'Functor V c i).map f.inv hom_inv_id := by - rw [← Functor.map_comp, homology_map_eq_of_homotopy f.homotopyHomInvId, + rw [← Functor.map_comp, homology'_map_eq_of_homotopy f.homotopyHomInvId, CategoryTheory.Functor.map_id] inv_hom_id := by - rw [← Functor.map_comp, homology_map_eq_of_homotopy f.homotopyInvHomId, + rw [← Functor.map_comp, homology'_map_eq_of_homotopy f.homotopyInvHomId, CategoryTheory.Functor.map_id] #align homology_obj_iso_of_homotopy_equiv homologyObjIsoOfHomotopyEquiv diff --git a/Mathlib/Algebra/Homology/HomotopyCategory.lean b/Mathlib/Algebra/Homology/HomotopyCategory.lean index ce6aae6ceb9a4..03c4a76e79ac8 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory.lean @@ -144,34 +144,34 @@ variable (V c) variable [HasEqualizers V] [HasImages V] [HasImageMaps V] [HasCokernels V] /-- The `i`-th homology, as a functor from the homotopy category. -/ -def homologyFunctor (i : ι) : HomotopyCategory V c ⥤ V := - CategoryTheory.Quotient.lift _ (_root_.homologyFunctor V c i) fun _ _ _ _ ⟨h⟩ => - homology_map_eq_of_homotopy h i -#align homotopy_category.homology_functor HomotopyCategory.homologyFunctor +def homology'Functor (i : ι) : HomotopyCategory V c ⥤ V := + CategoryTheory.Quotient.lift _ (_root_.homology'Functor V c i) fun _ _ _ _ ⟨h⟩ => + homology'_map_eq_of_homotopy h i +#align homotopy_category.homology_functor HomotopyCategory.homology'Functor /-- The homology functor on the homotopy category is just the usual homology functor. -/ -def homologyFactors (i : ι) : - quotient V c ⋙ homologyFunctor V c i ≅ _root_.homologyFunctor V c i := +def homology'Factors (i : ι) : + quotient V c ⋙ homology'Functor V c i ≅ _root_.homology'Functor V c i := CategoryTheory.Quotient.lift.isLift _ _ _ -#align homotopy_category.homology_factors HomotopyCategory.homologyFactors +#align homotopy_category.homology_factors HomotopyCategory.homology'Factors @[simp] -theorem homologyFactors_hom_app (i : ι) (C : HomologicalComplex V c) : - (homologyFactors V c i).hom.app C = 𝟙 _ := +theorem homology'Factors_hom_app (i : ι) (C : HomologicalComplex V c) : + (homology'Factors V c i).hom.app C = 𝟙 _ := rfl -#align homotopy_category.homology_factors_hom_app HomotopyCategory.homologyFactors_hom_app +#align homotopy_category.homology_factors_hom_app HomotopyCategory.homology'Factors_hom_app @[simp] -theorem homologyFactors_inv_app (i : ι) (C : HomologicalComplex V c) : - (homologyFactors V c i).inv.app C = 𝟙 _ := +theorem homology'Factors_inv_app (i : ι) (C : HomologicalComplex V c) : + (homology'Factors V c i).inv.app C = 𝟙 _ := rfl -#align homotopy_category.homology_factors_inv_app HomotopyCategory.homologyFactors_inv_app +#align homotopy_category.homology_factors_inv_app HomotopyCategory.homology'Factors_inv_app -theorem homologyFunctor_map_factors (i : ι) {C D : HomologicalComplex V c} (f : C ⟶ D) : - (_root_.homologyFunctor V c i).map f = - ((homologyFunctor V c i).map ((quotient V c).map f) : _) := - (CategoryTheory.Quotient.lift_map_functor_map _ (_root_.homologyFunctor V c i) _ f).symm -#align homotopy_category.homology_functor_map_factors HomotopyCategory.homologyFunctor_map_factors +theorem homology'Functor_map_factors (i : ι) {C D : HomologicalComplex V c} (f : C ⟶ D) : + (_root_.homology'Functor V c i).map f = + ((homology'Functor V c i).map ((quotient V c).map f) : _) := + (CategoryTheory.Quotient.lift_map_functor_map _ (_root_.homology'Functor V c i) _ f).symm +#align homotopy_category.homology_functor_map_factors HomotopyCategory.homology'Functor_map_factors end HomotopyCategory diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Shift.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Shift.lean index 234e7c07c8996..187a9855b8d84 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Shift.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Shift.lean @@ -80,7 +80,7 @@ def shiftFunctorZero' (n : ℤ) (h : n = 0) : /-- The compatibility of the shift functors on `CochainComplex C ℤ` with respect to the addition of integers. -/ @[simps!] -def shiftFunctorAdd' (n₁ n₂ n₁₂ : ℤ) (h : n₁ + n₂ = n₁₂ ) : +def shiftFunctorAdd' (n₁ n₂ n₁₂ : ℤ) (h : n₁ + n₂ = n₁₂) : shiftFunctor C n₁₂ ≅ shiftFunctor C n₁ ⋙ shiftFunctor C n₂ := NatIso.ofComponents (fun K => Hom.isoOfComponents (fun i => K.shiftFunctorObjXIso _ _ _ (by linarith)) diff --git a/Mathlib/Algebra/Homology/ImageToKernel.lean b/Mathlib/Algebra/Homology/ImageToKernel.lean index 11f736382376d..6d38e1591b0f2 100644 --- a/Mathlib/Algebra/Homology/ImageToKernel.lean +++ b/Mathlib/Algebra/Homology/ImageToKernel.lean @@ -16,7 +16,12 @@ we have `image_le_kernel f g w : imageSubobject f ≤ kernelSubobject g` `imageToKernel f g w` is the corresponding morphism between objects in `C`. -We define `homology f g w` of such a pair as the cokernel of `imageToKernel f g w`. +We define `homology' f g w` of such a pair as the cokernel of `imageToKernel f g w`. + +Note: As part of the transition to the new homology API, `homology` is temporarily +renamed `homology'`. It is planned that this definition shall be removed and replaced by +`ShortComplex.homology`. + -/ set_option autoImplicit true @@ -181,72 +186,72 @@ variable {A B C : V} (f : A ⟶ B) [HasImage f] (g : B ⟶ C) [HasKernel g] /-- The homology of a pair of morphisms `f : A ⟶ B` and `g : B ⟶ C` satisfying `f ≫ g = 0` is the cokernel of the `imageToKernel` morphism for `f` and `g`. -/ -def homology {A B C : V} (f : A ⟶ B) [HasImage f] (g : B ⟶ C) [HasKernel g] (w : f ≫ g = 0) +def homology' {A B C : V} (f : A ⟶ B) [HasImage f] (g : B ⟶ C) [HasKernel g] (w : f ≫ g = 0) [HasCokernel (imageToKernel f g w)] : V := cokernel (imageToKernel f g w) -#align homology homology +#align homology homology' section variable (w : f ≫ g = 0) [HasCokernel (imageToKernel f g w)] /-- The morphism from cycles to homology. -/ -def homology.π : (kernelSubobject g : V) ⟶ homology f g w := +def homology'.π : (kernelSubobject g : V) ⟶ homology' f g w := cokernel.π _ -#align homology.π homology.π +#align homology.π homology'.π @[simp] -theorem homology.condition : imageToKernel f g w ≫ homology.π f g w = 0 := +theorem homology'.condition : imageToKernel f g w ≫ homology'.π f g w = 0 := cokernel.condition _ -#align homology.condition homology.condition +#align homology.condition homology'.condition /-- To construct a map out of homology, it suffices to construct a map out of the cycles which vanishes on boundaries. -/ -def homology.desc {D : V} (k : (kernelSubobject g : V) ⟶ D) (p : imageToKernel f g w ≫ k = 0) : - homology f g w ⟶ D := +def homology'.desc {D : V} (k : (kernelSubobject g : V) ⟶ D) (p : imageToKernel f g w ≫ k = 0) : + homology' f g w ⟶ D := cokernel.desc _ k p -#align homology.desc homology.desc +#align homology.desc homology'.desc -- porting note: removed elementwise attribute which does not seem to be helpful here @[reassoc (attr := simp)] -theorem homology.π_desc {D : V} (k : (kernelSubobject g : V) ⟶ D) - (p : imageToKernel f g w ≫ k = 0) : homology.π f g w ≫ homology.desc f g w k p = k := by - simp [homology.π, homology.desc] -#align homology.π_desc homology.π_desc +theorem homology'.π_desc {D : V} (k : (kernelSubobject g : V) ⟶ D) + (p : imageToKernel f g w ≫ k = 0) : homology'.π f g w ≫ homology'.desc f g w k p = k := by + simp [homology'.π, homology'.desc] +#align homology.π_desc homology'.π_desc -/-- To check two morphisms out of `homology f g w` are equal, it suffices to check on cycles. -/ +/-- To check two morphisms out of `homology' f g w` are equal, it suffices to check on cycles. -/ @[ext] -theorem homology.ext {D : V} {k k' : homology f g w ⟶ D} - (p : homology.π f g w ≫ k = homology.π f g w ≫ k') : k = k' := +theorem homology'.ext {D : V} {k k' : homology' f g w ⟶ D} + (p : homology'.π f g w ≫ k = homology'.π f g w ≫ k') : k = k' := coequalizer.hom_ext p -#align homology.ext homology.ext +#align homology.ext homology'.ext /-- The cokernel of the map `Im f ⟶ Ker 0` is isomorphic to the cokernel of `f.` -/ -def homologyOfZeroRight [HasCokernel (imageToKernel f (0 : B ⟶ C) comp_zero)] [HasCokernel f] +def homology'OfZeroRight [HasCokernel (imageToKernel f (0 : B ⟶ C) comp_zero)] [HasCokernel f] [HasCokernel (image.ι f)] [Epi (factorThruImage f)] : - homology f (0 : B ⟶ C) comp_zero ≅ cokernel f := + homology' f (0 : B ⟶ C) comp_zero ≅ cokernel f := (cokernel.mapIso _ _ (imageSubobjectIso _) ((kernelSubobjectIso 0).trans kernelZeroIsoSource) (by simp)).trans (cokernelImageι _) -#align homology_of_zero_right homologyOfZeroRight +#align homology_of_zero_right homology'OfZeroRight /-- The kernel of the map `Im 0 ⟶ Ker f` is isomorphic to the kernel of `f.` -/ -def homologyOfZeroLeft [HasZeroObject V] [HasKernels V] [HasImage (0 : A ⟶ B)] +def homology'OfZeroLeft [HasZeroObject V] [HasKernels V] [HasImage (0 : A ⟶ B)] [HasCokernel (imageToKernel (0 : A ⟶ B) g zero_comp)] : - homology (0 : A ⟶ B) g zero_comp ≅ kernel g := + homology' (0 : A ⟶ B) g zero_comp ≅ kernel g := ((cokernelIsoOfEq <| imageToKernel_zero_left _).trans cokernelZeroIsoTarget).trans (kernelSubobjectIso _) -#align homology_of_zero_left homologyOfZeroLeft +#align homology_of_zero_left homology'OfZeroLeft /-- `homology 0 0 _` is just the middle object. -/ @[simps] -def homologyZeroZero [HasZeroObject V] [HasImage (0 : A ⟶ B)] +def homology'ZeroZero [HasZeroObject V] [HasImage (0 : A ⟶ B)] [HasCokernel (imageToKernel (0 : A ⟶ B) (0 : B ⟶ C) zero_comp)] : - homology (0 : A ⟶ B) (0 : B ⟶ C) zero_comp ≅ B where - hom := homology.desc (0 : A ⟶ B) (0 : B ⟶ C) zero_comp (kernelSubobject 0).arrow (by simp) - inv := inv (kernelSubobject 0).arrow ≫ homology.π _ _ _ -#align homology_zero_zero homologyZeroZero + homology' (0 : A ⟶ B) (0 : B ⟶ C) zero_comp ≅ B where + hom := homology'.desc (0 : A ⟶ B) (0 : B ⟶ C) zero_comp (kernelSubobject 0).arrow (by simp) + inv := inv (kernelSubobject 0).arrow ≫ homology'.π _ _ _ +#align homology_zero_zero homology'ZeroZero end @@ -281,83 +286,84 @@ variable [HasCokernel (imageToKernel f₃ g₃ w₃)] a pair `f g` and a pair `f' g'` satisfying `f ≫ g = 0` and `f' ≫ g' = 0`, we get a morphism on homology. -/ -def homology.map (p : α.right = β.left) : homology f g w ⟶ homology f' g' w' := +def homology'.map (p : α.right = β.left) : homology' f g w ⟶ homology' f' g' w' := cokernel.desc _ (kernelSubobjectMap β ≫ cokernel.π _) <| by rw [imageSubobjectMap_comp_imageToKernel_assoc w w' α β p] simp only [cokernel.condition, comp_zero] -#align homology.map homology.map +#align homology.map homology'.map -- porting note: removed elementwise attribute which does not seem to be helpful here, -- the correct lemma is stated below @[reassoc (attr := simp)] -theorem homology.π_map (p : α.right = β.left) : - homology.π f g w ≫ homology.map w w' α β p = kernelSubobjectMap β ≫ homology.π f' g' w' := by - simp only [homology.π, homology.map, cokernel.π_desc] -#align homology.π_map homology.π_map +theorem homology'.π_map (p : α.right = β.left) : + homology'.π f g w ≫ homology'.map w w' α β p = + kernelSubobjectMap β ≫ homology'.π f' g' w' := by + simp only [homology'.π, homology'.map, cokernel.π_desc] +#align homology.π_map homology'.π_map section attribute [local instance] ConcreteCategory.funLike @[simp] -lemma homology.π_map_apply [ConcreteCategory.{w} V] (p : α.right = β.left) +lemma homology'.π_map_apply [ConcreteCategory.{w} V] (p : α.right = β.left) (x : (forget V).obj (Subobject.underlying.obj (kernelSubobject g))) : - homology.map w w' α β p (homology.π f g w x) = - homology.π f' g' w' (kernelSubobjectMap β x) := by - simp only [← comp_apply, homology.π_map w w' α β p] + homology'.map w w' α β p (homology'.π f g w x) = + homology'.π f' g' w' (kernelSubobjectMap β x) := by + simp only [← comp_apply, homology'.π_map w w' α β p] end @[reassoc (attr := simp), elementwise (attr := simp)] -theorem homology.map_desc (p : α.right = β.left) {D : V} (k : (kernelSubobject g' : V) ⟶ D) +theorem homology'.map_desc (p : α.right = β.left) {D : V} (k : (kernelSubobject g' : V) ⟶ D) (z : imageToKernel f' g' w' ≫ k = 0) : - homology.map w w' α β p ≫ homology.desc f' g' w' k z = - homology.desc f g w (kernelSubobjectMap β ≫ k) + homology'.map w w' α β p ≫ homology'.desc f' g' w' k z = + homology'.desc f g w (kernelSubobjectMap β ≫ k) (by simp only [imageSubobjectMap_comp_imageToKernel_assoc w w' α β p, z, comp_zero]) := by ext - simp only [homology.π_desc, homology.π_map_assoc] -#align homology.map_desc homology.map_desc + simp only [homology'.π_desc, homology'.π_map_assoc] +#align homology.map_desc homology'.map_desc @[simp] -theorem homology.map_id : homology.map w w (𝟙 _) (𝟙 _) rfl = 𝟙 _ := by +theorem homology'.map_id : homology'.map w w (𝟙 _) (𝟙 _) rfl = 𝟙 _ := by ext - simp only [homology.π_map, kernelSubobjectMap_id, Category.id_comp, Category.comp_id] -#align homology.map_id homology.map_id + simp only [homology'.π_map, kernelSubobjectMap_id, Category.id_comp, Category.comp_id] +#align homology.map_id homology'.map_id /-- Auxiliary lemma for homology computations. -/ -theorem homology.comp_right_eq_comp_left {V : Type*} [Category V] {A₁ B₁ C₁ A₂ B₂ C₂ A₃ B₃ C₃ : V} +theorem homology'.comp_right_eq_comp_left {V : Type*} [Category V] {A₁ B₁ C₁ A₂ B₂ C₂ A₃ B₃ C₃ : V} {f₁ : A₁ ⟶ B₁} {g₁ : B₁ ⟶ C₁} {f₂ : A₂ ⟶ B₂} {g₂ : B₂ ⟶ C₂} {f₃ : A₃ ⟶ B₃} {g₃ : B₃ ⟶ C₃} {α₁ : Arrow.mk f₁ ⟶ Arrow.mk f₂} {β₁ : Arrow.mk g₁ ⟶ Arrow.mk g₂} {α₂ : Arrow.mk f₂ ⟶ Arrow.mk f₃} {β₂ : Arrow.mk g₂ ⟶ Arrow.mk g₃} (p₁ : α₁.right = β₁.left) (p₂ : α₂.right = β₂.left) : (α₁ ≫ α₂).right = (β₁ ≫ β₂).left := by simp only [Arrow.comp_left, Arrow.comp_right, p₁, p₂] -#align homology.comp_right_eq_comp_left homology.comp_right_eq_comp_left +#align homology.comp_right_eq_comp_left homology'.comp_right_eq_comp_left @[reassoc] -theorem homology.map_comp (p₁ : α₁.right = β₁.left) (p₂ : α₂.right = β₂.left) : - homology.map w₁ w₂ α₁ β₁ p₁ ≫ homology.map w₂ w₃ α₂ β₂ p₂ = - homology.map w₁ w₃ (α₁ ≫ α₂) (β₁ ≫ β₂) (homology.comp_right_eq_comp_left p₁ p₂) := by +theorem homology'.map_comp (p₁ : α₁.right = β₁.left) (p₂ : α₂.right = β₂.left) : + homology'.map w₁ w₂ α₁ β₁ p₁ ≫ homology'.map w₂ w₃ α₂ β₂ p₂ = + homology'.map w₁ w₃ (α₁ ≫ α₂) (β₁ ≫ β₂) (homology'.comp_right_eq_comp_left p₁ p₂) := by ext - simp only [kernelSubobjectMap_comp, homology.π_map_assoc, homology.π_map, Category.assoc] -#align homology.map_comp homology.map_comp + simp only [kernelSubobjectMap_comp, homology'.π_map_assoc, homology'.π_map, Category.assoc] +#align homology.map_comp homology'.map_comp /-- An isomorphism between two three-term complexes induces an isomorphism on homology. -/ -def homology.mapIso (α : Arrow.mk f₁ ≅ Arrow.mk f₂) (β : Arrow.mk g₁ ≅ Arrow.mk g₂) - (p : α.hom.right = β.hom.left) : homology f₁ g₁ w₁ ≅ homology f₂ g₂ w₂ where - hom := homology.map w₁ w₂ α.hom β.hom p +def homology'.mapIso (α : Arrow.mk f₁ ≅ Arrow.mk f₂) (β : Arrow.mk g₁ ≅ Arrow.mk g₂) + (p : α.hom.right = β.hom.left) : homology' f₁ g₁ w₁ ≅ homology' f₂ g₂ w₂ where + hom := homology'.map w₁ w₂ α.hom β.hom p inv := - homology.map w₂ w₁ α.inv β.inv + homology'.map w₂ w₁ α.inv β.inv (by rw [← cancel_mono α.hom.right, ← Comma.comp_right, α.inv_hom_id, Comma.id_right, p, ← Comma.comp_left, β.inv_hom_id, Comma.id_left] rfl) hom_inv_id := by - rw [homology.map_comp, ← homology.map_id] + rw [homology'.map_comp, ← homology'.map_id] congr <;> simp only [Iso.hom_inv_id] inv_hom_id := by - rw [homology.map_comp, ← homology.map_id] + rw [homology'.map_comp, ← homology'.map_id] congr <;> simp only [Iso.inv_hom_id] -#align homology.map_iso homology.mapIso +#align homology.map_iso homology'.mapIso end @@ -377,20 +383,20 @@ variable {A B C : V} {f : A ⟶ B} {g : B ⟶ C} (w : f ≫ g = 0) {f' : A ⟶ B (Note the objects are not changing here.) -/ @[simps] -def homology.congr (pf : f = f') (pg : g = g') : homology f g w ≅ homology f' g' w' where - hom := homology.map w w' ⟨𝟙 _, 𝟙 _, by aesop_cat⟩ ⟨𝟙 _, 𝟙 _, by aesop_cat⟩ rfl - inv := homology.map w' w ⟨𝟙 _, 𝟙 _, by aesop_cat⟩ ⟨𝟙 _, 𝟙 _, by aesop_cat⟩ rfl +def homology'.congr (pf : f = f') (pg : g = g') : homology' f g w ≅ homology' f' g' w' where + hom := homology'.map w w' ⟨𝟙 _, 𝟙 _, by aesop_cat⟩ ⟨𝟙 _, 𝟙 _, by aesop_cat⟩ rfl + inv := homology'.map w' w ⟨𝟙 _, 𝟙 _, by aesop_cat⟩ ⟨𝟙 _, 𝟙 _, by aesop_cat⟩ rfl hom_inv_id := by obtain rfl := pf obtain rfl := pg - rw [homology.map_comp, ← homology.map_id] + rw [homology'.map_comp, ← homology'.map_id] congr <;> aesop_cat inv_hom_id := by obtain rfl := pf obtain rfl := pg - rw [homology.map_comp, ← homology.map_id] + rw [homology'.map_comp, ← homology'.map_id] congr <;> aesop_cat -#align homology.congr homology.congr +#align homology.congr homology'.congr end @@ -435,17 +441,17 @@ theorem imageToKernel'_kernelSubobjectIso (w : f ≫ g = 0) : variable [HasCokernels V] -/-- `homology f g w` can be computed as the cokernel of `imageToKernel' f g w`. +/-- `homology' f g w` can be computed as the cokernel of `imageToKernel' f g w`. -/ -def homologyIsoCokernelImageToKernel' (w : f ≫ g = 0) : - homology f g w ≅ cokernel (imageToKernel' f g w) where +def homology'IsoCokernelImageToKernel' (w : f ≫ g = 0) : + homology' f g w ≅ cokernel (imageToKernel' f g w) where hom := cokernel.map _ _ (imageSubobjectIso f).hom (kernelSubobjectIso g).hom (by simp only [imageSubobjectIso_imageToKernel']) inv := cokernel.map _ _ (imageSubobjectIso f).inv (kernelSubobjectIso g).inv (by simp only [imageToKernel'_kernelSubobjectIso]) hom_inv_id := by - -- Just calling `ext` here uses the higher priority `homology.ext`, - -- which precomposes with `homology.π`. + -- Just calling `ext` here uses the higher priority `homology'.ext`, + -- which precomposes with `homology'.π`. -- As we are trying to work in terms of `cokernel`, it is better to use `coequalizer.hom_ext`. apply coequalizer.hom_ext simp only [Iso.hom_inv_id_assoc, cokernel.π_desc, cokernel.π_desc_assoc, Category.assoc, @@ -455,18 +461,18 @@ def homologyIsoCokernelImageToKernel' (w : f ≫ g = 0) : ext simp only [Iso.inv_hom_id_assoc, cokernel.π_desc, Category.comp_id, cokernel.π_desc_assoc, Category.assoc] -#align homology_iso_cokernel_image_to_kernel' homologyIsoCokernelImageToKernel' +#align homology_iso_cokernel_image_to_kernel' homology'IsoCokernelImageToKernel' variable [HasEqualizers V] /-- `homology f g w` can be computed as the cokernel of `kernel.lift g f w`. -/ -def homologyIsoCokernelLift (w : f ≫ g = 0) : homology f g w ≅ cokernel (kernel.lift g f w) := by - refine' homologyIsoCokernelImageToKernel' f g w ≪≫ _ +def homology'IsoCokernelLift (w : f ≫ g = 0) : homology' f g w ≅ cokernel (kernel.lift g f w) := by + refine' homology'IsoCokernelImageToKernel' f g w ≪≫ _ have p : factorThruImage f ≫ imageToKernel' f g w = kernel.lift g f w := by ext simp [imageToKernel'] exact (cokernelEpiComp _ _).symm ≪≫ cokernelIsoOfEq p -#align homology_iso_cokernel_lift homologyIsoCokernelLift +#align homology_iso_cokernel_lift homology'IsoCokernelLift end imageToKernel' diff --git a/Mathlib/Algebra/Homology/ModuleCat.lean b/Mathlib/Algebra/Homology/ModuleCat.lean index ad46db4390a20..535724074f5e6 100644 --- a/Mathlib/Algebra/Homology/ModuleCat.lean +++ b/Mathlib/Algebra/Homology/ModuleCat.lean @@ -35,8 +35,8 @@ namespace ModuleCat /-- To prove that two maps out of a homology group are equal, it suffices to check they are equal on the images of cycles. -/ -theorem homology_ext {L M N K : ModuleCat R} {f : L ⟶ M} {g : M ⟶ N} (w : f ≫ g = 0) - {h k : homology f g w ⟶ K} +theorem homology'_ext {L M N K : ModuleCat R} {f : L ⟶ M} {g : M ⟶ N} (w : f ≫ g = 0) + {h k : homology' f g w ⟶ K} (w : ∀ x : LinearMap.ker g, h (cokernel.π (imageToKernel _ _ w) (toKernelSubobject x)) = @@ -49,51 +49,51 @@ theorem homology_ext {L M N K : ModuleCat R} {f : L ⟶ M} {g : M ⟶ N} (w : f ModuleCat.kernelIsoKer g).toLinearEquiv.toEquiv.symm.surjective n exact w n set_option linter.uppercaseLean3 false in -#align Module.homology_ext ModuleCat.homology_ext +#align Module.homology_ext ModuleCat.homology'_ext /-- Bundle an element `C.X i` such that `C.dFrom i x = 0` as a term of `C.cycles i`. -/ -abbrev toCycles {C : HomologicalComplex (ModuleCat.{u} R) c} {i : ι} - (x : LinearMap.ker (C.dFrom i)) : (C.cycles i : Type u) := +abbrev toCycles' {C : HomologicalComplex (ModuleCat.{u} R) c} {i : ι} + (x : LinearMap.ker (C.dFrom i)) : (C.cycles' i : Type u) := toKernelSubobject x set_option linter.uppercaseLean3 false in -#align Module.to_cycles ModuleCat.toCycles +#align Module.to_cycles ModuleCat.toCycles' @[ext] -theorem cycles_ext {C : HomologicalComplex (ModuleCat.{u} R) c} {i : ι} - {x y : (C.cycles i : Type u)} - (w : (C.cycles i).arrow x = (C.cycles i).arrow y) : x = y := by - apply_fun (C.cycles i).arrow using (ModuleCat.mono_iff_injective _).mp (cycles C i).arrow_mono +theorem cycles'_ext {C : HomologicalComplex (ModuleCat.{u} R) c} {i : ι} + {x y : (C.cycles' i : Type u)} + (w : (C.cycles' i).arrow x = (C.cycles' i).arrow y) : x = y := by + apply_fun (C.cycles' i).arrow using (ModuleCat.mono_iff_injective _).mp (cycles' C i).arrow_mono exact w set_option linter.uppercaseLean3 false in -#align Module.cycles_ext ModuleCat.cycles_ext +#align Module.cycles_ext ModuleCat.cycles'_ext -- porting note: both proofs by `rw` were proofs by `simp` which no longer worked -- see https://github.com/leanprover-community/mathlib4/issues/5026 @[simp] -theorem cyclesMap_toCycles (f : C ⟶ D) {i : ι} (x : LinearMap.ker (C.dFrom i)) : - (cyclesMap f i) (toCycles x) = toCycles ⟨f.f i x.1, by +theorem cycles'Map_toCycles' (f : C ⟶ D) {i : ι} (x : LinearMap.ker (C.dFrom i)) : + (cycles'Map f i) (toCycles' x) = toCycles' ⟨f.f i x.1, by -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 rw [LinearMap.mem_ker]; erw [Hom.comm_from_apply, x.2, map_zero]⟩ := by ext -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - erw [cyclesMap_arrow_apply, toKernelSubobject_arrow, toKernelSubobject_arrow] + erw [cycles'Map_arrow_apply, toKernelSubobject_arrow, toKernelSubobject_arrow] rfl set_option linter.uppercaseLean3 false in -#align Module.cycles_map_to_cycles ModuleCat.cyclesMap_toCycles +#align Module.cycles_map_to_cycles ModuleCat.cycles'Map_toCycles' /-- Build a term of `C.homology i` from an element `C.X i` such that `C.d_from i x = 0`. -/ -abbrev toHomology {C : HomologicalComplex (ModuleCat.{u} R) c} {i : ι} - (x : LinearMap.ker (C.dFrom i)) : C.homology i := - homology.π (C.dTo i) (C.dFrom i) _ (toCycles x) +abbrev toHomology' {C : HomologicalComplex (ModuleCat.{u} R) c} {i : ι} + (x : LinearMap.ker (C.dFrom i)) : C.homology' i := + homology'.π (C.dTo i) (C.dFrom i) _ (toCycles' x) set_option linter.uppercaseLean3 false in -#align Module.to_homology ModuleCat.toHomology +#align Module.to_homology ModuleCat.toHomology' @[ext] -theorem homology_ext' {M : ModuleCat R} (i : ι) {h k : C.homology i ⟶ M} - (w : ∀ x : LinearMap.ker (C.dFrom i), h (toHomology x) = k (toHomology x)) : h = k := - homology_ext _ w +theorem homology'_ext' {M : ModuleCat R} (i : ι) {h k : C.homology' i ⟶ M} + (w : ∀ x : LinearMap.ker (C.dFrom i), h (toHomology' x) = k (toHomology' x)) : h = k := + homology'_ext _ w set_option linter.uppercaseLean3 false in -#align Module.homology_ext' ModuleCat.homology_ext' +#align Module.homology_ext' ModuleCat.homology'_ext' -- porting note: `erw` had to be used instead of `simp` -- see https://github.com/leanprover-community/mathlib4/issues/5026 @@ -101,13 +101,13 @@ set_option linter.uppercaseLean3 false in specialized to the setting of `V = Module R`, to demonstrate the use of extensionality lemmas for homology in `Module R`. -/ example (f g : C ⟶ D) (h : Homotopy f g) (i : ι) : - (homologyFunctor (ModuleCat.{u} R) c i).map f = - (homologyFunctor (ModuleCat.{u} R) c i).map g := by + (homology'Functor (ModuleCat.{u} R) c i).map f = + (homology'Functor (ModuleCat.{u} R) c i).map g := by -- To check that two morphisms out of a homology group agree, it suffices to check on cycles: - apply homology_ext + apply homology'_ext intro x - simp only [homologyFunctor_map] - erw [homology.π_map_apply, homology.π_map_apply] + simp only [homology'Functor_map] + erw [homology'.π_map_apply, homology'.π_map_apply] -- To check that two elements are equal mod boundaries, it suffices to exhibit a boundary: refine' cokernel_π_imageSubobject_ext _ _ ((toPrev i h.hom) x.1) _ -- Moreover, to check that two cycles are equal, it suffices to check their underlying elements: diff --git a/Mathlib/Algebra/Homology/Opposite.lean b/Mathlib/Algebra/Homology/Opposite.lean index aab55ef6b27a1..aba70459c7234 100644 --- a/Mathlib/Algebra/Homology/Opposite.lean +++ b/Mathlib/Algebra/Homology/Opposite.lean @@ -64,23 +64,24 @@ theorem imageToKernel_unop {X Y Z : Vᵒᵖ} (f : X ⟶ Y) (g : Y ⟶ Z) (w : f /-- Given `f, g` with `f ≫ g = 0`, the homology of `g.op, f.op` is the opposite of the homology of `f, g`. -/ -def homologyOp {X Y Z : V} (f : X ⟶ Y) (g : Y ⟶ Z) (w : f ≫ g = 0) : - homology g.op f.op (by rw [← op_comp, w, op_zero]) ≅ Opposite.op (homology f g w) := +def homology'Op {X Y Z : V} (f : X ⟶ Y) (g : Y ⟶ Z) (w : f ≫ g = 0) : + homology' g.op f.op (by rw [← op_comp, w, op_zero]) ≅ Opposite.op (homology' f g w) := cokernelIsoOfEq (imageToKernel_op _ _ w) ≪≫ cokernelEpiComp _ _ ≪≫ cokernelCompIsIso _ _ ≪≫ - cokernelOpOp _ ≪≫ (homologyIsoKernelDesc _ _ _ ≪≫ + cokernelOpOp _ ≪≫ (homology'IsoKernelDesc _ _ _ ≪≫ kernelIsoOfEq (by ext; simp only [image.fac, cokernel.π_desc, cokernel.π_desc_assoc]) ≪≫ kernelCompMono _ (image.ι g)).op -#align homology_op homologyOp +#align homology_op homology'Op /-- Given morphisms `f, g` in `Vᵒᵖ` with `f ≫ g = 0`, the homology of `g.unop, f.unop` is the opposite of the homology of `f, g`. -/ -def homologyUnop {X Y Z : Vᵒᵖ} (f : X ⟶ Y) (g : Y ⟶ Z) (w : f ≫ g = 0) : - homology g.unop f.unop (by rw [← unop_comp, w, unop_zero]) ≅ Opposite.unop (homology f g w) := +def homology'Unop {X Y Z : Vᵒᵖ} (f : X ⟶ Y) (g : Y ⟶ Z) (w : f ≫ g = 0) : + homology' g.unop f.unop (by rw [← unop_comp, w, unop_zero]) ≅ + Opposite.unop (homology' f g w) := cokernelIsoOfEq (imageToKernel_unop _ _ w) ≪≫ cokernelEpiComp _ _ ≪≫ cokernelCompIsIso _ _ ≪≫ - cokernelUnopUnop _ ≪≫ (homologyIsoKernelDesc _ _ _ ≪≫ + cokernelUnopUnop _ ≪≫ (homology'IsoKernelDesc _ _ _ ≪≫ kernelIsoOfEq (by ext; simp only [image.fac, cokernel.π_desc, cokernel.π_desc_assoc]) ≪≫ kernelCompMono _ (image.ι g)).unop -#align homology_unop homologyUnop +#align homology_unop homology'Unop end @@ -255,30 +256,30 @@ end variable [Abelian V] (C : HomologicalComplex V c) (i : ι) /-- Auxiliary tautological definition for `homologyOp`. -/ -def homologyOpDef : C.op.homology i ≅ - _root_.homology (C.dFrom i).op (C.dTo i).op (by rw [← op_comp, C.dTo_comp_dFrom i, op_zero]) := +def homology'OpDef : C.op.homology' i ≅ + _root_.homology' (C.dFrom i).op (C.dTo i).op (by rw [← op_comp, C.dTo_comp_dFrom i, op_zero]) := Iso.refl _ -#align homological_complex.homology_op_def HomologicalComplex.homologyOpDef +#align homological_complex.homology_op_def HomologicalComplex.homology'OpDef /-- Given a complex `C` of objects in `V`, the `i`th homology of its 'opposite' complex (with objects in `Vᵒᵖ`) is the opposite of the `i`th homology of `C`. -/ -nonrec def homologyOp : C.op.homology i ≅ Opposite.op (C.homology i) := - homologyOpDef _ _ ≪≫ homologyOp _ _ _ -#align homological_complex.homology_op HomologicalComplex.homologyOp +nonrec def homology'Op : C.op.homology' i ≅ Opposite.op (C.homology' i) := + homology'OpDef _ _ ≪≫ homology'Op _ _ _ +#align homological_complex.homology_op HomologicalComplex.homology'Op /-- Auxiliary tautological definition for `homologyUnop`. -/ -def homologyUnopDef (C : HomologicalComplex Vᵒᵖ c) : - C.unop.homology i ≅ - _root_.homology (C.dFrom i).unop (C.dTo i).unop +def homology'UnopDef (C : HomologicalComplex Vᵒᵖ c) : + C.unop.homology' i ≅ + _root_.homology' (C.dFrom i).unop (C.dTo i).unop (by rw [← unop_comp, C.dTo_comp_dFrom i, unop_zero]) := Iso.refl _ -#align homological_complex.homology_unop_def HomologicalComplex.homologyUnopDef +#align homological_complex.homology_unop_def HomologicalComplex.homology'UnopDef /-- Given a complex `C` of objects in `Vᵒᵖ`, the `i`th homology of its 'opposite' complex (with objects in `V`) is the opposite of the `i`th homology of `C`. -/ -nonrec def homologyUnop (C : HomologicalComplex Vᵒᵖ c) : - C.unop.homology i ≅ Opposite.unop (C.homology i) := - homologyUnopDef _ _ ≪≫ homologyUnop _ _ _ -#align homological_complex.homology_unop HomologicalComplex.homologyUnop +nonrec def homology'Unop (C : HomologicalComplex Vᵒᵖ c) : + C.unop.homology' i ≅ Opposite.unop (C.homology' i) := + homology'UnopDef _ _ ≪≫ homology'Unop _ _ _ +#align homological_complex.homology_unop HomologicalComplex.homology'Unop end HomologicalComplex diff --git a/Mathlib/Algebra/Homology/QuasiIso.lean b/Mathlib/Algebra/Homology/QuasiIso.lean index c533ac1c07343..eb70dc818a55d 100644 --- a/Mathlib/Algebra/Homology/QuasiIso.lean +++ b/Mathlib/Algebra/Homology/QuasiIso.lean @@ -35,33 +35,34 @@ variable {c : ComplexShape ι} {C D E : HomologicalComplex V c} /-- A chain map is a quasi-isomorphism if it induces isomorphisms on homology. -/ -class QuasiIso (f : C ⟶ D) : Prop where - isIso : ∀ i, IsIso ((homologyFunctor V c i).map f) -#align quasi_iso QuasiIso +class QuasiIso' (f : C ⟶ D) : Prop where + isIso : ∀ i, IsIso ((homology'Functor V c i).map f) +#align quasi_iso QuasiIso' -attribute [instance] QuasiIso.isIso +attribute [instance] QuasiIso'.isIso -instance (priority := 100) quasiIso_of_iso (f : C ⟶ D) [IsIso f] : QuasiIso f where +instance (priority := 100) quasiIso'_of_iso (f : C ⟶ D) [IsIso f] : QuasiIso' f where isIso i := by - change IsIso ((homologyFunctor V c i).mapIso (asIso f)).hom + change IsIso ((homology'Functor V c i).mapIso (asIso f)).hom infer_instance -#align quasi_iso_of_iso quasiIso_of_iso +#align quasi_iso_of_iso quasiIso'_of_iso -instance quasiIso_comp (f : C ⟶ D) [QuasiIso f] (g : D ⟶ E) [QuasiIso g] : QuasiIso (f ≫ g) where +instance quasiIso'_comp (f : C ⟶ D) [QuasiIso' f] (g : D ⟶ E) [QuasiIso' g] : + QuasiIso' (f ≫ g) where isIso i := by rw [Functor.map_comp] infer_instance -#align quasi_iso_comp quasiIso_comp +#align quasi_iso_comp quasiIso'_comp -theorem quasiIso_of_comp_left (f : C ⟶ D) [QuasiIso f] (g : D ⟶ E) [QuasiIso (f ≫ g)] : - QuasiIso g := - { isIso := fun i => IsIso.of_isIso_fac_left ((homologyFunctor V c i).map_comp f g).symm } -#align quasi_iso_of_comp_left quasiIso_of_comp_left +theorem quasiIso'_of_comp_left (f : C ⟶ D) [QuasiIso' f] (g : D ⟶ E) [QuasiIso' (f ≫ g)] : + QuasiIso' g := + { isIso := fun i => IsIso.of_isIso_fac_left ((homology'Functor V c i).map_comp f g).symm } +#align quasi_iso_of_comp_left quasiIso'_of_comp_left -theorem quasiIso_of_comp_right (f : C ⟶ D) (g : D ⟶ E) [QuasiIso g] [QuasiIso (f ≫ g)] : - QuasiIso f := - { isIso := fun i => IsIso.of_isIso_fac_right ((homologyFunctor V c i).map_comp f g).symm } -#align quasi_iso_of_comp_right quasiIso_of_comp_right +theorem quasiIso'_of_comp_right (f : C ⟶ D) (g : D ⟶ E) [QuasiIso' g] [QuasiIso' (f ≫ g)] : + QuasiIso' f := + { isIso := fun i => IsIso.of_isIso_fac_right ((homology'Functor V c i).map_comp f g).symm } +#align quasi_iso_of_comp_right quasiIso'_of_comp_right namespace HomotopyEquiv @@ -71,21 +72,21 @@ variable {W : Type*} [Category W] [Preadditive W] [HasCokernels W] [HasImages W] [HasZeroObject W] [HasImageMaps W] /-- A homotopy equivalence is a quasi-isomorphism. -/ -theorem toQuasiIso {C D : HomologicalComplex W c} (e : HomotopyEquiv C D) : QuasiIso e.hom := +theorem toQuasiIso' {C D : HomologicalComplex W c} (e : HomotopyEquiv C D) : QuasiIso' e.hom := ⟨fun i => by - refine' ⟨⟨(homologyFunctor W c i).map e.inv, _⟩⟩ - simp only [← Functor.map_comp, ← (homologyFunctor W c i).map_id] - constructor <;> apply homology_map_eq_of_homotopy + refine' ⟨⟨(homology'Functor W c i).map e.inv, _⟩⟩ + simp only [← Functor.map_comp, ← (homology'Functor W c i).map_id] + constructor <;> apply homology'_map_eq_of_homotopy exacts [e.homotopyHomInvId, e.homotopyInvHomId]⟩ -#align homotopy_equiv.to_quasi_iso HomotopyEquiv.toQuasiIso +#align homotopy_equiv.to_quasi_iso HomotopyEquiv.toQuasiIso' -theorem toQuasiIso_inv {C D : HomologicalComplex W c} (e : HomotopyEquiv C D) (i : ι) : - (@asIso _ _ _ _ _ (e.toQuasiIso.1 i)).inv = (homologyFunctor W c i).map e.inv := by +theorem toQuasiIso'_inv {C D : HomologicalComplex W c} (e : HomotopyEquiv C D) (i : ι) : + (@asIso _ _ _ _ _ (e.toQuasiIso'.1 i)).inv = (homology'Functor W c i).map e.inv := by symm - haveI := e.toQuasiIso.1 i -- Porting note: Added this to get `asIso_hom` to work. - simp only [← Iso.hom_comp_eq_id, asIso_hom, ← Functor.map_comp, ← (homologyFunctor W c i).map_id, - homology_map_eq_of_homotopy e.homotopyHomInvId _] -#align homotopy_equiv.to_quasi_iso_inv HomotopyEquiv.toQuasiIso_inv + haveI := e.toQuasiIso'.1 i -- Porting note: Added this to get `asIso_hom` to work. + simp only [← Iso.hom_comp_eq_id, asIso_hom, ← Functor.map_comp, + ← (homology'Functor W c i).map_id, homology'_map_eq_of_homotopy e.homotopyHomInvId _] +#align homotopy_equiv.to_quasi_iso_inv HomotopyEquiv.toQuasiIso'_inv end @@ -99,27 +100,27 @@ variable {W : Type*} [Category W] [Abelian W] section -variable {X : ChainComplex W ℕ} {Y : W} (f : X ⟶ (ChainComplex.single₀ _).obj Y) [hf : QuasiIso f] +variable {X : ChainComplex W ℕ} {Y : W} (f : X ⟶ (ChainComplex.single₀ _).obj Y) [hf : QuasiIso' f] /-- If a chain map `f : X ⟶ Y[0]` is a quasi-isomorphism, then the cokernel of the differential `d : X₁ → X₀` is isomorphic to `Y`. -/ noncomputable def toSingle₀CokernelAtZeroIso : cokernel (X.d 1 0) ≅ Y := - X.homologyZeroIso.symm.trans - ((@asIso _ _ _ _ _ (hf.1 0)).trans ((ChainComplex.homologyFunctor0Single₀ W).app Y)) + X.homology'ZeroIso.symm.trans + ((@asIso _ _ _ _ _ (hf.1 0)).trans ((ChainComplex.homology'Functor0Single₀ W).app Y)) #align homological_complex.hom.to_single₀_cokernel_at_zero_iso HomologicalComplex.Hom.toSingle₀CokernelAtZeroIso -theorem toSingle₀CokernelAtZeroIso_hom_eq [hf : QuasiIso f] : +theorem toSingle₀CokernelAtZeroIso_hom_eq [hf : QuasiIso' f] : f.toSingle₀CokernelAtZeroIso.hom = cokernel.desc (X.d 1 0) (f.f 0) (by rw [← f.2 1 0 rfl]; exact comp_zero) := by ext - dsimp only [toSingle₀CokernelAtZeroIso, ChainComplex.homologyZeroIso, homologyOfZeroRight, - homology.mapIso, ChainComplex.homologyFunctor0Single₀, cokernel.map] + dsimp only [toSingle₀CokernelAtZeroIso, ChainComplex.homology'ZeroIso, homology'OfZeroRight, + homology'.mapIso, ChainComplex.homology'Functor0Single₀, cokernel.map] dsimp [asIso] - simp only [cokernel.π_desc, Category.assoc, homology.map_desc, cokernel.π_desc_assoc] - simp [homology.desc, Iso.refl_inv (X.X 0)] + simp only [cokernel.π_desc, Category.assoc, homology'.map_desc, cokernel.π_desc_assoc] + simp [homology'.desc, Iso.refl_inv (X.X 0)] #align homological_complex.hom.to_single₀_cokernel_at_zero_iso_hom_eq HomologicalComplex.Hom.toSingle₀CokernelAtZeroIso_hom_eq -theorem to_single₀_epi_at_zero [hf : QuasiIso f] : Epi (f.f 0) := by +theorem to_single₀_epi_at_zero [hf : QuasiIso' f] : Epi (f.f 0) := by constructor intro Z g h Hgh rw [← cokernel.π_desc (X.d 1 0) (f.f 0) (by rw [← f.2 1 0 rfl]; exact comp_zero), @@ -127,22 +128,22 @@ theorem to_single₀_epi_at_zero [hf : QuasiIso f] : Epi (f.f 0) := by rw [(@cancel_epi _ _ _ _ _ _ (epi_comp _ _) _ _).1 Hgh] #align homological_complex.hom.to_single₀_epi_at_zero HomologicalComplex.Hom.to_single₀_epi_at_zero -theorem to_single₀_exact_d_f_at_zero [hf : QuasiIso f] : Exact (X.d 1 0) (f.f 0) := by - rw [Preadditive.exact_iff_homology_zero] +theorem to_single₀_exact_d_f_at_zero [hf : QuasiIso' f] : Exact (X.d 1 0) (f.f 0) := by + rw [Preadditive.exact_iff_homology'_zero] have h : X.d 1 0 ≫ f.f 0 = 0 := by simp only [← f.2 1 0 rfl, ChainComplex.single₀_obj_X_d, comp_zero] - refine' ⟨h, Nonempty.intro (homologyIsoKernelDesc _ _ _ ≪≫ _)⟩ + refine' ⟨h, Nonempty.intro (homology'IsoKernelDesc _ _ _ ≪≫ _)⟩ suffices IsIso (cokernel.desc _ _ h) by apply kernel.ofMono rw [← toSingle₀CokernelAtZeroIso_hom_eq] infer_instance #align homological_complex.hom.to_single₀_exact_d_f_at_zero HomologicalComplex.Hom.to_single₀_exact_d_f_at_zero -theorem to_single₀_exact_at_succ [hf : QuasiIso f] (n : ℕ) : +theorem to_single₀_exact_at_succ [hf : QuasiIso' f] (n : ℕ) : Exact (X.d (n + 2) (n + 1)) (X.d (n + 1) n) := - (Preadditive.exact_iff_homology_zero _ _).2 + (Preadditive.exact_iff_homology'_zero _ _).2 ⟨X.d_comp_d _ _ _, - ⟨(ChainComplex.homologySuccIso _ _).symm.trans - ((@asIso _ _ _ _ _ (hf.1 (n + 1))).trans homologyZeroZero)⟩⟩ + ⟨(ChainComplex.homology'SuccIso _ _).symm.trans + ((@asIso _ _ _ _ _ (hf.1 (n + 1))).trans homology'ZeroZero)⟩⟩ #align homological_complex.hom.to_single₀_exact_at_succ HomologicalComplex.Hom.to_single₀_exact_at_succ end @@ -153,26 +154,26 @@ variable {X : CochainComplex W ℕ} {Y : W} (f : (CochainComplex.single₀ _).ob /-- If a cochain map `f : Y[0] ⟶ X` is a quasi-isomorphism, then the kernel of the differential `d : X₀ → X₁` is isomorphic to `Y`. -/ -noncomputable def fromSingle₀KernelAtZeroIso [hf : QuasiIso f] : kernel (X.d 0 1) ≅ Y := - X.homologyZeroIso.symm.trans +noncomputable def fromSingle₀KernelAtZeroIso [hf : QuasiIso' f] : kernel (X.d 0 1) ≅ Y := + X.homology'ZeroIso.symm.trans ((@asIso _ _ _ _ _ (hf.1 0)).symm.trans ((CochainComplex.homologyFunctor0Single₀ W).app Y)) #align homological_complex.hom.from_single₀_kernel_at_zero_iso HomologicalComplex.Hom.fromSingle₀KernelAtZeroIso -theorem fromSingle₀KernelAtZeroIso_inv_eq [hf : QuasiIso f] : +theorem fromSingle₀KernelAtZeroIso_inv_eq [hf : QuasiIso' f] : f.fromSingle₀KernelAtZeroIso.inv = kernel.lift (X.d 0 1) (f.f 0) (by rw [f.2 0 1 rfl]; exact zero_comp) := by ext - dsimp only [fromSingle₀KernelAtZeroIso, CochainComplex.homologyZeroIso, homologyOfZeroLeft, - homology.mapIso, CochainComplex.homologyFunctor0Single₀, kernel.map] + dsimp only [fromSingle₀KernelAtZeroIso, CochainComplex.homology'ZeroIso, homology'OfZeroLeft, + homology'.mapIso, CochainComplex.homologyFunctor0Single₀, kernel.map] simp only [Iso.trans_inv, Iso.app_inv, Iso.symm_inv, Category.assoc, equalizer_as_kernel, kernel.lift_ι] dsimp [asIso] - simp only [Category.assoc, homology.π_map, cokernelZeroIsoTarget_hom, - cokernelIsoOfEq_hom_comp_desc, kernelSubobject_arrow, homology.π_map_assoc, IsIso.inv_comp_eq] - simp [homology.π, kernelSubobjectMap_comp, Iso.refl_hom (X.X 0), Category.comp_id] + simp only [Category.assoc, homology'.π_map, cokernelZeroIsoTarget_hom, + cokernelIsoOfEq_hom_comp_desc, kernelSubobject_arrow, homology'.π_map_assoc, IsIso.inv_comp_eq] + simp [homology'.π, kernelSubobjectMap_comp, Iso.refl_hom (X.X 0), Category.comp_id] #align homological_complex.hom.from_single₀_kernel_at_zero_iso_inv_eq HomologicalComplex.Hom.fromSingle₀KernelAtZeroIso_inv_eq -theorem from_single₀_mono_at_zero [hf : QuasiIso f] : Mono (f.f 0) := by +theorem from_single₀_mono_at_zero [hf : QuasiIso' f] : Mono (f.f 0) := by constructor intro Z g h Hgh rw [← kernel.lift_ι (X.d 0 1) (f.f 0) (by rw [f.2 0 1 rfl]; exact zero_comp), @@ -180,22 +181,22 @@ theorem from_single₀_mono_at_zero [hf : QuasiIso f] : Mono (f.f 0) := by rw [(@cancel_mono _ _ _ _ _ _ (mono_comp _ _) _ _).1 Hgh] #align homological_complex.hom.from_single₀_mono_at_zero HomologicalComplex.Hom.from_single₀_mono_at_zero -theorem from_single₀_exact_f_d_at_zero [hf : QuasiIso f] : Exact (f.f 0) (X.d 0 1) := by - rw [Preadditive.exact_iff_homology_zero] +theorem from_single₀_exact_f_d_at_zero [hf : QuasiIso' f] : Exact (f.f 0) (X.d 0 1) := by + rw [Preadditive.exact_iff_homology'_zero] have h : f.f 0 ≫ X.d 0 1 = 0 := by simp only [HomologicalComplex.Hom.comm, CochainComplex.single₀_obj_X_d, zero_comp] - refine' ⟨h, Nonempty.intro (homologyIsoCokernelLift _ _ _ ≪≫ _)⟩ + refine' ⟨h, Nonempty.intro (homology'IsoCokernelLift _ _ _ ≪≫ _)⟩ suffices IsIso (kernel.lift (X.d 0 1) (f.f 0) h) by apply cokernel.ofEpi rw [← fromSingle₀KernelAtZeroIso_inv_eq f] infer_instance #align homological_complex.hom.from_single₀_exact_f_d_at_zero HomologicalComplex.Hom.from_single₀_exact_f_d_at_zero -theorem from_single₀_exact_at_succ [hf : QuasiIso f] (n : ℕ) : +theorem from_single₀_exact_at_succ [hf : QuasiIso' f] (n : ℕ) : Exact (X.d n (n + 1)) (X.d (n + 1) (n + 2)) := - (Preadditive.exact_iff_homology_zero _ _).2 + (Preadditive.exact_iff_homology'_zero _ _).2 ⟨X.d_comp_d _ _ _, - ⟨(CochainComplex.homologySuccIso _ _).symm.trans - ((@asIso _ _ _ _ _ (hf.1 (n + 1))).symm.trans homologyZeroZero)⟩⟩ + ⟨(CochainComplex.homology'SuccIso _ _).symm.trans + ((@asIso _ _ _ _ _ (hf.1 (n + 1))).symm.trans homology'ZeroZero)⟩⟩ #align homological_complex.hom.from_single₀_exact_at_succ HomologicalComplex.Hom.from_single₀_exact_at_succ end @@ -207,11 +208,11 @@ end HomologicalComplex.Hom variable {A : Type*} [Category A] [Abelian A] {B : Type*} [Category B] [Abelian B] (F : A ⥤ B) [Functor.Additive F] [PreservesFiniteLimits F] [PreservesFiniteColimits F] [Faithful F] -theorem CategoryTheory.Functor.quasiIso_of_map_quasiIso {C D : HomologicalComplex A c} (f : C ⟶ D) - (hf : QuasiIso ((F.mapHomologicalComplex _).map f)) : QuasiIso f := +theorem CategoryTheory.Functor.quasiIso'_of_map_quasiIso' {C D : HomologicalComplex A c} + (f : C ⟶ D) (hf : QuasiIso' ((F.mapHomologicalComplex _).map f)) : QuasiIso' f := ⟨fun i => - haveI : IsIso (F.map ((homologyFunctor A c i).map f)) := by - rw [← Functor.comp_map, ← NatIso.naturality_2 (F.homologyFunctorIso i) f, Functor.comp_map] + haveI : IsIso (F.map ((homology'Functor A c i).map f)) := by + rw [← Functor.comp_map, ← NatIso.naturality_2 (F.homology'FunctorIso i) f, Functor.comp_map] infer_instance isIso_of_reflects_iso _ F⟩ -#align category_theory.functor.quasi_iso_of_map_quasi_iso CategoryTheory.Functor.quasiIso_of_map_quasiIso +#align category_theory.functor.quasi_iso_of_map_quasi_iso CategoryTheory.Functor.quasiIso'_of_map_quasiIso' diff --git a/Mathlib/Algebra/Homology/ShortComplex/Ab.lean b/Mathlib/Algebra/Homology/ShortComplex/Ab.lean new file mode 100644 index 0000000000000..822b100f595d3 --- /dev/null +++ b/Mathlib/Algebra/Homology/ShortComplex/Ab.lean @@ -0,0 +1,122 @@ +/- +Copyright (c) 2023 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Homology.ShortComplex.ShortExact +import Mathlib.Algebra.Category.GroupCat.Abelian + +/-! +# Homology and exactness of short complexes of abelian groups + +In this file, the homology of a short complex `S` of abelian groups is identified +with the quotient of `AddMonoidHom.ker S.g` by the image of the morphism +`S.abToCycles : S.X₁ →+ AddMonoidHom.ker S.g` induced by `S.f`. + +The definitions are made in the `ShortComplex` namespace so as to enable dot notation. +The names contain the prefix `ab` in order to allow similar constructions for +other categories like `ModuleCat`. + +## Main definitions +- `ShortComplex.abHomologyIso` identifies the homology of a short complex of abelian +groups to an explicit quotient. +- `ShortComplex.ab_exact_iff` expresses that a short complex of abelian groups `S` +is exact iff any element in the kernel of `S.g` belongs to the image of `S.f`. + +-/ + +universe u + +namespace CategoryTheory + +namespace ShortComplex + +variable (S : ShortComplex Ab.{u}) + +@[simp] +lemma ab_zero_apply (x : S.X₁) : S.g (S.f x) = 0 := by + erw [← comp_apply, S.zero] + rfl + +/-- The canonical additive morphism for `S.X₁ →+ AddMonoidHom.ker S.g` induced by `S.f`. -/ +@[simps!] +def abToCycles : S.X₁ →+ AddMonoidHom.ker S.g := + AddMonoidHom.mk' (fun x => ⟨S.f x, S.ab_zero_apply x⟩) (by aesop) + +/-- The explicit left homology data of a short complex of abelian group that is +given by a kernel and a quotient given by the `AddMonoidHom` API. -/ +@[simps] +def abLeftHomologyData : S.LeftHomologyData where + K := AddCommGroupCat.of (AddMonoidHom.ker S.g) + H := AddCommGroupCat.of ((AddMonoidHom.ker S.g) ⧸ AddMonoidHom.range S.abToCycles) + i := (AddMonoidHom.ker S.g).subtype + π := QuotientAddGroup.mk' _ + wi := by + ext ⟨_, hx⟩ + exact hx + hi := AddCommGroupCat.kernelIsLimit _ + wπ := by + ext (x : S.X₁) + erw [QuotientAddGroup.eq_zero_iff] + rw [AddMonoidHom.mem_range] + apply exists_apply_eq_apply + hπ := AddCommGroupCat.cokernelIsColimit (AddCommGroupCat.ofHom S.abToCycles) + +@[simp] +lemma abLeftHomologyData_f' : S.abLeftHomologyData.f' = S.abToCycles := rfl + +/-- Given a short complex `S` of abelian groups, this is the isomorphism between +the abstract `S.cycles` of the homology API and the more concrete description as +`AddMonoidHom.ker S.g`. -/ +noncomputable def abCyclesIso : S.cycles ≅ AddCommGroupCat.of (AddMonoidHom.ker S.g) := + S.abLeftHomologyData.cyclesIso + +/-- Given a short complex `S` of abelian groups, this is the isomorphism between +the abstract `S.homology` of the homology API and the more explicit +quotient of `AddMonoidHom.ker S.g` by the image of +`S.abToCycles : S.X₁ →+ AddMonoidHom.ker S.g`. -/ +noncomputable def abHomologyIso : S.homology ≅ + AddCommGroupCat.of ((AddMonoidHom.ker S.g) ⧸ AddMonoidHom.range S.abToCycles) := + S.abLeftHomologyData.homologyIso + +lemma exact_iff_surjective_abToCycles : + S.Exact ↔ Function.Surjective S.abToCycles := by + rw [S.abLeftHomologyData.exact_iff_epi_f', abLeftHomologyData_f', + AddCommGroupCat.epi_iff_surjective] + rfl + +lemma ab_exact_iff : + S.Exact ↔ ∀ (x₂ : S.X₂) (_ : S.g x₂ = 0), ∃ (x₁ : S.X₁), S.f x₁ = x₂ := by + rw [exact_iff_surjective_abToCycles] + constructor + · intro h x₂ hx₂ + obtain ⟨x₁, hx₁⟩ := h ⟨x₂, hx₂⟩ + exact ⟨x₁, by simpa only [Subtype.ext_iff, abToCycles_apply_coe] using hx₁⟩ + · rintro h ⟨x₂, hx₂⟩ + obtain ⟨x₁, rfl⟩ := h x₂ hx₂ + exact ⟨x₁, rfl⟩ + +lemma ab_exact_iff_ker_le_range : S.Exact ↔ S.g.ker ≤ S.f.range := S.ab_exact_iff + +lemma ab_exact_iff_range_eq_ker : S.Exact ↔ S.f.range = S.g.ker := by + rw [ab_exact_iff_ker_le_range] + constructor + · intro h + refine' le_antisymm _ h + rintro _ ⟨x₁, rfl⟩ + erw [AddMonoidHom.mem_ker, ← comp_apply, S.zero] + rfl + · intro h + rw [h] + +lemma ShortExact.ab_injective_f (hS : S.ShortExact) : + Function.Injective S.f := + (AddCommGroupCat.mono_iff_injective _).1 hS.mono_f + +lemma ShortExact.ab_surjective_g (hS : S.ShortExact) : + Function.Surjective S.g := + (AddCommGroupCat.epi_iff_surjective _).1 hS.epi_g + +end ShortComplex + +end CategoryTheory diff --git a/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean b/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean index 6e5243f46e888..1dffc9bed40a0 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ +import Mathlib.Algebra.Homology.Homology import Mathlib.Algebra.Homology.ShortComplex.Homology import Mathlib.CategoryTheory.Abelian.Basic @@ -181,4 +182,9 @@ instance _root_.CategoryTheory.categoryWithHomology_of_abelian : CategoryWithHomology C where hasHomology S := HasHomology.mk' (HomologyData.ofAbelian S) +/-- Comparison isomorphism between two definitions of homology. -/ +noncomputable def homology'IsoHomology : + _root_.homology' S.f S.g S.zero ≅ S.homology := + homology'IsoCokernelLift S.f S.g S.zero ≪≫ S.homologyIsoCokernelLift.symm + end ShortComplex diff --git a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean new file mode 100644 index 0000000000000..7905ad98b6a0a --- /dev/null +++ b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean @@ -0,0 +1,80 @@ +/- +Copyright (c) 2023 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Homology.ShortComplex.Ab + +/-! +# Exactness of short complexes in concrete abelian categories + +If an additive concrete category `C` has an additive forgetful functor to `Ab` +which preserves homology, then a short complex `S` in `C` is exact +if and only if it is so after applying the functor `forget₂ C Ab`. + +-/ + +namespace CategoryTheory + +open Limits + +variable {C : Type*} [Category C] [ConcreteCategory C] [HasForget₂ C Ab] + +@[simp] +lemma ShortComplex.zero_apply + [Limits.HasZeroMorphisms C] [(forget₂ C Ab).PreservesZeroMorphisms] + (S : ShortComplex C) (x : (forget₂ C Ab).obj S.X₁) : + ((forget₂ C Ab).map S.g) (((forget₂ C Ab).map S.f) x) = 0 := by + erw [← comp_apply, ← Functor.map_comp, S.zero, Functor.map_zero] + rfl + +section preadditive + +variable [Preadditive C] [(forget₂ C Ab).Additive] [(forget₂ C Ab).PreservesHomology] + [HasZeroObject C] (S : ShortComplex C) + +lemma Preadditive.mono_iff_injective {X Y : C} (f : X ⟶ Y) : + Mono f ↔ Function.Injective ((forget₂ C Ab).map f) := by + rw [← AddCommGroupCat.mono_iff_injective] + constructor + · intro + infer_instance + · apply Functor.mono_of_mono_map + +lemma Preadditive.epi_iff_injective {X Y : C} (f : X ⟶ Y) : + Epi f ↔ Function.Surjective ((forget₂ C Ab).map f) := by + rw [← AddCommGroupCat.epi_iff_surjective] + constructor + · intro + infer_instance + · apply Functor.epi_of_epi_map + +namespace ShortComplex + +lemma exact_iff_exact_map_forget₂ [S.HasHomology] : + S.Exact ↔ (S.map (forget₂ C Ab)).Exact := + (S.exact_map_iff_of_faithful (forget₂ C Ab)).symm + +lemma exact_iff_of_concreteCategory [S.HasHomology] : + S.Exact ↔ ∀ (x₂ : (forget₂ C Ab).obj S.X₂) (_ : ((forget₂ C Ab).map S.g) x₂ = 0), + ∃ (x₁ : (forget₂ C Ab).obj S.X₁), ((forget₂ C Ab).map S.f) x₁ = x₂ := by + rw [S.exact_iff_exact_map_forget₂, ab_exact_iff] + rfl + +variable {S} + +lemma ShortExact.injective_f (hS : S.ShortExact) : + Function.Injective ((forget₂ C Ab).map S.f) := by + rw [← Preadditive.mono_iff_injective] + exact hS.mono_f + +lemma ShortExact.surjective_g (hS : S.ShortExact) : + Function.Surjective ((forget₂ C Ab).map S.g) := by + rw [← Preadditive.epi_iff_injective] + exact hS.epi_g + +end ShortComplex + +end preadditive + +end CategoryTheory diff --git a/Mathlib/Algebra/Homology/ShortComplex/Exact.lean b/Mathlib/Algebra/Homology/ShortComplex/Exact.lean index f64f9b7cfe397..e9226c5930386 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Exact.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Exact.lean @@ -5,6 +5,7 @@ Authors: Joël Riou -/ import Mathlib.Algebra.Homology.ShortComplex.PreservesHomology import Mathlib.Algebra.Homology.ShortComplex.Abelian +import Mathlib.Algebra.Homology.ShortComplex.QuasiIso import Mathlib.CategoryTheory.Abelian.Exact import Mathlib.CategoryTheory.MorphismProperty @@ -390,8 +391,534 @@ lemma exact_iff_mono_cokernel_desc [S.HasHomology] [HasCokernel S.f] : refine' (MorphismProperty.RespectsIso.monomorphisms C).arrow_mk_iso_iff (Iso.symm _) exact Arrow.isoMk S.opcyclesIsoCokernel.symm (Iso.refl _) (by aesop_cat) +lemma QuasiIso.exact_iff {S₁ S₂ : ShortComplex C} (φ : S₁ ⟶ S₂) + [S₁.HasHomology] [S₂.HasHomology] [QuasiIso φ] : S₁.Exact ↔ S₂.Exact := by + simp only [exact_iff_isZero_homology] + exact Iso.isZero_iff (asIso (homologyMap φ)) + +lemma exact_of_f_is_kernel (hS : IsLimit (KernelFork.ofι S.f S.zero)) + [S.HasHomology] : S.Exact := by + rw [exact_iff_epi_toCycles] + have : IsSplitEpi S.toCycles := + ⟨⟨{ section_ := hS.lift (KernelFork.ofι S.iCycles S.iCycles_g) + id := by + rw [← cancel_mono S.iCycles, assoc, toCycles_i, id_comp] + exact Fork.IsLimit.lift_ι hS }⟩⟩ + infer_instance + +lemma exact_of_g_is_cokernel (hS : IsColimit (CokernelCofork.ofπ S.g S.zero)) + [S.HasHomology] : S.Exact := by + rw [exact_iff_mono_fromOpcycles] + have : IsSplitMono S.fromOpcycles := + ⟨⟨{ retraction := hS.desc (CokernelCofork.ofπ S.pOpcycles S.f_pOpcycles) + id := by + rw [← cancel_epi S.pOpcycles, p_fromOpcycles_assoc, comp_id] + exact Cofork.IsColimit.π_desc hS }⟩⟩ + infer_instance + +variable {S} + +lemma Exact.mono_g (hS : S.Exact) (hf : S.f = 0) : Mono S.g := by + have := hS.hasHomology + have := hS.epi_toCycles + have : S.iCycles = 0 := by rw [← cancel_epi S.toCycles, comp_zero, toCycles_i, hf] + apply Preadditive.mono_of_cancel_zero + intro A x₂ hx₂ + rw [← S.liftCycles_i x₂ hx₂, this, comp_zero] + +lemma Exact.epi_f (hS : S.Exact) (hg : S.g = 0) : Epi S.f := by + have := hS.hasHomology + have := hS.mono_fromOpcycles + have : S.pOpcycles = 0 := by rw [← cancel_mono S.fromOpcycles, zero_comp, p_fromOpcycles, hg] + apply Preadditive.epi_of_cancel_zero + intro A x₂ hx₂ + rw [← S.p_descOpcycles x₂ hx₂, this, zero_comp] + +lemma Exact.mono_g_iff (hS : S.Exact) : Mono S.g ↔ S.f = 0 := by + constructor + · intro + rw [← cancel_mono S.g, zero, zero_comp] + · exact hS.mono_g + +lemma Exact.epi_f_iff (hS : S.Exact) : Epi S.f ↔ S.g = 0 := by + constructor + · intro + rw [← cancel_epi S.f, zero, comp_zero] + · exact hS.epi_f + +lemma Exact.isZero_X₂ (hS : S.Exact) (hf : S.f = 0) (hg : S.g = 0) : IsZero S.X₂ := by + have := hS.mono_g hf + rw [IsZero.iff_id_eq_zero, ← cancel_mono S.g, hg, comp_zero, comp_zero] + +lemma Exact.isZero_X₂_iff (hS : S.Exact) : IsZero S.X₂ ↔ S.f = 0 ∧ S.g = 0 := by + constructor + · intro h + exact ⟨h.eq_of_tgt _ _, h.eq_of_src _ _⟩ + · rintro ⟨hf, hg⟩ + exact hS.isZero_X₂ hf hg + +variable (S) + +/-- A splitting for a short complex `S` consists of the data of a retraction `r : X₂ ⟶ X₁` +of `S.f` and section `s : X₃ ⟶ X₂` of `S.g` which satisfy `r ≫ S.f + S.g ≫ s = 𝟙 _` -/ +structure Splitting (S : ShortComplex C) where + /-- a retraction of `S.f` -/ + r : S.X₂ ⟶ S.X₁ + /-- a section of `S.g` -/ + s : S.X₃ ⟶ S.X₂ + /-- the condition that `r` is a retraction of `S.f` -/ + f_r : S.f ≫ r = 𝟙 _ := by aesop_cat + /-- the condition that `s` is a section of `S.g` -/ + s_g : s ≫ S.g = 𝟙 _ := by aesop_cat + /-- the compatibility between the given section and retraction -/ + id : r ≫ S.f + S.g ≫ s = 𝟙 _ := by aesop_cat + +namespace Splitting + +attribute [reassoc (attr := simp)] f_r s_g + +variable {S} + +@[reassoc] +lemma r_f (s : S.Splitting) : s.r ≫ S.f = 𝟙 _ - S.g ≫ s.s := by rw [← s.id, add_sub_cancel] + +@[reassoc] +lemma g_s (s : S.Splitting) : S.g ≫ s.s = 𝟙 _ - s.r ≫ S.f := by rw [← s.id, add_sub_cancel'] + +/-- Given a splitting of a short complex `S`, this shows that `S.f` is a split monomorphism. -/ +@[simps] def splitMono_f (s : S.Splitting) : SplitMono S.f := ⟨s.r, s.f_r⟩ + +lemma isSplitMono_f (s : S.Splitting) : IsSplitMono S.f := ⟨⟨s.splitMono_f⟩⟩ + +lemma mono_f (s : S.Splitting) : Mono S.f := by + have := s.isSplitMono_f + infer_instance + +/-- Given a splitting of a short complex `S`, this shows that `S.g` is a split epimorphism. -/ +@[simps] def splitEpi_g (s : S.Splitting) : SplitEpi S.g := ⟨s.s, s.s_g⟩ + +lemma isSplitEpi_g (s : S.Splitting) : IsSplitEpi S.g := ⟨⟨s.splitEpi_g⟩⟩ + +lemma epi_g (s : S.Splitting) : Epi S.g := by + have := s.isSplitEpi_g + infer_instance + +@[reassoc (attr := simp)] +lemma s_r (s : S.Splitting) : s.s ≫ s.r = 0 := by + have := s.epi_g + simp only [← cancel_epi S.g, comp_zero, g_s_assoc, sub_comp, id_comp, + assoc, f_r, comp_id, sub_self] + +lemma ext_r (s s' : S.Splitting) (h : s.r = s'.r) : s = s' := by + have := s.epi_g + have eq := s.id + rw [← s'.id, h, add_right_inj, cancel_epi S.g] at eq + cases s + cases s' + obtain rfl := eq + obtain rfl := h + rfl + +lemma ext_s (s s' : S.Splitting) (h : s.s = s'.s) : s = s' := by + have := s.mono_f + have eq := s.id + rw [← s'.id, h, add_left_inj, cancel_mono S.f] at eq + cases s + cases s' + obtain rfl := eq + obtain rfl := h + rfl + +/-- The left homology data on a short complex equipped with a splitting. -/ +@[simps] +noncomputable def leftHomologyData [HasZeroObject C] (s : S.Splitting) : + LeftHomologyData S := by + have hi := KernelFork.IsLimit.ofι S.f S.zero + (fun x _ => x ≫ s.r) + (fun x hx => by simp only [assoc, s.r_f, comp_sub, comp_id, + sub_eq_self, reassoc_of% hx, zero_comp]) + (fun x _ b hb => by simp only [← hb, assoc, f_r, comp_id]) + let f' := hi.lift (KernelFork.ofι S.f S.zero) + have hf' : f' = 𝟙 _ := by + apply Fork.IsLimit.hom_ext hi + dsimp + erw [Fork.IsLimit.lift_ι hi] + simp only [Fork.ι_ofι, id_comp] + have wπ : f' ≫ (0 : S.X₁ ⟶ 0) = 0 := comp_zero + have hπ : IsColimit (CokernelCofork.ofπ 0 wπ) := CokernelCofork.IsColimit.ofEpiOfIsZero _ + (by rw [hf']; infer_instance) (isZero_zero _) + exact + { K := S.X₁ + H := 0 + i := S.f + wi := S.zero + hi := hi + π := 0 + wπ := wπ + hπ := hπ } + +/-- The right homology data on a short complex equipped with a splitting. -/ +@[simps] +noncomputable def rightHomologyData [HasZeroObject C] (s : S.Splitting) : + RightHomologyData S := by + have hp := CokernelCofork.IsColimit.ofπ S.g S.zero + (fun x _ => s.s ≫ x) + (fun x hx => by simp only [s.g_s_assoc, sub_comp, id_comp, sub_eq_self, assoc, hx, comp_zero]) + (fun x _ b hb => by simp only [← hb, s.s_g_assoc]) + let g' := hp.desc (CokernelCofork.ofπ S.g S.zero) + have hg' : g' = 𝟙 _ := by + apply Cofork.IsColimit.hom_ext hp + dsimp + erw [Cofork.IsColimit.π_desc hp] + simp only [Cofork.π_ofπ, comp_id] + have wι : (0 : 0 ⟶ S.X₃) ≫ g' = 0 := zero_comp + have hι : IsLimit (KernelFork.ofι 0 wι) := KernelFork.IsLimit.ofMonoOfIsZero _ + (by rw [hg']; dsimp; infer_instance) (isZero_zero _) + exact + { Q := S.X₃ + H := 0 + p := S.g + wp := S.zero + hp := hp + ι := 0 + wι := wι + hι := hι } + +/-- The homology data on a short complex equipped with a splitting. -/ +@[simps] +noncomputable def homologyData [HasZeroObject C] (s : S.Splitting) : S.HomologyData where + left := s.leftHomologyData + right := s.rightHomologyData + iso := Iso.refl 0 + +/-- A short complex equipped with a splitting is exact. -/ +lemma exact [HasZeroObject C] (s : S.Splitting) : S.Exact := + ⟨s.homologyData, isZero_zero _⟩ + +/-- If a short complex `S` is equipped with a splitting, then `S.X₁` is the kernel of `S.g`. -/ +noncomputable def fIsKernel [HasZeroObject C] (s : S.Splitting) : + IsLimit (KernelFork.ofι S.f S.zero) := + s.homologyData.left.hi + +/-- If a short complex `S` is equipped with a splitting, then `S.X₃` is the cokernel of `S.f`. -/ +noncomputable def gIsCokernel [HasZeroObject C] (s : S.Splitting) : + IsColimit (CokernelCofork.ofπ S.g S.zero) := + s.homologyData.right.hp + +/-- If a short complex `S` has a splitting and `F` is an additive functor, then +`S.map F` also has a splitting. -/ +@[simps] +def map (s : S.Splitting) (F : C ⥤ D) [F.Additive] : (S.map F).Splitting where + r := F.map s.r + s := F.map s.s + f_r := by + dsimp [ShortComplex.map] + rw [← F.map_comp, f_r, F.map_id] + s_g := by + dsimp [ShortComplex.map] + simp only [← F.map_comp, s_g, F.map_id] + id := by + dsimp [ShortComplex.map] + simp only [← F.map_id, ← s.id, Functor.map_comp, Functor.map_add] + +/-- A splitting on a short complex induces splittings on isomorphic short complexes. -/ +@[simps] +def ofIso {S₁ S₂ : ShortComplex C} (s : S₁.Splitting) (e : S₁ ≅ S₂) : S₂.Splitting where + r := e.inv.τ₂ ≫ s.r ≫ e.hom.τ₁ + s := e.inv.τ₃ ≫ s.s ≫ e.hom.τ₂ + f_r := by rw [← e.inv.comm₁₂_assoc, s.f_r_assoc, ← comp_τ₁, e.inv_hom_id, id_τ₁] + s_g := by rw [assoc, assoc, e.hom.comm₂₃, s.s_g_assoc, ← comp_τ₃, e.inv_hom_id, id_τ₃] + id := by + have eq := e.inv.τ₂ ≫= s.id =≫ e.hom.τ₂ + rw [id_comp, ← comp_τ₂, e.inv_hom_id, id_τ₂] at eq + rw [← eq, assoc, assoc, add_comp, assoc, assoc, comp_add, + e.hom.comm₁₂, e.inv.comm₂₃_assoc] + +/-- The obvious splitting of the short complex `X₁ ⟶ X₁ ⊞ X₂ ⟶ X₂`. -/ +noncomputable def ofHasBinaryBiproduct (X₁ X₂ : C) [HasBinaryBiproduct X₁ X₂] : + Splitting (ShortComplex.mk (biprod.inl : X₁ ⟶ _) (biprod.snd : _ ⟶ X₂) (by simp)) where + r := biprod.fst + s := biprod.inr + +variable (S) + +/-- The obvious splitting of a short complex when `S.X₁` is zero and `S.g` is an isomorphism. -/ +noncomputable def ofIsZeroOfIsIso (hf : IsZero S.X₁) (hg : IsIso S.g) : Splitting S where + r := 0 + s := inv S.g + f_r := hf.eq_of_src _ _ + +/-- The obvious splitting of a short complex when `S.f` is an isomorphism and `S.X₃` is zero. -/ +noncomputable def ofIsIsoOfIsZero (hf : IsIso S.f) (hg : IsZero S.X₃) : Splitting S where + r := inv S.f + s := 0 + s_g := hg.eq_of_src _ _ + +variable {S} + +/-- The splitting of the short complex `S.op` deduced from a splitting of `S`. -/ +@[simps] +def op (h : Splitting S) : Splitting S.op where + r := h.s.op + s := h.r.op + f_r := Quiver.Hom.unop_inj (by simp) + s_g := Quiver.Hom.unop_inj (by simp) + id := Quiver.Hom.unop_inj (by + simp only [op_X₂, Opposite.unop_op, op_X₁, op_f, op_X₃, op_g, unop_add, unop_comp, + Quiver.Hom.unop_op, unop_id, ← h.id] + abel) + +/-- The splitting of the short complex `S.unop` deduced from a splitting of `S`. -/ +@[simps] +def unop {S : ShortComplex Cᵒᵖ} (h : Splitting S) : Splitting S.unop where + r := h.s.unop + s := h.r.unop + f_r := Quiver.Hom.op_inj (by simp) + s_g := Quiver.Hom.op_inj (by simp) + id := Quiver.Hom.op_inj (by + simp only [unop_X₂, Opposite.op_unop, unop_X₁, unop_f, unop_X₃, unop_g, op_add, + op_comp, Quiver.Hom.op_unop, op_id, ← h.id] + abel) + +/-- The isomorphism `S.X₂ ≅ S.X₁ ⊞ S.X₃` induced by a splitting of the short complex `S`. -/ +@[simps] +noncomputable def isoBinaryBiproduct (h : Splitting S) [HasBinaryBiproduct S.X₁ S.X₃] : + S.X₂ ≅ S.X₁ ⊞ S.X₃ where + hom := biprod.lift h.r S.g + inv := biprod.desc S.f h.s + hom_inv_id := by simp [h.id] + +end Splitting + +section Balanced + +variable {S} +variable [Balanced C] + +namespace Exact + +variable (hS : S.Exact) + +lemma isIso_f' (h : S.LeftHomologyData) [Mono S.f] : + IsIso h.f' := by + have := hS.epi_f' h + have := mono_of_mono_fac h.f'_i + exact isIso_of_mono_of_epi h.f' + +lemma isIso_toCycles [Mono S.f] [S.HasLeftHomology] : + IsIso S.toCycles := + hS.isIso_f' _ + +lemma isIso_g' (h : S.RightHomologyData) [Epi S.g] : + IsIso h.g' := by + have := hS.mono_g' h + have := epi_of_epi_fac h.p_g' + exact isIso_of_mono_of_epi h.g' + +lemma isIso_fromOpcycles [Epi S.g] [S.HasRightHomology] : + IsIso S.fromOpcycles := + hS.isIso_g' _ + +/-- In a balanced category, if a short complex `S` is exact and `S.f` is a mono, then +`S.X₁` is the kernel of `S.g`. -/ +noncomputable def fIsKernel [Mono S.f] : IsLimit (KernelFork.ofι S.f S.zero) := by + have := hS.hasHomology + have := hS.isIso_toCycles + exact IsLimit.ofIsoLimit S.cyclesIsKernel + (Fork.ext (asIso S.toCycles).symm (by simp)) + +lemma map_of_mono_of_preservesKernel (F : C ⥤ D) + [F.PreservesZeroMorphisms] [(S.map F).HasHomology] (_ : Mono S.f) + (_ : PreservesLimit (parallelPair S.g 0) F) : + (S.map F).Exact := + exact_of_f_is_kernel _ (KernelFork.mapIsLimit _ hS.fIsKernel F) + +/-- In a balanced category, if a short complex `S` is exact and `S.g` is an epi, then +`S.X₃` is the cokernel of `S.g`. -/ +noncomputable def gIsCokernel [Epi S.g] : IsColimit (CokernelCofork.ofπ S.g S.zero) := by + have := hS.hasHomology + have := hS.isIso_fromOpcycles + exact IsColimit.ofIsoColimit S.opcyclesIsCokernel + (Cofork.ext (asIso S.fromOpcycles) (by simp)) + +lemma map_of_epi_of_preservesCokernel (F : C ⥤ D) + [F.PreservesZeroMorphisms] [(S.map F).HasHomology] (_ : Epi S.g) + (_ : PreservesColimit (parallelPair S.f 0) F) : + (S.map F).Exact := + exact_of_g_is_cokernel _ (CokernelCofork.mapIsColimit _ hS.gIsCokernel F) + +/-- If a short complex `S` in a balanced category is exact and such that `S.f` is a mono, +then a morphism `k : A ⟶ S.X₂` such that `k ≫ S.g = 0` lifts to a morphism `A ⟶ S.X₁`. -/ +noncomputable def lift {A : C} (k : A ⟶ S.X₂) (hk : k ≫ S.g = 0) [Mono S.f] : + A ⟶ S.X₁ := hS.fIsKernel.lift (KernelFork.ofι k hk) + +@[reassoc (attr := simp)] +lemma lift_f {A : C} (k : A ⟶ S.X₂) (hk : k ≫ S.g = 0) [Mono S.f] : + hS.lift k hk ≫ S.f = k := + Fork.IsLimit.lift_ι _ + +lemma lift' {A : C} (k : A ⟶ S.X₂) (hk : k ≫ S.g = 0) [Mono S.f] : + ∃ (l : A ⟶ S.X₁), l ≫ S.f = k := + ⟨hS.lift k hk, by simp⟩ + +/-- If a short complex `S` in a balanced category is exact and such that `S.g` is an epi, +then a morphism `k : S.X₂ ⟶ A` such that `S.f ≫ k = 0` descends to a morphism `S.X₃ ⟶ A`. -/ +noncomputable def desc {A : C} (k : S.X₂ ⟶ A) (hk : S.f ≫ k = 0) [Epi S.g] : + S.X₃ ⟶ A := hS.gIsCokernel.desc (CokernelCofork.ofπ k hk) + +@[reassoc (attr := simp)] +lemma g_desc {A : C} (k : S.X₂ ⟶ A) (hk : S.f ≫ k = 0) [Epi S.g] : + S.g ≫ hS.desc k hk = k := + Cofork.IsColimit.π_desc (hS.gIsCokernel) + +lemma desc' {A : C} (k : S.X₂ ⟶ A) (hk : S.f ≫ k = 0) [Epi S.g] : + ∃ (l : S.X₃ ⟶ A), S.g ≫ l = k := + ⟨hS.desc k hk, by simp⟩ + +end Exact + +lemma mono_τ₂_of_exact_of_mono {S₁ S₂ : ShortComplex C} (φ : S₁ ⟶ S₂) + (h₁ : S₁.Exact) [Mono S₁.f] [Mono S₂.f] [Mono φ.τ₁] [Mono φ.τ₃] : Mono φ.τ₂ := by + rw [mono_iff_cancel_zero] + intro A x₂ hx₂ + obtain ⟨x₁, hx₁⟩ : ∃ x₁, x₁ ≫ S₁.f = x₂ := ⟨_, h₁.lift_f x₂ + (by simp only [← cancel_mono φ.τ₃, assoc, zero_comp, ← φ.comm₂₃, reassoc_of% hx₂])⟩ + suffices x₁ = 0 by rw [← hx₁, this, zero_comp] + simp only [← cancel_mono φ.τ₁, ← cancel_mono S₂.f, assoc, φ.comm₁₂, zero_comp, + reassoc_of% hx₁, hx₂] + +attribute [local instance] balanced_opposite + +lemma epi_τ₂_of_exact_of_epi {S₁ S₂ : ShortComplex C} (φ : S₁ ⟶ S₂) + (h₂ : S₂.Exact) [Epi S₁.g] [Epi S₂.g] [Epi φ.τ₁] [Epi φ.τ₃] : Epi φ.τ₂ := by + have : Mono S₁.op.f := by dsimp; infer_instance + have : Mono S₂.op.f := by dsimp; infer_instance + have : Mono (opMap φ).τ₁ := by dsimp; infer_instance + have : Mono (opMap φ).τ₃ := by dsimp; infer_instance + have := mono_τ₂_of_exact_of_mono (opMap φ) h₂.op + exact unop_epi_of_mono (opMap φ).τ₂ + +end Balanced + end Preadditive +section Abelian + +variable [Abelian C] + +/-- Given a morphism of short complexes `φ : S₁ ⟶ S₂` in an abelian category, if `S₁.f` +and `S₁.g` are zero (e.g. when `S₁` is of the form `0 ⟶ S₁.X₂ ⟶ 0`) and `S₂.f = 0` +(e.g when `S₂` is of the form `0 ⟶ S₂.X₂ ⟶ S₂.X₃`), then `φ` is a quasi-isomorphism iff +the obvious short complex `S₁.X₂ ⟶ S₂.X₂ ⟶ S₂.X₃` is exact and `φ.τ₂` is a mono). -/ +lemma quasiIso_iff_of_zeros {S₁ S₂ : ShortComplex C} (φ : S₁ ⟶ S₂) + (hf₁ : S₁.f = 0) (hg₁ : S₁.g = 0) (hf₂ : S₂.f = 0) : + QuasiIso φ ↔ + (ShortComplex.mk φ.τ₂ S₂.g (by rw [φ.comm₂₃, hg₁, zero_comp])).Exact ∧ Mono φ.τ₂ := by + have w : φ.τ₂ ≫ S₂.g = 0 := by rw [φ.comm₂₃, hg₁, zero_comp] + rw [quasiIso_iff_isIso_liftCycles φ hf₁ hg₁ hf₂] + constructor + · intro h + have : Mono φ.τ₂ := by + rw [← S₂.liftCycles_i φ.τ₂ w] + apply mono_comp + refine' ⟨_, this⟩ + apply exact_of_f_is_kernel + exact IsLimit.ofIsoLimit S₂.cyclesIsKernel + (Fork.ext (asIso (S₂.liftCycles φ.τ₂ w)).symm (by simp)) + · rintro ⟨h₁, h₂⟩ + refine' ⟨⟨h₁.lift S₂.iCycles (by simp), _, _⟩⟩ + · rw [← cancel_mono φ.τ₂, assoc, h₁.lift_f, liftCycles_i, id_comp] + · rw [← cancel_mono S₂.iCycles, assoc, liftCycles_i, h₁.lift_f, id_comp] + +/-- Given a morphism of short complexes `φ : S₁ ⟶ S₂` in an abelian category, if `S₁.g = 0` +(e.g when `S₁` is of the form `S₁.X₁ ⟶ S₁.X₂ ⟶ 0`) and both `S₂.f` and `S₂.g` are zero +(e.g when `S₂` is of the form `0 ⟶ S₂.X₂ ⟶ 0`), then `φ` is a quasi-isomorphism iff +the obvious short complex `S₁.X₂ ⟶ S₁.X₂ ⟶ S₂.X₂` is exact and `φ.τ₂` is an epi). -/ +lemma quasiIso_iff_of_zeros' {S₁ S₂ : ShortComplex C} (φ : S₁ ⟶ S₂) + (hg₁ : S₁.g = 0) (hf₂ : S₂.f = 0) (hg₂ : S₂.g = 0) : + QuasiIso φ ↔ + (ShortComplex.mk S₁.f φ.τ₂ (by rw [← φ.comm₁₂, hf₂, comp_zero])).Exact ∧ Epi φ.τ₂ := by + rw [← quasiIso_opMap_iff, quasiIso_iff_of_zeros] + rotate_left + · dsimp + rw [hg₂, op_zero] + · dsimp + rw [hf₂, op_zero] + · dsimp + rw [hg₁, op_zero] + rw [← exact_unop_iff] + have : Mono φ.τ₂.op ↔ Epi φ.τ₂ := + ⟨fun _ => unop_epi_of_mono φ.τ₂.op, fun _ => op_mono_of_epi _⟩ + tauto + +end Abelian + +end ShortComplex + +namespace Functor + +variable (F : C ⥤ D) [Preadditive C] [Preadditive D] [HasZeroObject C] + [HasZeroObject D] [F.PreservesZeroMorphisms] [F.PreservesHomology] + +instance : F.PreservesMonomorphisms where + preserves {X Y} f hf := by + let S := ShortComplex.mk (0 : X ⟶ X) f zero_comp + exact ((S.map F).exact_iff_mono (by simp)).1 + (((S.exact_iff_mono rfl).2 hf).map F) + +instance [Faithful F] [CategoryWithHomology C] : F.ReflectsMonomorphisms where + reflects {X Y} f hf := by + let S := ShortComplex.mk (0 : X ⟶ X) f zero_comp + exact (S.exact_iff_mono rfl).1 + ((ShortComplex.exact_map_iff_of_faithful S F).1 + (((S.map F).exact_iff_mono (by simp)).2 hf)) + +instance : F.PreservesEpimorphisms where + preserves {X Y} f hf := by + let S := ShortComplex.mk f (0 : Y ⟶ Y) comp_zero + exact ((S.map F).exact_iff_epi (by simp)).1 + (((S.exact_iff_epi rfl).2 hf).map F) + +instance [Faithful F] [CategoryWithHomology C] : F.ReflectsEpimorphisms where + reflects {X Y} f hf := by + let S := ShortComplex.mk f (0 : Y ⟶ Y) comp_zero + exact (S.exact_iff_epi rfl).1 + ((ShortComplex.exact_map_iff_of_faithful S F).1 + (((S.map F).exact_iff_epi (by simp)).2 hf)) + +end Functor + +namespace ShortComplex + +namespace Splitting + +variable [Preadditive C] [Balanced C] + +/-- This is the splitting of a short complex `S` in a balanced category induced by +a section of the morphism `S.g : S.X₂ ⟶ S.X₃` -/ +noncomputable def ofExactOfSection (S : ShortComplex C) (hS : S.Exact) (s : S.X₃ ⟶ S.X₂) + (s_g : s ≫ S.g = 𝟙 S.X₃) (hf : Mono S.f) : + S.Splitting where + r := hS.lift (𝟙 S.X₂ - S.g ≫ s) (by simp [s_g]) + s := s + f_r := by rw [← cancel_mono S.f, assoc, Exact.lift_f, comp_sub, comp_id, + zero_assoc, zero_comp, sub_zero, id_comp] + s_g := s_g + +/-- This is the splitting of a short complex `S` in a balanced category induced by +a retraction of the morphism `S.f : S.X₁ ⟶ S.X₂` -/ +noncomputable def ofExactOfRetraction (S : ShortComplex C) (hS : S.Exact) (r : S.X₂ ⟶ S.X₁) + (f_r : S.f ≫ r = 𝟙 S.X₁) (hg : Epi S.g) : + S.Splitting where + r := r + s := hS.desc (𝟙 S.X₂ - r ≫ S.f) (by simp [reassoc_of% f_r]) + f_r := f_r + s_g := by + rw [← cancel_epi S.g, Exact.g_desc_assoc, sub_comp, id_comp, assoc, zero, + comp_zero, sub_zero, comp_id] + +end Splitting + end ShortComplex end CategoryTheory diff --git a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean index 075c73870f6bc..b8ee8675f7c24 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ import Mathlib.Algebra.Homology.HomologicalComplex -import Mathlib.Algebra.Homology.ShortComplex.Homology +import Mathlib.Algebra.Homology.ShortComplex.Exact +import Mathlib.Tactic.Linarith /-! # The short complexes attached to homological complexes @@ -14,9 +15,8 @@ In this file, we define a functor By definition, the image of a homological complex `K` by this functor is the short complex `K.X (c.prev i) ⟶ K.X i ⟶ K.X (c.next i)`. -When the homology refactor is completed (TODO @joelriou), the homology -of a homological complex `K` in degree `i` shall be the homology -of the short complex `(shortComplexFunctor C c i).obj K`, which can be +The homology `K.homology i` of a homological complex `K` in degree `i` is defined as +the homology of the short complex `(shortComplexFunctor C c i).obj K`, which can be abbreviated as `K.sc i`. -/ @@ -52,16 +52,552 @@ noncomputable def natIsoSc' (i j k : ι) (hi : c.prev j = i) (hk : c.next j = k) (by aesop_cat) (by aesop_cat)) (by aesop_cat) variable {C c} -variable (K L M : HomologicalComplex C c) (φ : K ⟶ L) (ψ : L ⟶ M) + +section + +variable (K L M : HomologicalComplex C c) (φ : K ⟶ L) (ψ : L ⟶ M) (i j k : ι) /-- The short complex `K.X i ⟶ K.X j ⟶ K.X k` for arbitrary indices `i`, `j` and `k`. -/ -abbrev sc' (i j k : ι) := (shortComplexFunctor' C c i j k).obj K +abbrev sc' := (shortComplexFunctor' C c i j k).obj K /-- The short complex `K.X (c.prev i) ⟶ K.X i ⟶ K.X (c.next i)`. -/ -noncomputable abbrev sc (i : ι) := (shortComplexFunctor C c i).obj K +noncomputable abbrev sc := (shortComplexFunctor C c i).obj K /-- The canonical isomorphism `K.sc j ≅ K.sc' i j k` when `c.prev j = i` and `c.next j = k`. -/ -noncomputable abbrev isoSc' (i j k : ι) (hi : c.prev j = i) (hk : c.next j = k) : +noncomputable abbrev isoSc' (hi : c.prev j = i) (hk : c.next j = k) : K.sc j ≅ K.sc' i j k := (natIsoSc' C c i j k hi hk).app K +/-- A homological complex `K` has homology in degree `i` if the associated +short complex `K.sc i` has. -/ +abbrev HasHomology := (K.sc i).HasHomology + +variable [K.HasHomology i] + +/-- The homology in degree `i` of a homological complex. -/ +noncomputable def homology := (K.sc i).homology + +/-- Comparison isomorphism between the homology for the two homology API. -/ +noncomputable def homology'IsoHomology {A : Type*} [Category A] [Abelian A] + (K : HomologicalComplex A c) (i : ι) : + K.homology' i ≅ K.homology i := + (K.sc i).homology'IsoHomology + +/-- The cycles in degree `i` of a homological complex. -/ +noncomputable def cycles := (K.sc i).cycles + +/-- The inclusion of the cycles of a homological complex. -/ +noncomputable def iCycles : K.cycles i ⟶ K.X i := (K.sc i).iCycles + +/-- The homology class map from cycles to the homology of a homological complex. -/ +noncomputable def homologyπ : K.cycles i ⟶ K.homology i := (K.sc i).homologyπ + +variable {i} + +/-- The morphism to `K.cycles i` that is induced by a "cycle", i.e. a morphism +to `K.X i` whose postcomposition with the differential is zero. -/ +noncomputable def liftCycles {A : C} (k : A ⟶ K.X i) (j : ι) (hj : c.next i = j) + (hk : k ≫ K.d i j = 0) : A ⟶ K.cycles i := + (K.sc i).liftCycles k (by subst hj; exact hk) + +/-- The morphism to `K.cycles i` that is induced by a "cycle", i.e. a morphism +to `K.X i` whose postcomposition with the differential is zero. -/ +@[reducible] +noncomputable def liftCycles' {A : C} (k : A ⟶ K.X i) (j : ι) (hj : c.Rel i j) + (hk : k ≫ K.d i j = 0) : A ⟶ K.cycles i := + K.liftCycles k j (c.next_eq' hj) hk + +@[reassoc (attr := simp)] +lemma liftCycles_i {A : C} (k : A ⟶ K.X i) (j : ι) (hj : c.next i = j) + (hk : k ≫ K.d i j = 0) : K.liftCycles k j hj hk ≫ K.iCycles i = k := by + dsimp [liftCycles, iCycles] + simp + +variable (i) + +/-- The map `K.X i ⟶ K.cycles j` induced by the differential `K.d i j`. -/ +noncomputable def toCycles [K.HasHomology j] : + K.X i ⟶ K.cycles j := + K.liftCycles (K.d i j) (c.next j) rfl (K.d_comp_d _ _ _) + +@[reassoc (attr := simp)] +lemma iCycles_d : K.iCycles i ≫ K.d i j = 0 := by + by_cases hij : c.Rel i j + · obtain rfl := c.next_eq' hij + exact (K.sc i).iCycles_g + · rw [K.shape _ _ hij, comp_zero] + +/-- `K.cycles i` is the kernel of `K.d i j` when `c.next i = j`. -/ +noncomputable def cyclesIsKernel (hj : c.next i = j) : + IsLimit (KernelFork.ofι (K.iCycles i) (K.iCycles_d i j)) := by + obtain rfl := hj + exact (K.sc i).cyclesIsKernel + +@[reassoc (attr := simp)] +lemma toCycles_i [K.HasHomology j] : + K.toCycles i j ≫ K.iCycles j = K.d i j := + liftCycles_i _ _ _ _ _ + +instance : Mono (K.iCycles i) := by + dsimp only [iCycles] + infer_instance + +instance : Epi (K.homologyπ i) := by + dsimp only [homologyπ] + infer_instance + +@[reassoc (attr := simp)] +lemma d_toCycles [K.HasHomology k] : + K.d i j ≫ K.toCycles j k = 0 := by + simp only [← cancel_mono (K.iCycles k), assoc, toCycles_i, d_comp_d, zero_comp] + +variable {i} + +@[reassoc] +lemma comp_liftCycles {A' A : C} (k : A ⟶ K.X i) (j : ι) (hj : c.next i = j) + (hk : k ≫ K.d i j = 0) (α : A' ⟶ A) : + α ≫ K.liftCycles k j hj hk = K.liftCycles (α ≫ k) j hj (by rw [assoc, hk, comp_zero]) := by + simp only [← cancel_mono (K.iCycles i), assoc, liftCycles_i] + +@[reassoc] +lemma liftCycles_homologyπ_eq_zero_of_boundary {A : C} (k : A ⟶ K.X i) (j : ι) + (hj : c.next i = j) {i' : ι} (x : A ⟶ K.X i') (hx : k = x ≫ K.d i' i) : + K.liftCycles k j hj (by rw [hx, assoc, K.d_comp_d, comp_zero]) ≫ K.homologyπ i = 0 := by + by_cases c.Rel i' i + · obtain rfl := c.prev_eq' h + exact (K.sc i).liftCycles_homologyπ_eq_zero_of_boundary _ x hx + · have : liftCycles K k j hj (by rw [hx, assoc, K.d_comp_d, comp_zero]) = 0 := by + rw [K.shape _ _ h, comp_zero] at hx + rw [← cancel_mono (K.iCycles i), zero_comp, liftCycles_i, hx] + rw [this, zero_comp] + +variable (i) + +@[reassoc (attr := simp)] +lemma toCycles_comp_homologyπ [K.HasHomology j] : + K.toCycles i j ≫ K.homologyπ j = 0 := + K.liftCycles_homologyπ_eq_zero_of_boundary (K.d i j) (c.next j) rfl (𝟙 _) (by simp) + +/-- `K.homology j` is the cokernel of `K.toCycles i j : K.X i ⟶ K.cycles j` +when `c.prev j = i`. -/ +noncomputable def homologyIsCokernel (hi : c.prev j = i) [K.HasHomology j] : + IsColimit (CokernelCofork.ofπ (K.homologyπ j) (K.toCycles_comp_homologyπ i j)) := by + subst hi + exact ((K.sc j).homologyIsCokernel) + +/-- The opcycles in degree `i` of a homological complex. -/ +noncomputable def opcycles := (K.sc i).opcycles + +/-- The projection to the opcycles of a homological complex. -/ +noncomputable def pOpcycles : K.X i ⟶ K.opcycles i := (K.sc i).pOpcycles + +/-- The inclusion map of the homology of a homological complex into its opcycles. -/ +noncomputable def homologyι : K.homology i ⟶ K.opcycles i := (K.sc i).homologyι + +variable {i} + +/-- The morphism from `K.opcycles i` that is induced by an "opcycle", i.e. a morphism +from `K.X i` whose precomposition with the differential is zero. -/ +noncomputable def descOpcycles {A : C} (k : K.X i ⟶ A) (j : ι) (hj : c.prev i = j) + (hk : K.d j i ≫ k = 0) : K.opcycles i ⟶ A := + (K.sc i).descOpcycles k (by subst hj; exact hk) + +/-- The morphism from `K.opcycles i` that is induced by an "opcycle", i.e. a morphism +from `K.X i` whose precomposition with the differential is zero. -/ +@[reducible] +noncomputable def descOpcycles' {A : C} (k : K.X i ⟶ A) (j : ι) (hj : c.Rel j i) + (hk : K.d j i ≫ k = 0) : K.opcycles i ⟶ A := + K.descOpcycles k j (c.prev_eq' hj) hk + +@[reassoc (attr := simp)] +lemma p_descOpcycles {A : C} (k : K.X i ⟶ A) (j : ι) (hj : c.prev i = j) + (hk : K.d j i ≫ k = 0) : K.pOpcycles i ≫ K.descOpcycles k j hj hk = k := by + dsimp [descOpcycles, pOpcycles] + simp + +variable (i) + +/-- The map `K.opcycles i ⟶ K.X j` induced by the differential `K.d i j`. -/ +noncomputable def fromOpcycles : + K.opcycles i ⟶ K.X j := + K.descOpcycles (K.d i j) (c.prev i) rfl (K.d_comp_d _ _ _) + +@[reassoc (attr := simp)] +lemma d_pOpcycles [K.HasHomology j] : K.d i j ≫ K.pOpcycles j = 0 := by + by_cases hij : c.Rel i j + · obtain rfl := c.prev_eq' hij + exact (K.sc j).f_pOpcycles + · rw [K.shape _ _ hij, zero_comp] + +/-- `K.opcycles j` is the cokernel of `K.d i j` when `c.prev j = i`. -/ +noncomputable def opcyclesIsCokernel (hi : c.prev j = i) [K.HasHomology j] : + IsColimit (CokernelCofork.ofπ (K.pOpcycles j) (K.d_pOpcycles i j)) := by + obtain rfl := hi + exact (K.sc j).opcyclesIsCokernel + +@[reassoc (attr := simp)] +lemma p_fromOpcycles : + K.pOpcycles i ≫ K.fromOpcycles i j = K.d i j := + p_descOpcycles _ _ _ _ _ + +instance : Epi (K.pOpcycles i) := by + dsimp only [pOpcycles] + infer_instance + +instance : Mono (K.homologyι i) := by + dsimp only [homologyι] + infer_instance + +@[reassoc (attr := simp)] +lemma fromOpcycles_d : + K.fromOpcycles i j ≫ K.d j k = 0 := by + simp only [← cancel_epi (K.pOpcycles i), p_fromOpcycles_assoc, d_comp_d, comp_zero] + +variable {i} + +@[reassoc] +lemma descOpcycles_comp {A A' : C} (k : K.X i ⟶ A) (j : ι) (hj : c.prev i = j) + (hk : K.d j i ≫ k = 0) (α : A ⟶ A') : + K.descOpcycles k j hj hk ≫ α = K.descOpcycles (k ≫ α) j hj + (by rw [reassoc_of% hk, zero_comp]) := by + simp only [← cancel_epi (K.pOpcycles i), p_descOpcycles_assoc, p_descOpcycles] + +@[reassoc] +lemma homologyι_descOpcycles_eq_zero_of_boundary {A : C} (k : K.X i ⟶ A) (j : ι) + (hj : c.prev i = j) {i' : ι} (x : K.X i' ⟶ A) (hx : k = K.d i i' ≫ x) : + K.homologyι i ≫ K.descOpcycles k j hj (by rw [hx, K.d_comp_d_assoc, zero_comp]) = 0 := by + by_cases c.Rel i i' + · obtain rfl := c.next_eq' h + exact (K.sc i).homologyι_descOpcycles_eq_zero_of_boundary _ x hx + · have : K.descOpcycles k j hj (by rw [hx, K.d_comp_d_assoc, zero_comp]) = 0 := by + rw [K.shape _ _ h, zero_comp] at hx + rw [← cancel_epi (K.pOpcycles i), comp_zero, p_descOpcycles, hx] + rw [this, comp_zero] + +variable (i) + +@[reassoc (attr := simp)] +lemma homologyι_comp_fromOpcycles : + K.homologyι i ≫ K.fromOpcycles i j = 0 := + K.homologyι_descOpcycles_eq_zero_of_boundary (K.d i j) _ rfl (𝟙 _) (by simp) + +/-- `K.homology i` is the kernel of `K.fromOpcycles i j : K.opcycles i ⟶ K.X j` +when `c.next i = j`. -/ +noncomputable def homologyIsKernel (hi : c.next i = j) : + IsLimit (KernelFork.ofι (K.homologyι i) (K.homologyι_comp_fromOpcycles i j)) := by + subst hi + exact (K.sc i).homologyIsKernel + +variable {K L M} +variable [L.HasHomology i] [M.HasHomology i] + +/-- The map `K.homology i ⟶ L.homology i` induced by a morphism in `HomologicalComplex`. -/ +noncomputable def homologyMap : K.homology i ⟶ L.homology i := + ShortComplex.homologyMap ((shortComplexFunctor C c i).map φ) + +/-- The map `K.cycles i ⟶ L.cycles i` induced by a morphism in `HomologicalComplex`. -/ +noncomputable def cyclesMap : K.cycles i ⟶ L.cycles i := + ShortComplex.cyclesMap ((shortComplexFunctor C c i).map φ) + +/-- The map `K.opcycles i ⟶ L.opcycles i` induced by a morphism in `HomologicalComplex`. -/ +noncomputable def opcyclesMap : K.opcycles i ⟶ L.opcycles i := + ShortComplex.opcyclesMap ((shortComplexFunctor C c i).map φ) + +@[reassoc (attr := simp)] +lemma cyclesMap_i : cyclesMap φ i ≫ L.iCycles i = K.iCycles i ≫ φ.f i := + ShortComplex.cyclesMap_i _ + +@[reassoc (attr := simp)] +lemma p_opcyclesMap : K.pOpcycles i ≫ opcyclesMap φ i = φ.f i ≫ L.pOpcycles i := + ShortComplex.p_opcyclesMap _ + +variable (K) + +@[simp] +lemma homologyMap_id : homologyMap (𝟙 K) i = 𝟙 _ := + ShortComplex.homologyMap_id _ + +@[simp] +lemma cyclesMap_id : cyclesMap (𝟙 K) i = 𝟙 _ := + ShortComplex.cyclesMap_id _ + +@[simp] +lemma opcyclesMap_id : opcyclesMap (𝟙 K) i = 𝟙 _ := + ShortComplex.opcyclesMap_id _ + +variable {K} + +@[reassoc] +lemma homologyMap_comp : homologyMap (φ ≫ ψ) i = homologyMap φ i ≫ homologyMap ψ i := by + dsimp [homologyMap] + rw [Functor.map_comp, ShortComplex.homologyMap_comp] + +@[reassoc] +lemma cyclesMap_comp : cyclesMap (φ ≫ ψ) i = cyclesMap φ i ≫ cyclesMap ψ i := by + dsimp [cyclesMap] + rw [Functor.map_comp, ShortComplex.cyclesMap_comp] + +@[reassoc] +lemma opcyclesMap_comp : opcyclesMap (φ ≫ ψ) i = opcyclesMap φ i ≫ opcyclesMap ψ i := by + dsimp [opcyclesMap] + rw [Functor.map_comp, ShortComplex.opcyclesMap_comp] + +variable (K L) + +@[simp] +lemma homologyMap_zero : homologyMap (0 : K ⟶ L) i = 0 := + ShortComplex.homologyMap_zero _ _ + +@[simp] +lemma cyclesMap_zero : cyclesMap (0 : K ⟶ L) i = 0 := + ShortComplex.cyclesMap_zero _ _ + +@[simp] +lemma opcyclesMap_zero : opcyclesMap (0 : K ⟶ L) i = 0 := + ShortComplex.opcyclesMap_zero _ _ + +variable {K L} + +@[reassoc (attr := simp)] +lemma homologyπ_naturality : + K.homologyπ i ≫ homologyMap φ i = cyclesMap φ i ≫ L.homologyπ i := + ShortComplex.homologyπ_naturality _ + +@[reassoc (attr := simp)] +lemma homologyι_naturality : + homologyMap φ i ≫ L.homologyι i = K.homologyι i ≫ opcyclesMap φ i := + ShortComplex.homologyι_naturality _ + +@[reassoc (attr := simp)] +lemma homology_π_ι : + K.homologyπ i ≫ K.homologyι i = K.iCycles i ≫ K.pOpcycles i := + (K.sc i).homology_π_ι + +variable {i} + +@[reassoc (attr := simp)] +lemma opcyclesMap_comp_descOpcycles {A : C} (k : L.X i ⟶ A) (j : ι) (hj : c.prev i = j) + (hk : L.d j i ≫ k = 0) (φ : K ⟶ L) : + opcyclesMap φ i ≫ L.descOpcycles k j hj hk = K.descOpcycles (φ.f i ≫ k) j hj + (by rw [← φ.comm_assoc, hk, comp_zero]) := by + simp only [← cancel_epi (K.pOpcycles i), p_opcyclesMap_assoc, p_descOpcycles] + +@[reassoc (attr := simp)] +lemma liftCycles_comp_cyclesMap {A : C} (k : A ⟶ K.X i) (j : ι) (hj : c.next i = j) + (hk : k ≫ K.d i j = 0) (φ : K ⟶ L) : + K.liftCycles k j hj hk ≫ cyclesMap φ i = L.liftCycles (k ≫ φ.f i) j hj + (by rw [assoc, φ.comm, reassoc_of% hk, zero_comp]) := by + simp only [← cancel_mono (L.iCycles i), assoc, cyclesMap_i, liftCycles_i_assoc, liftCycles_i] + +section + +variable (C c i) + +attribute [local simp] homologyMap_comp cyclesMap_comp opcyclesMap_comp + +/-- The `i`th homology functor `HomologicalComplex C c ⥤ C`. -/ +@[simps] +noncomputable def homologyFunctor [CategoryWithHomology C] : HomologicalComplex C c ⥤ C where + obj K := K.homology i + map f := homologyMap f i + +/-- The homology functor to graded objects. -/ +@[simps] +noncomputable def gradedHomologyFunctor [CategoryWithHomology C] : + HomologicalComplex C c ⥤ GradedObject ι C where + obj K i := K.homology i + map f i := homologyMap f i + +/-- The `i`th cycles functor `HomologicalComplex C c ⥤ C`. -/ +@[simps] +noncomputable def cyclesFunctor [CategoryWithHomology C] : HomologicalComplex C c ⥤ C where + obj K := K.cycles i + map f := cyclesMap f i + +/-- The `i`th opcycles functor `HomologicalComplex C c ⥤ C`. -/ +@[simps] +noncomputable def opcyclesFunctor [CategoryWithHomology C] : HomologicalComplex C c ⥤ C where + obj K := K.opcycles i + map f := opcyclesMap f i + +/-- The natural isomorphism `K.homology i ≅ (K.sc i).homology` +for all homological complexes `K`. -/ +@[simps!] +noncomputable def homologyFunctorIso [CategoryWithHomology C] : + homologyFunctor C c i ≅ + shortComplexFunctor C c i ⋙ ShortComplex.homologyFunctor C := + Iso.refl _ + +/-- The natural isomorphism `K.homology j ≅ (K.sc' i j k).homology` +for all homological complexes `K` when `c.prev j = i` and `c.next j = k`. -/ +noncomputable def homologyFunctorIso' [CategoryWithHomology C] + (hi : c.prev j = i) (hk : c.next j = k) : + homologyFunctor C c j ≅ + shortComplexFunctor' C c i j k ⋙ ShortComplex.homologyFunctor C := + homologyFunctorIso C c j ≪≫ isoWhiskerRight (natIsoSc' C c i j k hi hk) _ + +end + +end + +variable (K : HomologicalComplex C c) (i j k : ι) + +section + +variable (hj : c.next i = j) (h : K.d i j = 0) [K.HasHomology i] + +lemma isIso_iCycles : IsIso (K.iCycles i) := by + subst hj + exact ShortComplex.isIso_iCycles _ h + +/-- The canonical isomorphism `K.cycles i ≅ K.X i` when the differential from `i` is zero. -/ +@[simps! hom] +noncomputable def iCyclesIso : K.cycles i ≅ K.X i := + have := K.isIso_iCycles i j hj h + asIso (K.iCycles i) + +@[reassoc (attr := simp)] +lemma iCyclesIso_hom_inv_id : + K.iCycles i ≫ (K.iCyclesIso i j hj h).inv = 𝟙 _ := + (K.iCyclesIso i j hj h).hom_inv_id + +@[reassoc (attr := simp)] +lemma iCyclesIso_inv_hom_id : + (K.iCyclesIso i j hj h).inv ≫ K.iCycles i = 𝟙 _ := + (K.iCyclesIso i j hj h).inv_hom_id + +lemma isIso_homologyι : IsIso (K.homologyι i) := + ShortComplex.isIso_homologyι _ (by aesop_cat) + +/-- The canonical isomorphism `K.homology i ≅ K.opcycles i` +when the differential from `i` is zero. -/ +@[simps! hom] +noncomputable def isoHomologyι : K.homology i ≅ K.opcycles i := + have := K.isIso_homologyι i j hj h + asIso (K.homologyι i) + +@[reassoc (attr := simp)] +lemma isoHomologyι_hom_inv_id : + K.homologyι i ≫ (K.isoHomologyι i j hj h).inv = 𝟙 _ := + (K.isoHomologyι i j hj h).hom_inv_id + +@[reassoc (attr := simp)] +lemma isoHomologyι_inv_hom_id : + (K.isoHomologyι i j hj h).inv ≫ K.homologyι i = 𝟙 _ := + (K.isoHomologyι i j hj h).inv_hom_id + +end + +section + +variable (hi : c.prev j = i) (h : K.d i j = 0) [K.HasHomology j] + +lemma isIso_pOpcycles : IsIso (K.pOpcycles j) := by + obtain rfl := hi + exact ShortComplex.isIso_pOpcycles _ h + +/-- The canonical isomorphism `K.X j ≅ K.opCycles j` when the differential to `j` is zero. -/ +@[simps! hom] +noncomputable def pOpcyclesIso : K.X j ≅ K.opcycles j := + have := K.isIso_pOpcycles i j hi h + asIso (K.pOpcycles j) + +@[reassoc (attr := simp)] +lemma pOpcyclesIso_hom_inv_id : + K.pOpcycles j ≫ (K.pOpcyclesIso i j hi h).inv = 𝟙 _ := + (K.pOpcyclesIso i j hi h).hom_inv_id + +@[reassoc (attr := simp)] +lemma pOpcyclesIso_inv_hom_id : + (K.pOpcyclesIso i j hi h).inv ≫ K.pOpcycles j = 𝟙 _ := + (K.pOpcyclesIso i j hi h).inv_hom_id + +lemma isIso_homologyπ : IsIso (K.homologyπ j) := + ShortComplex.isIso_homologyπ _ (by aesop_cat) + +/-- The canonical isomorphism `K.cycles j ≅ K.homology j` +when the differential to `j` is zero. -/ +@[simps! hom] +noncomputable def isoHomologyπ : K.cycles j ≅ K.homology j := + have := K.isIso_homologyπ i j hi h + asIso (K.homologyπ j) + +@[reassoc (attr := simp)] +lemma isoHomologyπ_hom_inv_id : + K.homologyπ j ≫ (K.isoHomologyπ i j hi h).inv = 𝟙 _ := + (K.isoHomologyπ i j hi h).hom_inv_id + +@[reassoc (attr := simp)] +lemma isoHomologyπ_inv_hom_id : + (K.isoHomologyπ i j hi h).inv ≫ K.homologyπ j = 𝟙 _ := + (K.isoHomologyπ i j hi h).inv_hom_id + +end + +/-- A homological complex `K` is exact at `i` if the short complex `K.sc i` is exact. -/ +def ExactAt := (K.sc i).Exact + +lemma exactAt_iff : + K.ExactAt i ↔ (K.sc i).Exact := by rfl + +lemma exactAt_iff' (hi : c.prev j = i) (hk : c.next j = k) : + K.ExactAt j ↔ (K.sc' i j k).Exact := + ShortComplex.exact_iff_of_iso (K.isoSc' i j k hi hk) + +lemma exactAt_iff_isZero_homology [K.HasHomology i] : + K.ExactAt i ↔ IsZero (K.homology i) := by + dsimp [homology] + rw [exactAt_iff, ShortComplex.exact_iff_isZero_homology] + end HomologicalComplex + +namespace ChainComplex + +variable {C : Type*} [Category C] [HasZeroMorphisms C] + (K L : ChainComplex C ℕ) (φ : K ⟶ L) [K.HasHomology 0] + +instance isIso_homologyι₀ : + IsIso (K.homologyι 0) := + K.isIso_homologyι 0 _ rfl (by simp) + +/-- The canonical isomorphism `K.homology 0 ≅ K.opcycles 0` for a chain complex `K` +indexed by `ℕ`. -/ +noncomputable abbrev isoHomologyι₀ : + K.homology 0 ≅ K.opcycles 0 := K.isoHomologyι 0 _ rfl (by simp) + +variable {K L} + +@[reassoc (attr := simp)] +lemma isoHomologyι₀_inv_naturality [L.HasHomology 0] : + K.isoHomologyι₀.inv ≫ HomologicalComplex.homologyMap φ 0 = + HomologicalComplex.opcyclesMap φ 0 ≫ L.isoHomologyι₀.inv := by + simp only [assoc, ← cancel_mono (L.homologyι 0), + HomologicalComplex.homologyι_naturality, HomologicalComplex.isoHomologyι_inv_hom_id_assoc, + HomologicalComplex.isoHomologyι_inv_hom_id, comp_id] + +end ChainComplex + +namespace CochainComplex + +variable {C : Type*} [Category C] [HasZeroMorphisms C] + (K L : CochainComplex C ℕ) (φ : K ⟶ L) [K.HasHomology 0] + +instance isIso_homologyπ₀ : + IsIso (K.homologyπ 0) := + K.isIso_homologyπ _ 0 rfl (by simp) + +/-- The canonical isomorphism `K.cycles 0 ≅ K.homology 0` for a cochain complex `K` +indexed by `ℕ`. -/ +noncomputable abbrev isoHomologyπ₀ : + K.cycles 0 ≅ K.homology 0 := K.isoHomologyπ _ 0 rfl (by simp) + +variable {K L} + +@[reassoc (attr := simp)] +lemma isoHomologyπ₀_inv_naturality [L.HasHomology 0] : + HomologicalComplex.homologyMap φ 0 ≫ L.isoHomologyπ₀.inv = + K.isoHomologyπ₀.inv ≫ HomologicalComplex.cyclesMap φ 0 := by + simp only [← cancel_epi (K.homologyπ 0), HomologicalComplex.homologyπ_naturality_assoc, + HomologicalComplex.isoHomologyπ_hom_inv_id, comp_id, + HomologicalComplex.isoHomologyπ_hom_inv_id_assoc] + +end CochainComplex diff --git a/Mathlib/Algebra/Homology/ShortComplex/Homology.lean b/Mathlib/Algebra/Homology/ShortComplex/Homology.lean index dd79d24c27ac1..5b37a37288d75 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Homology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Homology.lean @@ -1059,7 +1059,7 @@ lemma liftCycles_homologyπ_eq_zero_of_boundary [S.HasHomology] rw [S.liftCycles_leftHomologyπ_eq_zero_of_boundary_assoc k x hx, zero_comp] @[reassoc] -lemma homologyι_descOpcycles_π_eq_zero_of_boundary [S.HasHomology] +lemma homologyι_descOpcycles_eq_zero_of_boundary [S.HasHomology] (k : S.X₂ ⟶ A) (x : S.X₃ ⟶ A) (hx : k = S.g ≫ x) : S.homologyι ≫ S.descOpcycles k (by rw [hx, S.zero_assoc, zero_comp]) = 0 := by dsimp only [homologyι] @@ -1123,7 +1123,7 @@ lemma asIsoHomologyπ_inv_comp_homologyπ (hf : S.f = 0) [S.HasHomology] : @[reassoc (attr := simp)] lemma homologyπ_comp_asIsoHomologyπ_inv (hf : S.f = 0) [S.HasHomology] : - S.homologyπ ≫ (S.asIsoHomologyπ hf).inv = 𝟙 _ := (S.asIsoHomologyπ hf).hom_inv_id + S.homologyπ ≫ (S.asIsoHomologyπ hf).inv = 𝟙 _ := (S.asIsoHomologyπ hf).hom_inv_id /-- The canonical isomorphism `S.homology ≅ S.opcycles` when `S.g = 0`. -/ @[simps! hom] diff --git a/Mathlib/Algebra/Homology/ShortComplex/Preadditive.lean b/Mathlib/Algebra/Homology/ShortComplex/Preadditive.lean index 431a26dc9db07..d2082f8b3ac79 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Preadditive.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Preadditive.lean @@ -5,6 +5,7 @@ Authors: Joël Riou -/ import Mathlib.Algebra.Homology.ShortComplex.Homology import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor +import Mathlib.CategoryTheory.Preadditive.Opposite /-! # Homology of preadditive categories @@ -12,8 +13,6 @@ import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor In this file, it is shown that if `C` is a preadditive category, then `ShortComplex C` is a preadditive category. -TODO: Introduce the notion of homotopy of morphisms of short complexes. - -/ namespace CategoryTheory @@ -74,15 +73,15 @@ namespace LeftHomologyMapData variable (γ : LeftHomologyMapData φ h₁ h₂) (γ' : LeftHomologyMapData φ' h₁ h₂) -/-- Given a left homology map data for morphism `φ`, this is induced left homology +/-- Given a left homology map data for morphism `φ`, this is the induced left homology map data for `-φ`. -/ @[simps] def neg : LeftHomologyMapData (-φ) h₁ h₂ where φK := -γ.φK φH := -γ.φH -/-- Given left homology map data for morphisms `φ` and `φ'`, this is induced left homology -map data for `φ + φ'`. -/ +/-- Given left homology map data for morphisms `φ` and `φ'`, this is +the induced left homology map data for `φ + φ'`. -/ @[simps] def add : LeftHomologyMapData (φ + φ') h₁ h₂ where φK := γ.φK + γ'.φK @@ -158,7 +157,7 @@ lemma cyclesMap_add : cyclesMap (φ + φ') = cyclesMap φ + cyclesMap φ' := @[simp] lemma leftHomologyMap_sub : leftHomologyMap (φ - φ') = leftHomologyMap φ - leftHomologyMap φ' := - leftHomologyMap'_sub _ _ + leftHomologyMap'_sub _ _ @[simp] lemma cyclesMap_sub : cyclesMap (φ - φ') = cyclesMap φ - cyclesMap φ' := @@ -174,6 +173,555 @@ instance cyclesFunctor_additive [HasKernels C] [HasCokernels C] : end LeftHomology + +section RightHomology + +variable {φ φ' : S₁ ⟶ S₂} {h₁ : S₁.RightHomologyData} {h₂ : S₂.RightHomologyData} + +namespace RightHomologyMapData + +variable (γ : RightHomologyMapData φ h₁ h₂) (γ' : RightHomologyMapData φ' h₁ h₂) + +/-- Given a right homology map data for morphism `φ`, this is the induced right homology +map data for `-φ`. -/ +@[simps] +def neg : RightHomologyMapData (-φ) h₁ h₂ where + φQ := -γ.φQ + φH := -γ.φH + +/-- Given right homology map data for morphisms `φ` and `φ'`, this is the induced +right homology map data for `φ + φ'`. -/ +@[simps] +def add : RightHomologyMapData (φ + φ') h₁ h₂ where + φQ := γ.φQ + γ'.φQ + φH := γ.φH + γ'.φH + +end RightHomologyMapData + +variable (h₁ h₂) + +@[simp] +lemma rightHomologyMap'_neg : + rightHomologyMap' (-φ) h₁ h₂ = -rightHomologyMap' φ h₁ h₂ := by + have γ : RightHomologyMapData φ h₁ h₂ := default + simp only [γ.rightHomologyMap'_eq, γ.neg.rightHomologyMap'_eq, RightHomologyMapData.neg_φH] + +@[simp] +lemma opcyclesMap'_neg : + opcyclesMap' (-φ) h₁ h₂ = -opcyclesMap' φ h₁ h₂ := by + have γ : RightHomologyMapData φ h₁ h₂ := default + simp only [γ.opcyclesMap'_eq, γ.neg.opcyclesMap'_eq, RightHomologyMapData.neg_φQ] + +@[simp] +lemma rightHomologyMap'_add : + rightHomologyMap' (φ + φ') h₁ h₂ = rightHomologyMap' φ h₁ h₂ + + rightHomologyMap' φ' h₁ h₂ := by + have γ : RightHomologyMapData φ h₁ h₂ := default + have γ' : RightHomologyMapData φ' h₁ h₂ := default + simp only [γ.rightHomologyMap'_eq, γ'.rightHomologyMap'_eq, + (γ.add γ').rightHomologyMap'_eq, RightHomologyMapData.add_φH] + +@[simp] +lemma opcyclesMap'_add : + opcyclesMap' (φ + φ') h₁ h₂ = opcyclesMap' φ h₁ h₂ + + opcyclesMap' φ' h₁ h₂ := by + have γ : RightHomologyMapData φ h₁ h₂ := default + have γ' : RightHomologyMapData φ' h₁ h₂ := default + simp only [γ.opcyclesMap'_eq, γ'.opcyclesMap'_eq, + (γ.add γ').opcyclesMap'_eq, RightHomologyMapData.add_φQ] + +@[simp] +lemma rightHomologyMap'_sub : + rightHomologyMap' (φ - φ') h₁ h₂ = rightHomologyMap' φ h₁ h₂ - + rightHomologyMap' φ' h₁ h₂ := by + simp only [sub_eq_add_neg, rightHomologyMap'_add, rightHomologyMap'_neg] + +@[simp] +lemma opcyclesMap'_sub : + opcyclesMap' (φ - φ') h₁ h₂ = opcyclesMap' φ h₁ h₂ - + opcyclesMap' φ' h₁ h₂ := by + simp only [sub_eq_add_neg, opcyclesMap'_add, opcyclesMap'_neg] + +variable (φ φ') + +section + +variable [S₁.HasRightHomology] [S₂.HasRightHomology] + +@[simp] +lemma rightHomologyMap_neg : rightHomologyMap (-φ) = -rightHomologyMap φ := + rightHomologyMap'_neg _ _ + +@[simp] +lemma opcyclesMap_neg : opcyclesMap (-φ) = -opcyclesMap φ := + opcyclesMap'_neg _ _ + +@[simp] +lemma rightHomologyMap_add : + rightHomologyMap (φ + φ') = rightHomologyMap φ + rightHomologyMap φ' := + rightHomologyMap'_add _ _ + +@[simp] +lemma opcyclesMap_add : opcyclesMap (φ + φ') = opcyclesMap φ + opcyclesMap φ' := + opcyclesMap'_add _ _ + +@[simp] +lemma rightHomologyMap_sub : + rightHomologyMap (φ - φ') = rightHomologyMap φ - rightHomologyMap φ' := + rightHomologyMap'_sub _ _ + +@[simp] +lemma opcyclesMap_sub : opcyclesMap (φ - φ') = opcyclesMap φ - opcyclesMap φ' := + opcyclesMap'_sub _ _ + +end + +instance rightHomologyFunctor_additive [HasKernels C] [HasCokernels C] : + (rightHomologyFunctor C).Additive where + +instance opcyclesFunctor_additive [HasKernels C] [HasCokernels C] : + (opcyclesFunctor C).Additive where + +end RightHomology + +section Homology + +variable {φ φ' : S₁ ⟶ S₂} {h₁ : S₁.HomologyData} {h₂ : S₂.HomologyData} + +namespace HomologyMapData + +variable (γ : HomologyMapData φ h₁ h₂) (γ' : HomologyMapData φ' h₁ h₂) + +/-- Given a homology map data for a morphism `φ`, this is the induced homology +map data for `-φ`. -/ +@[simps] +def neg : HomologyMapData (-φ) h₁ h₂ where + left := γ.left.neg + right := γ.right.neg + +/-- Given homology map data for morphisms `φ` and `φ'`, this is the induced homology +map data for `φ + φ'`. -/ +@[simps] +def add : HomologyMapData (φ + φ') h₁ h₂ where + left := γ.left.add γ'.left + right := γ.right.add γ'.right + +end HomologyMapData + +variable (h₁ h₂) + +@[simp] +lemma homologyMap'_neg : + homologyMap' (-φ) h₁ h₂ = -homologyMap' φ h₁ h₂ := + leftHomologyMap'_neg _ _ + +@[simp] +lemma homologyMap'_add : + homologyMap' (φ + φ') h₁ h₂ = homologyMap' φ h₁ h₂ + homologyMap' φ' h₁ h₂ := + leftHomologyMap'_add _ _ + +@[simp] +lemma homologyMap'_sub : + homologyMap' (φ - φ') h₁ h₂ = homologyMap' φ h₁ h₂ - homologyMap' φ' h₁ h₂ := + leftHomologyMap'_sub _ _ + +variable (φ φ') + +section + +variable [S₁.HasHomology] [S₂.HasHomology] + +@[simp] +lemma homologyMap_neg : homologyMap (-φ) = -homologyMap φ := + homologyMap'_neg _ _ + +@[simp] +lemma homologyMap_add : homologyMap (φ + φ') = homologyMap φ + homologyMap φ' := + homologyMap'_add _ _ + +@[simp] +lemma homologyMap_sub : homologyMap (φ - φ') = homologyMap φ - homologyMap φ' := + homologyMap'_sub _ _ + +end + +instance homologyFunctor_additive [CategoryWithHomology C] : + (homologyFunctor C).Additive where + +end Homology + +section Homotopy + +variable (φ₁ φ₂ φ₃ φ₄ : S₁ ⟶ S₂) + +/-- A homotopy between two morphisms of short complexes `S₁ ⟶ S₂` consists of various +maps and conditions which will be sufficient to show that they induce the same morphism +in homology. -/ +@[ext] +structure Homotopy where + /-- a morphism `S₁.X₁ ⟶ S₂.X₁` -/ + h₀ : S₁.X₁ ⟶ S₂.X₁ + h₀_f : h₀ ≫ S₂.f = 0 := by aesop_cat + /-- a morphism `S₁.X₂ ⟶ S₂.X₁` -/ + h₁ : S₁.X₂ ⟶ S₂.X₁ + /-- a morphism `S₁.X₃ ⟶ S₂.X₂` -/ + h₂ : S₁.X₃ ⟶ S₂.X₂ + /-- a morphism `S₁.X₃ ⟶ S₂.X₃` -/ + h₃ : S₁.X₃ ⟶ S₂.X₃ + g_h₃ : S₁.g ≫ h₃ = 0 := by aesop_cat + comm₁ : φ₁.τ₁ = S₁.f ≫ h₁ + h₀ + φ₂.τ₁ := by aesop_cat + comm₂ : φ₁.τ₂ = S₁.g ≫ h₂ + h₁ ≫ S₂.f + φ₂.τ₂ := by aesop_cat + comm₃ : φ₁.τ₃ = h₃ + h₂ ≫ S₂.g + φ₂.τ₃ := by aesop_cat + +attribute [reassoc (attr := simp)] Homotopy.h₀_f Homotopy.g_h₃ + +variable (S₁ S₂) + +/-- Constructor for null homotopic morphisms, see also `Homotopy.ofNullHomotopic` +and `Homotopy.eq_add_nullHomotopic`. -/ +@[simps] +def nullHomotopic (h₀ : S₁.X₁ ⟶ S₂.X₁) (h₀_f : h₀ ≫ S₂.f = 0) + (h₁ : S₁.X₂ ⟶ S₂.X₁) (h₂ : S₁.X₃ ⟶ S₂.X₂) (h₃ : S₁.X₃ ⟶ S₂.X₃) (g_h₃ : S₁.g ≫ h₃ = 0) : + S₁ ⟶ S₂ where + τ₁ := h₀ + S₁.f ≫ h₁ + τ₂ := h₁ ≫ S₂.f + S₁.g ≫ h₂ + τ₃ := h₂ ≫ S₂.g + h₃ + +namespace Homotopy + +attribute [local simp] neg_comp + +variable {S₁ S₂ φ₁ φ₂ φ₃ φ₄} + +/-- The obvious homotopy between two equal morphisms of short complexes. -/ +@[simps] +def ofEq (h : φ₁ = φ₂) : Homotopy φ₁ φ₂ where + h₀ := 0 + h₁ := 0 + h₂ := 0 + h₃ := 0 + +/-- The obvious homotopy between a morphism of short complexes and itself. -/ +@[simps!] +def refl (φ : S₁ ⟶ S₂) : Homotopy φ φ := ofEq rfl + +/-- The symmetry of homotopy between morphisms of short complexes. -/ +@[simps] +def symm (h : Homotopy φ₁ φ₂) : Homotopy φ₂ φ₁ where + h₀ := -h.h₀ + h₁ := -h.h₁ + h₂ := -h.h₂ + h₃ := -h.h₃ + comm₁ := by rw [h.comm₁, comp_neg]; abel + comm₂ := by rw [h.comm₂, comp_neg, neg_comp]; abel + comm₃ := by rw [h.comm₃, neg_comp]; abel + +/-- If two maps of short complexes are homotopic, their opposites also are. -/ +@[simps] +def neg (h : Homotopy φ₁ φ₂) : Homotopy (-φ₁) (-φ₂) where + h₀ := -h.h₀ + h₁ := -h.h₁ + h₂ := -h.h₂ + h₃ := -h.h₃ + comm₁ := by rw [neg_τ₁, neg_τ₁, h.comm₁, neg_add_rev, comp_neg]; abel + comm₂ := by rw [neg_τ₂, neg_τ₂, h.comm₂, neg_add_rev, comp_neg, neg_comp]; abel + comm₃ := by rw [neg_τ₃, neg_τ₃, h.comm₃, neg_comp]; abel + +/-- The transitivity of homotopy between morphisms of short complexes. -/ +@[simps] +def trans (h₁₂ : Homotopy φ₁ φ₂) (h₂₃ : Homotopy φ₂ φ₃) : Homotopy φ₁ φ₃ where + h₀ := h₁₂.h₀ + h₂₃.h₀ + h₁ := h₁₂.h₁ + h₂₃.h₁ + h₂ := h₁₂.h₂ + h₂₃.h₂ + h₃ := h₁₂.h₃ + h₂₃.h₃ + comm₁ := by rw [h₁₂.comm₁, h₂₃.comm₁, comp_add]; abel + comm₂ := by rw [h₁₂.comm₂, h₂₃.comm₂, comp_add, add_comp]; abel + comm₃ := by rw [h₁₂.comm₃, h₂₃.comm₃, add_comp]; abel + +/-- Homotopy between morphisms of short complexes is compatible withe addition. -/ +@[simps] +def add (h : Homotopy φ₁ φ₂) (h' : Homotopy φ₃ φ₄) : Homotopy (φ₁ + φ₃) (φ₂ + φ₄) where + h₀ := h.h₀ + h'.h₀ + h₁ := h.h₁ + h'.h₁ + h₂ := h.h₂ + h'.h₂ + h₃ := h.h₃ + h'.h₃ + comm₁ := by rw [add_τ₁, add_τ₁, h.comm₁, h'.comm₁, comp_add]; abel + comm₂ := by rw [add_τ₂, add_τ₂, h.comm₂, h'.comm₂, comp_add, add_comp]; abel + comm₃ := by rw [add_τ₃, add_τ₃, h.comm₃, h'.comm₃, add_comp]; abel + +/-- Homotopy between morphisms of short complexes is compatible withe substraction. -/ +@[simps] +def sub (h : Homotopy φ₁ φ₂) (h' : Homotopy φ₃ φ₄) : Homotopy (φ₁ - φ₃) (φ₂ - φ₄) where + h₀ := h.h₀ - h'.h₀ + h₁ := h.h₁ - h'.h₁ + h₂ := h.h₂ - h'.h₂ + h₃ := h.h₃ - h'.h₃ + comm₁ := by rw [sub_τ₁, sub_τ₁, h.comm₁, h'.comm₁, comp_sub]; abel + comm₂ := by rw [sub_τ₂, sub_τ₂, h.comm₂, h'.comm₂, comp_sub, sub_comp]; abel + comm₃ := by rw [sub_τ₃, sub_τ₃, h.comm₃, h'.comm₃, sub_comp]; abel + +/-- Homotopy between morphisms of short complexes is compatible with precomposition. -/ +@[simps] +def compLeft (h : Homotopy φ₁ φ₂) (ψ : S₃ ⟶ S₁) : Homotopy (ψ ≫ φ₁) (ψ ≫ φ₂) where + h₀ := ψ.τ₁ ≫ h.h₀ + h₁ := ψ.τ₂ ≫ h.h₁ + h₂ := ψ.τ₃ ≫ h.h₂ + h₃ := ψ.τ₃ ≫ h.h₃ + g_h₃ := by rw [← ψ.comm₂₃_assoc, h.g_h₃, comp_zero] + comm₁ := by rw [comp_τ₁, comp_τ₁, h.comm₁, comp_add, comp_add, add_left_inj, ψ.comm₁₂_assoc] + comm₂ := by rw [comp_τ₂, comp_τ₂, h.comm₂, comp_add, comp_add, assoc, ψ.comm₂₃_assoc] + comm₃ := by rw [comp_τ₃, comp_τ₃, h.comm₃, comp_add, comp_add, assoc] + +/-- Homotopy between morphisms of short complexes is compatible with postcomposition. -/ +@[simps] +def compRight (h : Homotopy φ₁ φ₂) (ψ : S₂ ⟶ S₃) : Homotopy (φ₁ ≫ ψ) (φ₂ ≫ ψ) where + h₀ := h.h₀ ≫ ψ.τ₁ + h₁ := h.h₁ ≫ ψ.τ₁ + h₂ := h.h₂ ≫ ψ.τ₂ + h₃ := h.h₃ ≫ ψ.τ₃ + comm₁ := by rw [comp_τ₁, comp_τ₁, h.comm₁, add_comp, add_comp, assoc] + comm₂ := by rw [comp_τ₂, comp_τ₂, h.comm₂, add_comp, add_comp, assoc, assoc, assoc, ψ.comm₁₂] + comm₃ := by rw [comp_τ₃, comp_τ₃, h.comm₃, add_comp, add_comp, assoc, assoc, ψ.comm₂₃] + +/-- Homotopy between morphisms of short complexes is compatible with composition. -/ +@[simps!] +def comp (h : Homotopy φ₁ φ₂) {ψ₁ ψ₂ : S₂ ⟶ S₃} (h' : Homotopy ψ₁ ψ₂) : + Homotopy (φ₁ ≫ ψ₁) (φ₂ ≫ ψ₂) := + (h.compRight ψ₁).trans (h'.compLeft φ₂) + +/-- The homotopy between morphisms in `ShortComplex Cᵒᵖ` that is induced by a homotopy +between morphisms in `ShortComplex C`. -/ +@[simps] +def op (h : Homotopy φ₁ φ₂) : Homotopy (opMap φ₁) (opMap φ₂) where + h₀ := h.h₃.op + h₁ := h.h₂.op + h₂ := h.h₁.op + h₃ := h.h₀.op + h₀_f := Quiver.Hom.unop_inj h.g_h₃ + g_h₃ := Quiver.Hom.unop_inj h.h₀_f + comm₁ := Quiver.Hom.unop_inj (by dsimp; rw [h.comm₃]; abel) + comm₂ := Quiver.Hom.unop_inj (by dsimp; rw [h.comm₂]; abel) + comm₃ := Quiver.Hom.unop_inj (by dsimp; rw [h.comm₁]; abel) + +/-- The homotopy between morphisms in `ShortComplex C` that is induced by a homotopy +between morphisms in `ShortComplex Cᵒᵖ`. -/ +@[simps] +def unop {S₁ S₂ : ShortComplex Cᵒᵖ} {φ₁ φ₂ : S₁ ⟶ S₂} (h : Homotopy φ₁ φ₂) : + Homotopy (unopMap φ₁) (unopMap φ₂) where + h₀ := h.h₃.unop + h₁ := h.h₂.unop + h₂ := h.h₁.unop + h₃ := h.h₀.unop + h₀_f := Quiver.Hom.op_inj h.g_h₃ + g_h₃ := Quiver.Hom.op_inj h.h₀_f + comm₁ := Quiver.Hom.op_inj (by dsimp; rw [h.comm₃]; abel) + comm₂ := Quiver.Hom.op_inj (by dsimp; rw [h.comm₂]; abel) + comm₃ := Quiver.Hom.op_inj (by dsimp; rw [h.comm₁]; abel) + +variable (φ₁ φ₂) + +/-- Equivalence expressing that two morphisms are homotopic iff +their difference is homotopic to zero. -/ +@[simps] +def equivSubZero : Homotopy φ₁ φ₂ ≃ Homotopy (φ₁ - φ₂) 0 where + toFun h := (h.sub (refl φ₂)).trans (ofEq (sub_self φ₂)) + invFun h := ((ofEq (sub_add_cancel φ₁ φ₂).symm).trans + (h.add (refl φ₂))).trans (ofEq (zero_add φ₂)) + left_inv := by aesop_cat + right_inv := by aesop_cat + +variable {φ₁ φ₂} + +lemma eq_add_nullHomotopic (h : Homotopy φ₁ φ₂) : + φ₁ = φ₂ + nullHomotopic _ _ h.h₀ h.h₀_f h.h₁ h.h₂ h.h₃ h.g_h₃ := by + ext + · dsimp; rw [h.comm₁]; abel + · dsimp; rw [h.comm₂]; abel + · dsimp; rw [h.comm₃]; abel + +variable (S₁ S₂) + +/-- A morphism constructed with `nullHomotopic` is homotopic to zero. -/ +@[simps] +def ofNullHomotopic (h₀ : S₁.X₁ ⟶ S₂.X₁) (h₀_f : h₀ ≫ S₂.f = 0) + (h₁ : S₁.X₂ ⟶ S₂.X₁) (h₂ : S₁.X₃ ⟶ S₂.X₂) (h₃ : S₁.X₃ ⟶ S₂.X₃) (g_h₃ : S₁.g ≫ h₃ = 0) : + Homotopy (nullHomotopic _ _ h₀ h₀_f h₁ h₂ h₃ g_h₃) 0 where + h₀ := h₀ + h₁ := h₁ + h₂ := h₂ + h₃ := h₃ + h₀_f := h₀_f + g_h₃ := g_h₃ + comm₁ := by rw [nullHomotopic_τ₁, zero_τ₁, add_zero]; abel + comm₂ := by rw [nullHomotopic_τ₂, zero_τ₂, add_zero]; abel + comm₃ := by rw [nullHomotopic_τ₃, zero_τ₃, add_zero]; abel + +end Homotopy + +variable {S₁ S₂} + +/-- The left homology map data expressing that null homotopic maps induce the zero +morphism in left homology. -/ +def LeftHomologyMapData.ofNullHomotopic + (H₁ : S₁.LeftHomologyData) (H₂ : S₂.LeftHomologyData) + (h₀ : S₁.X₁ ⟶ S₂.X₁) (h₀_f : h₀ ≫ S₂.f = 0) + (h₁ : S₁.X₂ ⟶ S₂.X₁) (h₂ : S₁.X₃ ⟶ S₂.X₂) (h₃ : S₁.X₃ ⟶ S₂.X₃) (g_h₃ : S₁.g ≫ h₃ = 0) : + LeftHomologyMapData (nullHomotopic _ _ h₀ h₀_f h₁ h₂ h₃ g_h₃) H₁ H₂ where + φK := H₂.liftK (H₁.i ≫ h₁ ≫ S₂.f) (by simp) + φH := 0 + commf' := by + rw [← cancel_mono H₂.i, assoc, LeftHomologyData.liftK_i, LeftHomologyData.f'_i_assoc, + nullHomotopic_τ₁, add_comp, add_comp, assoc, assoc, assoc, LeftHomologyData.f'_i, + self_eq_add_left, h₀_f] + commπ := by + rw [H₂.liftK_π_eq_zero_of_boundary (H₁.i ≫ h₁ ≫ S₂.f) (H₁.i ≫ h₁) (by rw [assoc]), comp_zero] + +/-- The right homology map data expressing that null homotopic maps induce the zero +morphism in right homology. -/ +def RightHomologyMapData.ofNullHomotopic + (H₁ : S₁.RightHomologyData) (H₂ : S₂.RightHomologyData) + (h₀ : S₁.X₁ ⟶ S₂.X₁) (h₀_f : h₀ ≫ S₂.f = 0) + (h₁ : S₁.X₂ ⟶ S₂.X₁) (h₂ : S₁.X₃ ⟶ S₂.X₂) (h₃ : S₁.X₃ ⟶ S₂.X₃) (g_h₃ : S₁.g ≫ h₃ = 0) : + RightHomologyMapData (nullHomotopic _ _ h₀ h₀_f h₁ h₂ h₃ g_h₃) H₁ H₂ where + φQ := H₁.descQ (S₁.g ≫ h₂ ≫ H₂.p) (by simp) + φH := 0 + commg' := by + rw [← cancel_epi H₁.p, RightHomologyData.p_descQ_assoc, RightHomologyData.p_g'_assoc, + nullHomotopic_τ₃, comp_add, assoc, assoc, RightHomologyData.p_g', g_h₃, add_zero] + commι := by + rw [H₁.ι_descQ_eq_zero_of_boundary (S₁.g ≫ h₂ ≫ H₂.p) (h₂ ≫ H₂.p) rfl, zero_comp] + +@[simp] +lemma leftHomologyMap'_nullHomotopic + (H₁ : S₁.LeftHomologyData) (H₂ : S₂.LeftHomologyData) + (h₀ : S₁.X₁ ⟶ S₂.X₁) (h₀_f : h₀ ≫ S₂.f = 0) + (h₁ : S₁.X₂ ⟶ S₂.X₁) (h₂ : S₁.X₃ ⟶ S₂.X₂) (h₃ : S₁.X₃ ⟶ S₂.X₃) (g_h₃ : S₁.g ≫ h₃ = 0) : + leftHomologyMap' (nullHomotopic _ _ h₀ h₀_f h₁ h₂ h₃ g_h₃) H₁ H₂ = 0 := + (LeftHomologyMapData.ofNullHomotopic H₁ H₂ h₀ h₀_f h₁ h₂ h₃ g_h₃).leftHomologyMap'_eq + +@[simp] +lemma rightHomologyMap'_nullHomotopic + (H₁ : S₁.RightHomologyData) (H₂ : S₂.RightHomologyData) + (h₀ : S₁.X₁ ⟶ S₂.X₁) (h₀_f : h₀ ≫ S₂.f = 0) + (h₁ : S₁.X₂ ⟶ S₂.X₁) (h₂ : S₁.X₃ ⟶ S₂.X₂) (h₃ : S₁.X₃ ⟶ S₂.X₃) (g_h₃ : S₁.g ≫ h₃ = 0) : + rightHomologyMap' (nullHomotopic _ _ h₀ h₀_f h₁ h₂ h₃ g_h₃) H₁ H₂ = 0 := + (RightHomologyMapData.ofNullHomotopic H₁ H₂ h₀ h₀_f h₁ h₂ h₃ g_h₃).rightHomologyMap'_eq + +@[simp] +lemma homologyMap'_nullHomotopic + (H₁ : S₁.HomologyData) (H₂ : S₂.HomologyData) + (h₀ : S₁.X₁ ⟶ S₂.X₁) (h₀_f : h₀ ≫ S₂.f = 0) + (h₁ : S₁.X₂ ⟶ S₂.X₁) (h₂ : S₁.X₃ ⟶ S₂.X₂) (h₃ : S₁.X₃ ⟶ S₂.X₃) (g_h₃ : S₁.g ≫ h₃ = 0) : + homologyMap' (nullHomotopic _ _ h₀ h₀_f h₁ h₂ h₃ g_h₃) H₁ H₂ = 0 := + by apply leftHomologyMap'_nullHomotopic + +variable (S₁ S₂) + +@[simp] +lemma leftHomologyMap_nullHomotopic [S₁.HasLeftHomology] [S₂.HasLeftHomology] + (h₀ : S₁.X₁ ⟶ S₂.X₁) (h₀_f : h₀ ≫ S₂.f = 0) + (h₁ : S₁.X₂ ⟶ S₂.X₁) (h₂ : S₁.X₃ ⟶ S₂.X₂) (h₃ : S₁.X₃ ⟶ S₂.X₃) (g_h₃ : S₁.g ≫ h₃ = 0) : + leftHomologyMap (nullHomotopic _ _ h₀ h₀_f h₁ h₂ h₃ g_h₃) = 0 := + by apply leftHomologyMap'_nullHomotopic + +@[simp] +lemma rightHomologyMap_nullHomotopic [S₁.HasRightHomology] [S₂.HasRightHomology] + (h₀ : S₁.X₁ ⟶ S₂.X₁) (h₀_f : h₀ ≫ S₂.f = 0) + (h₁ : S₁.X₂ ⟶ S₂.X₁) (h₂ : S₁.X₃ ⟶ S₂.X₂) (h₃ : S₁.X₃ ⟶ S₂.X₃) (g_h₃ : S₁.g ≫ h₃ = 0) : + rightHomologyMap (nullHomotopic _ _ h₀ h₀_f h₁ h₂ h₃ g_h₃) = 0 := + by apply rightHomologyMap'_nullHomotopic + +@[simp] +lemma homologyMap_nullHomotopic [S₁.HasHomology] [S₂.HasHomology] + (h₀ : S₁.X₁ ⟶ S₂.X₁) (h₀_f : h₀ ≫ S₂.f = 0) + (h₁ : S₁.X₂ ⟶ S₂.X₁) (h₂ : S₁.X₃ ⟶ S₂.X₂) (h₃ : S₁.X₃ ⟶ S₂.X₃) (g_h₃ : S₁.g ≫ h₃ = 0) : + homologyMap (nullHomotopic _ _ h₀ h₀_f h₁ h₂ h₃ g_h₃) = 0 := + by apply homologyMap'_nullHomotopic + +namespace Homotopy + +variable {φ₁ φ₂ S₁ S₂} + +lemma leftHomologyMap'_congr (h : Homotopy φ₁ φ₂) (h₁ : S₁.LeftHomologyData) + (h₂ : S₂.LeftHomologyData) : leftHomologyMap' φ₁ h₁ h₂ = leftHomologyMap' φ₂ h₁ h₂ := by + rw [h.eq_add_nullHomotopic, leftHomologyMap'_add, leftHomologyMap'_nullHomotopic, add_zero] + +lemma rightHomologyMap'_congr (h : Homotopy φ₁ φ₂) (h₁ : S₁.RightHomologyData) + (h₂ : S₂.RightHomologyData) : rightHomologyMap' φ₁ h₁ h₂ = rightHomologyMap' φ₂ h₁ h₂ := by + rw [h.eq_add_nullHomotopic, rightHomologyMap'_add, rightHomologyMap'_nullHomotopic, add_zero] + +lemma homologyMap'_congr (h : Homotopy φ₁ φ₂) (h₁ : S₁.HomologyData) + (h₂ : S₂.HomologyData) : homologyMap' φ₁ h₁ h₂ = homologyMap' φ₂ h₁ h₂ := by + rw [h.eq_add_nullHomotopic, homologyMap'_add, homologyMap'_nullHomotopic, add_zero] + +lemma leftHomologyMap_congr (h : Homotopy φ₁ φ₂) [S₁.HasLeftHomology] [S₂.HasLeftHomology] : + leftHomologyMap φ₁ = leftHomologyMap φ₂ := + h.leftHomologyMap'_congr _ _ + +lemma rightHomologyMap_congr (h : Homotopy φ₁ φ₂) [S₁.HasRightHomology] [S₂.HasRightHomology] : + rightHomologyMap φ₁ = rightHomologyMap φ₂ := + h.rightHomologyMap'_congr _ _ + +lemma homologyMap_congr (h : Homotopy φ₁ φ₂) [S₁.HasHomology] [S₂.HasHomology] : + homologyMap φ₁ = homologyMap φ₂ := + h.homologyMap'_congr _ _ + +end Homotopy + +/-- An homotopy equivalence between two short complexes `S₁` and `S₂` consists +of morphisms `hom : S₁ ⟶ S₂` and `inv : S₂ ⟶ S₁` such that both compositions +`hom ≫ inv` and `inv ≫ hom` are homotopic to the identity. -/ +@[ext] +structure HomotopyEquiv where + /-- the forward direction of a homotopy equivalence. -/ + hom : S₁ ⟶ S₂ + /-- the backwards direction of a homotopy equivalence. -/ + inv : S₂ ⟶ S₁ + /-- the composition of the two directions of a homotopy equivalence is + homotopic to the identity of the source -/ + homotopyHomInvId : Homotopy (hom ≫ inv) (𝟙 S₁) + /-- the composition of the two directions of a homotopy equivalence is + homotopic to the identity of the target -/ + homotopyInvHomId : Homotopy (inv ≫ hom) (𝟙 S₂) + +namespace HomotopyEquiv + +variable {S₁ S₂} + +/-- The homotopy equivalence from a short complex to itself that is induced +by the identity. -/ +@[simps] +def refl (S : ShortComplex C) : HomotopyEquiv S S where + hom := 𝟙 S + inv := 𝟙 S + homotopyHomInvId := Homotopy.ofEq (by simp) + homotopyInvHomId := Homotopy.ofEq (by simp) + +/-- The inverse of a homotopy equivalence. -/ +@[simps] +def symm (e : HomotopyEquiv S₁ S₂) : HomotopyEquiv S₂ S₁ where + hom := e.inv + inv := e.hom + homotopyHomInvId := e.homotopyInvHomId + homotopyInvHomId := e.homotopyHomInvId + +/-- The composition of homotopy equivalences. -/ +@[simps] +def trans (e : HomotopyEquiv S₁ S₂) (e' : HomotopyEquiv S₂ S₃) : + HomotopyEquiv S₁ S₃ where + hom := e.hom ≫ e'.hom + inv := e'.inv ≫ e.inv + homotopyHomInvId := (Homotopy.ofEq (by simp)).trans + (((e'.homotopyHomInvId.compRight e.inv).compLeft e.hom).trans + ((Homotopy.ofEq (by simp)).trans e.homotopyHomInvId)) + homotopyInvHomId := (Homotopy.ofEq (by simp)).trans + (((e.homotopyInvHomId.compRight e'.hom).compLeft e'.inv).trans + ((Homotopy.ofEq (by simp)).trans e'.homotopyInvHomId)) + +end HomotopyEquiv + +end Homotopy + end ShortComplex end CategoryTheory diff --git a/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean index 3fe0875899d1d..1070cb41d872b 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.Algebra.Homology.ShortComplex.Homology +import Mathlib.Algebra.Homology.ShortComplex.QuasiIso import Mathlib.CategoryTheory.Limits.Preserves.Finite import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Kernels @@ -720,6 +720,127 @@ noncomputable def homologyFunctorIso NatIso.ofComponents (fun S => S.mapHomologyIso F) (fun f => ShortComplex.mapHomologyIso_hom_naturality f F) +section + +variable + {S₁ S₂ : ShortComplex C} {φ : S₁ ⟶ S₂} + {hl₁ : S₁.LeftHomologyData} {hr₁ : S₁.RightHomologyData} + {hl₂ : S₂.LeftHomologyData} {hr₂ : S₂.RightHomologyData} + (ψl : LeftHomologyMapData φ hl₁ hl₂) + (ψr : RightHomologyMapData φ hr₁ hr₂) + +lemma LeftHomologyMapData.quasiIso_map_iff + [(F.mapShortComplex.obj S₁).HasHomology] + [(F.mapShortComplex.obj S₂).HasHomology] + [hl₁.IsPreservedBy F] [hl₂.IsPreservedBy F] : + QuasiIso (F.mapShortComplex.map φ) ↔ IsIso (F.map ψl.φH) := + (ψl.map F).quasiIso_iff + +lemma RightHomologyMapData.quasiIso_map_iff + [(F.mapShortComplex.obj S₁).HasHomology] + [(F.mapShortComplex.obj S₂).HasHomology] + [hr₁.IsPreservedBy F] [hr₂.IsPreservedBy F] : + QuasiIso (F.mapShortComplex.map φ) ↔ IsIso (F.map ψr.φH) := + (ψr.map F).quasiIso_iff + +variable (φ) [S₁.HasHomology] [S₂.HasHomology] + [(F.mapShortComplex.obj S₁).HasHomology] [(F.mapShortComplex.obj S₂).HasHomology] + +instance quasiIso_map_of_preservesLeftHomology + [F.PreservesLeftHomologyOf S₁] [F.PreservesLeftHomologyOf S₂] + [QuasiIso φ] : QuasiIso (F.mapShortComplex.map φ) := by + have γ : LeftHomologyMapData φ S₁.leftHomologyData S₂.leftHomologyData := default + have : IsIso γ.φH := by + rw [← γ.quasiIso_iff] + infer_instance + rw [(γ.map F).quasiIso_iff, LeftHomologyMapData.map_φH] + infer_instance + +instance quasiIso_map_of_preservesRightHomology + [F.PreservesRightHomologyOf S₁] [F.PreservesRightHomologyOf S₂] + [QuasiIso φ] : QuasiIso (F.mapShortComplex.map φ) := by + have γ : RightHomologyMapData φ S₁.rightHomologyData S₂.rightHomologyData := default + have : IsIso γ.φH := by + rw [← γ.quasiIso_iff] + infer_instance + rw [(γ.map F).quasiIso_iff, RightHomologyMapData.map_φH] + infer_instance + +lemma quasiIso_map_iff_of_preservesRightHomology + [F.PreservesRightHomologyOf S₁] [F.PreservesRightHomologyOf S₂] + [ReflectsIsomorphisms F] : + QuasiIso (F.mapShortComplex.map φ) ↔ QuasiIso φ := by + have γ : RightHomologyMapData φ S₁.rightHomologyData S₂.rightHomologyData := default + rw [γ.quasiIso_iff, (γ.map F).quasiIso_iff, RightHomologyMapData.map_φH] + constructor + · intro + exact isIso_of_reflects_iso _ F + · intro + infer_instance + +end + end ShortComplex +namespace Functor + +variable (F : C ⥤ D) [F.PreservesZeroMorphisms] (S : ShortComplex C) + +/-- If a short complex `S` is such that `S.f = 0` and that the kernel of `S.g` is preserved +by a functor `F`, then `F` preserves the left homology of `S`. -/ +noncomputable def preservesLeftHomologyOfZerof (hf : S.f = 0) + [PreservesLimit (parallelPair S.g 0) F] : + F.PreservesLeftHomologyOf S := ⟨fun h => + { g := by infer_instance + f' := Limits.preservesCokernelZero' _ _ + (by rw [← cancel_mono h.i, h.f'_i, zero_comp, hf]) }⟩ + +/-- If a short complex `S` is such that `S.g = 0` and that the cokernel of `S.f` is preserved +by a functor `F`, then `F` preserves the right homology of `S`. -/ +noncomputable def preservesRightHomologyOfZerog (hg : S.g = 0) + [PreservesColimit (parallelPair S.f 0) F] : + F.PreservesRightHomologyOf S := ⟨fun h => + { f := by infer_instance + g' := Limits.preservesKernelZero' _ _ + (by rw [← cancel_epi h.p, h.p_g', comp_zero, hg]) }⟩ + +/-- If a short complex `S` is such that `S.g = 0` and that the cokernel of `S.f` is preserved +by a functor `F`, then `F` preserves the left homology of `S`. -/ +noncomputable def preservesLeftHomologyOfZerog (hg : S.g = 0) + [PreservesColimit (parallelPair S.f 0) F] : + F.PreservesLeftHomologyOf S := ⟨fun h => + { g := by + rw [hg] + infer_instance + f' := by + have := h.isIso_i hg + let e : parallelPair h.f' 0 ≅ parallelPair S.f 0 := + parallelPair.ext (Iso.refl _) (asIso h.i) (by aesop_cat) (by aesop_cat) + exact Limits.preservesColimitOfIsoDiagram F e.symm}⟩ + +/-- If a short complex `S` is such that `S.f = 0` and that the kernel of `S.g` is preserved +by a functor `F`, then `F` preserves the right homology of `S`. -/ +noncomputable def preservesRightHomologyOfZerof (hf : S.f = 0) + [PreservesLimit (parallelPair S.g 0) F] : + F.PreservesRightHomologyOf S := ⟨fun h => + { f := by + rw [hf] + infer_instance + g' := by + have := h.isIso_p hf + let e : parallelPair S.g 0 ≅ parallelPair h.g' 0 := + parallelPair.ext (asIso h.p) (Iso.refl _) (by aesop_cat) (by aesop_cat) + exact Limits.preservesLimitOfIsoDiagram F e }⟩ + +end Functor + +lemma NatTrans.app_homology {F G : C ⥤ D} (τ : F ⟶ G) + (S : ShortComplex C) [S.HasHomology] [F.PreservesZeroMorphisms] [G.PreservesZeroMorphisms] + [F.PreservesLeftHomologyOf S] [G.PreservesLeftHomologyOf S] [F.PreservesRightHomologyOf S] + [G.PreservesRightHomologyOf S] : + τ.app S.homology = (S.mapHomologyIso F).inv ≫ + ShortComplex.homologyMap (S.mapNatTrans τ) ≫ (S.mapHomologyIso G).hom := by + rw [ShortComplex.homologyMap_mapNatTrans, assoc, assoc, Iso.inv_hom_id, + comp_id, Iso.inv_hom_id_assoc] + end CategoryTheory diff --git a/Mathlib/Algebra/Homology/ShortComplex/ShortExact.lean b/Mathlib/Algebra/Homology/ShortComplex/ShortExact.lean new file mode 100644 index 0000000000000..4a564b0bf84a7 --- /dev/null +++ b/Mathlib/Algebra/Homology/ShortComplex/ShortExact.lean @@ -0,0 +1,201 @@ +/- +Copyright (c) 2023 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Homology.ShortComplex.Exact +import Mathlib.CategoryTheory.Preadditive.Injective + +/-! +# Short exact short complexes + +A short complex `S : ShortComplex C` is short exact (`S.ShortExact`) when it is exact, +`S.f` is a mono and `S.g` is an epi. + +-/ + +namespace CategoryTheory + +open Category Limits ZeroObject + +variable {C D : Type*} [Category C] [Category D] + +namespace ShortComplex + +section + +variable [HasZeroMorphisms C] [HasZeroMorphisms D] + (S : ShortComplex C) {S₁ S₂ : ShortComplex C} + +/-- A short complex `S` is short exact if it is exact, `S.f` is a mono and `S.g` is an epi. -/ +structure ShortExact : Prop where + exact : S.Exact + [mono_f : Mono S.f] + [epi_g : Epi S.g] + +variable {S} + +lemma ShortExact.mk' (h : S.Exact) (_ : Mono S.f) (_ : Epi S.g) : S.ShortExact where + exact := h + +lemma shortExact_of_iso (e : S₁ ≅ S₂) (h : S₁.ShortExact) : S₂.ShortExact where + exact := exact_of_iso e h.exact + mono_f := by + suffices Mono (S₂.f ≫ e.inv.τ₂) by + exact mono_of_mono _ e.inv.τ₂ + have := h.mono_f + rw [← e.inv.comm₁₂] + infer_instance + epi_g := by + suffices Epi (e.hom.τ₂ ≫ S₂.g) by + exact epi_of_epi e.hom.τ₂ _ + have := h.epi_g + rw [e.hom.comm₂₃] + apply epi_comp + +lemma shortExact_iff_of_iso (e : S₁ ≅ S₂) : S₁.ShortExact ↔ S₂.ShortExact := by + constructor + · exact shortExact_of_iso e + · exact shortExact_of_iso e.symm + +lemma ShortExact.op (h : S.ShortExact) : S.op.ShortExact where + exact := h.exact.op + mono_f := by + have := h.epi_g + dsimp + infer_instance + epi_g := by + have := h.mono_f + dsimp + infer_instance + +lemma ShortExact.unop {S : ShortComplex Cᵒᵖ} (h : S.ShortExact) : S.unop.ShortExact where + exact := h.exact.unop + mono_f := by + have := h.epi_g + dsimp + infer_instance + epi_g := by + have := h.mono_f + dsimp + infer_instance + +variable (S) + +lemma shortExact_iff_op : S.ShortExact ↔ S.op.ShortExact := + ⟨ShortExact.op, ShortExact.unop⟩ + +lemma shortExact_iff_unop (S : ShortComplex Cᵒᵖ) : S.ShortExact ↔ S.unop.ShortExact := + S.unop.shortExact_iff_op.symm + +variable {S} + +lemma ShortExact.map (h : S.ShortExact) (F : C ⥤ D) + [F.PreservesZeroMorphisms] [F.PreservesLeftHomologyOf S] + [F.PreservesRightHomologyOf S] [Mono (F.map S.f)] [Epi (F.map S.g)] : + (S.map F).ShortExact where + exact := h.exact.map F + mono_f := (inferInstance : Mono (F.map S.f)) + epi_g := (inferInstance : Epi (F.map S.g)) + +lemma ShortExact.map_of_exact (hS : S.ShortExact) + (F : C ⥤ D) [F.PreservesZeroMorphisms] [PreservesFiniteLimits F] + [PreservesFiniteColimits F] : (S.map F).ShortExact := by + have := hS.mono_f + have := hS.epi_g + have := preserves_mono_of_preservesLimit F S.f + have := preserves_epi_of_preservesColimit F S.g + exact hS.map F + +end + +section Preadditive + +variable [Preadditive C] + +lemma ShortExact.isIso_f_iff {S : ShortComplex C} (hS : S.ShortExact) [Balanced C] : + IsIso S.f ↔ IsZero S.X₃ := by + have := hS.exact.hasZeroObject + have := hS.mono_f + have := hS.epi_g + constructor + · intro hf + simp only [IsZero.iff_id_eq_zero, ← cancel_epi S.g, ← cancel_epi S.f, + S.zero_assoc, zero_comp] + · intro hX₃ + have : Epi S.f := (S.exact_iff_epi (hX₃.eq_of_tgt _ _)).1 hS.exact + apply isIso_of_mono_of_epi + +lemma ShortExact.isIso_g_iff {S : ShortComplex C} (hS : S.ShortExact) [Balanced C] : + IsIso S.g ↔ IsZero S.X₁ := by + have := hS.exact.hasZeroObject + have := hS.mono_f + have := hS.epi_g + constructor + · intro hf + simp only [IsZero.iff_id_eq_zero, ← cancel_mono S.f, ← cancel_mono S.g, + S.zero, zero_comp, assoc, comp_zero] + · intro hX₁ + have : Mono S.g := (S.exact_iff_mono (hX₁.eq_of_src _ _)).1 hS.exact + apply isIso_of_mono_of_epi + +lemma isIso₂_of_shortExact_of_isIso₁₃ [Balanced C] {S₁ S₂ : ShortComplex C} (φ : S₁ ⟶ S₂) + (h₁ : S₁.ShortExact) (h₂ : S₂.ShortExact) [IsIso φ.τ₁] [IsIso φ.τ₃] : IsIso φ.τ₂ := by + have := h₁.mono_f + have := h₂.mono_f + have := h₁.epi_g + have := h₂.epi_g + have := mono_τ₂_of_exact_of_mono φ h₁.exact + have := epi_τ₂_of_exact_of_epi φ h₂.exact + apply isIso_of_mono_of_epi + +lemma isIso₂_of_shortExact_of_isIso₁₃' [Balanced C] {S₁ S₂ : ShortComplex C} (φ : S₁ ⟶ S₂) + (h₁ : S₁.ShortExact) (h₂ : S₂.ShortExact) (_ : IsIso φ.τ₁) (_ : IsIso φ.τ₃) : IsIso φ.τ₂ := + isIso₂_of_shortExact_of_isIso₁₃ φ h₁ h₂ + +/-- If `S` is a short exact short complex in a balanced category, +then `S.X₁` is the kernel of `S.g`. -/ +noncomputable def ShortExact.fIsKernel [Balanced C] {S : ShortComplex C} (hS : S.ShortExact) : + IsLimit (KernelFork.ofι S.f S.zero) := by + have := hS.mono_f + exact hS.exact.fIsKernel + +/-- If `S` is a short exact short complex in a balanced category, +then `S.X₃` is the cokernel of `S.f`. -/ +noncomputable def ShortExact.gIsCokernel [Balanced C] {S : ShortComplex C} (hS : S.ShortExact) : + IsColimit (CokernelCofork.ofπ S.g S.zero) := by + have := hS.epi_g + exact hS.exact.gIsCokernel + +/-- A split short complex is short exact. -/ +lemma Splitting.shortExact {S : ShortComplex C} [HasZeroObject C] (s : S.Splitting) : + S.ShortExact where + exact := s.exact + mono_f := s.mono_f + epi_g := s.epi_g + +namespace ShortExact + +/-- A choice of splitting for a short exact short complex `S` in a balanced category +such that `S.X₁` is injective. -/ +noncomputable def splittingOfInjective {S : ShortComplex C} (hS : S.ShortExact) + [Injective S.X₁] [Balanced C] : + S.Splitting := + have := hS.mono_f + Splitting.ofExactOfRetraction S hS.exact (Injective.factorThru (𝟙 S.X₁) S.f) (by simp) hS.epi_g + +/-- A choice of splitting for a short exact short complex `S` in a balanced category +such that `S.X₃` is projective. -/ +noncomputable def splittingOfProjective {S : ShortComplex C} (hS : S.ShortExact) + [Projective S.X₃] [Balanced C] : + S.Splitting := + have := hS.epi_g + Splitting.ofExactOfSection S hS.exact (Projective.factorThru (𝟙 S.X₃) S.g) (by simp) hS.mono_f + +end ShortExact + +end Preadditive + +end ShortComplex + +end CategoryTheory diff --git a/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean b/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean new file mode 100644 index 0000000000000..dd2f0b2a879a2 --- /dev/null +++ b/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean @@ -0,0 +1,353 @@ +/- +Copyright (c) 2023 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Homology.ShortComplex.Limits +import Mathlib.CategoryTheory.Abelian.Refinements + +/-! +# The snake lemma + +The snake lemma is a standard tool in homological algebra. The basic situation +is when we have a diagram as follows in an abelian category `C`, with exact rows: + + L₁.X₁ ⟶ L₁.X₂ ⟶ L₁.X₃ ⟶ 0 + | | | + |v₁₂.τ₁ |v₁₂.τ₂ |v₁₂.τ₃ + v v v +0 ⟶ L₂.X₁ ⟶ L₂.X₂ ⟶ L₂.X₃ + +We shall think of this diagram as the datum of a morphism `v₁₂ : L₁ ⟶ L₂` in the +category `ShortComplex C` such that both `L₁` and `L₂` are exact, and `L₁.g` is epi +and `L₂.f` is a mono (which is equivalent to saying that `L₁.X₃` is the cokernel +of `L₁.f` and `L₂.X₁` is the kernel of `L₂.g`). Then, we may introduce the kernels +and cokernels of the vertical maps. In other words, we may introduce short complexes +`L₀` and `L₃` that are respectively the kernel and the cokernel of `v₁₂`. All these +data constitute a `SnakeInput C`. + +Given such a `S : SnakeInput C`, we shall define a connecting homomorphism +`S.δ : L₀.X₃ ⟶ L₃.X₁` and show that it is part of an exact sequence +`L₀.X₁ ⟶ L₀.X₂ ⟶ L₀.X₃ ⟶ L₃.X₁ ⟶ L₃.X₂ ⟶ L₃.X₃`. This is stated as lemmas +`L₀_exact`, `L₁'_exact`, `L₂'_exact` and `L₃_exact`. This sequence can even +be extended with an extra `0` on the left (see `mono_L₀_f`) +if `L₁.X₁ ⟶ L₁.X₂` is a mono (i.e. `L₁` is short exact), +and similarly an extra `0` can be added on the right (`epi_L₃_g`) +if `L₂.X₂ ⟶ L₂.X₃` is an epi (i.e. `L₂` is short exact). + +These results were also obtained in the Liquid Tensor Experiment. The code and the proof +here are slightly easier because of the use of the category `ShortComplex C`, +the use of duality (which allows to construct only half of the sequence, and deducing +the other half by arguing in the opposite category), and the use of "refinements" +(see `CategoryTheory.Abelian.Refinements`) instead of a weak form of pseudo-elements. + +-/ + +namespace CategoryTheory + +open Category Limits Preadditive + +variable (C : Type*) [Category C] [Abelian C] + +namespace ShortComplex + +/-- A snake input in an abelian category `C` consists of morphisms +of short complexes `L₀ ⟶ L₁ ⟶ L₂ ⟶ L₃` (which should be visualized vertically) such +that `L₀` and `L₃` are respectively the kernel and the cokernel of `L₁ ⟶ L₂`, +`L₁` and `L₂` are exact, `L₁.g` is epi and `L₂.f` is mono. -/ +structure SnakeInput where + /-- the zeroth row -/ + L₀ : ShortComplex C + /-- the first row -/ + L₁ : ShortComplex C + /-- the second row -/ + L₂ : ShortComplex C + /-- the third row -/ + L₃ : ShortComplex C + /-- the morphism from the zeroth row to the first row -/ + v₀₁ : L₀ ⟶ L₁ + /-- the morphism from the first row to the second row -/ + v₁₂ : L₁ ⟶ L₂ + /-- the morphism from the second row to the third row -/ + v₂₃ : L₂ ⟶ L₃ + w₀₂ : v₀₁ ≫ v₁₂ = 0 := by aesop_cat + w₁₃ : v₁₂ ≫ v₂₃ = 0 := by aesop_cat + /-- `L₀` is the kernel of `v₁₂ : L₁ ⟶ L₂`. -/ + h₀ : IsLimit (KernelFork.ofι _ w₀₂) + /-- `L₃` is the cokernel of `v₁₂ : L₁ ⟶ L₂`. -/ + h₃ : IsColimit (CokernelCofork.ofπ _ w₁₃) + L₁_exact : L₁.Exact + epi_L₁_g : Epi L₁.g + L₂_exact : L₂.Exact + mono_L₂_f : Mono L₂.f + +initialize_simps_projections SnakeInput (-h₀, -h₃) + +namespace SnakeInput + +attribute [reassoc (attr := simp)] w₀₂ w₁₃ +attribute [instance] epi_L₁_g +attribute [instance] mono_L₂_f + +variable {C} +variable (S : SnakeInput C) + +/-- The snake input in the opposite category that is deduced from a snake input. -/ +@[simps] +noncomputable def op : SnakeInput Cᵒᵖ where + L₀ := S.L₃.op + L₁ := S.L₂.op + L₂ := S.L₁.op + L₃ := S.L₀.op + epi_L₁_g := by dsimp; infer_instance + mono_L₂_f := by dsimp; infer_instance + v₀₁ := opMap S.v₂₃ + v₁₂ := opMap S.v₁₂ + v₂₃ := opMap S.v₀₁ + w₀₂ := congr_arg opMap S.w₁₃ + w₁₃ := congr_arg opMap S.w₀₂ + h₀ := isLimitForkMapOfIsLimit' (ShortComplex.opEquiv C).functor _ + (CokernelCofork.IsColimit.ofπOp _ _ S.h₃) + h₃ := isColimitCoforkMapOfIsColimit' (ShortComplex.opEquiv C).functor _ + (KernelFork.IsLimit.ofιOp _ _ S.h₀) + L₁_exact := S.L₂_exact.op + L₂_exact := S.L₁_exact.op + +@[reassoc (attr := simp)] lemma w₀₂_τ₁ : S.v₀₁.τ₁ ≫ S.v₁₂.τ₁ = 0 := by + rw [← comp_τ₁, S.w₀₂, zero_τ₁] +@[reassoc (attr := simp)] lemma w₀₂_τ₂ : S.v₀₁.τ₂ ≫ S.v₁₂.τ₂ = 0 := by + rw [← comp_τ₂, S.w₀₂, zero_τ₂] +@[reassoc (attr := simp)] lemma w₀₂_τ₃ : S.v₀₁.τ₃ ≫ S.v₁₂.τ₃ = 0 := by + rw [← comp_τ₃, S.w₀₂, zero_τ₃] +@[reassoc (attr := simp)] lemma w₁₃_τ₁ : S.v₁₂.τ₁ ≫ S.v₂₃.τ₁ = 0 := by + rw [← comp_τ₁, S.w₁₃, zero_τ₁] +@[reassoc (attr := simp)] lemma w₁₃_τ₂ : S.v₁₂.τ₂ ≫ S.v₂₃.τ₂ = 0 := by + rw [← comp_τ₂, S.w₁₃, zero_τ₂] +@[reassoc (attr := simp)] lemma w₁₃_τ₃ : S.v₁₂.τ₃ ≫ S.v₂₃.τ₃ = 0 := by + rw [← comp_τ₃, S.w₁₃, zero_τ₃] + +/-- `L₀.X₁` is the kernel of `v₁₂.τ₁ : L₁.X₁ ⟶ L₂.X₁`. -/ +noncomputable def h₀τ₁ : IsLimit (KernelFork.ofι S.v₀₁.τ₁ S.w₀₂_τ₁) := + isLimitForkMapOfIsLimit' π₁ S.w₀₂ S.h₀ + +/-- `L₀.X₂` is the kernel of `v₁₂.τ₂ : L₁.X₂ ⟶ L₂.X₂`. -/ +noncomputable def h₀τ₂ : IsLimit (KernelFork.ofι S.v₀₁.τ₂ S.w₀₂_τ₂) := + isLimitForkMapOfIsLimit' π₂ S.w₀₂ S.h₀ + +/-- `L₀.X₃` is the kernel of `v₁₂.τ₃ : L₁.X₃ ⟶ L₂.X₃`. -/ +noncomputable def h₀τ₃ : IsLimit (KernelFork.ofι S.v₀₁.τ₃ S.w₀₂_τ₃) := + isLimitForkMapOfIsLimit' π₃ S.w₀₂ S.h₀ + +instance mono_v₀₁_τ₁ : Mono S.v₀₁.τ₁ := mono_of_isLimit_fork S.h₀τ₁ +instance mono_v₀₁_τ₂ : Mono S.v₀₁.τ₂ := mono_of_isLimit_fork S.h₀τ₂ +instance mono_v₀₁_τ₃ : Mono S.v₀₁.τ₃ := mono_of_isLimit_fork S.h₀τ₃ + +/-- The upper part of the first column of the snake diagram is exact. -/ +lemma exact_C₁_up : (ShortComplex.mk S.v₀₁.τ₁ S.v₁₂.τ₁ + (by rw [← comp_τ₁, S.w₀₂, zero_τ₁])).Exact := + exact_of_f_is_kernel _ S.h₀τ₁ + +/-- The upper part of the second column of the snake diagram is exact. -/ +lemma exact_C₂_up : (ShortComplex.mk S.v₀₁.τ₂ S.v₁₂.τ₂ + (by rw [← comp_τ₂, S.w₀₂, zero_τ₂])).Exact := + exact_of_f_is_kernel _ S.h₀τ₂ + +/-- The upper part of the third column of the snake diagram is exact. -/ +lemma exact_C₃_up : (ShortComplex.mk S.v₀₁.τ₃ S.v₁₂.τ₃ + (by rw [← comp_τ₃, S.w₀₂, zero_τ₃])).Exact := + exact_of_f_is_kernel _ S.h₀τ₃ + +instance mono_L₀_f [Mono S.L₁.f] : Mono S.L₀.f := by + have : Mono (S.L₀.f ≫ S.v₀₁.τ₂) := by + rw [← S.v₀₁.comm₁₂] + apply mono_comp + exact mono_of_mono _ S.v₀₁.τ₂ + +/-- `L₃.X₁` is the cokernel of `v₁₂.τ₁ : L₁.X₁ ⟶ L₂.X₁`. -/ +noncomputable def h₃τ₁ : IsColimit (CokernelCofork.ofπ S.v₂₃.τ₁ S.w₁₃_τ₁) := + isColimitCoforkMapOfIsColimit' π₁ S.w₁₃ S.h₃ + +/-- `L₃.X₂` is the cokernel of `v₁₂.τ₂ : L₁.X₂ ⟶ L₂.X₂`. -/ +noncomputable def h₃τ₂ : IsColimit (CokernelCofork.ofπ S.v₂₃.τ₂ S.w₁₃_τ₂) := + isColimitCoforkMapOfIsColimit' π₂ S.w₁₃ S.h₃ + +/-- `L₃.X₃` is the cokernel of `v₁₂.τ₃ : L₁.X₃ ⟶ L₂.X₃`. -/ +noncomputable def h₃τ₃ : IsColimit (CokernelCofork.ofπ S.v₂₃.τ₃ S.w₁₃_τ₃) := + isColimitCoforkMapOfIsColimit' π₃ S.w₁₃ S.h₃ + +instance epi_v₂₃_τ₁ : Epi S.v₂₃.τ₁ := epi_of_isColimit_cofork S.h₃τ₁ +instance epi_v₂₃_τ₂ : Epi S.v₂₃.τ₂ := epi_of_isColimit_cofork S.h₃τ₂ +instance epi_v₂₃_τ₃ : Epi S.v₂₃.τ₃ := epi_of_isColimit_cofork S.h₃τ₃ + +/-- The lower part of the first column of the snake diagram is exact. -/ +lemma exact_C₁_down: (ShortComplex.mk S.v₁₂.τ₁ S.v₂₃.τ₁ + (by rw [← comp_τ₁, S.w₁₃, zero_τ₁])).Exact := + exact_of_g_is_cokernel _ S.h₃τ₁ + +/-- The lower part of the second column of the snake diagram is exact. -/ +lemma exact_C₂_down : (ShortComplex.mk S.v₁₂.τ₂ S.v₂₃.τ₂ + (by rw [← comp_τ₂, S.w₁₃, zero_τ₂])).Exact := + exact_of_g_is_cokernel _ S.h₃τ₂ + +/-- The lower part of the third column of the snake diagram is exact. -/ +lemma exact_C₃_down : (ShortComplex.mk S.v₁₂.τ₃ S.v₂₃.τ₃ + (by rw [← comp_τ₃, S.w₁₃, zero_τ₃])).Exact := + exact_of_g_is_cokernel _ S.h₃τ₃ + +instance epi_L₃_g [Epi S.L₂.g] : Epi S.L₃.g := by + have : Epi (S.v₂₃.τ₂ ≫ S.L₃.g) := by + rw [S.v₂₃.comm₂₃] + apply epi_comp + exact epi_of_epi S.v₂₃.τ₂ _ + +lemma L₀_exact : S.L₀.Exact := by + rw [ShortComplex.exact_iff_exact_up_to_refinements] + intro A x₂ hx₂ + obtain ⟨A₁, π₁, hπ₁, y₁, hy₁⟩ := S.L₁_exact.exact_up_to_refinements (x₂ ≫ S.v₀₁.τ₂) + (by rw [assoc, S.v₀₁.comm₂₃, reassoc_of% hx₂, zero_comp]) + have hy₁' : y₁ ≫ S.v₁₂.τ₁ = 0 := by + simp only [← cancel_mono S.L₂.f, assoc, zero_comp, S.v₁₂.comm₁₂, + ← reassoc_of% hy₁, w₀₂_τ₂, comp_zero] + obtain ⟨x₁, hx₁⟩ : ∃ x₁, x₁ ≫ S.v₀₁.τ₁ = y₁ := ⟨_, S.exact_C₁_up.lift_f y₁ hy₁'⟩ + refine' ⟨A₁, π₁, hπ₁, x₁, _⟩ + simp only [← cancel_mono S.v₀₁.τ₂, assoc, ← S.v₀₁.comm₁₂, reassoc_of% hx₁, hy₁] + +lemma L₃_exact : S.L₃.Exact := S.op.L₀_exact.unop + +/-- The fiber product of `L₁.X₂` and `L₀.X₃` over `L₁.X₃`. This is an auxiliary +object in the construction of the morphism `δ : L₀.X₃ ⟶ L₃.X₁`. -/ +noncomputable def P := pullback S.L₁.g S.v₀₁.τ₃ + +/-- The canonical map `P ⟶ L₂.X₂`. -/ +noncomputable def φ₂ : S.P ⟶ S.L₂.X₂ := pullback.fst ≫ S.v₁₂.τ₂ + +/-- The canonical map `P ⟶ L₂.X₁`. -/ +noncomputable def φ₁ : S.P ⟶ S.L₂.X₁ := + S.L₂_exact.lift S.φ₂ + (by simp only [φ₂, assoc, S.v₁₂.comm₂₃, pullback.condition_assoc, w₀₂_τ₃, comp_zero]) + +@[reassoc (attr := simp)] lemma φ₁_L₂_f : S.φ₁ ≫ S.L₂.f = S.φ₂ := S.L₂_exact.lift_f _ _ + +/-- The short complex that is part of an exact sequence `L₁.X₁ ⟶ P ⟶ L₀.X₃ ⟶ 0`. -/ +noncomputable def L₀' : ShortComplex C where + X₁ := S.L₁.X₁ + X₂ := S.P + X₃ := S.L₀.X₃ + f := pullback.lift S.L₁.f 0 (by simp) + g := pullback.snd + zero := by simp + +@[reassoc (attr := simp)] lemma L₁_f_φ₁ : S.L₀'.f ≫ S.φ₁ = S.v₁₂.τ₁ := by + dsimp only [L₀'] + simp only [← cancel_mono S.L₂.f, assoc, φ₁_L₂_f, φ₂, pullback.lift_fst_assoc, + S.v₁₂.comm₁₂] + +instance : Epi S.L₀'.g := by dsimp only [L₀']; infer_instance + +instance [Mono S.L₁.f] : Mono S.L₀'.f := + mono_of_mono_fac (show S.L₀'.f ≫ pullback.fst = S.L₁.f by simp [L₀']) + +lemma L₀'_exact : S.L₀'.Exact := by + rw [ShortComplex.exact_iff_exact_up_to_refinements] + intro A x₂ hx₂ + dsimp [L₀'] at x₂ hx₂ + obtain ⟨A', π, hπ, x₁, fac⟩ := S.L₁_exact.exact_up_to_refinements (x₂ ≫ pullback.fst) + (by rw [assoc, pullback.condition, reassoc_of% hx₂, zero_comp]) + exact ⟨A', π, hπ, x₁, pullback.hom_ext (by simpa [L₀'] using fac) (by simp [L₀', hx₂])⟩ + +/-- The connecting homomorphism `δ : L₀.X₃ ⟶ L₃.X₁`. -/ +noncomputable def δ : S.L₀.X₃ ⟶ S.L₃.X₁ := + S.L₀'_exact.desc (S.φ₁ ≫ S.v₂₃.τ₁) (by simp only [L₁_f_φ₁_assoc, w₁₃_τ₁]) + +@[reassoc (attr := simp)] +lemma snd_δ : (pullback.snd : S.P ⟶ _) ≫ S.δ = S.φ₁ ≫ S.v₂₃.τ₁ := + S.L₀'_exact.g_desc _ _ + +/-- The pushout of `L₂.X₂` and `L₃.X₁` along `L₂.X₁`. -/ +noncomputable def P' := pushout S.L₂.f S.v₂₃.τ₁ + +lemma snd_δ_inr : (pullback.snd : S.P ⟶ _) ≫ S.δ ≫ (pushout.inr : _ ⟶ S.P') = + pullback.fst ≫ S.v₁₂.τ₂ ≫ pushout.inl := by + simp only [snd_δ_assoc, ← pushout.condition, φ₂, φ₁_L₂_f_assoc, assoc] + +/-- The canonical morphism `L₀.X₂ ⟶ P`. -/ +@[simp] +noncomputable def L₀X₂ToP : S.L₀.X₂ ⟶ S.P := pullback.lift S.v₀₁.τ₂ S.L₀.g S.v₀₁.comm₂₃ + +@[reassoc] +lemma L₀X₂ToP_comp_pullback_snd : S.L₀X₂ToP ≫ pullback.snd = S.L₀.g := by simp + +@[reassoc] +lemma L₀X₂ToP_comp_φ₁ : S.L₀X₂ToP ≫ S.φ₁ = 0 := by + simp only [← cancel_mono S.L₂.f, L₀X₂ToP, assoc, φ₂, φ₁_L₂_f, + pullback.lift_fst_assoc, w₀₂_τ₂, zero_comp] + +lemma L₀_g_δ : S.L₀.g ≫ S.δ = 0 := by + erw [← L₀X₂ToP_comp_pullback_snd, assoc, S.L₀'_exact.g_desc, + L₀X₂ToP_comp_φ₁_assoc, zero_comp] + +lemma δ_L₃_f : S.δ ≫ S.L₃.f = 0 := by + erw [← cancel_epi S.L₀'.g, S.L₀'_exact.g_desc_assoc, assoc, S.v₂₃.comm₁₂, S.φ₁_L₂_f_assoc, + φ₂, assoc, w₁₃_τ₂, comp_zero, comp_zero] + +/-- The short complex `L₀.X₂ ⟶ L₀.X₃ ⟶ L₃.X₁`. -/ +@[simps] +noncomputable def L₁' : ShortComplex C := ShortComplex.mk _ _ S.L₀_g_δ + +/-- The short complex `L₀.X₃ ⟶ L₃.X₁ ⟶ L₃.X₂`. -/ +@[simps] +noncomputable def L₂' : ShortComplex C := ShortComplex.mk _ _ S.δ_L₃_f + +/-- Exactness of `L₀.X₂ ⟶ L₀.X₃ ⟶ L₃.X₁`. -/ +lemma L₁'_exact : S.L₁'.Exact := by + rw [ShortComplex.exact_iff_exact_up_to_refinements] + intro A₀ x₃ hx₃ + dsimp at x₃ hx₃ + obtain ⟨A₁, π₁, hπ₁, p, hp⟩ := surjective_up_to_refinements_of_epi S.L₀'.g x₃ + dsimp [L₀'] at p hp + have hp' : (p ≫ S.φ₁) ≫ S.v₂₃.τ₁ = 0 := by + rw [assoc, ← S.snd_δ, ← reassoc_of% hp, hx₃, comp_zero] + obtain ⟨A₂, π₂, hπ₂, x₁, hx₁⟩ := S.exact_C₁_down.exact_up_to_refinements (p ≫ S.φ₁) hp' + dsimp at x₁ hx₁ + let x₂' := x₁ ≫ S.L₁.f + let x₂ := π₂ ≫ p ≫ pullback.fst + have hx₂' : (x₂ - x₂') ≫ S.v₁₂.τ₂ = 0 := by + simp only [sub_comp, assoc, ← S.v₁₂.comm₁₂, ← reassoc_of% hx₁, φ₂, φ₁_L₂_f, sub_self] + let k₂ : A₂ ⟶ S.L₀.X₂ := S.exact_C₂_up.lift _ hx₂' + have hk₂ : k₂ ≫ S.v₀₁.τ₂ = x₂ - x₂' := S.exact_C₂_up.lift_f _ _ + have hk₂' : k₂ ≫ S.L₀.g = π₂ ≫ p ≫ pullback.snd := by + simp only [← cancel_mono S.v₀₁.τ₃, assoc, ← S.v₀₁.comm₂₃, reassoc_of% hk₂, + sub_comp, S.L₁.zero, comp_zero, sub_zero, pullback.condition] + exact ⟨A₂, π₂ ≫ π₁, epi_comp _ _, k₂, by simp only [assoc, L₁'_f, ← hk₂', hp]⟩ + +/-- The duality isomorphism `S.P ≅ Opposite.unop S.op.P'`. -/ +noncomputable def PIsoUnopOpP' : S.P ≅ Opposite.unop S.op.P' := pullbackIsoUnopPushout _ _ + +/-- The duality isomorphism `S.P' ≅ Opposite.unop S.op.P`. -/ +noncomputable def P'IsoUnopOpP : S.P' ≅ Opposite.unop S.op.P := pushoutIsoUnopPullback _ _ + +lemma op_δ : S.op.δ = S.δ.op := Quiver.Hom.unop_inj (by + rw [Quiver.Hom.unop_op, ← cancel_mono (pushout.inr : _ ⟶ S.P'), + ← cancel_epi (pullback.snd : S.P ⟶ _), S.snd_δ_inr, + ← cancel_mono S.P'IsoUnopOpP.hom, ← cancel_epi S.PIsoUnopOpP'.inv, + P'IsoUnopOpP, PIsoUnopOpP', assoc, assoc, assoc, assoc, + pushoutIsoUnopPullback_inr_hom, pullbackIsoUnopPushout_inv_snd_assoc, + pushoutIsoUnopPullback_inl_hom, pullbackIsoUnopPushout_inv_fst_assoc] + apply Quiver.Hom.op_inj + simpa only [op_comp, Quiver.Hom.op_unop, assoc] using S.op.snd_δ_inr) + +/-- The duality isomorphism `S.L₂'.op ≅ S.op.L₁'`. -/ +noncomputable def L₂'OpIso : S.L₂'.op ≅ S.op.L₁' := + ShortComplex.isoMk (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) + (by dsimp; simp only [id_comp, comp_id, S.op_δ]) + +/-- Exactness of `L₀.X₃ ⟶ L₃.X₁ ⟶ L₃.X₂`. -/ +lemma L₂'_exact : S.L₂'.Exact := by + rw [← exact_op_iff, exact_iff_of_iso S.L₂'OpIso] + exact S.op.L₁'_exact + +end SnakeInput + +end ShortComplex + +end CategoryTheory diff --git a/Mathlib/Algebra/Homology/Single.lean b/Mathlib/Algebra/Homology/Single.lean index 4df4d18925d9e..56aa1491d15bb 100644 --- a/Mathlib/Algebra/Homology/Single.lean +++ b/Mathlib/Algebra/Homology/Single.lean @@ -131,14 +131,6 @@ def single₀ : V ⥤ ChainComplex V ℕ where match n with | 0 => f | n + 1 => 0 } - map_id X := by - ext (_|_) - · rfl - · simp - map_comp f g := by - ext (_|_) - · rfl - · simp #align chain_complex.single₀ ChainComplex.single₀ @[simp] @@ -193,26 +185,26 @@ variable [HasEqualizers V] [HasCokernels V] [HasImages V] [HasImageMaps V] /-- Sending objects to chain complexes supported at `0` then taking `0`-th homology is the same as doing nothing. -/ -noncomputable def homologyFunctor0Single₀ : single₀ V ⋙ homologyFunctor V _ 0 ≅ 𝟭 V := - NatIso.ofComponents (fun X => homology.congr _ _ (by simp) (by simp) ≪≫ homologyZeroZero) +noncomputable def homology'Functor0Single₀ : single₀ V ⋙ homology'Functor V _ 0 ≅ 𝟭 V := + NatIso.ofComponents (fun X => homology'.congr _ _ (by simp) (by simp) ≪≫ homology'ZeroZero) fun f => by -- Porting note: why can't `aesop_cat` do this? dsimp ext simp -#align chain_complex.homology_functor_0_single₀ ChainComplex.homologyFunctor0Single₀ +#align chain_complex.homology_functor_0_single₀ ChainComplex.homology'Functor0Single₀ /-- Sending objects to chain complexes supported at `0` then taking `(n+1)`-st homology is the same as the zero functor. -/ -noncomputable def homologyFunctorSuccSingle₀ (n : ℕ) : - single₀ V ⋙ homologyFunctor V _ (n + 1) ≅ 0 := +noncomputable def homology'FunctorSuccSingle₀ (n : ℕ) : + single₀ V ⋙ homology'Functor V _ (n + 1) ≅ 0 := NatIso.ofComponents (fun X => - homology.congr _ _ (by simp) (by simp) ≪≫ - homologyZeroZero ≪≫ (Functor.zero_obj _).isoZero.symm) + homology'.congr _ _ (by simp) (by simp) ≪≫ + homology'ZeroZero ≪≫ (Functor.zero_obj _).isoZero.symm) fun f => (Functor.zero_obj _).eq_of_tgt _ _ -#align chain_complex.homology_functor_succ_single₀ ChainComplex.homologyFunctorSuccSingle₀ +#align chain_complex.homology_functor_succ_single₀ ChainComplex.homology'FunctorSuccSingle₀ end @@ -273,7 +265,8 @@ def fromSingle₀Equiv (C : ChainComplex V ℕ) (X : V) : ((single₀ V).obj X comm' := fun i j h => by cases i <;> cases j <;> simp only [shape, ComplexShape.down_Rel, Nat.one_ne_zero, not_false_iff, - zero_comp, single₀_obj_X_d, Nat.zero_eq, add_eq_zero, comp_zero] } + zero_comp, single₀_obj_X_d, Nat.zero_eq, add_eq_zero, comp_zero, not_false_eq_true, + Nat.succ_ne_zero, and_self] } left_inv f := by ext i cases i @@ -328,14 +321,6 @@ def single₀ : V ⥤ CochainComplex V ℕ where match n with | 0 => f | n + 1 => 0 } - map_id X := by - ext (_|_) - · rfl - · simp - map_comp f g := by - ext (_|_) - · rfl - · simp #align cochain_complex.single₀ CochainComplex.single₀ @[simp] @@ -390,8 +375,8 @@ variable [HasEqualizers V] [HasCokernels V] [HasImages V] [HasImageMaps V] /-- Sending objects to cochain complexes supported at `0` then taking `0`-th homology is the same as doing nothing. -/ -noncomputable def homologyFunctor0Single₀ : single₀ V ⋙ homologyFunctor V _ 0 ≅ 𝟭 V := - NatIso.ofComponents (fun X => homology.congr _ _ (by simp) (by simp) ≪≫ homologyZeroZero) +noncomputable def homologyFunctor0Single₀ : single₀ V ⋙ homology'Functor V _ 0 ≅ 𝟭 V := + NatIso.ofComponents (fun X => homology'.congr _ _ (by simp) (by simp) ≪≫ homology'ZeroZero) fun f => by -- Porting note: why can't `aesop_cat` do this? dsimp @@ -402,14 +387,14 @@ noncomputable def homologyFunctor0Single₀ : single₀ V ⋙ homologyFunctor V /-- Sending objects to cochain complexes supported at `0` then taking `(n+1)`-st homology is the same as the zero functor. -/ -noncomputable def homologyFunctorSuccSingle₀ (n : ℕ) : - single₀ V ⋙ homologyFunctor V _ (n + 1) ≅ 0 := +noncomputable def homology'FunctorSuccSingle₀ (n : ℕ) : + single₀ V ⋙ homology'Functor V _ (n + 1) ≅ 0 := NatIso.ofComponents (fun X => - homology.congr _ _ (by simp) (by simp) ≪≫ - homologyZeroZero ≪≫ (Functor.zero_obj _).isoZero.symm) + homology'.congr _ _ (by simp) (by simp) ≪≫ + homology'ZeroZero ≪≫ (Functor.zero_obj _).isoZero.symm) fun f => (Functor.zero_obj _).eq_of_tgt _ _ -#align cochain_complex.homology_functor_succ_single₀ CochainComplex.homologyFunctorSuccSingle₀ +#align cochain_complex.homology_functor_succ_single₀ CochainComplex.homology'FunctorSuccSingle₀ end diff --git a/Mathlib/Algebra/IsPrimePow.lean b/Mathlib/Algebra/IsPrimePow.lean index 8050955c70d8c..72ac069276671 100644 --- a/Mathlib/Algebra/IsPrimePow.lean +++ b/Mathlib/Algebra/IsPrimePow.lean @@ -101,7 +101,7 @@ theorem IsPrimePow.dvd {n m : ℕ} (hn : IsPrimePow n) (hm : m ∣ n) (hm₁ : m refine' ⟨p, i, hp, _, rfl⟩ apply Nat.pos_of_ne_zero rintro rfl - simp only [pow_zero, ne_eq] at hm₁ + simp only [pow_zero, ne_eq, not_true_eq_false] at hm₁ #align is_prime_pow.dvd IsPrimePow.dvd theorem Nat.disjoint_divisors_filter_isPrimePow {a b : ℕ} (hab : a.Coprime b) : diff --git a/Mathlib/Algebra/Jordan/Basic.lean b/Mathlib/Algebra/Jordan/Basic.lean index 745fedf19a000..10b66852e895f 100644 --- a/Mathlib/Algebra/Jordan/Basic.lean +++ b/Mathlib/Algebra/Jordan/Basic.lean @@ -194,6 +194,7 @@ private theorem aux0 {a b c : A} : ⁅L (a + b + c), L ((a + b + c) * (a + b + c simp only [lie_add, add_lie, commute_lmul_lmul_sq, zero_add, add_zero] abel +set_option maxHeartbeats 300000 in private theorem aux1 {a b c : A} : ⁅L a + L b + L c, L (a * a) + L (b * b) + L (c * c) + 2 • L (a * b) + 2 • L (c * a) + 2 • L (b * c)⁆ diff --git a/Mathlib/Algebra/Lie/NonUnitalNonAssocAlgebra.lean b/Mathlib/Algebra/Lie/NonUnitalNonAssocAlgebra.lean index 8f755767122d0..45c61709abb2b 100644 --- a/Mathlib/Algebra/Lie/NonUnitalNonAssocAlgebra.lean +++ b/Mathlib/Algebra/Lie/NonUnitalNonAssocAlgebra.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ -import Mathlib.Algebra.Hom.NonUnitalAlg +import Mathlib.Algebra.Algebra.NonUnitalHom import Mathlib.Algebra.Lie.Basic #align_import algebra.lie.non_unital_non_assoc_algebra from "leanprover-community/mathlib"@"841ac1a3d9162bf51c6327812ecb6e5e71883ac4" diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 2f836528f6afd..0723e225aabb4 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -320,8 +320,7 @@ theorem exists_lieIdeal_coe_eq_iff : (∃ I : LieIdeal R L, ↑I = K) ↔ ∀ x y : L, y ∈ K → ⁅x, y⁆ ∈ K := by simp only [← coe_to_submodule_eq_iff, LieIdeal.coe_to_lieSubalgebra_to_submodule, Submodule.exists_lieSubmodule_coe_eq_iff L] - -- Porting note: was `exact Iff.rfl` - simp only [mem_coe_submodule] + exact Iff.rfl #align lie_subalgebra.exists_lie_ideal_coe_eq_iff LieSubalgebra.exists_lieIdeal_coe_eq_iff theorem exists_nested_lieIdeal_coe_eq_iff {K' : LieSubalgebra R L} (h : K ≤ K') : diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index f56d9f900a200..425b2cb718597 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -55,8 +55,7 @@ open scoped BigOperators TensorProduct section notation_weight_space_of /-- Until we define `LieModule.weightSpaceOf`, it is useful to have some notation as follows: -/ -local notation3 (prettyPrint := false) "𝕎("M"," χ"," x")" => - (toEndomorphism R L M x).maximalGeneralizedEigenspace χ +local notation3 "𝕎("M", " χ", " x")" => (toEndomorphism R L M x).maximalGeneralizedEigenspace χ /-- See also `bourbaki1975b` Chapter VII §1.1, Proposition 2 (ii). -/ protected theorem weight_vector_multiplication (M₁ M₂ M₃ : Type*) @@ -125,7 +124,7 @@ protected theorem weight_vector_multiplication (M₁ M₂ M₃ : Type*) -- Eliminate the binomial coefficients from the picture. suffices (f₁ ^ i * f₂ ^ j) (m₁ ⊗ₜ m₂) = 0 by rw [this]; apply smul_zero -- Finish off with appropriate case analysis. - cases' Nat.le_or_le_of_add_eq_add_pred (Finset.Nat.mem_antidiagonal.mp hij) with hi hj + cases' Nat.le_or_le_of_add_eq_add_pred (Finset.mem_antidiagonal.mp hij) with hi hj · rw [(hf_comm.pow_pow i j).eq, LinearMap.mul_apply, LinearMap.pow_map_zero_of_le hi hf₁, LinearMap.map_zero] · rw [LinearMap.mul_apply, LinearMap.pow_map_zero_of_le hj hf₂, LinearMap.map_zero] diff --git a/Mathlib/Algebra/Module/LinearMap.lean b/Mathlib/Algebra/Module/LinearMap.lean index bcbe02f3a94f7..9ef262f585814 100644 --- a/Mathlib/Algebra/Module/LinearMap.lean +++ b/Mathlib/Algebra/Module/LinearMap.lean @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Baanen, Frédéric Dupuis, Heather Macbeth -/ -import Mathlib.Algebra.Hom.GroupAction import Mathlib.Algebra.Module.Pi +import Mathlib.Algebra.Ring.CompTypeclasses import Mathlib.Algebra.Star.Basic import Mathlib.Data.Set.Pointwise.SMul -import Mathlib.Algebra.Ring.CompTypeclasses +import Mathlib.GroupTheory.GroupAction.Hom #align_import algebra.module.linear_map from "leanprover-community/mathlib"@"cc8e88c7c8c7bc80f91f84d11adb584bf9bd658f" diff --git a/Mathlib/Algebra/Module/Submodule/Lattice.lean b/Mathlib/Algebra/Module/Submodule/Lattice.lean index 3be2e8ff64048..c1bafa28014d3 100644 --- a/Mathlib/Algebra/Module/Submodule/Lattice.lean +++ b/Mathlib/Algebra/Module/Submodule/Lattice.lean @@ -130,12 +130,18 @@ def botEquivPUnit : (⊥ : Submodule R M) ≃ₗ[R] PUnit.{v+1} where right_inv _ := rfl #align submodule.bot_equiv_punit Submodule.botEquivPUnit -theorem eq_bot_of_subsingleton (p : Submodule R M) [Subsingleton p] : p = ⊥ := by - rw [eq_bot_iff] - intro v hv - exact congr_arg Subtype.val (Subsingleton.elim (⟨v, hv⟩ : p) 0) +theorem subsingleton_iff_eq_bot : Subsingleton p ↔ p = ⊥ := by + rw [subsingleton_iff, Submodule.eq_bot_iff] + refine ⟨fun h x hx ↦ by simpa using h ⟨x, hx⟩ ⟨0, p.zero_mem⟩, + fun h ⟨x, hx⟩ ⟨y, hy⟩ ↦ by simp [h x hx, h y hy]⟩ + +theorem eq_bot_of_subsingleton [Subsingleton p] : p = ⊥ := + subsingleton_iff_eq_bot.mp inferInstance #align submodule.eq_bot_of_subsingleton Submodule.eq_bot_of_subsingleton +theorem nontrivial_iff_ne_bot : Nontrivial p ↔ p ≠ ⊥ := by + rw [iff_not_comm, not_nontrivial_iff_subsingleton, subsingleton_iff_eq_bot] + /-! ## Top element of a submodule -/ diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean new file mode 100644 index 0000000000000..974269f74ffb7 --- /dev/null +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -0,0 +1,621 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov, Frédéric Dupuis, + Heather Macbeth +-/ + +import Mathlib.Algebra.Module.Submodule.Lattice + +/-! +# `map` and `comap` for `Submodule`s + +## Main declarations + +* `Submodule.map`: The pushforward of a submodule `p ⊆ M` by `f : M → M₂` +* `Submodule.comap`: The pullback of a submodule `p ⊆ M₂` along `f : M → M₂` +* `Submodule.giMapComap`: `map f` and `comap f` form a `GaloisInsertion` when `f` is surjective. +* `Submodule.gciMapComap`: `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. + +## Tags + +submodule, subspace, linear map, pushforward, pullback +-/ + +open Function BigOperators Pointwise Set + +variable {R : Type*} {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} +variable {S : Type*} +variable {M : Type*} {M₁ : Type*} {M₂ : Type*} {M₃ : Type*} + +namespace Submodule + +section AddCommMonoid + +variable [Semiring R] [Semiring R₂] [Semiring R₃] +variable [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] +variable [Module R M] [Module R₂ M₂] [Module R₃ M₃] +variable {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} +variable {σ₂₁ : R₂ →+* R} +variable [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] +variable [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] +variable (p p' : Submodule R M) (q q' : Submodule R₂ M₂) +variable {x : M} + +section + +variable [RingHomSurjective σ₁₂] {F : Type*} [sc : SemilinearMapClass F σ₁₂ M M₂] + +/-- The pushforward of a submodule `p ⊆ M` by `f : M → M₂` -/ +def map (f : F) (p : Submodule R M) : Submodule R₂ M₂ := + { p.toAddSubmonoid.map f with + carrier := f '' p + smul_mem' := by + rintro c x ⟨y, hy, rfl⟩ + obtain ⟨a, rfl⟩ := σ₁₂.surjective c + exact ⟨_, p.smul_mem a hy, map_smulₛₗ f _ _⟩ } +#align submodule.map Submodule.map + +@[simp] +theorem map_coe (f : F) (p : Submodule R M) : (map f p : Set M₂) = f '' p := + rfl +#align submodule.map_coe Submodule.map_coe + +theorem map_toAddSubmonoid (f : M →ₛₗ[σ₁₂] M₂) (p : Submodule R M) : + (p.map f).toAddSubmonoid = p.toAddSubmonoid.map (f : M →+ M₂) := + SetLike.coe_injective rfl +#align submodule.map_to_add_submonoid Submodule.map_toAddSubmonoid + +theorem map_toAddSubmonoid' (f : M →ₛₗ[σ₁₂] M₂) (p : Submodule R M) : + (p.map f).toAddSubmonoid = p.toAddSubmonoid.map f := + SetLike.coe_injective rfl +#align submodule.map_to_add_submonoid' Submodule.map_toAddSubmonoid' + +@[simp] +theorem _root_.AddMonoidHom.coe_toIntLinearMap_map {A A₂ : Type*} [AddCommGroup A] [AddCommGroup A₂] + (f : A →+ A₂) (s : AddSubgroup A) : + (AddSubgroup.toIntSubmodule s).map f.toIntLinearMap = + AddSubgroup.toIntSubmodule (s.map f) := rfl + +@[simp] +theorem _root_.MonoidHom.coe_toAdditive_map {G G₂ : Type*} [Group G] [Group G₂] (f : G →* G₂) + (s : Subgroup G) : + s.toAddSubgroup.map (MonoidHom.toAdditive f) = Subgroup.toAddSubgroup (s.map f) := rfl + +@[simp] +theorem _root_.AddMonoidHom.coe_toMultiplicative_map {G G₂ : Type*} [AddGroup G] [AddGroup G₂] + (f : G →+ G₂) (s : AddSubgroup G) : + s.toSubgroup.map (AddMonoidHom.toMultiplicative f) = AddSubgroup.toSubgroup (s.map f) := rfl + +@[simp] +theorem mem_map {f : F} {p : Submodule R M} {x : M₂} : x ∈ map f p ↔ ∃ y, y ∈ p ∧ f y = x := + Iff.rfl +#align submodule.mem_map Submodule.mem_map + +theorem mem_map_of_mem {f : F} {p : Submodule R M} {r} (h : r ∈ p) : f r ∈ map f p := + Set.mem_image_of_mem _ h +#align submodule.mem_map_of_mem Submodule.mem_map_of_mem + +theorem apply_coe_mem_map (f : F) {p : Submodule R M} (r : p) : f r ∈ map f p := + mem_map_of_mem r.prop +#align submodule.apply_coe_mem_map Submodule.apply_coe_mem_map + +@[simp] +theorem map_id : map (LinearMap.id : M →ₗ[R] M) p = p := + Submodule.ext fun a => by simp +#align submodule.map_id Submodule.map_id + +theorem map_comp [RingHomSurjective σ₂₃] [RingHomSurjective σ₁₃] (f : M →ₛₗ[σ₁₂] M₂) + (g : M₂ →ₛₗ[σ₂₃] M₃) (p : Submodule R M) : map (g.comp f : M →ₛₗ[σ₁₃] M₃) p = map g (map f p) := + SetLike.coe_injective <| by simp only [← image_comp, map_coe, LinearMap.coe_comp, comp_apply] +#align submodule.map_comp Submodule.map_comp + +theorem map_mono {f : F} {p p' : Submodule R M} : p ≤ p' → map f p ≤ map f p' := + image_subset _ +#align submodule.map_mono Submodule.map_mono + +@[simp] +theorem map_zero : map (0 : M →ₛₗ[σ₁₂] M₂) p = ⊥ := + have : ∃ x : M, x ∈ p := ⟨0, p.zero_mem⟩ + ext <| by simp [this, eq_comm] +#align submodule.map_zero Submodule.map_zero + +theorem map_add_le (f g : M →ₛₗ[σ₁₂] M₂) : map (f + g) p ≤ map f p ⊔ map g p := by + rintro x ⟨m, hm, rfl⟩ + exact add_mem_sup (mem_map_of_mem hm) (mem_map_of_mem hm) +#align submodule.map_add_le Submodule.map_add_le + +theorem map_inf_le (f : F) {p q : Submodule R M} : + (p ⊓ q).map f ≤ p.map f ⊓ q.map f := + image_inter_subset f p q + +theorem map_inf (f : F) {p q : Submodule R M} (hf : Injective f) : + (p ⊓ q).map f = p.map f ⊓ q.map f := + SetLike.coe_injective <| Set.image_inter hf + +theorem range_map_nonempty (N : Submodule R M) : + (Set.range (fun ϕ => Submodule.map ϕ N : (M →ₛₗ[σ₁₂] M₂) → Submodule R₂ M₂)).Nonempty := + ⟨_, Set.mem_range.mpr ⟨0, rfl⟩⟩ +#align submodule.range_map_nonempty Submodule.range_map_nonempty + +end + +section SemilinearMap + +variable {F : Type*} [sc : SemilinearMapClass F σ₁₂ M M₂] + +/-- The pushforward of a submodule by an injective linear map is +linearly equivalent to the original submodule. See also `LinearEquiv.submoduleMap` for a +computable version when `f` has an explicit inverse. -/ +noncomputable def equivMapOfInjective (f : F) (i : Injective f) (p : Submodule R M) : + p ≃ₛₗ[σ₁₂] p.map f := + { Equiv.Set.image f p i with + map_add' := by + intros + simp only [coe_add, map_add, Equiv.toFun_as_coe, Equiv.Set.image_apply] + rfl + map_smul' := by + intros + simp only [coe_smul_of_tower, map_smulₛₗ, Equiv.toFun_as_coe, Equiv.Set.image_apply] + rfl } +#align submodule.equiv_map_of_injective Submodule.equivMapOfInjective + +@[simp] +theorem coe_equivMapOfInjective_apply (f : F) (i : Injective f) (p : Submodule R M) (x : p) : + (equivMapOfInjective f i p x : M₂) = f x := + rfl +#align submodule.coe_equiv_map_of_injective_apply Submodule.coe_equivMapOfInjective_apply + +/-- The pullback of a submodule `p ⊆ M₂` along `f : M → M₂` -/ +def comap (f : F) (p : Submodule R₂ M₂) : Submodule R M := + { p.toAddSubmonoid.comap f with + carrier := f ⁻¹' p + smul_mem' := fun a x h => by simp [p.smul_mem (σ₁₂ a) h] } +#align submodule.comap Submodule.comap + +@[simp] +theorem comap_coe (f : F) (p : Submodule R₂ M₂) : (comap f p : Set M) = f ⁻¹' p := + rfl +#align submodule.comap_coe Submodule.comap_coe + +@[simp] +theorem AddMonoidHom.coe_toIntLinearMap_comap {A A₂ : Type*} [AddCommGroup A] [AddCommGroup A₂] + (f : A →+ A₂) (s : AddSubgroup A₂) : + (AddSubgroup.toIntSubmodule s).comap f.toIntLinearMap = + AddSubgroup.toIntSubmodule (s.comap f) := rfl + +@[simp] +theorem mem_comap {f : F} {p : Submodule R₂ M₂} : x ∈ comap f p ↔ f x ∈ p := + Iff.rfl +#align submodule.mem_comap Submodule.mem_comap + +@[simp] +theorem comap_id : comap (LinearMap.id : M →ₗ[R] M) p = p := + SetLike.coe_injective rfl +#align submodule.comap_id Submodule.comap_id + +theorem comap_comp (f : M →ₛₗ[σ₁₂] M₂) (g : M₂ →ₛₗ[σ₂₃] M₃) (p : Submodule R₃ M₃) : + comap (g.comp f : M →ₛₗ[σ₁₃] M₃) p = comap f (comap g p) := + rfl +#align submodule.comap_comp Submodule.comap_comp + +theorem comap_mono {f : F} {q q' : Submodule R₂ M₂} : q ≤ q' → comap f q ≤ comap f q' := + preimage_mono +#align submodule.comap_mono Submodule.comap_mono + +theorem le_comap_pow_of_le_comap (p : Submodule R M) {f : M →ₗ[R] M} (h : p ≤ p.comap f) (k : ℕ) : + p ≤ p.comap (f ^ k) := by + induction' k with k ih + · simp [LinearMap.one_eq_id] + · simp [LinearMap.iterate_succ, comap_comp, h.trans (comap_mono ih)] +#align submodule.le_comap_pow_of_le_comap Submodule.le_comap_pow_of_le_comap + +section + +variable [RingHomSurjective σ₁₂] + +theorem map_le_iff_le_comap {f : F} {p : Submodule R M} {q : Submodule R₂ M₂} : + map f p ≤ q ↔ p ≤ comap f q := + image_subset_iff +#align submodule.map_le_iff_le_comap Submodule.map_le_iff_le_comap + +theorem gc_map_comap (f : F) : GaloisConnection (map f) (comap f) + | _, _ => map_le_iff_le_comap +#align submodule.gc_map_comap Submodule.gc_map_comap + +@[simp] +theorem map_bot (f : F) : map f ⊥ = ⊥ := + (gc_map_comap f).l_bot +#align submodule.map_bot Submodule.map_bot + +@[simp] +theorem map_sup (f : F) : map f (p ⊔ p') = map f p ⊔ map f p' := + (gc_map_comap f : GaloisConnection (map f) (comap f)).l_sup +#align submodule.map_sup Submodule.map_sup + +@[simp] +theorem map_iSup {ι : Sort*} (f : F) (p : ι → Submodule R M) : + map f (⨆ i, p i) = ⨆ i, map f (p i) := + (gc_map_comap f : GaloisConnection (map f) (comap f)).l_iSup +#align submodule.map_supr Submodule.map_iSup + +end + +@[simp] +theorem comap_top (f : F) : comap f ⊤ = ⊤ := + rfl +#align submodule.comap_top Submodule.comap_top + +@[simp] +theorem comap_inf (f : F) : comap f (q ⊓ q') = comap f q ⊓ comap f q' := + rfl +#align submodule.comap_inf Submodule.comap_inf + +@[simp] +theorem comap_iInf [RingHomSurjective σ₁₂] {ι : Sort*} (f : F) (p : ι → Submodule R₂ M₂) : + comap f (⨅ i, p i) = ⨅ i, comap f (p i) := + (gc_map_comap f : GaloisConnection (map f) (comap f)).u_iInf +#align submodule.comap_infi Submodule.comap_iInf + +@[simp] +theorem comap_zero : comap (0 : M →ₛₗ[σ₁₂] M₂) q = ⊤ := + ext <| by simp +#align submodule.comap_zero Submodule.comap_zero + +theorem map_comap_le [RingHomSurjective σ₁₂] (f : F) (q : Submodule R₂ M₂) : + map f (comap f q) ≤ q := + (gc_map_comap f).l_u_le _ +#align submodule.map_comap_le Submodule.map_comap_le + +theorem le_comap_map [RingHomSurjective σ₁₂] (f : F) (p : Submodule R M) : p ≤ comap f (map f p) := + (gc_map_comap f).le_u_l _ +#align submodule.le_comap_map Submodule.le_comap_map + +section GaloisInsertion + +variable {f : F} (hf : Surjective f) + +variable [RingHomSurjective σ₁₂] + +/-- `map f` and `comap f` form a `GaloisInsertion` when `f` is surjective. -/ +def giMapComap : GaloisInsertion (map f) (comap f) := + (gc_map_comap f).toGaloisInsertion fun S x hx => by + rcases hf x with ⟨y, rfl⟩ + simp only [mem_map, mem_comap] + exact ⟨y, hx, rfl⟩ +#align submodule.gi_map_comap Submodule.giMapComap + +theorem map_comap_eq_of_surjective (p : Submodule R₂ M₂) : (p.comap f).map f = p := + (giMapComap hf).l_u_eq _ +#align submodule.map_comap_eq_of_surjective Submodule.map_comap_eq_of_surjective + +theorem map_surjective_of_surjective : Function.Surjective (map f) := + (giMapComap hf).l_surjective +#align submodule.map_surjective_of_surjective Submodule.map_surjective_of_surjective + +theorem comap_injective_of_surjective : Function.Injective (comap f) := + (giMapComap hf).u_injective +#align submodule.comap_injective_of_surjective Submodule.comap_injective_of_surjective + +theorem map_sup_comap_of_surjective (p q : Submodule R₂ M₂) : + (p.comap f ⊔ q.comap f).map f = p ⊔ q := + (giMapComap hf).l_sup_u _ _ +#align submodule.map_sup_comap_of_surjective Submodule.map_sup_comap_of_surjective + +theorem map_iSup_comap_of_sujective {ι : Sort*} (S : ι → Submodule R₂ M₂) : + (⨆ i, (S i).comap f).map f = iSup S := + (giMapComap hf).l_iSup_u _ +#align submodule.map_supr_comap_of_sujective Submodule.map_iSup_comap_of_sujective + +theorem map_inf_comap_of_surjective (p q : Submodule R₂ M₂) : + (p.comap f ⊓ q.comap f).map f = p ⊓ q := + (giMapComap hf).l_inf_u _ _ +#align submodule.map_inf_comap_of_surjective Submodule.map_inf_comap_of_surjective + +theorem map_iInf_comap_of_surjective {ι : Sort*} (S : ι → Submodule R₂ M₂) : + (⨅ i, (S i).comap f).map f = iInf S := + (giMapComap hf).l_iInf_u _ +#align submodule.map_infi_comap_of_surjective Submodule.map_iInf_comap_of_surjective + +theorem comap_le_comap_iff_of_surjective (p q : Submodule R₂ M₂) : p.comap f ≤ q.comap f ↔ p ≤ q := + (giMapComap hf).u_le_u_iff +#align submodule.comap_le_comap_iff_of_surjective Submodule.comap_le_comap_iff_of_surjective + +theorem comap_strictMono_of_surjective : StrictMono (comap f) := + (giMapComap hf).strictMono_u +#align submodule.comap_strict_mono_of_surjective Submodule.comap_strictMono_of_surjective + +end GaloisInsertion + +section GaloisCoinsertion + +variable [RingHomSurjective σ₁₂] {f : F} (hf : Injective f) + +/-- `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. -/ +def gciMapComap : GaloisCoinsertion (map f) (comap f) := + (gc_map_comap f).toGaloisCoinsertion fun S x => by + simp [mem_comap, mem_map, forall_exists_index, and_imp] + intro y hy hxy + rw [hf.eq_iff] at hxy + rwa [← hxy] +#align submodule.gci_map_comap Submodule.gciMapComap + +theorem comap_map_eq_of_injective (p : Submodule R M) : (p.map f).comap f = p := + (gciMapComap hf).u_l_eq _ +#align submodule.comap_map_eq_of_injective Submodule.comap_map_eq_of_injective + +theorem comap_surjective_of_injective : Function.Surjective (comap f) := + (gciMapComap hf).u_surjective +#align submodule.comap_surjective_of_injective Submodule.comap_surjective_of_injective + +theorem map_injective_of_injective : Function.Injective (map f) := + (gciMapComap hf).l_injective +#align submodule.map_injective_of_injective Submodule.map_injective_of_injective + +theorem comap_inf_map_of_injective (p q : Submodule R M) : (p.map f ⊓ q.map f).comap f = p ⊓ q := + (gciMapComap hf).u_inf_l _ _ +#align submodule.comap_inf_map_of_injective Submodule.comap_inf_map_of_injective + +theorem comap_iInf_map_of_injective {ι : Sort*} (S : ι → Submodule R M) : + (⨅ i, (S i).map f).comap f = iInf S := + (gciMapComap hf).u_iInf_l _ +#align submodule.comap_infi_map_of_injective Submodule.comap_iInf_map_of_injective + +theorem comap_sup_map_of_injective (p q : Submodule R M) : (p.map f ⊔ q.map f).comap f = p ⊔ q := + (gciMapComap hf).u_sup_l _ _ +#align submodule.comap_sup_map_of_injective Submodule.comap_sup_map_of_injective + +theorem comap_iSup_map_of_injective {ι : Sort*} (S : ι → Submodule R M) : + (⨆ i, (S i).map f).comap f = iSup S := + (gciMapComap hf).u_iSup_l _ +#align submodule.comap_supr_map_of_injective Submodule.comap_iSup_map_of_injective + +theorem map_le_map_iff_of_injective (p q : Submodule R M) : p.map f ≤ q.map f ↔ p ≤ q := + (gciMapComap hf).l_le_l_iff +#align submodule.map_le_map_iff_of_injective Submodule.map_le_map_iff_of_injective + +theorem map_strictMono_of_injective : StrictMono (map f) := + (gciMapComap hf).strictMono_l +#align submodule.map_strict_mono_of_injective Submodule.map_strictMono_of_injective + +end GaloisCoinsertion + +end SemilinearMap + +section OrderIso + +variable {F : Type*} [SemilinearEquivClass F σ₁₂ M M₂] + +/-- A linear isomorphism induces an order isomorphism of submodules. -/ +@[simps symm_apply apply] +def orderIsoMapComap (f : F) : Submodule R M ≃o Submodule R₂ M₂ where + toFun := map f + invFun := comap f + left_inv := comap_map_eq_of_injective (EquivLike.injective f) + right_inv := map_comap_eq_of_surjective (EquivLike.surjective f) + map_rel_iff' := map_le_map_iff_of_injective (EquivLike.injective f) _ _ +#align submodule.order_iso_map_comap Submodule.orderIsoMapComap + +end OrderIso + +variable {F : Type*} [sc : SemilinearMapClass F σ₁₂ M M₂] + +--TODO(Mario): is there a way to prove this from order properties? +theorem map_inf_eq_map_inf_comap [RingHomSurjective σ₁₂] {f : F} {p : Submodule R M} + {p' : Submodule R₂ M₂} : map f p ⊓ p' = map f (p ⊓ comap f p') := + le_antisymm (by rintro _ ⟨⟨x, h₁, rfl⟩, h₂⟩; exact ⟨_, ⟨h₁, h₂⟩, rfl⟩) + (le_inf (map_mono inf_le_left) (map_le_iff_le_comap.2 inf_le_right)) +#align submodule.map_inf_eq_map_inf_comap Submodule.map_inf_eq_map_inf_comap + +theorem map_comap_subtype : map p.subtype (comap p.subtype p') = p ⊓ p' := + ext fun x => ⟨by rintro ⟨⟨_, h₁⟩, h₂, rfl⟩; exact ⟨h₁, h₂⟩, fun ⟨h₁, h₂⟩ => ⟨⟨_, h₁⟩, h₂, rfl⟩⟩ +#align submodule.map_comap_subtype Submodule.map_comap_subtype + +theorem eq_zero_of_bot_submodule : ∀ b : (⊥ : Submodule R M), b = 0 + | ⟨b', hb⟩ => Subtype.eq <| show b' = 0 from (mem_bot R).1 hb +#align submodule.eq_zero_of_bot_submodule Submodule.eq_zero_of_bot_submodule + +/-- The infimum of a family of invariant submodule of an endomorphism is also an invariant +submodule. -/ +theorem _root_.LinearMap.iInf_invariant {σ : R →+* R} [RingHomSurjective σ] {ι : Sort*} + (f : M →ₛₗ[σ] M) {p : ι → Submodule R M} (hf : ∀ i, ∀ v ∈ p i, f v ∈ p i) : + ∀ v ∈ iInf p, f v ∈ iInf p := by + have : ∀ i, (p i).map f ≤ p i := by + rintro i - ⟨v, hv, rfl⟩ + exact hf i v hv + suffices (iInf p).map f ≤ iInf p by exact fun v hv => this ⟨v, hv, rfl⟩ + exact le_iInf fun i => (Submodule.map_mono (iInf_le p i)).trans (this i) +#align linear_map.infi_invariant LinearMap.iInf_invariant + +end AddCommMonoid + +section AddCommGroup + +variable [Ring R] [AddCommGroup M] [Module R M] (p : Submodule R M) + +variable [AddCommGroup M₂] [Module R M₂] + +@[simp] +protected theorem map_neg (f : M →ₗ[R] M₂) : map (-f) p = map f p := + ext fun _ => + ⟨fun ⟨x, hx, hy⟩ => hy ▸ ⟨-x, show -x ∈ p from neg_mem hx, map_neg f x⟩, fun ⟨x, hx, hy⟩ => + hy ▸ ⟨-x, show -x ∈ p from neg_mem hx, (map_neg (-f) _).trans (neg_neg (f x))⟩⟩ +#align submodule.map_neg Submodule.map_neg + +end AddCommGroup + +end Submodule + +namespace Submodule + +variable {K : Type*} {V : Type*} {V₂ : Type*} +variable [Semifield K] +variable [AddCommMonoid V] [Module K V] +variable [AddCommMonoid V₂] [Module K V₂] + +theorem comap_smul (f : V →ₗ[K] V₂) (p : Submodule K V₂) (a : K) (h : a ≠ 0) : + p.comap (a • f) = p.comap f := by + ext b; simp only [Submodule.mem_comap, p.smul_mem_iff h, LinearMap.smul_apply] +#align submodule.comap_smul Submodule.comap_smul + +protected theorem map_smul (f : V →ₗ[K] V₂) (p : Submodule K V) (a : K) (h : a ≠ 0) : + p.map (a • f) = p.map f := + le_antisymm (by rw [map_le_iff_le_comap, comap_smul f _ a h, ← map_le_iff_le_comap]) + (by rw [map_le_iff_le_comap, ← comap_smul f _ a h, ← map_le_iff_le_comap]) +#align submodule.map_smul Submodule.map_smul + +theorem comap_smul' (f : V →ₗ[K] V₂) (p : Submodule K V₂) (a : K) : + p.comap (a • f) = ⨅ _ : a ≠ 0, p.comap f := by + classical by_cases h : a = 0 <;> simp [h, comap_smul] +#align submodule.comap_smul' Submodule.comap_smul' + +theorem map_smul' (f : V →ₗ[K] V₂) (p : Submodule K V) (a : K) : + p.map (a • f) = ⨆ _ : a ≠ 0, map f p := by + classical by_cases h : a = 0 <;> simp [h, Submodule.map_smul] +#align submodule.map_smul' Submodule.map_smul' + +end Submodule + +namespace Submodule + +section Module + +variable [Semiring R] [AddCommMonoid M] [Module R M] + +/-- If `s ≤ t`, then we can view `s` as a submodule of `t` by taking the comap +of `t.subtype`. -/ +@[simps symm_apply] +def comapSubtypeEquivOfLe {p q : Submodule R M} (hpq : p ≤ q) : comap q.subtype p ≃ₗ[R] p where + toFun x := ⟨x, x.2⟩ + invFun x := ⟨⟨x, hpq x.2⟩, x.2⟩ + left_inv x := by simp only [coe_mk, SetLike.eta, LinearEquiv.coe_coe] + right_inv x := by simp only [Subtype.coe_mk, SetLike.eta, LinearEquiv.coe_coe] + map_add' x y := rfl + map_smul' c x := rfl +#align submodule.comap_subtype_equiv_of_le Submodule.comapSubtypeEquivOfLe +#align submodule.comap_subtype_equiv_of_le_symm_apply_coe_coe Submodule.comapSubtypeEquivOfLe_symm_apply + +-- Porting note: The original theorem generated by `simps` was using `LinearEquiv.toLinearMap`, +-- different from the theorem on Lean 3, and not simp-normal form. +@[simp] +theorem comapSubtypeEquivOfLe_apply_coe {p q : Submodule R M} (hpq : p ≤ q) + (x : comap q.subtype p) : + (comapSubtypeEquivOfLe hpq x : M) = (x : M) := + rfl +#align submodule.comap_subtype_equiv_of_le_apply_coe Submodule.comapSubtypeEquivOfLe_apply_coe + +end Module + +end Submodule + +namespace Submodule + +variable [Semiring R] [Semiring R₂] + +variable [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] + +variable {τ₁₂ : R →+* R₂} {τ₂₁ : R₂ →+* R} + +variable [RingHomInvPair τ₁₂ τ₂₁] [RingHomInvPair τ₂₁ τ₁₂] + +variable (p : Submodule R M) (q : Submodule R₂ M₂) + +-- Porting note: Was `@[simp]`. +@[simp high] +theorem mem_map_equiv {e : M ≃ₛₗ[τ₁₂] M₂} {x : M₂} : + x ∈ p.map (e : M →ₛₗ[τ₁₂] M₂) ↔ e.symm x ∈ p := by + rw [Submodule.mem_map]; constructor + · rintro ⟨y, hy, hx⟩ + simp [← hx, hy] + · intro hx + refine' ⟨e.symm x, hx, by simp⟩ +#align submodule.mem_map_equiv Submodule.mem_map_equiv + +theorem map_equiv_eq_comap_symm (e : M ≃ₛₗ[τ₁₂] M₂) (K : Submodule R M) : + K.map (e : M →ₛₗ[τ₁₂] M₂) = K.comap (e.symm : M₂ →ₛₗ[τ₂₁] M) := + Submodule.ext fun _ => by rw [mem_map_equiv, mem_comap, LinearEquiv.coe_coe] +#align submodule.map_equiv_eq_comap_symm Submodule.map_equiv_eq_comap_symm + +theorem comap_equiv_eq_map_symm (e : M ≃ₛₗ[τ₁₂] M₂) (K : Submodule R₂ M₂) : + K.comap (e : M →ₛₗ[τ₁₂] M₂) = K.map (e.symm : M₂ →ₛₗ[τ₂₁] M) := + (map_equiv_eq_comap_symm e.symm K).symm +#align submodule.comap_equiv_eq_map_symm Submodule.comap_equiv_eq_map_symm + +variable {p} + +theorem map_symm_eq_iff (e : M ≃ₛₗ[τ₁₂] M₂) {K : Submodule R₂ M₂} : + K.map e.symm = p ↔ p.map e = K := by + constructor <;> rintro rfl + · calc + map e (map e.symm K) = comap e.symm (map e.symm K) := map_equiv_eq_comap_symm _ _ + _ = K := comap_map_eq_of_injective e.symm.injective _ + · calc + map e.symm (map e p) = comap e (map e p) := (comap_equiv_eq_map_symm _ _).symm + _ = p := comap_map_eq_of_injective e.injective _ +#align submodule.map_symm_eq_iff Submodule.map_symm_eq_iff + +theorem orderIsoMapComap_apply' (e : M ≃ₛₗ[τ₁₂] M₂) (p : Submodule R M) : + orderIsoMapComap e p = comap e.symm p := + p.map_equiv_eq_comap_symm _ +#align submodule.order_iso_map_comap_apply' Submodule.orderIsoMapComap_apply' + +theorem orderIsoMapComap_symm_apply' (e : M ≃ₛₗ[τ₁₂] M₂) (p : Submodule R₂ M₂) : + (orderIsoMapComap e).symm p = map e.symm p := + p.comap_equiv_eq_map_symm _ +#align submodule.order_iso_map_comap_symm_apply' Submodule.orderIsoMapComap_symm_apply' + +theorem inf_comap_le_comap_add (f₁ f₂ : M →ₛₗ[τ₁₂] M₂) : + comap f₁ q ⊓ comap f₂ q ≤ comap (f₁ + f₂) q := by + rw [SetLike.le_def] + intro m h + change f₁ m + f₂ m ∈ q + change f₁ m ∈ q ∧ f₂ m ∈ q at h + apply q.add_mem h.1 h.2 +#align submodule.inf_comap_le_comap_add Submodule.inf_comap_le_comap_add + +end Submodule + +namespace Submodule + +variable {N N₂ : Type*} + +variable [CommSemiring R] [CommSemiring R₂] + +variable [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] + +variable [AddCommMonoid N] [AddCommMonoid N₂] [Module R N] [Module R N₂] + +variable {τ₁₂ : R →+* R₂} {τ₂₁ : R₂ →+* R} + +variable [RingHomInvPair τ₁₂ τ₂₁] [RingHomInvPair τ₂₁ τ₁₂] + +variable (p : Submodule R M) (q : Submodule R₂ M₂) + +variable (pₗ : Submodule R N) (qₗ : Submodule R N₂) + +theorem comap_le_comap_smul (fₗ : N →ₗ[R] N₂) (c : R) : comap fₗ qₗ ≤ comap (c • fₗ) qₗ := by + rw [SetLike.le_def] + intro m h + change c • fₗ m ∈ qₗ + change fₗ m ∈ qₗ at h + apply qₗ.smul_mem _ h +#align submodule.comap_le_comap_smul Submodule.comap_le_comap_smul + +/-- Given modules `M`, `M₂` over a commutative ring, together with submodules `p ⊆ M`, `q ⊆ M₂`, +the set of maps $\{f ∈ Hom(M, M₂) | f(p) ⊆ q \}$ is a submodule of `Hom(M, M₂)`. -/ +def compatibleMaps : Submodule R (N →ₗ[R] N₂) where + carrier := { fₗ | pₗ ≤ comap fₗ qₗ } + zero_mem' := by + change pₗ ≤ comap (0 : N →ₗ[R] N₂) qₗ + rw [comap_zero] + refine' le_top + add_mem' {f₁ f₂} h₁ h₂ := by + apply le_trans _ (inf_comap_le_comap_add qₗ f₁ f₂) + rw [le_inf_iff] + exact ⟨h₁, h₂⟩ + smul_mem' c fₗ h := by + dsimp at h + exact le_trans h (comap_le_comap_smul qₗ fₗ c) +#align submodule.compatible_maps Submodule.compatibleMaps + +end Submodule diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index 04a4c0c4882a6..55b89da97b872 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury G. Kudryashov, Scott Morrison -/ import Mathlib.Algebra.Algebra.Equiv +import Mathlib.Algebra.Algebra.NonUnitalHom import Mathlib.Algebra.BigOperators.Finsupp -import Mathlib.Algebra.Hom.NonUnitalAlg import Mathlib.Algebra.Module.BigOperators import Mathlib.LinearAlgebra.Finsupp diff --git a/Mathlib/Algebra/NeZero.lean b/Mathlib/Algebra/NeZero.lean index 10870d796632e..121211ebb14f2 100644 --- a/Mathlib/Algebra/NeZero.lean +++ b/Mathlib/Algebra/NeZero.lean @@ -39,6 +39,9 @@ theorem neZero_iff {n : R} : NeZero n ↔ n ≠ 0 := ⟨fun h ↦ h.out, NeZero.mk⟩ #align ne_zero_iff neZero_iff +@[simp] lemma neZero_zero_iff_false {α : Type*} [Zero α] : NeZero (0 : α) ↔ False := + ⟨fun h ↦ h.ne rfl, fun h ↦ h.elim⟩ + theorem not_neZero {n : R} : ¬NeZero n ↔ n = 0 := by simp [neZero_iff] #align not_ne_zero not_neZero diff --git a/Mathlib/Algebra/Order/CompleteField.lean b/Mathlib/Algebra/Order/CompleteField.lean index 0debf923fce4b..b290f004d4725 100644 --- a/Mathlib/Algebra/Order/CompleteField.lean +++ b/Mathlib/Algebra/Order/CompleteField.lean @@ -155,7 +155,8 @@ theorem cutMap_add (a b : α) : cutMap β (a + b) = cutMap β a + cutMap β b := rw [coe_mem_cutMap_iff] exact_mod_cast sub_lt_comm.mp hq₁q · rintro _ ⟨_, _, ⟨qa, ha, rfl⟩, ⟨qb, hb, rfl⟩, rfl⟩ - refine' ⟨qa + qb, _, by norm_cast⟩ + -- After leanprover/lean4#2734, `norm_cast` needs help with beta reduction. + refine' ⟨qa + qb, _, by beta_reduce; norm_cast⟩ rw [mem_setOf_eq, cast_add] exact add_lt_add ha hb #align linear_ordered_field.cut_map_add LinearOrderedField.cutMap_add diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index a47e675ea4234..78764f36a702a 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -658,6 +658,7 @@ notation "⌊" a "⌋" => Int.floor a notation "⌈" a "⌉" => Int.ceil a -- Mathematical notation for `fract a` is usually `{a}`. Let's not even go there. + @[simp] theorem floorRing_floor_eq : @FloorRing.floor = @Int.floor := rfl diff --git a/Mathlib/Algebra/Order/Group/Abs.lean b/Mathlib/Algebra/Order/Group/Abs.lean index 018135c67a63f..99a578fa0f03b 100644 --- a/Mathlib/Algebra/Order/Group/Abs.lean +++ b/Mathlib/Algebra/Order/Group/Abs.lean @@ -368,4 +368,32 @@ theorem eq_of_abs_sub_nonpos (h : |a - b| ≤ 0) : a = b := eq_of_abs_sub_eq_zero (le_antisymm h (abs_nonneg (a - b))) #align eq_of_abs_sub_nonpos eq_of_abs_sub_nonpos +@[simp] +theorem abs_eq_self : |a| = a ↔ 0 ≤ a := by + rw [abs_eq_max_neg, max_eq_left_iff, neg_le_self_iff] +#align abs_eq_self abs_eq_self + +@[simp] +theorem abs_eq_neg_self : |a| = -a ↔ a ≤ 0 := by + rw [abs_eq_max_neg, max_eq_right_iff, le_neg_self_iff] +#align abs_eq_neg_self abs_eq_neg_self + +/-- For an element `a` of a linear ordered ring, either `abs a = a` and `0 ≤ a`, + or `abs a = -a` and `a < 0`. + Use cases on this lemma to automate linarith in inequalities -/ +theorem abs_cases (a : α) : |a| = a ∧ 0 ≤ a ∨ |a| = -a ∧ a < 0 := by + by_cases h : 0 ≤ a + · left + exact ⟨abs_eq_self.mpr h, h⟩ + · right + push_neg at h + exact ⟨abs_eq_neg_self.mpr (le_of_lt h), h⟩ +#align abs_cases abs_cases + +@[simp] +theorem max_zero_add_max_neg_zero_eq_abs_self (a : α) : max a 0 + max (-a) 0 = |a| := by + symm + rcases le_total 0 a with (ha | ha) <;> simp [ha] +#align max_zero_add_max_neg_zero_eq_abs_self max_zero_add_max_neg_zero_eq_abs_self + end LinearOrderedAddCommGroup diff --git a/Mathlib/Algebra/Order/Group/Defs.lean b/Mathlib/Algebra/Order/Group/Defs.lean index c28c15b42ea6b..f89249f355ae5 100644 --- a/Mathlib/Algebra/Order/Group/Defs.lean +++ b/Mathlib/Algebra/Order/Group/Defs.lean @@ -1159,6 +1159,22 @@ instance (priority := 100) LinearOrderedCommGroup.toLinearOrderedCancelCommMonoi #align linear_ordered_comm_group.to_linear_ordered_cancel_comm_monoid LinearOrderedCommGroup.toLinearOrderedCancelCommMonoid #align linear_ordered_add_comm_group.to_linear_ordered_cancel_add_comm_monoid LinearOrderedAddCommGroup.toLinearOrderedAddCancelCommMonoid +@[to_additive, simp] +theorem inv_le_self_iff : a⁻¹ ≤ a ↔ 1 ≤ a := by simp [inv_le_iff_one_le_mul'] +#align neg_le_self_iff neg_le_self_iff + +@[to_additive, simp] +theorem inv_lt_self_iff : a⁻¹ < a ↔ 1 < a := by simp [inv_lt_iff_one_lt_mul] +#align neg_lt_self_iff neg_lt_self_iff + +@[to_additive, simp] +theorem le_inv_self_iff : a ≤ a⁻¹ ↔ a ≤ 1 := by simp [← not_iff_not] +#align le_neg_self_iff le_neg_self_iff + +@[to_additive, simp] +theorem lt_inv_self_iff : a < a⁻¹ ↔ a < 1 := by simp [← not_iff_not] +#align lt_neg_self_iff lt_neg_self_iff + end LinearOrderedCommGroup namespace AddCommGroup diff --git a/Mathlib/Algebra/Order/Group/OrderIso.lean b/Mathlib/Algebra/Order/Group/OrderIso.lean index 545090a55366c..09c9a2c85a3ac 100644 --- a/Mathlib/Algebra/Order/Group/OrderIso.lean +++ b/Mathlib/Algebra/Order/Group/OrderIso.lean @@ -3,8 +3,8 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ +import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Algebra.Order.Group.Defs -import Mathlib.Algebra.Hom.Equiv.Units.Basic #align_import algebra.order.group.order_iso from "leanprover-community/mathlib"@"6632ca2081e55ff5cf228ca63011979a0efb495b" diff --git a/Mathlib/Algebra/Order/Hom/Monoid.lean b/Mathlib/Algebra/Order/Hom/Monoid.lean index cf6c0080f027c..7619d3caba290 100644 --- a/Mathlib/Algebra/Order/Hom/Monoid.lean +++ b/Mathlib/Algebra/Order/Hom/Monoid.lean @@ -3,10 +3,10 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Pi.Algebra -import Mathlib.Algebra.Hom.Group.Basic +import Mathlib.Algebra.Group.Hom.Basic import Mathlib.Algebra.Order.Group.Instances import Mathlib.Algebra.Order.Monoid.WithZero.Defs +import Mathlib.Data.Pi.Algebra import Mathlib.Order.Hom.Basic #align_import algebra.order.hom.monoid from "leanprover-community/mathlib"@"3342d1b2178381196f818146ff79bc0e7ccd9e2d" diff --git a/Mathlib/Algebra/Order/Module.lean b/Mathlib/Algebra/Order/Module.lean index 38413e037c287..029187d36b8a7 100644 --- a/Mathlib/Algebra/Order/Module.lean +++ b/Mathlib/Algebra/Order/Module.lean @@ -24,7 +24,7 @@ ordered module, ordered scalar, ordered smul, ordered action, ordered vector spa open Pointwise -variable {k M N : Type*} +variable {ι k M N : Type*} instance instModuleOrderDual [Semiring k] [OrderedAddCommMonoid M] [Module k M] : Module k Mᵒᵈ where @@ -48,7 +48,7 @@ end Semiring section Ring -variable [OrderedRing k] [OrderedAddCommGroup M] [Module k M] [OrderedSMul k M] {a b : M} {c : k} +variable [OrderedRing k] [OrderedAddCommGroup M] [Module k M] [OrderedSMul k M] {a b : M} {c d : k} theorem smul_lt_smul_of_neg (h : a < b) (hc : c < 0) : c • b < c • a := by rw [← neg_neg c, neg_smul, neg_smul (-c), neg_lt_neg_iff] @@ -70,6 +70,12 @@ theorem lt_of_smul_lt_smul_of_nonpos (h : c • a < c • b) (hc : c ≤ 0) : b exact lt_of_smul_lt_smul_of_nonneg h (neg_nonneg_of_nonpos hc) #align lt_of_smul_lt_smul_of_nonpos lt_of_smul_lt_smul_of_nonpos +lemma smul_le_smul_of_nonneg_right (h : c ≤ d) (hb : 0 ≤ b) : c • b ≤ d • b := by + rw [←sub_nonneg, ←sub_smul]; exact smul_nonneg (sub_nonneg.2 h) hb + +lemma smul_le_smul (hcd : c ≤ d) (hab : a ≤ b) (ha : 0 ≤ a) (hd : 0 ≤ d) : c • a ≤ d • b := + (smul_le_smul_of_nonneg_right hcd ha).trans $ smul_le_smul_of_nonneg_left hab hd + theorem smul_lt_smul_iff_of_neg (hc : c < 0) : c • a < c • b ↔ b < a := by rw [← neg_neg c, neg_smul, neg_smul (-c), neg_lt_neg_iff] exact smul_lt_smul_iff_of_pos (neg_pos_of_neg hc) @@ -218,7 +224,8 @@ theorem BddAbove.smul_of_nonpos (hc : c ≤ 0) (hs : BddAbove s) : BddBelow (c end OrderedRing section LinearOrderedRing -variable [LinearOrderedRing k] [LinearOrderedAddCommGroup M] [Module k M] [OrderedSMul k M] {a : k} +variable [LinearOrderedRing k] [LinearOrderedAddCommGroup M] [Module k M] [OrderedSMul k M] + {f : ι → k} {g : ι → M} {s : Set ι} {a a₁ a₂ : k} {b b₁ b₂ : M} theorem smul_max_of_nonpos (ha : a ≤ 0) (b₁ b₂ : M) : a • max b₁ b₂ = min (a • b₁) (a • b₂) := (antitone_smul_left ha : Antitone (_ : M → M)).map_max @@ -228,6 +235,37 @@ theorem smul_min_of_nonpos (ha : a ≤ 0) (b₁ b₂ : M) : a • min b₁ b₂ (antitone_smul_left ha : Antitone (_ : M → M)).map_min #align smul_min_of_nonpos smul_min_of_nonpos +lemma nonneg_and_nonneg_or_nonpos_and_nonpos_of_smul_nonneg (hab : 0 ≤ a • b) : + 0 ≤ a ∧ 0 ≤ b ∨ a ≤ 0 ∧ b ≤ 0 := by + simp only [Decidable.or_iff_not_and_not, not_and, not_le] + refine fun ab nab ↦ hab.not_lt ?_ + obtain ha | rfl | ha := lt_trichotomy 0 a + exacts [smul_neg_of_pos_of_neg ha (ab ha.le), ((ab le_rfl).asymm (nab le_rfl)).elim, + smul_neg_of_neg_of_pos ha (nab ha.le)] + +lemma smul_nonneg_iff : 0 ≤ a • b ↔ 0 ≤ a ∧ 0 ≤ b ∨ a ≤ 0 ∧ b ≤ 0 := + ⟨nonneg_and_nonneg_or_nonpos_and_nonpos_of_smul_nonneg, + fun h ↦ h.elim (and_imp.2 smul_nonneg) (and_imp.2 smul_nonneg_of_nonpos_of_nonpos)⟩ + +lemma smul_nonpos_iff : a • b ≤ 0 ↔ 0 ≤ a ∧ b ≤ 0 ∨ a ≤ 0 ∧ 0 ≤ b := by + rw [←neg_nonneg, ←smul_neg, smul_nonneg_iff, neg_nonneg, neg_nonpos] + +lemma smul_nonneg_iff_pos_imp_nonneg : 0 ≤ a • b ↔ (0 < a → 0 ≤ b) ∧ (0 < b → 0 ≤ a) := by + refine smul_nonneg_iff.trans ?_ + simp_rw [←not_le, ←or_iff_not_imp_left] + have := le_total a 0 + have := le_total b 0 + tauto + +lemma smul_nonneg_iff_neg_imp_nonpos : 0 ≤ a • b ↔ (a < 0 → b ≤ 0) ∧ (b < 0 → a ≤ 0) := by + rw [←neg_smul_neg, smul_nonneg_iff_pos_imp_nonneg]; simp only [neg_pos, neg_nonneg] + +lemma smul_nonpos_iff_pos_imp_nonpos : a • b ≤ 0 ↔ (0 < a → b ≤ 0) ∧ (b < 0 → 0 ≤ a) := by + rw [←neg_nonneg, ←smul_neg, smul_nonneg_iff_pos_imp_nonneg]; simp only [neg_pos, neg_nonneg] + +lemma smul_nonpos_iff_neg_imp_nonneg : a • b ≤ 0 ↔ (a < 0 → 0 ≤ b) ∧ (0 < b → a ≤ 0) := by + rw [←neg_nonneg, ←neg_smul, smul_nonneg_iff_pos_imp_nonneg]; simp only [neg_pos, neg_nonneg] + end LinearOrderedRing section LinearOrderedField diff --git a/Mathlib/Algebra/Order/Monoid/Defs.lean b/Mathlib/Algebra/Order/Monoid/Defs.lean index d7cbc25e79f4a..9daf15aeffd50 100644 --- a/Mathlib/Algebra/Order/Monoid/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Defs.lean @@ -112,3 +112,19 @@ theorem add_top (a : α) : a + ⊤ = ⊤ := #align add_top add_top end LinearOrderedAddCommMonoidWithTop + +variable [LinearOrderedCommMonoid α] {a : α} + +@[to_additive, simp] +theorem one_le_mul_self_iff : 1 ≤ a * a ↔ 1 ≤ a := + ⟨(fun h ↦ by push_neg at h ⊢; exact mul_lt_one' h h).mtr, fun h ↦ one_le_mul h h⟩ + +@[to_additive, simp] +theorem one_lt_mul_self_iff : 1 < a * a ↔ 1 < a := + ⟨(fun h ↦ by push_neg at h ⊢; exact mul_le_one' h h).mtr, fun h ↦ one_lt_mul'' h h⟩ + +@[to_additive, simp] +theorem mul_self_le_one_iff : a * a ≤ 1 ↔ a ≤ 1 := by simp [← not_iff_not] + +@[to_additive, simp] +theorem mul_self_lt_one_iff : a * a < 1 ↔ a < 1 := by simp [← not_iff_not] diff --git a/Mathlib/Algebra/Order/Monoid/WithTop.lean b/Mathlib/Algebra/Order/Monoid/WithTop.lean index 05fe395b28c27..26d379e4b5fb0 100644 --- a/Mathlib/Algebra/Order/Monoid/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/WithTop.lean @@ -3,12 +3,12 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ -import Mathlib.Algebra.Hom.Group.Defs +import Mathlib.Algebra.CharZero.Defs +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Algebra.Order.Monoid.OrderDual import Mathlib.Algebra.Order.Monoid.WithZero.Basic -import Mathlib.Data.Nat.Cast.Defs import Mathlib.Algebra.Order.ZeroLEOne -import Mathlib.Algebra.CharZero.Defs +import Mathlib.Data.Nat.Cast.Defs #align_import algebra.order.monoid.with_top from "leanprover-community/mathlib"@"0111834459f5d7400215223ea95ae38a1265a907" diff --git a/Mathlib/Algebra/Order/Nonneg/Module.lean b/Mathlib/Algebra/Order/Nonneg/Module.lean index 96b07b3d335e9..02b6dba734143 100644 --- a/Mathlib/Algebra/Order/Nonneg/Module.lean +++ b/Mathlib/Algebra/Order/Nonneg/Module.lean @@ -21,8 +21,7 @@ variable {𝕜 𝕜' E : Type*} variable [OrderedSemiring 𝕜] --- TODO: remove `prettyPrint := false` once #6833 is merged -local notation3 (prettyPrint := false) "𝕜≥0" => {c : 𝕜 // 0 ≤ c} +local notation3 "𝕜≥0" => {c : 𝕜 // 0 ≤ c} namespace Nonneg diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index d62e2a3219a26..bad88dc014d28 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -54,32 +54,6 @@ theorem abs_mul_abs_self (a : α) : |a| * |a| = a * a := theorem abs_mul_self (a : α) : |a * a| = a * a := by rw [abs_mul, abs_mul_abs_self] #align abs_mul_self abs_mul_self -@[simp] -theorem abs_eq_self : |a| = a ↔ 0 ≤ a := by simp [abs_eq_max_neg] -#align abs_eq_self abs_eq_self - -@[simp] -theorem abs_eq_neg_self : |a| = -a ↔ a ≤ 0 := by simp [abs_eq_max_neg] -#align abs_eq_neg_self abs_eq_neg_self - -/-- For an element `a` of a linear ordered ring, either `abs a = a` and `0 ≤ a`, - or `abs a = -a` and `a < 0`. - Use cases on this lemma to automate linarith in inequalities -/ -theorem abs_cases (a : α) : |a| = a ∧ 0 ≤ a ∨ |a| = -a ∧ a < 0 := by - by_cases h : 0 ≤ a - · left - exact ⟨abs_eq_self.mpr h, h⟩ - · right - push_neg at h - exact ⟨abs_eq_neg_self.mpr (le_of_lt h), h⟩ -#align abs_cases abs_cases - -@[simp] -theorem max_zero_add_max_neg_zero_eq_abs_self (a : α) : max a 0 + max (-a) 0 = |a| := by - symm - rcases le_total 0 a with (ha | ha) <;> simp [ha] -#align max_zero_add_max_neg_zero_eq_abs_self max_zero_add_max_neg_zero_eq_abs_self - theorem abs_eq_iff_mul_self_eq : |a| = |b| ↔ a * a = b * b := by rw [← abs_mul_abs_self, ← abs_mul_abs_self b] exact (mul_self_inj (abs_nonneg a) (abs_nonneg b)).symm diff --git a/Mathlib/Algebra/Order/Ring/Defs.lean b/Mathlib/Algebra/Order/Ring/Defs.lean index 361c96450c564..99cc43176c4fc 100644 --- a/Mathlib/Algebra/Order/Ring/Defs.lean +++ b/Mathlib/Algebra/Order/Ring/Defs.lean @@ -800,12 +800,9 @@ theorem nonneg_and_nonneg_or_nonpos_and_nonpos_of_mul_nonneg (hab : 0 ≤ a * b) 0 ≤ a ∧ 0 ≤ b ∨ a ≤ 0 ∧ b ≤ 0 := by refine' Decidable.or_iff_not_and_not.2 _ simp only [not_and, not_le]; intro ab nab; apply not_lt_of_le hab _ - -- Porting note: for the middle case, we used to have `rfl`, but it is now rejected. - -- https://github.com/leanprover/std4/issues/62 - rcases lt_trichotomy 0 a with (ha | ha | ha) + rcases lt_trichotomy 0 a with (ha | rfl | ha) · exact mul_neg_of_pos_of_neg ha (ab ha.le) - · subst ha - exact ((ab le_rfl).asymm (nab le_rfl)).elim + · exact ((ab le_rfl).asymm (nab le_rfl)).elim · exact mul_neg_of_neg_of_pos ha (nab ha.le) #align nonneg_and_nonneg_or_nonpos_and_nonpos_of_mul_nnonneg nonneg_and_nonneg_or_nonpos_and_nonpos_of_mul_nonneg @@ -835,11 +832,8 @@ theorem nonpos_of_mul_nonpos_right (h : a * b ≤ 0) (ha : 0 < a) : b ≤ 0 := @[simp] theorem zero_le_mul_left (h : 0 < c) : 0 ≤ c * b ↔ 0 ≤ b := by - -- Porting note: this used to be by: - -- convert mul_le_mul_left h - -- simp - -- but the `convert` no longer works. - simpa using (mul_le_mul_left h : c * 0 ≤ c * b ↔ 0 ≤ b) + convert mul_le_mul_left h + simp #align zero_le_mul_left zero_le_mul_left @[simp] @@ -1097,32 +1091,6 @@ theorem mul_self_nonneg (a : α) : 0 ≤ a * a := (le_total 0 a).elim (fun h => mul_nonneg h h) fun h => mul_nonneg_of_nonpos_of_nonpos h h #align mul_self_nonneg mul_self_nonneg -@[simp] -theorem neg_le_self_iff : -a ≤ a ↔ 0 ≤ a := by - simp [neg_le_iff_add_nonneg, ← two_mul, mul_nonneg_iff, zero_le_one, (zero_lt_two' α).not_le] -#align neg_le_self_iff neg_le_self_iff - -@[simp] -theorem neg_lt_self_iff : -a < a ↔ 0 < a := by - simp [neg_lt_iff_pos_add, ← two_mul, mul_pos_iff, zero_lt_one, (zero_lt_two' α).not_lt] -#align neg_lt_self_iff neg_lt_self_iff - -@[simp] -theorem le_neg_self_iff : a ≤ -a ↔ a ≤ 0 := - calc - a ≤ -a ↔ - -a ≤ -a := by rw [neg_neg] - _ ↔ 0 ≤ -a := neg_le_self_iff - _ ↔ a ≤ 0 := neg_nonneg -#align le_neg_self_iff le_neg_self_iff - -@[simp] -theorem lt_neg_self_iff : a < -a ↔ a < 0 := - calc - a < -a ↔ - -a < -a := by rw [neg_neg] - _ ↔ 0 < -a := neg_lt_self_iff - _ ↔ a < 0 := neg_pos -#align lt_neg_self_iff lt_neg_self_iff - theorem neg_one_lt_zero : -1 < (0 : α) := neg_lt_zero.2 zero_lt_one #align neg_one_lt_zero neg_one_lt_zero diff --git a/Mathlib/Algebra/Order/Ring/WithTop.lean b/Mathlib/Algebra/Order/Ring/WithTop.lean index 4ff04cd98451a..06d498ee8fb2e 100644 --- a/Mathlib/Algebra/Order/Ring/WithTop.lean +++ b/Mathlib/Algebra/Order/Ring/WithTop.lean @@ -3,11 +3,11 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro -/ -import Mathlib.Algebra.Hom.Ring.Defs import Mathlib.Algebra.Order.Monoid.WithTop import Mathlib.Algebra.Order.Ring.Canonical -import Std.Data.Option.Lemmas +import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Tactic.Tauto +import Std.Data.Option.Lemmas #align_import algebra.order.ring.with_top from "leanprover-community/mathlib"@"0111834459f5d7400215223ea95ae38a1265a907" @@ -67,7 +67,7 @@ theorem mul_eq_top_iff {a b : WithTop α} : a * b = ⊤ ↔ a ≠ 0 ∧ b = ⊤ theorem mul_lt_top' [LT α] {a b : WithTop α} (ha : a < ⊤) (hb : b < ⊤) : a * b < ⊤ := by rw [WithTop.lt_top_iff_ne_top] at * - simp only [Ne.def, mul_eq_top_iff, *, and_false, false_and, false_or] + simp only [Ne.def, mul_eq_top_iff, *, and_false, false_and, or_self, not_false_eq_true] #align with_top.mul_lt_top' WithTop.mul_lt_top' theorem mul_lt_top [LT α] {a b : WithTop α} (ha : a ≠ ⊤) (hb : b ≠ ⊤) : a * b < ⊤ := diff --git a/Mathlib/Algebra/Order/SMul.lean b/Mathlib/Algebra/Order/SMul.lean index 933dfff133ca6..0fb038c5d4f71 100644 --- a/Mathlib/Algebra/Order/SMul.lean +++ b/Mathlib/Algebra/Order/SMul.lean @@ -53,7 +53,7 @@ class OrderedSMul (R M : Type*) [OrderedSemiring R] [OrderedAddCommMonoid M] [SM protected lt_of_smul_lt_smul_of_pos : ∀ {a b : M}, ∀ {c : R}, c • a < c • b → 0 < c → a < b #align ordered_smul OrderedSMul -variable {ι 𝕜 R M N : Type*} +variable {ι α β γ 𝕜 R M N : Type*} namespace OrderDual @@ -63,11 +63,36 @@ instance OrderDual.instSMulWithZero [Zero R] [AddZeroClass M] [SMulWithZero R M] zero_smul := fun m => OrderDual.rec (zero_smul _) m smul_zero := fun r => OrderDual.rec (@smul_zero R M _ _) r } +@[to_additive] instance OrderDual.instMulAction [Monoid R] [MulAction R M] : MulAction R Mᵒᵈ := { OrderDual.instSMul with one_smul := fun m => OrderDual.rec (one_smul _) m mul_smul := fun r => OrderDual.rec (@mul_smul R M _ _) r } +@[to_additive] +instance OrderDual.instSMulCommClass [SMul β γ] [SMul α γ] [SMulCommClass α β γ] : + SMulCommClass αᵒᵈ β γ := ‹SMulCommClass α β γ› + +@[to_additive] +instance OrderDual.instSMulCommClass' [SMul β γ] [SMul α γ] [SMulCommClass α β γ] : + SMulCommClass α βᵒᵈ γ := ‹SMulCommClass α β γ› + +@[to_additive] +instance OrderDual.instSMulCommClass'' [SMul β γ] [SMul α γ] [SMulCommClass α β γ] : + SMulCommClass α β γᵒᵈ := ‹SMulCommClass α β γ› + +@[to_additive OrderDual.instVAddAssocClass] +instance OrderDual.instIsScalarTower [SMul α β] [SMul β γ] [SMul α γ] [IsScalarTower α β γ] : + IsScalarTower αᵒᵈ β γ := ‹IsScalarTower α β γ› + +@[to_additive OrderDual.instVAddAssocClass'] +instance OrderDual.instIsScalarTower' [SMul α β] [SMul β γ] [SMul α γ] [IsScalarTower α β γ] : + IsScalarTower α βᵒᵈ γ := ‹IsScalarTower α β γ› + +@[to_additive OrderDual.instVAddAssocClass''] +instance OrderDual.IsScalarTower'' [SMul α β] [SMul β γ] [SMul α γ] [IsScalarTower α β γ] : + IsScalarTower α β γᵒᵈ := ‹IsScalarTower α β γ› + instance [MonoidWithZero R] [AddMonoid M] [MulActionWithZero R M] : MulActionWithZero R Mᵒᵈ := { OrderDual.instMulAction, OrderDual.instSMulWithZero with } @@ -99,6 +124,9 @@ variable [OrderedSemiring R] [OrderedAddCommMonoid M] [SMulWithZero R M] [Ordere · exact (smul_lt_smul_of_pos hab hc).le #align smul_le_smul_of_nonneg smul_le_smul_of_nonneg +-- TODO: Remove `smul_le_smul_of_nonneg` completely +alias smul_le_smul_of_nonneg_left := smul_le_smul_of_nonneg + theorem smul_nonneg (hc : 0 ≤ c) (ha : 0 ≤ a) : 0 ≤ c • a := calc (0 : M) = c • (0 : M) := (smul_zero c).symm diff --git a/Mathlib/Algebra/Order/Sub/Basic.lean b/Mathlib/Algebra/Order/Sub/Basic.lean index 855c9658bbd86..787da8d6a83b7 100644 --- a/Mathlib/Algebra/Order/Sub/Basic.lean +++ b/Mathlib/Algebra/Order/Sub/Basic.lean @@ -3,10 +3,10 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Mathlib.Order.Hom.Basic -import Mathlib.Algebra.Hom.Equiv.Basic +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Ring.Basic import Mathlib.Algebra.Order.Sub.Defs +import Mathlib.Order.Hom.Basic #align_import algebra.order.sub.basic from "leanprover-community/mathlib"@"10b4e499f43088dd3bb7b5796184ad5216648ab1" diff --git a/Mathlib/Algebra/Order/Sub/WithTop.lean b/Mathlib/Algebra/Order/Sub/WithTop.lean index 4baf09510cae5..dedf9ca024409 100644 --- a/Mathlib/Algebra/Order/Sub/WithTop.lean +++ b/Mathlib/Algebra/Order/Sub/WithTop.lean @@ -47,8 +47,8 @@ theorem sub_top {a : WithTop α} : a - ⊤ = 0 := by cases a <;> rfl @[simp] theorem sub_eq_top_iff {a b : WithTop α} : a - b = ⊤ ↔ a = ⊤ ∧ b ≠ ⊤ := by induction a using recTopCoe <;> induction b using recTopCoe <;> - simp only [← coe_sub, coe_ne_top, sub_top, zero_ne_top, coe_ne_top, top_sub_coe, false_and, - Ne.def] + simp only [← coe_sub, coe_ne_top, sub_top, zero_ne_top, top_sub_coe, false_and, + Ne.def, not_true_eq_false, not_false_eq_true, and_false, and_self] #align with_top.sub_eq_top_iff WithTop.sub_eq_top_iff theorem map_sub [Sub β] [Zero β] {f : α → β} (h : ∀ x y, f (x - y) = f x - f y) (h₀ : f 0 = 0) : diff --git a/Mathlib/Algebra/Order/WithZero.lean b/Mathlib/Algebra/Order/WithZero.lean index 518aba381de38..b10b17185653b 100644 --- a/Mathlib/Algebra/Order/WithZero.lean +++ b/Mathlib/Algebra/Order/WithZero.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Johan Commelin, Patrick Massot -/ -import Mathlib.Algebra.Hom.Equiv.Units.GroupWithZero import Mathlib.Algebra.GroupWithZero.InjSurj +import Mathlib.Algebra.GroupWithZero.Units.Equiv import Mathlib.Algebra.Order.Group.Units import Mathlib.Algebra.Order.Monoid.Basic import Mathlib.Algebra.Order.Monoid.WithZero.Defs diff --git a/Mathlib/Algebra/Parity.lean b/Mathlib/Algebra/Parity.lean index ffac5fd37016e..b12908269c6f4 100644 --- a/Mathlib/Algebra/Parity.lean +++ b/Mathlib/Algebra/Parity.lean @@ -58,13 +58,44 @@ theorem isSquare_mul_self (m : α) : IsSquare (m * m) := #align even_add_self even_add_self @[to_additive] -theorem isSquare_op_iff (a : α) : IsSquare (op a) ↔ IsSquare a := - ⟨fun ⟨c, hc⟩ => ⟨unop c, by rw [← unop_mul, ← hc, unop_op]⟩, fun ⟨c, hc⟩ => by simp [hc]⟩ +theorem isSquare_op_iff {a : α} : IsSquare (op a) ↔ IsSquare a := + ⟨fun ⟨c, hc⟩ => ⟨unop c, congr_arg unop hc⟩, fun ⟨c, hc⟩ => ⟨op c, congr_arg op hc⟩⟩ #align is_square_op_iff isSquare_op_iff #align even_op_iff even_op_iff +@[to_additive] +theorem isSquare_unop_iff {a : αᵐᵒᵖ} : IsSquare (unop a) ↔ IsSquare a := isSquare_op_iff.symm + +@[to_additive] +instance [DecidablePred (IsSquare : α → Prop)] : DecidablePred (IsSquare : αᵐᵒᵖ → Prop) := + fun _ => decidable_of_iff _ isSquare_unop_iff + +@[simp] +theorem even_ofMul_iff {a : α} : Even (Additive.ofMul a) ↔ IsSquare a := Iff.rfl + +@[simp] +theorem isSquare_toMul_iff {a : Additive α} : IsSquare (Additive.toMul a) ↔ Even a := Iff.rfl + +instance [DecidablePred (IsSquare : α → Prop)] : DecidablePred (Even : Additive α → Prop) := + fun _ => decidable_of_iff _ isSquare_toMul_iff + end Mul +section Add +variable [Add α] + +@[simp] +theorem isSquare_ofAdd_iff {a : α} : IsSquare (Multiplicative.ofAdd a) ↔ Even a := Iff.rfl + +@[simp] +theorem even_toAdd_iff {a : Multiplicative α} : + Even (Multiplicative.toAdd a) ↔ IsSquare a := Iff.rfl + +instance [DecidablePred (Even : α → Prop)] : DecidablePred (IsSquare : Multiplicative α → Prop) := + fun _ => decidable_of_iff _ even_toAdd_iff + +end Add + @[to_additive (attr := simp)] theorem isSquare_one [MulOneClass α] : IsSquare (1 : α) := ⟨1, (mul_one _).symm⟩ @@ -165,7 +196,7 @@ theorem isSquare_inv : IsSquare a⁻¹ ↔ IsSquare a := by refine' ⟨fun h => _, fun h => _⟩ · rw [← isSquare_op_iff, ← inv_inv a] exact h.map (MulEquiv.inv' α) - · exact ((isSquare_op_iff a).mpr h).map (MulEquiv.inv' α).symm + · exact (isSquare_op_iff.mpr h).map (MulEquiv.inv' α).symm #align is_square_inv isSquare_inv #align even_neg even_neg diff --git a/Mathlib/Algebra/Polynomial/GroupRingAction.lean b/Mathlib/Algebra/Polynomial/GroupRingAction.lean index c3cd0f5ed7723..4e48008dc6786 100644 --- a/Mathlib/Algebra/Polynomial/GroupRingAction.lean +++ b/Mathlib/Algebra/Polynomial/GroupRingAction.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import Mathlib.Algebra.GroupRingAction.Basic -import Mathlib.Algebra.Hom.GroupAction import Mathlib.Data.Polynomial.AlgebraMap import Mathlib.Data.Polynomial.Monic +import Mathlib.GroupTheory.GroupAction.Hom import Mathlib.GroupTheory.GroupAction.Quotient #align_import algebra.polynomial.group_ring_action from "leanprover-community/mathlib"@"afad8e438d03f9d89da2914aa06cb4964ba87a18" diff --git a/Mathlib/Algebra/Quandle.lean b/Mathlib/Algebra/Quandle.lean index dbb43d4c67b9a..4143d6ddcc789 100644 --- a/Mathlib/Algebra/Quandle.lean +++ b/Mathlib/Algebra/Quandle.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ -import Mathlib.Algebra.Hom.Equiv.Basic -import Mathlib.Algebra.Hom.Aut +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Algebra.Group.Aut import Mathlib.Data.ZMod.Defs import Mathlib.Tactic.Ring diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 0c4b786455629..747d7790b4cc6 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -1454,7 +1454,7 @@ section QuaternionAlgebra variable {R : Type*} (c₁ c₂ : R) private theorem pow_four [Infinite R] : #R ^ℕ 4 = #R := - power_nat_eq (aleph0_le_mk R) <| by simp + power_nat_eq (aleph0_le_mk R) <| by decide /-- The cardinality of a quaternion algebra, as a type. -/ theorem mk_quaternionAlgebra : #(ℍ[R,c₁,c₂]) = #R ^ℕ 4 := by diff --git a/Mathlib/Algebra/Regular/Pow.lean b/Mathlib/Algebra/Regular/Pow.lean index cadd463f93896..b6d13e79ae811 100644 --- a/Mathlib/Algebra/Regular/Pow.lean +++ b/Mathlib/Algebra/Regular/Pow.lean @@ -3,9 +3,9 @@ Copyright (c) 2021 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ -import Mathlib.Algebra.Hom.Iterate -import Mathlib.Algebra.Regular.Basic import Mathlib.Algebra.BigOperators.Basic +import Mathlib.Algebra.GroupPower.IterateHom +import Mathlib.Algebra.Regular.Basic #align_import algebra.regular.pow from "leanprover-community/mathlib"@"46a64b5b4268c594af770c44d9e502afc6a515cb" diff --git a/Mathlib/Algebra/Ring/Aut.lean b/Mathlib/Algebra/Ring/Aut.lean index 5a7eda1f74914..4594c00b79dbc 100644 --- a/Mathlib/Algebra/Ring/Aut.lean +++ b/Mathlib/Algebra/Ring/Aut.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ +import Mathlib.Algebra.Group.Aut import Mathlib.Algebra.GroupRingAction.Basic -import Mathlib.Algebra.Hom.Aut import Mathlib.Algebra.Ring.Equiv #align_import algebra.ring.aut from "leanprover-community/mathlib"@"207cfac9fcd06138865b5d04f7091e46d9320432" diff --git a/Mathlib/Algebra/Ring/Basic.lean b/Mathlib/Algebra/Ring/Basic.lean index c08c0a274b9d5..9e19057e9b74e 100644 --- a/Mathlib/Algebra/Ring/Basic.lean +++ b/Mathlib/Algebra/Ring/Basic.lean @@ -3,11 +3,11 @@ Copyright (c) 2014 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Neil Strickland -/ -import Mathlib.Algebra.Ring.Defs import Mathlib.Algebra.Group.Basic +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Algebra.GroupWithZero.NeZero -import Mathlib.Algebra.Hom.Group.Defs import Mathlib.Algebra.Opposites +import Mathlib.Algebra.Ring.Defs #align_import algebra.ring.basic from "leanprover-community/mathlib"@"2ed7e4aec72395b6a7c3ac4ac7873a7a43ead17c" diff --git a/Mathlib/Algebra/Hom/Centroid.lean b/Mathlib/Algebra/Ring/CentroidHom.lean similarity index 99% rename from Mathlib/Algebra/Hom/Centroid.lean rename to Mathlib/Algebra/Ring/CentroidHom.lean index f91a62f3b68f7..be0a9ad3965eb 100644 --- a/Mathlib/Algebra/Hom/Centroid.lean +++ b/Mathlib/Algebra/Ring/CentroidHom.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Christopher Hoskin -/ +import Mathlib.Algebra.Group.Hom.Instances import Mathlib.Algebra.GroupPower.Lemmas -import Mathlib.Algebra.Hom.GroupInstances #align_import algebra.hom.centroid from "leanprover-community/mathlib"@"6cb77a8eaff0ddd100e87b1591c6d3ad319514ff" @@ -509,7 +509,7 @@ def commRing (h : ∀ a b : α, (∀ r : α, a * r * b = 0) → a = 0 ∨ b = 0) { CentroidHom.instRing with mul_comm := fun f g ↦ by ext - refine' sub_eq_zero.1 ((or_self_iff _).1 <| (h _ _) fun r ↦ _) + refine' sub_eq_zero.1 (or_self_iff.1 <| (h _ _) fun r ↦ _) rw [mul_assoc, sub_mul, sub_eq_zero, ← map_mul_right, ← map_mul_right, coe_mul, coe_mul, comp_mul_comm] } #align centroid_hom.comm_ring CentroidHom.commRing diff --git a/Mathlib/Algebra/Ring/Divisibility/Basic.lean b/Mathlib/Algebra/Ring/Divisibility/Basic.lean index 8cf88a0003ece..761a8504eaccf 100644 --- a/Mathlib/Algebra/Ring/Divisibility/Basic.lean +++ b/Mathlib/Algebra/Ring/Divisibility/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Neil Strickland -/ import Mathlib.Algebra.Divisibility.Basic -import Mathlib.Algebra.Hom.Equiv.Basic +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Ring.Defs #align_import algebra.ring.divisibility from "leanprover-community/mathlib"@"e8638a0fcaf73e4500469f368ef9494e495099b3" diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index 64b4e8c10e2b5..bc348ff97752d 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -5,12 +5,12 @@ Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ import Mathlib.Init.CCLemmas import Mathlib.Algebra.Field.IsField +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.GroupWithZero.InjSurj -import Mathlib.Algebra.Hom.Ring.Defs +import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Logic.Equiv.Set import Mathlib.Util.AssertExists -import Mathlib.Algebra.Hom.Equiv.Basic #align_import algebra.ring.equiv from "leanprover-community/mathlib"@"00f91228655eecdcd3ac97a7fd8dbcb139fe990a" diff --git a/Mathlib/Algebra/Hom/Ring/Basic.lean b/Mathlib/Algebra/Ring/Hom/Basic.lean similarity index 96% rename from Mathlib/Algebra/Hom/Ring/Basic.lean rename to Mathlib/Algebra/Ring/Hom/Basic.lean index be3bfea5f6657..db3fb05f106f3 100644 --- a/Mathlib/Algebra/Hom/Ring/Basic.lean +++ b/Mathlib/Algebra/Ring/Hom/Basic.lean @@ -3,10 +3,10 @@ Copyright (c) 2019 Amelia Livingston. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Jireh Loreaux -/ -import Mathlib.Algebra.Hom.Ring.Defs -import Mathlib.Algebra.GroupWithZero.InjSurj import Mathlib.Algebra.Divisibility.Basic -import Mathlib.Algebra.Hom.Units +import Mathlib.Algebra.Group.Units.Hom +import Mathlib.Algebra.GroupWithZero.InjSurj +import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Data.Set.Image #align_import algebra.hom.ring from "leanprover-community/mathlib"@"cf9386b56953fb40904843af98b7a80757bbe7f9" diff --git a/Mathlib/Algebra/Hom/Ring/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean similarity index 100% rename from Mathlib/Algebra/Hom/Ring/Defs.lean rename to Mathlib/Algebra/Ring/Hom/Defs.lean diff --git a/Mathlib/Algebra/Ring/Opposite.lean b/Mathlib/Algebra/Ring/Opposite.lean index 6b49fc47cc334..935c235ce10ad 100644 --- a/Mathlib/Algebra/Ring/Opposite.lean +++ b/Mathlib/Algebra/Ring/Opposite.lean @@ -3,9 +3,9 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.GroupWithZero.Basic import Mathlib.Algebra.Group.Opposite -import Mathlib.Algebra.Hom.Ring.Defs +import Mathlib.Algebra.GroupWithZero.Basic +import Mathlib.Algebra.Ring.Hom.Defs #align_import algebra.ring.opposite from "leanprover-community/mathlib"@"76de8ae01554c3b37d66544866659ff174e66e1f" diff --git a/Mathlib/Algebra/Ring/Pi.lean b/Mathlib/Algebra/Ring/Pi.lean index 96f803d09b5b3..ee7108d8c2560 100644 --- a/Mathlib/Algebra/Ring/Pi.lean +++ b/Mathlib/Algebra/Ring/Pi.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot -/ import Mathlib.Algebra.Group.Pi -import Mathlib.Algebra.Hom.Ring.Defs +import Mathlib.Algebra.Ring.Hom.Defs #align_import algebra.ring.pi from "leanprover-community/mathlib"@"ba2245edf0c8bb155f1569fd9b9492a9b384cde6" diff --git a/Mathlib/Algebra/Star/StarAlgHom.lean b/Mathlib/Algebra/Star/StarAlgHom.lean index 5bf50043481e0..3732d611a92d6 100644 --- a/Mathlib/Algebra/Star/StarAlgHom.lean +++ b/Mathlib/Algebra/Star/StarAlgHom.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Algebra.Algebra.Equiv +import Mathlib.Algebra.Algebra.NonUnitalHom import Mathlib.Algebra.Algebra.Prod -import Mathlib.Algebra.Hom.NonUnitalAlg import Mathlib.Algebra.Star.Prod #align_import algebra.star.star_alg_hom from "leanprover-community/mathlib"@"35882ddc66524b6980532a123a4ad4166db34c81" diff --git a/Mathlib/Algebra/Star/Subalgebra.lean b/Mathlib/Algebra/Star/Subalgebra.lean index d304034e93675..f5f51ebc32ac1 100644 --- a/Mathlib/Algebra/Star/Subalgebra.lean +++ b/Mathlib/Algebra/Star/Subalgebra.lean @@ -612,8 +612,10 @@ instance adjoinCommRingOfIsStarNormal (R : Type u) {A : Type v} [CommRing R] [St variable {R} -- porting note: redundant binder annotation update -instance completeLattice : CompleteLattice (StarSubalgebra R A) := - GaloisInsertion.liftCompleteLattice StarSubalgebra.gi +instance completeLattice : CompleteLattice (StarSubalgebra R A) where + __ := GaloisInsertion.liftCompleteLattice StarSubalgebra.gi + bot := { toSubalgebra := ⊥, star_mem' := fun ⟨r, hr⟩ => ⟨star r, hr ▸ algebraMap_star_comm _⟩ } + bot_le S := (bot_le : ⊥ ≤ S.toSubalgebra) instance inhabited : Inhabited (StarSubalgebra R A) := ⟨⊤⟩ @@ -702,18 +704,14 @@ theorem iInf_toSubalgebra {ι : Sort*} (S : ι → StarSubalgebra R A) : SetLike.coe_injective <| by simp #align star_subalgebra.infi_to_subalgebra StarSubalgebra.iInf_toSubalgebra -theorem bot_toSubalgebra : (⊥ : StarSubalgebra R A).toSubalgebra = ⊥ := by - change Algebra.adjoin R (∅ ∪ star ∅) = Algebra.adjoin R ∅ - simp +theorem bot_toSubalgebra : (⊥ : StarSubalgebra R A).toSubalgebra = ⊥ := rfl #align star_subalgebra.bot_to_subalgebra StarSubalgebra.bot_toSubalgebra -theorem mem_bot {x : A} : x ∈ (⊥ : StarSubalgebra R A) ↔ x ∈ Set.range (algebraMap R A) := by - rw [← mem_toSubalgebra, bot_toSubalgebra, Algebra.mem_bot] +theorem mem_bot {x : A} : x ∈ (⊥ : StarSubalgebra R A) ↔ x ∈ Set.range (algebraMap R A) := Iff.rfl #align star_subalgebra.mem_bot StarSubalgebra.mem_bot @[simp] -theorem coe_bot : ((⊥ : StarSubalgebra R A) : Set A) = Set.range (algebraMap R A) := by - simp [Set.ext_iff, mem_bot] +theorem coe_bot : ((⊥ : StarSubalgebra R A) : Set A) = Set.range (algebraMap R A) := rfl #align star_subalgebra.coe_bot StarSubalgebra.coe_bot theorem eq_top_iff {S : StarSubalgebra R A} : S = ⊤ ↔ ∀ x : A, x ∈ S := diff --git a/Mathlib/Algebra/Support.lean b/Mathlib/Algebra/Support.lean index d9148f9ecc96a..a5fe356e778b7 100644 --- a/Mathlib/Algebra/Support.lean +++ b/Mathlib/Algebra/Support.lean @@ -10,6 +10,7 @@ import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Pi import Mathlib.Algebra.Module.Basic import Mathlib.GroupTheory.GroupAction.Pi +import Mathlib.Order.Cover #align_import algebra.support from "leanprover-community/mathlib"@"29cb56a7b35f72758b05a30490e1f10bd62c35c1" @@ -124,6 +125,12 @@ theorem range_subset_insert_image_mulSupport (f : α → M) : #align function.range_subset_insert_image_mul_support Function.range_subset_insert_image_mulSupport #align function.range_subset_insert_image_support Function.range_subset_insert_image_support +@[to_additive] +lemma range_eq_image_or_of_mulSupport_subset {f : α → M} {k : Set α} (h : mulSupport f ⊆ k) : + range f = f '' k ∨ range f = insert 1 (f '' k) := by + apply (wcovby_insert _ _).eq_or_eq (image_subset_range _ _) + exact (range_subset_insert_image_mulSupport f).trans (insert_subset_insert (image_subset f h)) + @[to_additive (attr := simp)] theorem mulSupport_one' : mulSupport (1 : α → M) = ∅ := mulSupport_eq_empty_iff.2 rfl diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index 37d8c01123cde..ebaafc6d57233 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -824,7 +824,7 @@ theorem liftAux_comp_inrHom (f : M →ₗ[R'] A) (hf : ∀ x y, f x * f y = 0) : LinearMap.ext <| liftAux_apply_inr f hf #align triv_sq_zero_ext.lift_aux_comp_inr_hom TrivSqZeroExt.liftAux_comp_inrHom --- When applied to `inr` itself, `lift_aux` is the identity. +/-- When applied to `inr` itself, `liftAux` is the identity. -/ @[simp] theorem liftAux_inrHom : liftAux (inrHom R' M) (inr_mul_inr R') = AlgHom.id R' (tsze R' M) := algHom_ext' <| liftAux_comp_inrHom _ _ diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean index 90e2e1ece6371..552e73d07efe7 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean @@ -638,8 +638,8 @@ lemma irreducible_polynomial [IsDomain R] : Irreducible W.polynomial := by apply (h1.symm.le.trans Cubic.degree_of_b_eq_zero').not_lt rcases Nat.WithBot.add_eq_three_iff.mp h0.symm with h | h | h | h -- porting note: replaced two `any_goals` proofs with two `iterate 2` proofs - iterate 2 rw [degree_add_eq_right_of_degree_lt] <;> simp only [h] - iterate 2 rw [degree_add_eq_left_of_degree_lt] <;> simp only [h] + iterate 2 rw [degree_add_eq_right_of_degree_lt] <;> simp only [h] <;> decide + iterate 2 rw [degree_add_eq_left_of_degree_lt] <;> simp only [h] <;> decide #align weierstrass_curve.irreducible_polynomial WeierstrassCurve.irreducible_polynomial -- porting note: removed `@[simp]` to avoid a `simpNF` linter error @@ -1029,7 +1029,7 @@ lemma norm_smul_basis (p q : R[X]) : simp_rw [Algebra.norm_eq_matrix_det <| CoordinateRing.basis W, Matrix.det_fin_two, Algebra.leftMulMatrix_eq_repr_mul, basis_zero, mul_one, basis_one, smul_basis_mul_Y, map_add, Finsupp.add_apply, map_smul, Finsupp.smul_apply, ← basis_zero, ← basis_one, - Basis.repr_self_apply, if_pos, if_neg, smul_eq_mul] + Basis.repr_self_apply, if_pos, one_ne_zero, if_false, smul_eq_mul] ring1 #align weierstrass_curve.coordinate_ring.norm_smul_basis WeierstrassCurve.CoordinateRing.norm_smul_basis diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 41b4c8fd4738c..97298a1fae244 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -199,7 +199,6 @@ theorem toΓSpecSheafedSpace_app_eq : dsimp at this rw [←this] dsimp - congr #align algebraic_geometry.LocallyRingedSpace.to_Γ_Spec_SheafedSpace_app_eq AlgebraicGeometry.LocallyRingedSpace.toΓSpecSheafedSpace_app_eq diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 52498ae6bf29b..2b5075ed6adc2 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -298,7 +298,7 @@ theorem zeroLocus_empty_iff_eq_top {I : Ideal R} : zeroLocus (I : Set R) = ∅ · contrapose! intro h rcases Ideal.exists_le_maximal I h with ⟨M, hM, hIM⟩ - exact Set.Nonempty.ne_empty ⟨⟨M, hM.isPrime⟩, hIM⟩ + exact ⟨⟨M, hM.isPrime⟩, hIM⟩ · rintro rfl apply zeroLocus_empty_of_one_mem trivial diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index 0815c5087d6bc..ca601bd40a510 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -37,14 +37,11 @@ variable (X : Scheme) /-- `f ⁻¹ᵁ U` is notation for `(Opens.map f.1.base).obj U`, the preimage of an open set `U` under `f`. -/ -notation3:90 f:91 "⁻¹ᵁ " U:90 => Prefunctor.obj - (Functor.toPrefunctor <| Opens.map (PresheafedSpace.Hom.base (LocallyRingedSpace.Hom.val f))) U +notation3:90 f:91 "⁻¹ᵁ " U:90 => (Opens.map (f : LocallyRingedSpace.Hom _ _).val.base).obj U /-- `X ∣_ᵤ U` is notation for `X.restrict U.openEmbedding`, the restriction of `X` to an open set `U` of `X`. -/ -notation3:60 X:60 " ∣_ᵤ " U:61 => (Scheme.restrict X (Opens.openEmbedding U)) - -attribute [nolint docBlame] «term_⁻¹ᵁ_».delab «term_∣_ᵤ_».delab +notation3:60 X:60 " ∣_ᵤ " U:61 => Scheme.restrict X (U : Opens X).openEmbedding /-- The restriction of a scheme to an open subset. -/ abbrev Scheme.ιOpens {X : Scheme} (U : Opens X.carrier) : X ∣_ᵤ U ⟶ X := X.ofRestrict _ @@ -182,7 +179,7 @@ lemma Scheme.restrictRestrict_inv_restrict_restrict (X : Scheme) (U : Opens X.ca noncomputable def Scheme.restrictIsoOfEq (X : Scheme) {U V : Opens X.carrier} (e : U = V) : X ∣_ᵤ U ≅ X ∣_ᵤ V := by - refine IsOpenImmersion.isoOfRangeEq (ιOpens U) (ιOpens V) (by rw [e]) + exact IsOpenImmersion.isoOfRangeEq (ιOpens U) (ιOpens V) (by rw [e]) end @@ -429,7 +426,7 @@ def morphismRestrictStalkMap {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (x) : instance {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) [IsOpenImmersion f] : IsOpenImmersion (f ∣_ U) := by delta morphismRestrict - refine PresheafedSpace.IsOpenImmersion.comp _ _ + exact PresheafedSpace.IsOpenImmersion.comp _ _ end MorphismRestrict diff --git a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean index 83f06a33a3ffc..eca5f1b303e24 100644 --- a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean +++ b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean @@ -112,7 +112,7 @@ theorem d_squared (n : ℕ) : objD X (n + 1) ≫ objD X n = 0 := by · rintro rfl simp only [Fin.val_zero, not_lt_zero'] at hij' · simpa only [Finset.mem_univ, forall_true_left, Prod.forall, ge_iff_le, Finset.mem_filter, - Fin.coe_castSucc, Fin.coe_pred, true_and] using Nat.le_pred_of_lt hij' + Fin.coe_castSucc, Fin.coe_pred, true_and] using Nat.le_sub_one_of_lt hij' · simp only [Fin.castLT_castSucc, Fin.succ_pred] #align algebraic_topology.alternating_face_map_complex.d_squared AlgebraicTopology.AlternatingFaceMapComplex.d_squared diff --git a/Mathlib/AlgebraicTopology/DoldKan/EquivalencePseudoabelian.lean b/Mathlib/AlgebraicTopology/DoldKan/EquivalencePseudoabelian.lean index 9361f1afb2e68..79b8d4df22d42 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/EquivalencePseudoabelian.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/EquivalencePseudoabelian.lean @@ -53,14 +53,14 @@ open AlgebraicTopology.DoldKan `N' : SimplicialObject C ⥤ Karoubi (ChainComplex C ℕ)` and the inverse of the equivalence `ChainComplex C ℕ ≌ Karoubi (ChainComplex C ℕ)`. -/ @[simps!, nolint unusedArguments] -def N [IsIdempotentComplete C] [HasFiniteCoproducts C] : SimplicialObject C ⥤ ChainComplex C ℕ := +def N [IsIdempotentComplete C] [HasFiniteCoproducts C] : SimplicialObject C ⥤ ChainComplex C ℕ := N₁ ⋙ (toKaroubiEquivalence _).inverse set_option linter.uppercaseLean3 false in #align category_theory.idempotents.dold_kan.N CategoryTheory.Idempotents.DoldKan.N /-- The functor `Γ` for the equivalence is `Γ'`. -/ @[simps!, nolint unusedArguments] -def Γ [IsIdempotentComplete C] [HasFiniteCoproducts C] : ChainComplex C ℕ ⥤ SimplicialObject C := +def Γ [IsIdempotentComplete C] [HasFiniteCoproducts C] : ChainComplex C ℕ ⥤ SimplicialObject C := Γ₀ #align category_theory.idempotents.dold_kan.Γ CategoryTheory.Idempotents.DoldKan.Γ diff --git a/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean b/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean index 25f486dc08c1f..9425f8d75b786 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean @@ -63,7 +63,7 @@ instance : ReflectsIsomorphisms (N₁ : SimplicialObject C ⥤ Karoubi (ChainCom b := fun i => inv (f.app (op [n])) ≫ X.σ i } simp only [MorphComponents.id, ← id_φ, ← preComp_φ, preComp, ← postComp_φ, postComp, PInfty_f_naturality_assoc, IsIso.hom_inv_id_assoc, assoc, IsIso.inv_hom_id_assoc, - SimplicialObject.σ_naturality, h₁, h₂, h₃]⟩ + SimplicialObject.σ_naturality, h₁, h₂, h₃, and_self]⟩ theorem compatibility_N₂_N₁_karoubi : N₂ ⋙ (karoubiChainComplexEquivalence C ℕ).functor = diff --git a/Mathlib/AlgebraicTopology/Nerve.lean b/Mathlib/AlgebraicTopology/Nerve.lean index 5104921194d1e..537cb6c7e41ab 100644 --- a/Mathlib/AlgebraicTopology/Nerve.lean +++ b/Mathlib/AlgebraicTopology/Nerve.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ import Mathlib.AlgebraicTopology.SimplicialSet +import Mathlib.CategoryTheory.ComposableArrows #align_import algebraic_topology.nerve from "leanprover-community/mathlib"@"841aef25c9d7a5a5d63a3dcf7bc43386b2c206d6" @@ -13,14 +14,15 @@ import Mathlib.AlgebraicTopology.SimplicialSet This file provides the definition of the nerve of a category `C`, which is a simplicial set `nerve C` (see [goerss-jardine-2009], Example I.1.4). +By definition, the type of `n`-simplices of `nerve C` is `ComposableArrows C n`, +which is the category `Fin (n + 1) ⥤ C`. ## References * [Paul G. Goerss, John F. Jardine, *Simplical Homotopy Theory*][goerss-jardine-2009] -/ - -open CategoryTheory.Category +open CategoryTheory.Category Simplicial universe v u @@ -29,18 +31,26 @@ namespace CategoryTheory /-- The nerve of a category -/ @[simps] def nerve (C : Type u) [Category.{v} C] : SSet.{max u v} where - obj Δ := SimplexCategory.toCat.obj Δ.unop ⥤ C - map f x := SimplexCategory.toCat.map f.unop ⋙ x + obj Δ := ComposableArrows C (Δ.unop.len) + map f x := x.whiskerLeft (SimplexCategory.toCat.map f.unop) #align category_theory.nerve CategoryTheory.nerve instance {C : Type*} [Category C] {Δ : SimplexCategoryᵒᵖ} : Category ((nerve C).obj Δ) := - (inferInstance : Category (SimplexCategory.toCat.obj Δ.unop ⥤ C)) + (inferInstance : Category (ComposableArrows C (Δ.unop.len))) /-- The nerve of a category, as a functor `Cat ⥤ SSet` -/ @[simps] def nerveFunctor : Cat ⥤ SSet where obj C := nerve C - map F := { app := fun Δ x => x ⋙ F } + map F := { app := fun Δ => (F.mapComposableArrows _).obj } #align category_theory.nerve_functor CategoryTheory.nerveFunctor +namespace Nerve + +variable {C : Type*} [Category C] {n : ℕ} + +lemma δ₀_eq {x : nerve C _[n + 1]} : (nerve C).δ (0 : Fin (n + 2)) x = x.δ₀ := rfl + +end Nerve + end CategoryTheory diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index ee274dd44afc2..6b15312710384 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -696,7 +696,7 @@ theorem eq_comp_δ_of_not_surjective' {n : ℕ} {Δ : SimplexCategory} (θ : Δ conv_rhs => dsimp erw [Fin.succ_pred] simpa only [Fin.le_iff_val_le_val, Fin.coe_castSucc, Fin.coe_pred] using - Nat.le_pred_of_lt (Fin.lt_iff_val_lt_val.mp h') + Nat.le_sub_one_of_lt (Fin.lt_iff_val_lt_val.mp h') · obtain rfl := le_antisymm (Fin.le_last i) (not_lt.mp h) use θ ≫ σ (Fin.last _) ext x : 4 diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index b685e0aa88f9a..d1ef22e6de2c0 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -1414,6 +1414,11 @@ theorem AnalyticAt.exists_mem_nhds_analyticOn {f : E → F} {x : E} (h : Analyti ∃ s ∈ 𝓝 x, AnalyticOn 𝕜 f s := h.eventually_analyticAt.exists_mem +/-- If we're analytic at a point, we're analytic in a nonempty ball -/ +theorem AnalyticAt.exists_ball_analyticOn {f : E → F} {x : E} (h : AnalyticAt 𝕜 f x) : + ∃ r : ℝ, 0 < r ∧ AnalyticOn 𝕜 f (Metric.ball x r) := + Metric.isOpen_iff.mp (isOpen_analyticAt _ _) _ h + end section diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index a73c82f4f47d1..0000e9c39a588 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -260,7 +260,7 @@ theorem comp_coeff_zero (q : FormalMultilinearSeries 𝕜 F G) (p : FormalMultil have : {c} = (Finset.univ : Finset (Composition 0)) := by apply Finset.eq_of_subset_of_card_le <;> simp [Finset.card_univ, composition_card 0] rw [← this, Finset.sum_singleton, compAlongComposition_apply] - symm; congr! -- porting note: needed the stronger version of `congr` here + symm; congr! -- porting note: needed the stronger `congr!`! #align formal_multilinear_series.comp_coeff_zero FormalMultilinearSeries.comp_coeff_zero @[simp] @@ -596,7 +596,6 @@ theorem compChangeOfVariables_length (m M N : ℕ) {i : Σ n, Fin n → ℕ} simp only [Composition.length, map_ofFn, length_ofFn] #align formal_multilinear_series.comp_change_of_variables_length FormalMultilinearSeries.compChangeOfVariables_length -set_option linter.deprecated false in theorem compChangeOfVariables_blocksFun (m M N : ℕ) {i : Σ n, Fin n → ℕ} (hi : i ∈ compPartialSumSource m M N) (j : Fin i.1) : (compChangeOfVariables m M N i hi).2.blocksFun @@ -604,7 +603,9 @@ theorem compChangeOfVariables_blocksFun (m M N : ℕ) {i : Σ n, Fin n → ℕ} i.2 j := by rcases i with ⟨n, f⟩ dsimp [Composition.blocksFun, Composition.blocks, compChangeOfVariables] - simp only [map_ofFn, List.nthLe_ofFn, Function.comp_apply] + simp only [map_ofFn, List.get_ofFn, Function.comp_apply] + -- Porting note: didn't used to need `rfl` + rfl #align formal_multilinear_series.comp_change_of_variables_blocks_fun FormalMultilinearSeries.compChangeOfVariables_blocksFun /-- Target set in the change of variables to compute the composition of partial sums of formal @@ -643,7 +644,6 @@ theorem mem_compPartialSumTarget_iff {m M N : ℕ} {a : Σ n, Composition n} : by simp [compPartialSumTarget, compPartialSumTargetSet] #align formal_multilinear_series.mem_comp_partial_sum_target_iff FormalMultilinearSeries.mem_compPartialSumTarget_iff -set_option linter.deprecated false in /-- `comp_change_of_variables m M N` is a bijection between `comp_partial_sum_source m M N` and `comp_partial_sum_target m M N`, yielding equal sums for functions that correspond to each other under the bijection. As `comp_change_of_variables m M N` is a dependent function, stating @@ -664,13 +664,7 @@ theorem compChangeOfVariables_sum {α : Type*} [AddCommMonoid α] (m M N : ℕ) simp only [mem_compPartialSumTarget_iff, Composition.length, Composition.blocks, H.left, map_ofFn, length_ofFn, true_and_iff, compChangeOfVariables] intro j - -- Porting note: the following `simp` was sufficient in lean 3. - simp only [Composition.blocksFun, (H.right _).right, List.nthLe_ofFn] - convert (H.right ⟨j, ?_⟩).right - · convert List.nthLe_ofFn _ _ using 2 - rfl - · apply j.prop.trans_eq - simp [Composition.length] + simp only [Composition.blocksFun, (H.right _).right, List.get_ofFn] -- 2 - show that the composition gives the `comp_along_composition` application · rintro ⟨k, blocks_fun⟩ H rw [h] @@ -1017,8 +1011,6 @@ theorem length_gather (a : Composition n) (b : Composition a.length) : rw [length_map, length_splitWrtComposition] #align composition.length_gather Composition.length_gather --- porting note: this needs `Composition.blocksFun` to be refactored in order to remove `nthLe` -set_option linter.deprecated false in /-- An auxiliary function used in the definition of `sigmaEquivSigmaPi` below, associating to two compositions `a` of `n` and `b` of `a.length`, and an index `i` bounded by the length of `a.gather b`, the subcomposition of `a` made of those blocks belonging to the `i`-th block of @@ -1033,31 +1025,26 @@ def sigmaCompositionAux (a : Composition n) (b : Composition a.length) (by rw [← a.blocks.join_splitWrtComposition b] exact mem_join_of_mem (List.get_mem _ _ _) hi) - blocks_sum := by simp only [Composition.blocksFun, nthLe_map', Composition.gather]; rfl + blocks_sum := by simp only [Composition.blocksFun, get_map, Composition.gather] #align composition.sigma_composition_aux Composition.sigmaCompositionAux --- porting note: this needs `Composition.blocksFun` to be refactored in order to remove `nthLe` -set_option linter.deprecated false in theorem length_sigmaCompositionAux (a : Composition n) (b : Composition a.length) (i : Fin b.length) : Composition.length (Composition.sigmaCompositionAux a b ⟨i, (length_gather a b).symm ▸ i.2⟩) = Composition.blocksFun b i := - show List.length (nthLe (splitWrtComposition a.blocks b) i _) = blocksFun b i by - rw [nthLe_map_rev List.length, nthLe_of_eq (map_length_splitWrtComposition _ _)]; rfl + show List.length ((splitWrtComposition a.blocks b).get ⟨i, _⟩) = blocksFun b i by + rw [get_map_rev List.length, get_of_eq (map_length_splitWrtComposition _ _)]; rfl #align composition.length_sigma_composition_aux Composition.length_sigmaCompositionAux --- porting note: this needs `Composition.blocksFun` to be refactored in order to remove `nthLe` -set_option linter.deprecated false in theorem blocksFun_sigmaCompositionAux (a : Composition n) (b : Composition a.length) (i : Fin b.length) (j : Fin (blocksFun b i)) : blocksFun (sigmaCompositionAux a b ⟨i, (length_gather a b).symm ▸ i.2⟩) ⟨j, (length_sigmaCompositionAux a b i).symm ▸ j.2⟩ = blocksFun a (embedding b i j) := - show nthLe (nthLe _ _ _) _ _ = nthLe a.blocks _ _ by - rw [nthLe_of_eq (nthLe_splitWrtComposition _ _ _), nthLe_drop', nthLe_take']; rfl + show get (get _ ⟨_, _⟩) ⟨_, _⟩ = a.blocks.get ⟨_, _⟩ by + rw [get_of_eq (get_splitWrtComposition _ _ _), get_drop', get_take']; rfl #align composition.blocks_fun_sigma_composition_aux Composition.blocksFun_sigmaCompositionAux -set_option linter.deprecated false in /-- Auxiliary lemma to prove that the composition of formal multilinear series is associative. Consider a composition `a` of `n` and a composition `b` of `a.length`. Grouping together some @@ -1088,7 +1075,7 @@ theorem sizeUpTo_sizeUpTo_add (a : Composition n) (b : Composition a.length) {i take (sum (take i b.blocks)) (take (sum (take (i + 1) b.blocks)) a.blocks) := by rw [take_take, min_eq_left] apply monotone_sum_take _ (Nat.le_succ _) - rw [this, nthLe_map', nthLe_splitWrtComposition, ← + rw [this, get_map, get_splitWrtComposition, ← take_append_drop (sum (take i b.blocks)) (take (sum (take (Nat.succ i) b.blocks)) a.blocks), sum_append] congr @@ -1104,8 +1091,7 @@ theorem sizeUpTo_sizeUpTo_add (a : Composition n) (b : Composition a.length) {i have : sizeUpTo b i + Nat.succ j = (sizeUpTo b i + j).succ := rfl rw [this, sizeUpTo_succ _ D, IHj A, sizeUpTo_succ _ B] simp only [sigmaCompositionAux, add_assoc, add_left_inj, Fin.val_mk] - simp_rw [← nthLe_eq] - rw [nthLe_of_eq (nthLe_splitWrtComposition _ _ _), nthLe_drop', nthLe_take _ _ C] + rw [get_of_eq (get_splitWrtComposition _ _ _), get_drop', get_take _ _ C] #align composition.size_up_to_size_up_to_add Composition.sizeUpTo_sizeUpTo_add /-- Natural equivalence between `(Σ (a : composition n), composition a.length)` and diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index acc3f2ad8fc3c..2e32092155058 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -153,27 +153,41 @@ lemma analyticAt_smul [NormedSpace 𝕝 E] [IsScalarTower 𝕜 𝕝 E] (z : 𝕝 lemma analyticAt_mul (z : A × A) : AnalyticAt 𝕜 (fun x : A × A ↦ x.1 * x.2) z := (ContinuousLinearMap.mul 𝕜 A).analyticAt_bilinear z -namespace AnalyticAt - /-- Scalar multiplication of one analytic function by another. -/ -lemma smul [NormedSpace 𝕝 F] [IsScalarTower 𝕜 𝕝 F] {f : E → 𝕝} {g : E → F} {z : E} +lemma AnalyticAt.smul [NormedSpace 𝕝 F] [IsScalarTower 𝕜 𝕝 F] {f : E → 𝕝} {g : E → F} {z : E} (hf : AnalyticAt 𝕜 f z) (hg : AnalyticAt 𝕜 g z) : - AnalyticAt 𝕜 (f • g) z := + AnalyticAt 𝕜 (fun x ↦ f x • g x) z := (analyticAt_smul _).comp₂ hf hg +/-- Scalar multiplication of one analytic function by another. -/ +lemma AnalyticOn.smul [NormedSpace 𝕝 F] [IsScalarTower 𝕜 𝕝 F] {f : E → 𝕝} {g : E → F} {s : Set E} + (hf : AnalyticOn 𝕜 f s) (hg : AnalyticOn 𝕜 g s) : + AnalyticOn 𝕜 (fun x ↦ f x • g x) s := + fun _ m ↦ (hf _ m).smul (hg _ m) + /-- Multiplication of analytic functions (valued in a normd `𝕜`-algebra) is analytic. -/ -lemma mul {f g : E → A} {z : E} (hf : AnalyticAt 𝕜 f z) (hg : AnalyticAt 𝕜 g z) : +lemma AnalyticAt.mul {f g : E → A} {z : E} (hf : AnalyticAt 𝕜 f z) (hg : AnalyticAt 𝕜 g z) : AnalyticAt 𝕜 (fun x ↦ f x * g x) z := (analyticAt_mul _).comp₂ hf hg +/-- Multiplication of analytic functions (valued in a normd `𝕜`-algebra) is analytic. -/ +lemma AnalyticOn.mul {f g : E → A} {s : Set E} (hf : AnalyticOn 𝕜 f s) (hg : AnalyticOn 𝕜 g s) : + AnalyticOn 𝕜 (fun x ↦ f x * g x) s := + fun _ m ↦ (hf _ m).mul (hg _ m) + /-- Powers of analytic functions (into a normed `𝕜`-algebra) are analytic. -/ -lemma pow {f : E → A} {z : E} (hf : AnalyticAt 𝕜 f z) (n : ℕ) : AnalyticAt 𝕜 (f ^ n) z := by +lemma AnalyticAt.pow {f : E → A} {z : E} (hf : AnalyticAt 𝕜 f z) (n : ℕ) : + AnalyticAt 𝕜 (fun x ↦ f x ^ n) z := by induction' n with m hm - · rw [pow_zero] - exact (analyticAt_const : AnalyticAt 𝕜 (fun _ ↦ (1 : A)) z) - · exact pow_succ f m ▸ hf.mul hm + · simp only [Nat.zero_eq, pow_zero] + apply analyticAt_const + · simp only [pow_succ] + exact hf.mul hm -end AnalyticAt +/-- Powers of analytic functions (into a normed `𝕜`-algebra) are analytic. -/ +lemma AnalyticOn.pow {f : E → A} {s : Set E} (hf : AnalyticOn 𝕜 f s) (n : ℕ) : + AnalyticOn 𝕜 (fun x ↦ f x ^ n) s := + fun _ m ↦ (hf _ m).pow n section Geometric diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index e64a7c0b3e6d0..31868233e51b2 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -563,7 +563,9 @@ theorem radius_rightInv_pos_of_radius_pos (p : FormalMultilinearSeries 𝕜 E F) -- conclude that all coefficients satisfy `aⁿ Qₙ ≤ (I + 1) a`. let a' : NNReal := ⟨a, apos.le⟩ suffices H : (a' : ENNReal) ≤ (p.rightInv i).radius - · apply lt_of_lt_of_le _ H; exact_mod_cast apos + · apply lt_of_lt_of_le _ H + -- Prior to leanprover/lean4#2734, this was `exact_mod_cast apos`. + simpa only [ENNReal.coe_pos] apply le_radius_of_bound _ ((I + 1) * a) fun n => ?_ by_cases hn : n = 0 · have : ‖p.rightInv i n‖ = ‖p.rightInv i 0‖ := by congr <;> try rw [hn] diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index e81b471249826..514f4c13030c0 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -861,9 +861,8 @@ namespace LocallyBoundedVariationOn `ae_differentiableWithinAt_of_mem`. -/ theorem ae_differentiableWithinAt_of_mem_real {f : ℝ → ℝ} {s : Set ℝ} (h : LocallyBoundedVariationOn f s) : ∀ᵐ x, x ∈ s → DifferentiableWithinAt ℝ f s x := by - obtain ⟨p, q, hp, hq, fpq⟩ : ∃ p q, MonotoneOn p s ∧ MonotoneOn q s ∧ f = p - q := + obtain ⟨p, q, hp, hq, rfl⟩ : ∃ p q, MonotoneOn p s ∧ MonotoneOn q s ∧ f = p - q := h.exists_monotoneOn_sub_monotoneOn - subst f -- porting note: TODO: `rfl` instead of `fpq` doesn't work filter_upwards [hp.ae_differentiableWithinAt_of_mem, hq.ae_differentiableWithinAt_of_mem] with x hxp hxq xs exact (hxp xs).sub (hxq xs) diff --git a/Mathlib/Analysis/BoxIntegral/Basic.lean b/Mathlib/Analysis/BoxIntegral/Basic.lean index 036a07ecfaa51..2fc0e3a4a2cbf 100644 --- a/Mathlib/Analysis/BoxIntegral/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Basic.lean @@ -819,8 +819,8 @@ theorem HasIntegral.of_le_Henstock_of_forall_isLittleO (hl : l ≤ Henstock) (B ∃ δ > 0, ∀ J ≤ I, Box.Icc J ⊆ Metric.closedBall x δ → x ∈ Box.Icc J → (l.bDistortion → J.distortion ≤ c) → dist (vol J (f x)) (g J) ≤ ε * B J) : HasIntegral I l f vol (g I) := - have A : l.bHenstock := hl.2.1.resolve_left (by decide) - HasIntegral.of_bRiemann_eq_false_of_forall_isLittleO (hl.1.resolve_right (by decide)) B hB0 _ s hs + have A : l.bHenstock := Bool.eq_true_of_true_le hl.2.1 + HasIntegral.of_bRiemann_eq_false_of_forall_isLittleO (Bool.eq_false_of_le_false hl.1) B hB0 _ s hs (fun _ => A) H₁ <| by simpa only [A, true_imp_iff] using H₂ set_option linter.uppercaseLean3 false in #align box_integral.has_integral_of_le_Henstock_of_forall_is_o BoxIntegral.HasIntegral.of_le_Henstock_of_forall_isLittleO diff --git a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean index d592dd748122e..da32ffc650d71 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean @@ -4,10 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Data.Set.Intervals.Monotone -import Mathlib.Tactic.GCongr -import Mathlib.Tactic.TFAE import Mathlib.Topology.Algebra.Order.MonotoneConvergence -import Mathlib.Topology.MetricSpace.Basic +import Mathlib.Topology.MetricSpace.Bounded #align_import analysis.box_integral.box.basic from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" /-! diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Split.lean b/Mathlib/Analysis/BoxIntegral/Partition/Split.lean index eb538b796d39d..e127b31b63fb5 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Split.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Split.lean @@ -89,7 +89,7 @@ theorem splitLower_def [DecidableEq ι] {i x} (h : x ∈ Ioo (I.lower i) (I.uppe (forall_update_iff I.upper fun j y => I.lower j < y).2 ⟨h.1, fun j _ => I.lower_lt_upper _⟩) : I.splitLower i x = (⟨I.lower, update I.upper i x, h'⟩ : Box ι) := by - simp only [splitLower, mk'_eq_coe, min_eq_left h.2.le, update] + simp only [splitLower, mk'_eq_coe, min_eq_left h.2.le, update, and_self] #align box_integral.box.split_lower_def BoxIntegral.Box.splitLower_def /-- Given a box `I` and `x ∈ (I.lower i, I.upper i)`, the hyperplane `{y : ι → ℝ | y i = x}` splits @@ -130,7 +130,7 @@ theorem splitUpper_def [DecidableEq ι] {i x} (h : x ∈ Ioo (I.lower i) (I.uppe (forall_update_iff I.lower fun j y => y < I.upper j).2 ⟨h.2, fun j _ => I.lower_lt_upper _⟩) : I.splitUpper i x = (⟨update I.lower i x, I.upper, h'⟩ : Box ι) := by - simp only [splitUpper, mk'_eq_coe, max_eq_left h.1.le, update] + simp only [splitUpper, mk'_eq_coe, max_eq_left h.1.le, update, and_self] #align box_integral.box.split_upper_def BoxIntegral.Box.splitUpper_def theorem disjoint_splitLower_splitUpper (I : Box ι) (i : ι) (x : ℝ) : diff --git a/Mathlib/Analysis/Calculus/Deriv/Support.lean b/Mathlib/Analysis/Calculus/Deriv/Support.lean index b381ea7c12f90..51a0c295463a5 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Support.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Support.lean @@ -43,9 +43,9 @@ theorem support_deriv_subset : support (deriv f) ⊆ tsupport f := by exact nmem_support.mpr (h2x.deriv_eq.trans (deriv_const x 0)) #align support_deriv_subset support_deriv_subset -theorem HasCompactSupport.deriv (hf : HasCompactSupport f) : HasCompactSupport (deriv f) := +protected theorem HasCompactSupport.deriv (hf : HasCompactSupport f) : + HasCompactSupport (deriv f) := hf.mono' support_deriv_subset #align has_compact_support.deriv HasCompactSupport.deriv end Support - diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index dac896f52f051..bdffbade52316 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -333,7 +333,7 @@ theorem D_subset_differentiable_set {K : Set (E →L[𝕜] F)} (hK : IsComplete rw [pow_lt_pow_iff_of_lt_one (by norm_num : (0 : ℝ) < 1 / 2) (by norm_num)] at this linarith set m := k - 1 - have m_ge : n e ≤ m := Nat.le_pred_of_lt k_gt + have m_ge : n e ≤ m := Nat.le_sub_one_of_lt k_gt have km : k = m + 1 := (Nat.succ_pred_eq_of_pos (lt_of_le_of_lt (zero_le _) k_gt)).symm rw [km] at hk h'k -- `f` is well approximated by `L e (n e) k` at the relevant scale @@ -701,7 +701,7 @@ theorem D_subset_differentiable_set {K : Set F} (hK : IsComplete K) : rw [pow_lt_pow_iff_of_lt_one (by norm_num : (0 : ℝ) < 1 / 2) (by norm_num)] at this linarith set m := k - 1 - have m_ge : n e ≤ m := Nat.le_pred_of_lt k_gt + have m_ge : n e ≤ m := Nat.le_sub_one_of_lt k_gt have km : k = m + 1 := (Nat.succ_pred_eq_of_pos (lt_of_le_of_lt (zero_le _) k_gt)).symm rw [km] at hk h'k -- `f` is well approximated by `L e (n e) k` at the relevant scale diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 973172c7b09ce..e0259f55c183b 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -158,7 +158,7 @@ theorem norm_rat (r : ℚ) : ‖(r : ℂ)‖ = |(r : ℝ)| := by @[simp 1100] theorem norm_nat (n : ℕ) : ‖(n : ℂ)‖ = n := - abs_of_nat _ + abs_natCast _ #align complex.norm_nat Complex.norm_nat @[simp 1100] diff --git a/Mathlib/Analysis/ConstantSpeed.lean b/Mathlib/Analysis/ConstantSpeed.lean index c939c34fae203..46d6c05471fc5 100644 --- a/Mathlib/Analysis/ConstantSpeed.lean +++ b/Mathlib/Analysis/ConstantSpeed.lean @@ -227,7 +227,7 @@ theorem unique_unit_speed_on_Icc_zero {s t : ℝ} (hs : 0 ≤ s) (ht : 0 ≤ t) rw [← φst] at hf convert unique_unit_speed φm hfφ hf ⟨le_rfl, hs⟩ using 1 have : φ 0 = 0 := by - have hm : 0 ∈ φ '' Icc 0 s := by simp only [mem_Icc, le_refl, ht, φst] + have hm : 0 ∈ φ '' Icc 0 s := by simp only [φst, ht, mem_Icc, le_refl, and_self] obtain ⟨x, xs, hx⟩ := hm apply le_antisymm ((φm ⟨le_rfl, hs⟩ xs xs.1).trans_eq hx) _ have := φst ▸ mapsTo_image φ (Icc 0 s) diff --git a/Mathlib/Analysis/Convex/Basic.lean b/Mathlib/Analysis/Convex/Basic.lean index a964161885568..124520374e378 100644 --- a/Mathlib/Analysis/Convex/Basic.lean +++ b/Mathlib/Analysis/Convex/Basic.lean @@ -552,6 +552,22 @@ end AddCommGroup end OrderedRing +section LinearOrderedRing + +variable [LinearOrderedRing 𝕜] [AddCommMonoid E] + +theorem Convex_subadditive_le [SMul 𝕜 E] {f : E → 𝕜} (hf1 : ∀ x y, f (x + y) ≤ (f x) + (f y)) + (hf2 : ∀ ⦃c⦄ x, 0 ≤ c → f (c • x) ≤ c * f x) (B : 𝕜) : + Convex 𝕜 { x | f x ≤ B } := by + rw [convex_iff_segment_subset] + rintro x hx y hy z ⟨a, b, ha, hb, hs, rfl⟩ + calc + _ ≤ a • (f x) + b • (f y) := le_trans (hf1 _ _) (add_le_add (hf2 x ha) (hf2 y hb)) + _ ≤ a • B + b • B := add_le_add (smul_le_smul_of_nonneg hx ha) (smul_le_smul_of_nonneg hy hb) + _ ≤ B := by rw [← add_smul, hs, one_smul] + +end LinearOrderedRing + section LinearOrderedField variable [LinearOrderedField 𝕜] diff --git a/Mathlib/Analysis/Convex/Between.lean b/Mathlib/Analysis/Convex/Between.lean index b6d4e16edbb92..1921785735593 100644 --- a/Mathlib/Analysis/Convex/Between.lean +++ b/Mathlib/Analysis/Convex/Between.lean @@ -152,6 +152,9 @@ def Sbtw (x y z : P) : Prop := variable {R} +lemma mem_segment_iff_wbtw {x y z : V} : y ∈ segment R x z ↔ Wbtw R x y z := by + rw [Wbtw, affineSegment_eq_segment] + theorem Wbtw.map {x y z : P} (h : Wbtw R x y z) (f : P →ᵃ[R] P') : Wbtw R (f x) (f y) (f z) := by rw [Wbtw, ← affineSegment_image] exact Set.mem_image_of_mem _ h @@ -608,7 +611,8 @@ theorem sbtw_of_sbtw_of_sbtw_of_mem_affineSpan_pair [NoZeroSMulDivisors R V] have hu : (Finset.univ : Finset (Fin 3)) = {i₁, i₂, i₃} := by clear h₁ h₂ h₁' h₂' -- Porting note: Originally `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> simp at h₁₂ h₁₃ h₂₃ ⊢ + fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ + <;> simp (config := {decide := true}) at h₁₂ h₁₃ h₂₃ ⊢ have hp : p ∈ affineSpan R (Set.range t.points) := by have hle : line[R, t.points i₁, p₁] ≤ affineSpan R (Set.range t.points) := by refine' affineSpan_pair_le_of_mem_of_mem (mem_affineSpan R (Set.mem_range_self _)) _ diff --git a/Mathlib/Analysis/Convex/Extrema.lean b/Mathlib/Analysis/Convex/Extrema.lean index f855a2eb2a715..7503228439ce3 100644 --- a/Mathlib/Analysis/Convex/Extrema.lean +++ b/Mathlib/Analysis/Convex/Extrema.lean @@ -5,8 +5,7 @@ Authors: Frédéric Dupuis -/ import Mathlib.Analysis.Convex.Function import Mathlib.Topology.Algebra.Affine -import Mathlib.Topology.LocalExtr -import Mathlib.Topology.MetricSpace.Basic +import Mathlib.Topology.MetricSpace.PseudoMetric #align_import analysis.convex.extrema from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" diff --git a/Mathlib/Analysis/Convex/Normed.lean b/Mathlib/Analysis/Convex/Normed.lean index 0cfe360a24511..a5b0c23a1f99b 100644 --- a/Mathlib/Analysis/Convex/Normed.lean +++ b/Mathlib/Analysis/Convex/Normed.lean @@ -3,9 +3,11 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alexander Bentkamp, Yury Kudryashov -/ +import Mathlib.Analysis.Convex.Between import Mathlib.Analysis.Convex.Jensen import Mathlib.Analysis.Convex.Topology import Mathlib.Analysis.Normed.Group.Pointwise +import Mathlib.Analysis.NormedSpace.AddTorsor import Mathlib.Analysis.NormedSpace.Ray #align_import analysis.convex.normed from "leanprover-community/mathlib"@"a63928c34ec358b5edcda2bf7513c50052a5230f" @@ -26,13 +28,14 @@ We prove the following facts: -/ -variable {ι : Type*} {E : Type*} +variable {ι : Type*} {E P : Type*} open Metric Set open Pointwise Convex -variable [SeminormedAddCommGroup E] [NormedSpace ℝ E] {s t : Set E} +variable [SeminormedAddCommGroup E] [NormedSpace ℝ E] [PseudoMetricSpace P] [NormedAddTorsor E P] +variable {s t : Set E} /-- The norm on a real normed space is convex on any convex set. See also `Seminorm.convexOn` and `convexOn_univ_norm`. -/ @@ -130,10 +133,14 @@ instance (priority := 100) NormedSpace.instLocPathConnectedSpace : LocPathConnec (convex_ball x r).isPathConnected <| by simp [r_pos] #align normed_space.loc_path_connected NormedSpace.instLocPathConnectedSpace -theorem dist_add_dist_of_mem_segment {x y z : E} (h : y ∈ [x -[ℝ] z]) : +theorem Wbtw.dist_add_dist {x y z : P} (h : Wbtw ℝ x y z) : dist x y + dist y z = dist x z := by - simp only [dist_eq_norm, mem_segment_iff_sameRay] at * - simpa only [sub_add_sub_cancel', norm_sub_rev] using h.norm_add.symm + obtain ⟨a, ⟨ha₀, ha₁⟩, rfl⟩ := h + simp [abs_of_nonneg, ha₀, ha₁, sub_mul] + +theorem dist_add_dist_of_mem_segment {x y z : E} (h : y ∈ [x -[ℝ] z]) : + dist x y + dist y z = dist x z := + (mem_segment_iff_wbtw.1 h).dist_add_dist #align dist_add_dist_of_mem_segment dist_add_dist_of_mem_segment /-- The set of vectors in the same ray as `x` is connected. -/ diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean index b1767ce282203..b76484f857883 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean @@ -105,7 +105,7 @@ theorem strictConvexOn_zpow {m : ℤ} (hm₀ : m ≠ 0) (hm₁ : m ≠ 1) : rw [iter_deriv_zpow] refine' mul_pos _ (zpow_pos_of_pos hx _) norm_cast - refine' int_prod_range_pos (by simp only) fun hm => _ + refine' int_prod_range_pos (by decide) fun hm => _ rw [← Finset.coe_Ico] at hm norm_cast at hm fin_cases hm <;> simp_all -- Porting note: `simp_all` was `cc` diff --git a/Mathlib/Analysis/Convex/StoneSeparation.lean b/Mathlib/Analysis/Convex/StoneSeparation.lean index c4417a4b90b59..0cd1081486832 100644 --- a/Mathlib/Analysis/Convex/StoneSeparation.lean +++ b/Mathlib/Analysis/Convex/StoneSeparation.lean @@ -107,6 +107,6 @@ theorem exists_convex_convex_compl_subset (hs : Convex 𝕜 s) (ht : Convex 𝕜 hCmax _ ⟨convex_convexHull _ _, h⟩ ((subset_insert _ _).trans <| subset_convexHull _ _)] at hc exact hc (subset_convexHull _ _ <| mem_insert _ _) rw [convexHull_insert ⟨z, hzC⟩, convexJoin_singleton_left] - refine' disjoint_iUnion₂_left.2 fun a ha => disjoint_iff_inf_le.mpr fun b hb => h a _ ⟨b, hb⟩ + refine disjoint_iUnion₂_left.2 fun a ha => disjoint_iff_inter_eq_empty.2 (h a ?_) rwa [← hC.1.convexHull_eq] #align exists_convex_convex_compl_subset exists_convex_convex_compl_subset diff --git a/Mathlib/Analysis/Convex/StrictConvexBetween.lean b/Mathlib/Analysis/Convex/StrictConvexBetween.lean index 761fe07019548..e4cf772256e97 100644 --- a/Mathlib/Analysis/Convex/StrictConvexBetween.lean +++ b/Mathlib/Analysis/Convex/StrictConvexBetween.lean @@ -16,10 +16,15 @@ space. -/ +open Metric +open scoped Convex -variable {V P : Type*} [NormedAddCommGroup V] [NormedSpace ℝ V] [PseudoMetricSpace P] +variable {V P : Type*} [NormedAddCommGroup V] [NormedSpace ℝ V] -variable [NormedAddTorsor V P] [StrictConvexSpace ℝ V] +variable [StrictConvexSpace ℝ V] + +section PseudoMetricSpace +variable [PseudoMetricSpace P] [NormedAddTorsor V P] theorem Sbtw.dist_lt_max_dist (p : P) {p₁ p₂ p₃ : P} (h : Sbtw ℝ p₁ p₂ p₃) : dist p₂ p < max (dist p₁ p) (dist p₃ p) := by @@ -78,3 +83,78 @@ theorem Collinear.sbtw_of_dist_eq_of_dist_lt {p p₁ p₂ p₃ : P} {r : ℝ} · rintro rfl exact hp₂.ne hp₃ #align collinear.sbtw_of_dist_eq_of_dist_lt Collinear.sbtw_of_dist_eq_of_dist_lt + +end PseudoMetricSpace + +section MetricSpace +variable [MetricSpace P] [NormedAddTorsor V P] {a b c : P} + +/-- In a strictly convex space, the triangle inequality turns into an equality if and only if the +middle point belongs to the segment joining two other points. -/ +lemma dist_add_dist_eq_iff : dist a b + dist b c = dist a c ↔ Wbtw ℝ a b c := by + have : + dist (a -ᵥ a) (b -ᵥ a) + dist (b -ᵥ a) (c -ᵥ a) = dist (a -ᵥ a) (c -ᵥ a) ↔ + b -ᵥ a ∈ segment ℝ (a -ᵥ a) (c -ᵥ a) := by + simp only [mem_segment_iff_sameRay, sameRay_iff_norm_add, dist_eq_norm', sub_add_sub_cancel', + eq_comm] + simp_rw [dist_vsub_cancel_right, ← affineSegment_eq_segment, ← affineSegment_vsub_const_image] + at this + rwa [(vsub_left_injective _).mem_set_image] at this +#align dist_add_dist_eq_iff dist_add_dist_eq_iff + +end MetricSpace + +variable {E F PE PF : Type*} [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedSpace ℝ E] + [NormedSpace ℝ F] [StrictConvexSpace ℝ E] [MetricSpace PE] [MetricSpace PF] [NormedAddTorsor E PE] + [NormedAddTorsor F PF] {r : ℝ} {f : PF → PE} {x y z : PE} + +lemma eq_lineMap_of_dist_eq_mul_of_dist_eq_mul (hxy : dist x y = r * dist x z) + (hyz : dist y z = (1 - r) * dist x z) : y = AffineMap.lineMap x z r := by + have : y -ᵥ x ∈ [(0 : E) -[ℝ] z -ᵥ x] := by + rw [mem_segment_iff_wbtw, ←dist_add_dist_eq_iff, dist_zero_left, dist_vsub_cancel_right, + ←dist_eq_norm_vsub', ←dist_eq_norm_vsub', hxy, hyz, ←add_mul, add_sub_cancel'_right, one_mul] + obtain rfl | hne := eq_or_ne x z + · obtain rfl : y = x := by simpa + simp + · rw [← dist_ne_zero] at hne + obtain ⟨a, b, _, hb, _, H⟩ := this + rw [smul_zero, zero_add] at H + have H' := congr_arg norm H + rw [norm_smul, Real.norm_of_nonneg hb, ← dist_eq_norm_vsub', ← dist_eq_norm_vsub', hxy, + mul_left_inj' hne] at H' + rw [AffineMap.lineMap_apply, ← H', H, vsub_vadd] +#align eq_line_map_of_dist_eq_mul_of_dist_eq_mul eq_lineMap_of_dist_eq_mul_of_dist_eq_mul + +lemma eq_midpoint_of_dist_eq_half (hx : dist x y = dist x z / 2) (hy : dist y z = dist x z / 2) : + y = midpoint ℝ x z := by + apply eq_lineMap_of_dist_eq_mul_of_dist_eq_mul + · rwa [invOf_eq_inv, ← div_eq_inv_mul] + · rwa [invOf_eq_inv, ← one_div, sub_half, one_div, ← div_eq_inv_mul] +#align eq_midpoint_of_dist_eq_half eq_midpoint_of_dist_eq_half + +namespace Isometry + +/-- An isometry of `NormedAddTorsor`s for real normed spaces, strictly convex in the case of the +codomain, is an affine isometry. Unlike Mazur-Ulam, this does not require the isometry to be +surjective. -/ +noncomputable def affineIsometryOfStrictConvexSpace (hi : Isometry f) : PF →ᵃⁱ[ℝ] PE := + { AffineMap.ofMapMidpoint f + (fun x y => by + apply eq_midpoint_of_dist_eq_half + · rw [hi.dist_eq, hi.dist_eq] + simp only [dist_left_midpoint, Real.norm_of_nonneg zero_le_two, div_eq_inv_mul] + · rw [hi.dist_eq, hi.dist_eq] + simp only [dist_midpoint_right, Real.norm_of_nonneg zero_le_two, div_eq_inv_mul]) + hi.continuous with + norm_map := fun x => by simp [AffineMap.ofMapMidpoint, ← dist_eq_norm_vsub E, hi.dist_eq] } +#align isometry.affine_isometry_of_strict_convex_space Isometry.affineIsometryOfStrictConvexSpace + +@[simp] lemma coe_affineIsometryOfStrictConvexSpace (hi : Isometry f) : + ⇑hi.affineIsometryOfStrictConvexSpace = f := rfl +#align isometry.coe_affine_isometry_of_strict_convex_space Isometry.coe_affineIsometryOfStrictConvexSpace + +@[simp] lemma affineIsometryOfStrictConvexSpace_apply (hi : Isometry f) (p : PF) : + hi.affineIsometryOfStrictConvexSpace p = f p := rfl +#align isometry.affine_isometry_of_strict_convex_space_apply Isometry.affineIsometryOfStrictConvexSpace_apply + +end Isometry diff --git a/Mathlib/Analysis/Convex/StrictConvexSpace.lean b/Mathlib/Analysis/Convex/StrictConvexSpace.lean index 2c85b502c6a46..3c9b6ab885409 100644 --- a/Mathlib/Analysis/Convex/StrictConvexSpace.lean +++ b/Mathlib/Analysis/Convex/StrictConvexSpace.lean @@ -232,77 +232,8 @@ theorem not_sameRay_iff_abs_lt_norm_sub : ¬SameRay ℝ x y ↔ |‖x‖ - ‖y sameRay_iff_norm_sub.not.trans <| ne_comm.trans (abs_norm_sub_norm_le _ _).lt_iff_ne.symm #align not_same_ray_iff_abs_lt_norm_sub not_sameRay_iff_abs_lt_norm_sub -/-- In a strictly convex space, the triangle inequality turns into an equality if and only if the -middle point belongs to the segment joining two other points. -/ -theorem dist_add_dist_eq_iff : dist x y + dist y z = dist x z ↔ y ∈ [x -[ℝ] z] := by - simp only [mem_segment_iff_sameRay, sameRay_iff_norm_add, dist_eq_norm', sub_add_sub_cancel', - eq_comm] -#align dist_add_dist_eq_iff dist_add_dist_eq_iff - theorem norm_midpoint_lt_iff (h : ‖x‖ = ‖y‖) : ‖(1 / 2 : ℝ) • (x + y)‖ < ‖x‖ ↔ x ≠ y := by rw [norm_smul, Real.norm_of_nonneg (one_div_nonneg.2 zero_le_two), ← inv_eq_one_div, ← div_eq_inv_mul, div_lt_iff (zero_lt_two' ℝ), mul_two, ← not_sameRay_iff_of_norm_eq h, not_sameRay_iff_norm_add_lt, h] #align norm_midpoint_lt_iff norm_midpoint_lt_iff - -variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] - -variable {PF : Type u} {PE : Type*} [MetricSpace PF] [MetricSpace PE] - -variable [NormedAddTorsor F PF] [NormedAddTorsor E PE] - -theorem eq_lineMap_of_dist_eq_mul_of_dist_eq_mul {x y z : PE} (hxy : dist x y = r * dist x z) - (hyz : dist y z = (1 - r) * dist x z) : y = AffineMap.lineMap x z r := by - have : y -ᵥ x ∈ [(0 : E) -[ℝ] z -ᵥ x] := by - rw [← dist_add_dist_eq_iff, dist_zero_left, dist_vsub_cancel_right, ← dist_eq_norm_vsub', ← - dist_eq_norm_vsub', hxy, hyz, ← add_mul, add_sub_cancel'_right, one_mul] - rcases eq_or_ne x z with (rfl | hne) - · obtain rfl : y = x := by simpa - simp - · rw [← dist_ne_zero] at hne - rcases this with ⟨a, b, _, hb, _, H⟩ - rw [smul_zero, zero_add] at H - have H' := congr_arg norm H - rw [norm_smul, Real.norm_of_nonneg hb, ← dist_eq_norm_vsub', ← dist_eq_norm_vsub', hxy, - mul_left_inj' hne] at H' - rw [AffineMap.lineMap_apply, ← H', H, vsub_vadd] -#align eq_line_map_of_dist_eq_mul_of_dist_eq_mul eq_lineMap_of_dist_eq_mul_of_dist_eq_mul - -theorem eq_midpoint_of_dist_eq_half {x y z : PE} (hx : dist x y = dist x z / 2) - (hy : dist y z = dist x z / 2) : y = midpoint ℝ x z := by - apply eq_lineMap_of_dist_eq_mul_of_dist_eq_mul - · rwa [invOf_eq_inv, ← div_eq_inv_mul] - · rwa [invOf_eq_inv, ← one_div, sub_half, one_div, ← div_eq_inv_mul] -#align eq_midpoint_of_dist_eq_half eq_midpoint_of_dist_eq_half - -namespace Isometry - -/-- An isometry of `NormedAddTorsor`s for real normed spaces, strictly convex in the case of -the codomain, is an affine isometry. Unlike Mazur-Ulam, this does not require the isometry to -be surjective. -/ -noncomputable def affineIsometryOfStrictConvexSpace {f : PF → PE} (hi : Isometry f) : - PF →ᵃⁱ[ℝ] PE := - { AffineMap.ofMapMidpoint f - (fun x y => by - apply eq_midpoint_of_dist_eq_half - · rw [hi.dist_eq, hi.dist_eq] - simp only [dist_left_midpoint, Real.norm_of_nonneg zero_le_two, div_eq_inv_mul] - · rw [hi.dist_eq, hi.dist_eq] - simp only [dist_midpoint_right, Real.norm_of_nonneg zero_le_two, div_eq_inv_mul]) - hi.continuous with - norm_map := fun x => by simp [AffineMap.ofMapMidpoint, ← dist_eq_norm_vsub E, hi.dist_eq] } -#align isometry.affine_isometry_of_strict_convex_space Isometry.affineIsometryOfStrictConvexSpace - -@[simp] -theorem coe_affineIsometryOfStrictConvexSpace {f : PF → PE} (hi : Isometry f) : - ⇑hi.affineIsometryOfStrictConvexSpace = f := - rfl -#align isometry.coe_affine_isometry_of_strict_convex_space Isometry.coe_affineIsometryOfStrictConvexSpace - -@[simp] -theorem affineIsometryOfStrictConvexSpace_apply {f : PF → PE} (hi : Isometry f) (p : PF) : - hi.affineIsometryOfStrictConvexSpace p = f p := - rfl -#align isometry.affine_isometry_of_strict_convex_space_apply Isometry.affineIsometryOfStrictConvexSpace_apply - -end Isometry diff --git a/Mathlib/Analysis/InnerProductSpace/GramSchmidtOrtho.lean b/Mathlib/Analysis/InnerProductSpace/GramSchmidtOrtho.lean index 9603ad9dc618f..35323888a0128 100644 --- a/Mathlib/Analysis/InnerProductSpace/GramSchmidtOrtho.lean +++ b/Mathlib/Analysis/InnerProductSpace/GramSchmidtOrtho.lean @@ -93,7 +93,7 @@ theorem gramSchmidt_orthogonal (f : ι → E) {a b : ι} (h₀ : a ≠ b) : clear h₀ a b intro a b h₀ revert a - apply WellFounded.induction (@IsWellFounded.wf ι (· < ·) _) b + apply wellFounded_lt.induction b intro b ih a h₀ simp only [gramSchmidt_def 𝕜 f b, inner_sub_right, inner_sum, orthogonalProjection_singleton, inner_smul_right] @@ -215,11 +215,7 @@ theorem gramSchmidt_ne_zero_coe {f : ι → E} (n : ι) have h₂ : (f ∘ ((↑) : Set.Iic n → ι)) ⟨n, le_refl n⟩ ∈ span 𝕜 (f ∘ ((↑) : Set.Iic n → ι) '' Set.Iio ⟨n, le_refl n⟩) := by rw [image_comp] - convert h₁ using 3 - ext i - apply Iff.intro <;> simp -- Porting note: was `simpa using @le_of_lt _ _ i n` - · intros; simp_all only - · intros q; use i; exact ⟨q, le_of_lt q, rfl⟩ + simpa using h₁ apply LinearIndependent.not_mem_span_image h₀ _ h₂ simp only [Set.mem_Iio, lt_self_iff_false, not_false_iff] #align gram_schmidt_ne_zero_coe gramSchmidt_ne_zero_coe diff --git a/Mathlib/Analysis/InnerProductSpace/Spectrum.lean b/Mathlib/Analysis/InnerProductSpace/Spectrum.lean index 75be24cea4c26..2e089c0552cec 100644 --- a/Mathlib/Analysis/InnerProductSpace/Spectrum.lean +++ b/Mathlib/Analysis/InnerProductSpace/Spectrum.lean @@ -129,7 +129,7 @@ theorem orthogonalComplement_iSup_eigenspaces_eq_bot : (⨆ μ, eigenspace T μ) -- a self-adjoint operator on a nontrivial inner product space has an eigenvalue haveI := hT'.subsingleton_of_no_eigenvalue_finiteDimensional hT.orthogonalComplement_iSup_eigenspaces - exact Submodule.eq_bot_of_subsingleton _ + exact Submodule.eq_bot_of_subsingleton #align linear_map.is_symmetric.orthogonal_supr_eigenspaces_eq_bot LinearMap.IsSymmetric.orthogonalComplement_iSup_eigenspaces_eq_bot theorem orthogonalComplement_iSup_eigenspaces_eq_bot' : diff --git a/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean b/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean index 4044deb4f31ca..9c2ca6088ebb9 100644 --- a/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean +++ b/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean @@ -234,7 +234,7 @@ protected theorem IsClosed.balancedCore (hU : IsClosed U) : IsClosed (balancedCo exact isClosedMap_smul_of_ne_zero ha' U hU · have : balancedCore 𝕜 U = ∅ := by contrapose! h - exact balancedCore_nonempty_iff.mp (Set.nonempty_iff_ne_empty.2 h) + exact balancedCore_nonempty_iff.mp h rw [this] exact isClosed_empty #align is_closed.balanced_core IsClosed.balancedCore diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index 2bb8327742321..fbb89e827f829 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -135,6 +135,19 @@ theorem geom_mean_le_arith_mean_weighted (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 · rw [exp_log hz] #align real.geom_mean_le_arith_mean_weighted Real.geom_mean_le_arith_mean_weighted +/-- AM-GM inequality: the **geometric mean is less than or equal to the arithmetic mean**. --/ +theorem geom_mean_le_arith_mean {ι : Type*} (s : Finset ι) (w : ι → ℝ) (z : ι → ℝ) + (hw : ∀ i ∈ s, 0 ≤ w i) (hw' : 0 < ∑ i in s, w i) (hz : ∀ i ∈ s, 0 ≤ z i) : + (∏ i in s, z i ^ w i) ^ (∑ i in s, w i)⁻¹ ≤ (∑ i in s, w i * z i) / (∑ i in s, w i) := by + convert geom_mean_le_arith_mean_weighted s (fun i => (w i) / ∑ i in s, w i) z ?_ ?_ hz using 2 + · rw [← finset_prod_rpow _ _ (fun i hi => rpow_nonneg_of_nonneg (hz _ hi) _) _] + refine Finset.prod_congr rfl (fun _ ih => ?_) + rw [div_eq_mul_inv, rpow_mul (hz _ ih)] + · simp_rw [div_eq_mul_inv, mul_assoc, mul_comm, ← mul_assoc, ← Finset.sum_mul, mul_comm] + · exact fun _ hi => div_nonneg (hw _ hi) (le_of_lt hw') + · simp_rw [div_eq_mul_inv, ← Finset.sum_mul] + exact mul_inv_cancel (by linarith) + theorem geom_mean_weighted_of_constant (w z : ι → ℝ) (x : ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) (hw' : ∑ i in s, w i = 1) (hz : ∀ i ∈ s, 0 ≤ z i) (hx : ∀ i ∈ s, w i ≠ 0 → z i = x) : ∏ i in s, z i ^ w i = x := @@ -609,6 +622,8 @@ theorem inner_le_Lp_mul_Lq_tsum_of_nonneg (hpq : p.IsConjugateExponent q) (hf : ∑' i, f i * g i ≤ (∑' i, f i ^ p) ^ (1 / p) * (∑' i, g i ^ q) ^ (1 / q) := by lift f to ι → ℝ≥0 using hf lift g to ι → ℝ≥0 using hg + -- After leanprover/lean4#2734, `norm_cast` needs help with beta reduction. + beta_reduce at * norm_cast at * exact NNReal.inner_le_Lp_mul_Lq_tsum hpq hf_sum hg_sum #align real.inner_le_Lp_mul_Lq_tsum_of_nonneg Real.inner_le_Lp_mul_Lq_tsum_of_nonneg @@ -637,6 +652,8 @@ theorem inner_le_Lp_mul_Lq_hasSum_of_nonneg (hpq : p.IsConjugateExponent q) {A B lift g to ι → ℝ≥0 using hg lift A to ℝ≥0 using hA lift B to ℝ≥0 using hB + -- After leanprover/lean4#2734, `norm_cast` needs help with beta reduction. + beta_reduce at * norm_cast at hf_sum hg_sum obtain ⟨C, hC, H⟩ := NNReal.inner_le_Lp_mul_Lq_hasSum hpq hf_sum hg_sum refine' ⟨C, C.prop, hC, _⟩ @@ -674,6 +691,8 @@ theorem Lp_add_le_tsum_of_nonneg (hp : 1 ≤ p) (hf : ∀ i, 0 ≤ f i) (hg : (∑' i, f i ^ p) ^ (1 / p) + (∑' i, g i ^ p) ^ (1 / p) := by lift f to ι → ℝ≥0 using hf lift g to ι → ℝ≥0 using hg + -- After leanprover/lean4#2734, `norm_cast` needs help with beta reduction. + beta_reduce at * norm_cast0 at * exact NNReal.Lp_add_le_tsum hp hf_sum hg_sum #align real.Lp_add_le_tsum_of_nonneg Real.Lp_add_le_tsum_of_nonneg @@ -702,9 +721,13 @@ theorem Lp_add_le_hasSum_of_nonneg (hp : 1 ≤ p) (hf : ∀ i, 0 ≤ f i) (hg : lift g to ι → ℝ≥0 using hg lift A to ℝ≥0 using hA lift B to ℝ≥0 using hB + -- After leanprover/lean4#2734, `norm_cast` needs help with beta reduction. + beta_reduce at hfA hgB norm_cast at hfA hgB obtain ⟨C, hC₁, hC₂⟩ := NNReal.Lp_add_le_hasSum hp hfA hgB use C + -- After leanprover/lean4#2734, `norm_cast` needs help with beta reduction. + beta_reduce norm_cast exact ⟨zero_le _, hC₁, hC₂⟩ #align real.Lp_add_le_has_sum_of_nonneg Real.Lp_add_le_hasSum_of_nonneg diff --git a/Mathlib/Analysis/Normed/Field/InfiniteSum.lean b/Mathlib/Analysis/Normed/Field/InfiniteSum.lean index 106a1cf5dc243..59b91c035881a 100644 --- a/Mathlib/Analysis/Normed/Field/InfiniteSum.lean +++ b/Mathlib/Analysis/Normed/Field/InfiniteSum.lean @@ -64,7 +64,7 @@ We prove two versions of the Cauchy product formula. The first one is In order to avoid `Nat` subtraction, we also provide `tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm`, where the `n`-th term is a sum over all pairs `(k, l)` such that `k+l=n`, which corresponds to the -`Finset` `Finset.Nat.antidiagonal n`. -/ +`Finset` `Finset.antidiagonal n`. -/ section Nat @@ -86,7 +86,7 @@ theorem summable_norm_sum_mul_antidiagonal_of_summable_norm {f g : ℕ → R} #align summable_norm_sum_mul_antidiagonal_of_summable_norm summable_norm_sum_mul_antidiagonal_of_summable_norm /-- The Cauchy product formula for the product of two infinite sums indexed by `ℕ`, - expressed by summing on `Finset.Nat.antidiagonal`. + expressed by summing on `Finset.antidiagonal`. See also `tsum_mul_tsum_eq_tsum_sum_antidiagonal` if `f` and `g` are *not* absolutely summable. -/ theorem tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm [CompleteSpace R] {f g : ℕ → R} diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index ad4a10e2f6b9f..cfe4030eff78d 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -5,7 +5,6 @@ Authors: Patrick Massot, Johannes Hölzl, Yaël Dillies -/ import Mathlib.Analysis.Normed.Group.Seminorm import Mathlib.Order.LiminfLimsup -import Mathlib.Topology.Algebra.UniformGroup import Mathlib.Topology.Instances.Rat import Mathlib.Topology.MetricSpace.Algebra import Mathlib.Topology.MetricSpace.IsometricSMul @@ -1104,7 +1103,7 @@ is phrased absolutely. -/ @[to_additive "Special case of the sandwich theorem: if the norm of `f` is eventually bounded by a real function `a` which tends to `0`, then `f` tends to `1`. In this pair of lemmas (`squeeze_zero_norm'` and `squeeze_zero_norm`), following a convention of similar lemmas in -`Topology.MetricSpace.Basic` and `Topology.Algebra.Order`, the `'` version is phrased using +`Topology.MetricSpace.PseudoMetric` and `Topology.Algebra.Order`, the `'` version is phrased using \"eventually\" and the non-`'` version is phrased absolutely."] theorem squeeze_one_norm' {f : α → E} {a : α → ℝ} {t₀ : Filter α} (h : ∀ᶠ n in t₀, ‖f n‖ ≤ a n) (h' : Tendsto a t₀ (𝓝 0)) : Tendsto f t₀ (𝓝 1) := diff --git a/Mathlib/Analysis/Normed/Group/Pointwise.lean b/Mathlib/Analysis/Normed/Group/Pointwise.lean index 3baef382f1d72..652c3b14b6159 100644 --- a/Mathlib/Analysis/Normed/Group/Pointwise.lean +++ b/Mathlib/Analysis/Normed/Group/Pointwise.lean @@ -5,8 +5,9 @@ Authors: Sébastien Gouëzel, Yaël Dillies -/ import Mathlib.Analysis.Normed.Group.Basic import Mathlib.Topology.MetricSpace.HausdorffDistance +import Mathlib.Topology.MetricSpace.IsometricSMul -#align_import analysis.normed.group.pointwise from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" +#align_import analysis.normed.group.pointwise from "leanprover-community/mathlib"@"c8f305514e0d47dfaa710f5a52f0d21b588e6328" /-! # Properties of pointwise addition of sets in normed groups @@ -24,6 +25,7 @@ section SeminormedGroup variable [SeminormedGroup E] {ε δ : ℝ} {s t : Set E} {x y : E} +-- note: we can't use `LipschitzOnWith.isBounded_image2` here without adding `[IsometricSMul E E]` @[to_additive] theorem Bornology.IsBounded.mul (hs : IsBounded s) (ht : IsBounded t) : IsBounded (s * t) := by obtain ⟨Rs, hRs⟩ : ∃ R, ∀ x ∈ s, ‖x‖ ≤ R := hs.exists_norm_le' @@ -34,6 +36,12 @@ theorem Bornology.IsBounded.mul (hs : IsBounded s) (ht : IsBounded t) : IsBounde #align metric.bounded.mul Bornology.IsBounded.mul #align metric.bounded.add Bornology.IsBounded.add +@[to_additive] +theorem Bornology.IsBounded.of_mul (hst : IsBounded (s * t)) : IsBounded s ∨ IsBounded t := + AntilipschitzWith.isBounded_of_image2_left _ (fun x => (isometry_mul_right x).antilipschitz) hst +#align metric.bounded.of_mul Bornology.IsBounded.of_mul +#align metric.bounded.of_add Bornology.IsBounded.of_add + @[to_additive] theorem Bornology.IsBounded.inv : IsBounded s → IsBounded s⁻¹ := by simp_rw [isBounded_iff_forall_norm_le', ← image_inv, ball_image_iff, norm_inv'] @@ -69,6 +77,15 @@ theorem infEdist_inv (x : E) (s : Set E) : infEdist x⁻¹ s = infEdist x s⁻¹ #align inf_edist_inv infEdist_inv #align inf_edist_neg infEdist_neg +@[to_additive] +theorem ediam_mul_le (x y : Set E) : EMetric.diam (x * y) ≤ EMetric.diam x + EMetric.diam y := + (LipschitzOnWith.ediam_image2_le (· * ·) _ _ + (fun _ _ => (isometry_mul_right _).lipschitz.lipschitzOnWith _) fun _ _ => + (isometry_mul_left _).lipschitz.lipschitzOnWith _).trans_eq <| + by simp only [ENNReal.coe_one, one_mul] +#align ediam_mul_le ediam_mul_le +#align ediam_add_le ediam_add_le + end EMetric variable (ε δ s t x y) diff --git a/Mathlib/Analysis/NormedSpace/Exponential.lean b/Mathlib/Analysis/NormedSpace/Exponential.lean index 54be6aea096dc..12cc16bc2f7bf 100644 --- a/Mathlib/Analysis/NormedSpace/Exponential.lean +++ b/Mathlib/Analysis/NormedSpace/Exponential.lean @@ -271,8 +271,8 @@ theorem exp_add_of_commute_of_mem_ball [CharZero 𝕂] {x y : 𝔸} (hxy : Commu ext rw [hxy.add_pow' _, Finset.smul_sum] refine' tsum_congr fun n => Finset.sum_congr rfl fun kl hkl => _ - rw [nsmul_eq_smul_cast 𝕂, smul_smul, smul_mul_smul, ← Finset.Nat.mem_antidiagonal.mp hkl, - Nat.cast_add_choose, Finset.Nat.mem_antidiagonal.mp hkl] + rw [nsmul_eq_smul_cast 𝕂, smul_smul, smul_mul_smul, ← Finset.mem_antidiagonal.mp hkl, + Nat.cast_add_choose, Finset.mem_antidiagonal.mp hkl] congr 1 have : (n ! : 𝕂) ≠ 0 := Nat.cast_ne_zero.mpr n.factorial_ne_zero field_simp [this] diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm.lean index c806e9a558524..244deb90a77c8 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm.lean @@ -607,7 +607,7 @@ theorem le_op_norm₂ [RingHomIsometric σ₁₃] (f : E →SL[σ₁₃] F →SL -- porting note: new theorem theorem le_of_op_norm₂_le_of_le [RingHomIsometric σ₁₃] (f : E →SL[σ₁₃] F →SL[σ₂₃] G) {x : E} {y : F} - {a b c : ℝ} (hf : ‖f‖ ≤ a) (hx : ‖x‖ ≤ b) (hy : ‖y‖ ≤ c) : + {a b c : ℝ} (hf : ‖f‖ ≤ a) (hx : ‖x‖ ≤ b) (hy : ‖y‖ ≤ c) : ‖f x y‖ ≤ a * b * c := (f x).le_of_op_norm_le_of_le (f.le_of_op_norm_le_of_le hf hx) hy @@ -1180,7 +1180,7 @@ class _root_.RegularNormedAlgebra : Prop := /-- Every (unital) normed algebra such that `‖1‖ = 1` is a `RegularNormedAlgebra`. -/ instance _root_.NormedAlgebra.instRegularNormedAlgebra {𝕜 𝕜' : Type*} [NontriviallyNormedField 𝕜] - [SeminormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] [NormOneClass 𝕜'] : RegularNormedAlgebra 𝕜 𝕜' where + [SeminormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] [NormOneClass 𝕜'] : RegularNormedAlgebra 𝕜 𝕜' where isometry_mul' := AddMonoidHomClass.isometry_of_norm (mul 𝕜 𝕜') <| fun x => le_antisymm (op_norm_mul_apply_le _ _ _) <| by convert ratio_le_op_norm ((mul 𝕜 𝕜') x) (1 : 𝕜') diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index ab393528aa5c8..90329943f5a89 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -382,7 +382,7 @@ def prodPseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : ← PseudoMetricSpace.edist_dist] exact le_sup_right · refine ENNReal.toReal_le_of_le_ofReal ?_ ?_ - · simp only [ge_iff_le, le_sup_iff, dist_nonneg] + · simp only [ge_iff_le, le_sup_iff, dist_nonneg, or_self] · simp [edist, PseudoMetricSpace.edist_dist, ENNReal.ofReal_le_ofReal] · have h1 : edist f.fst g.fst ^ p.toReal ≠ ⊤ := ENNReal.rpow_ne_top_of_nonneg (zero_le_one.trans h) (edist_ne_top _ _) diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean index 84b330a88367a..dc67b24bb5598 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean @@ -68,6 +68,14 @@ theorem abs_mul_cos_add_sin_mul_I (x : ℂ) : (abs x * (cos (arg x) + sin (arg x set_option linter.uppercaseLean3 false in #align complex.abs_mul_cos_add_sin_mul_I Complex.abs_mul_cos_add_sin_mul_I +@[simp] +lemma abs_mul_cos_arg (x : ℂ) : abs x * Real.cos (arg x) = x.re := by + simpa [-abs_mul_cos_add_sin_mul_I] using congr_arg re (abs_mul_cos_add_sin_mul_I x) + +@[simp] +lemma abs_mul_sin_arg (x : ℂ) : abs x * Real.sin (arg x) = x.im := by + simpa [-abs_mul_cos_add_sin_mul_I] using congr_arg im (abs_mul_cos_add_sin_mul_I x) + theorem abs_eq_one_iff (z : ℂ) : abs z = 1 ↔ ∃ θ : ℝ, exp (θ * I) = z := by refine' ⟨fun hz => ⟨arg z, _⟩, _⟩ · calc @@ -351,6 +359,13 @@ theorem neg_pi_div_two_le_arg_iff {z : ℂ} : -(π / 2) ≤ arg z ↔ 0 ≤ re z exacts [hre.ne, abs.pos <| ne_of_apply_ne re hre.ne] #align complex.neg_pi_div_two_le_arg_iff Complex.neg_pi_div_two_le_arg_iff +lemma neg_pi_div_two_lt_arg_iff {z : ℂ} : -(π / 2) < arg z ↔ 0 < re z ∨ 0 ≤ im z := by + rw [lt_iff_le_and_ne, neg_pi_div_two_le_arg_iff, ne_comm, Ne, arg_eq_neg_pi_div_two_iff] + rcases lt_trichotomy z.re 0 with hre | hre | hre + · simp [hre.ne, hre.not_le, hre.not_lt] + · simp [hre] + · simp [hre, hre.le, hre.ne'] + @[simp] theorem abs_arg_le_pi_div_two_iff {z : ℂ} : |arg z| ≤ π / 2 ↔ 0 ≤ re z := by rw [abs_le, arg_le_pi_div_two_iff, neg_pi_div_two_le_arg_iff, ← or_and_left, ← not_le, diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index a16ec64a81cf0..8a6043304a8d7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -543,15 +543,13 @@ open Lean.Meta Qq in /-- The `positivity` extension which identifies expressions of the form `Gamma a`. -/ @[positivity Gamma (_ : ℝ)] def _root_.Mathlib.Meta.Positivity.evalGamma : - Mathlib.Meta.Positivity.PositivityExt where eval {_ _α} zα pα e := do - let (.app _ (a : Q($_α))) ← withReducible (whnf e) | throwError "not Gamma ·" + Mathlib.Meta.Positivity.PositivityExt where eval {_ _α} zα pα (e : Q(ℝ)) := do + let ~q(Gamma $a) := e | throwError "failed to match on Gamma application" match ← Mathlib.Meta.Positivity.core zα pα a with - | .positive pa => - let pa' ← mkAppM ``Gamma_pos_of_pos #[pa] - pure (.positive pa') - | .nonnegative pa => - let pa' ← mkAppM ``Gamma_nonneg_of_nonneg #[pa] - pure (.nonnegative pa') + | .positive (pa : Q(0 < $a)) => + pure (.positive (q(Gamma_pos_of_pos $pa) : Q(0 < $e))) + | .nonnegative (pa : Q(0 ≤ $a)) => + pure (.nonnegative (q(Gamma_nonneg_of_nonneg $pa) : Q(0 ≤ $e))) | _ => pure .none /-- The Gamma function does not vanish on `ℝ` (except at non-positive integers, where the function diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index a062263bc9fba..17d3e4aeed916 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -71,13 +71,11 @@ def polarCoord : LocalHomeomorph (ℝ × ℝ) (ℝ × ℝ) where left_inv' := by rintro ⟨x, y⟩ _ have A : sqrt (x ^ 2 + y ^ 2) = Complex.abs (x + y * Complex.I) := by - simp [Complex.abs_def, Complex.normSq, pow_two, MonoidWithZeroHom.coe_mk, Complex.add_re, - Complex.ofReal_re, Complex.mul_re, Complex.I_re, mul_zero, Complex.ofReal_im, - Complex.I_im, sub_self, add_zero, Complex.add_im, Complex.mul_im, mul_one, zero_add] + rw [Complex.abs_apply, Complex.normSq_add_mul_I] have Z := Complex.abs_mul_cos_add_sin_mul_I (x + y * Complex.I) simp only [← Complex.ofReal_cos, ← Complex.ofReal_sin, mul_add, ← Complex.ofReal_mul, ← mul_assoc] at Z - simpa [A, -Complex.ofReal_cos, -Complex.ofReal_sin] using Complex.ext_iff.1 Z + simp [A] open_target := isOpen_Ioi.prod isOpen_Ioo open_source := (isOpen_lt continuous_const continuous_fst).union diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index 145f6aeaa6e2b..bac158beefdce 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -112,55 +112,99 @@ theorem cpow_sub {x : ℂ} (y z : ℂ) (hx : x ≠ 0) : x ^ (y - z) = x ^ y / x theorem cpow_neg_one (x : ℂ) : x ^ (-1 : ℂ) = x⁻¹ := by simpa using cpow_neg x 1 #align complex.cpow_neg_one Complex.cpow_neg_one +/-- See also `Complex.cpow_int_mul'`. -/ +lemma cpow_int_mul (x : ℂ) (n : ℤ) (y : ℂ) : x ^ (n * y) = (x ^ y) ^ n := by + rcases eq_or_ne x 0 with rfl | hx + · rcases eq_or_ne n 0 with rfl | hn + · simp + · rcases eq_or_ne y 0 with rfl | hy <;> simp [*, zero_zpow] + · rw [cpow_def_of_ne_zero hx, cpow_def_of_ne_zero hx, mul_left_comm, exp_int_mul] + +lemma cpow_mul_int (x y : ℂ) (n : ℤ) : x ^ (y * n) = (x ^ y) ^ n := by rw [mul_comm, cpow_int_mul] + +lemma cpow_nat_mul (x : ℂ) (n : ℕ) (y : ℂ) : x ^ (n * y) = (x ^ y) ^ n := by + exact_mod_cast cpow_int_mul x n y + +/-- See Note [no_index around OfNat.ofNat] -/ +lemma cpow_ofNat_mul (x : ℂ) (n : ℕ) [n.AtLeastTwo] (y : ℂ) : + x ^ (no_index (OfNat.ofNat n) * y) = (x ^ y) ^ (OfNat.ofNat n : ℕ) := + cpow_nat_mul x n y + +lemma cpow_mul_nat (x y : ℂ) (n : ℕ) : x ^ (y * n) = (x ^ y) ^ n := by + rw [mul_comm, cpow_nat_mul] + +/-- See Note [no_index around OfNat.ofNat] -/ +lemma cpow_mul_ofNat (x y : ℂ) (n : ℕ) [n.AtLeastTwo] : + x ^ (y * no_index (OfNat.ofNat n)) = (x ^ y) ^ (OfNat.ofNat n : ℕ) := + cpow_mul_nat x y n + @[simp, norm_cast] -theorem cpow_nat_cast (x : ℂ) : ∀ n : ℕ, x ^ (n : ℂ) = x ^ n - | 0 => by simp - | n + 1 => - if hx : x = 0 then by - simp only [hx, pow_succ, Complex.zero_cpow (Nat.cast_ne_zero.2 (Nat.succ_ne_zero _)), - zero_mul] - else by simp [cpow_add, hx, pow_add, cpow_nat_cast x n] +theorem cpow_nat_cast (x : ℂ) (n : ℕ) : x ^ (n : ℂ) = x ^ n := by simpa using cpow_nat_mul x n 1 #align complex.cpow_nat_cast Complex.cpow_nat_cast +/-- See Note [no_index around OfNat.ofNat] -/ @[simp] -theorem cpow_two (x : ℂ) : x ^ (2 : ℂ) = x ^ (2 : ℕ) := by - rw [← cpow_nat_cast] - simp only [Nat.cast_ofNat] +lemma cpow_ofNat (x : ℂ) (n : ℕ) [n.AtLeastTwo] : + x ^ (no_index (OfNat.ofNat n) : ℂ) = x ^ (OfNat.ofNat n : ℕ) := + cpow_nat_cast x n + +theorem cpow_two (x : ℂ) : x ^ (2 : ℂ) = x ^ (2 : ℕ) := cpow_ofNat x 2 #align complex.cpow_two Complex.cpow_two -open Int in @[simp, norm_cast] -theorem cpow_int_cast (x : ℂ) : ∀ n : ℤ, x ^ (n : ℂ) = x ^ n - | (n : ℕ) => by simp - | -[n+1] => by - rw [zpow_negSucc] - simp only [Int.negSucc_coe, Int.cast_neg, Complex.cpow_neg, inv_eq_one_div, Int.cast_ofNat, - cpow_nat_cast] +theorem cpow_int_cast (x : ℂ) (n : ℤ) : x ^ (n : ℂ) = x ^ n := by simpa using cpow_int_mul x n 1 #align complex.cpow_int_cast Complex.cpow_int_cast +@[simp] theorem cpow_nat_inv_pow (x : ℂ) {n : ℕ} (hn : n ≠ 0) : (x ^ (n⁻¹ : ℂ)) ^ n = x := by - suffices im (log x * (n⁻¹ : ℂ)) ∈ Ioc (-π) π by - rw [← cpow_nat_cast, ← cpow_mul _ this.1 this.2, inv_mul_cancel, cpow_one] - exact_mod_cast hn - rw [mul_comm, ← ofReal_nat_cast, ← ofReal_inv, ofReal_mul_im, ← div_eq_inv_mul] - rw [← pos_iff_ne_zero] at hn - have hn' : 0 < (n : ℝ) := by assumption_mod_cast - have hn1 : 1 ≤ (n : ℝ) := by exact_mod_cast Nat.succ_le_iff.2 hn - constructor - · rw [lt_div_iff hn'] - calc - -π * n ≤ -π * 1 := mul_le_mul_of_nonpos_left hn1 (neg_nonpos.2 Real.pi_pos.le) - _ = -π := (mul_one _) - _ < im (log x) := neg_pi_lt_log_im _ - - · rw [div_le_iff hn'] - calc - im (log x) ≤ π := log_im_le_pi _ - _ = π * 1 := (mul_one π).symm - _ ≤ π * n := mul_le_mul_of_nonneg_left hn1 Real.pi_pos.le - + rw [← cpow_nat_mul, mul_inv_cancel, cpow_one] + assumption_mod_cast #align complex.cpow_nat_inv_pow Complex.cpow_nat_inv_pow +/-- See Note [no_index around OfNat.ofNat] -/ +@[simp] +lemma cpow_ofNat_inv_pow (x : ℂ) (n : ℕ) [h : n.AtLeastTwo] : + (x ^ ((no_index (OfNat.ofNat n) : ℂ)⁻¹)) ^ (no_index (OfNat.ofNat n) : ℕ) = x := + cpow_nat_inv_pow _ (two_pos.trans_le h.1).ne' + +/-- A version of `Complex.cpow_int_mul` with RHS that matches `Complex.cpow_mul`. + +The assumptions on the arguments are needed +because the equality fails, e.g., for `x = -I`, `n = 2`, `y = 1/2`. -/ +lemma cpow_int_mul' {x : ℂ} {n : ℤ} (hlt : -π < n * x.arg) (hle : n * x.arg ≤ π) (y : ℂ) : + x ^ (n * y) = (x ^ n) ^ y := by + rw [mul_comm] at hlt hle + rw [cpow_mul, cpow_int_cast] <;> simpa [log_im] + +/-- A version of `Complex.cpow_nat_mul` with RHS that matches `Complex.cpow_mul`. + +The assumptions on the arguments are needed +because the equality fails, e.g., for `x = -I`, `n = 2`, `y = 1/2`. -/ +lemma cpow_nat_mul' {x : ℂ} {n : ℕ} (hlt : -π < n * x.arg) (hle : n * x.arg ≤ π) (y : ℂ) : + x ^ (n * y) = (x ^ n) ^ y := + cpow_int_mul' hlt hle y + +lemma cpow_ofNat_mul' {x : ℂ} {n : ℕ} [n.AtLeastTwo] (hlt : -π < OfNat.ofNat n * x.arg) + (hle : OfNat.ofNat n * x.arg ≤ π) (y : ℂ) : + x ^ (OfNat.ofNat n * y) = (x ^ (OfNat.ofNat n : ℕ)) ^ y := + cpow_nat_mul' hlt hle y + +lemma pow_cpow_nat_inv {x : ℂ} {n : ℕ} (h₀ : n ≠ 0) (hlt : -(π / n) < x.arg) (hle : x.arg ≤ π / n) : + (x ^ n) ^ (n⁻¹ : ℂ) = x := by + rw [← cpow_nat_mul', mul_inv_cancel (Nat.cast_ne_zero.2 h₀), cpow_one] + · rwa [← div_lt_iff' (Nat.cast_pos.2 h₀.bot_lt), neg_div] + · rwa [← le_div_iff' (Nat.cast_pos.2 h₀.bot_lt)] + +lemma pow_cpow_ofNat_inv {x : ℂ} {n : ℕ} [h : n.AtLeastTwo] (hlt : -(π / OfNat.ofNat n) < x.arg) + (hle : x.arg ≤ π / OfNat.ofNat n) : + (x ^ (OfNat.ofNat n : ℕ)) ^ ((OfNat.ofNat n : ℂ)⁻¹) = x := + pow_cpow_nat_inv (two_pos.trans_le h.1).ne' hlt hle + +/-- See also `Complex.pow_cpow_ofNat_inv` for a version that also works for `x * I`, `0 ≤ x`. -/ +lemma sq_cpow_two_inv {x : ℂ} (hx : 0 < x.re) : (x ^ (2 : ℕ)) ^ (2⁻¹ : ℂ) = x := + pow_cpow_ofNat_inv (neg_pi_div_two_lt_arg_iff.2 <| .inl hx) + (arg_le_pi_div_two_iff.2 <| .inl hx.le) + theorem mul_cpow_ofReal_nonneg {a b : ℝ} (ha : 0 ≤ a) (hb : 0 ≤ b) (r : ℂ) : ((a : ℂ) * (b : ℂ)) ^ r = (a : ℂ) ^ r * (b : ℂ) ^ r := by rcases eq_or_ne r 0 with (rfl | hr) diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index 501ec65e11821..7181855ce1ce4 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -587,6 +587,12 @@ theorem coe_mul_rpow (x y : ℝ≥0) (z : ℝ) : ((x : ℝ≥0∞) * y) ^ z = (x mul_rpow_of_ne_top coe_ne_top coe_ne_top z #align ennreal.coe_mul_rpow ENNReal.coe_mul_rpow +theorem prod_coe_rpow {ι} (s : Finset ι) (f : ι → ℝ≥0) (r : ℝ) : + ∏ i in s, (f i : ℝ≥0∞) ^ r = ((∏ i in s, f i : ℝ≥0) : ℝ≥0∞) ^ r := by + induction s using Finset.induction + case empty => simp + case insert i s hi ih => simp_rw [prod_insert hi, ih, ← coe_mul_rpow, coe_mul] + theorem mul_rpow_of_ne_zero {x y : ℝ≥0∞} (hx : x ≠ 0) (hy : y ≠ 0) (z : ℝ) : (x * y) ^ z = x ^ z * y ^ z := by simp [*, mul_rpow_eq_ite] #align ennreal.mul_rpow_of_ne_zero ENNReal.mul_rpow_of_ne_zero @@ -595,6 +601,21 @@ theorem mul_rpow_of_nonneg (x y : ℝ≥0∞) {z : ℝ} (hz : 0 ≤ z) : (x * y) simp [hz.not_lt, mul_rpow_eq_ite] #align ennreal.mul_rpow_of_nonneg ENNReal.mul_rpow_of_nonneg +theorem prod_rpow_of_ne_top {ι} {s : Finset ι} {f : ι → ℝ≥0∞} (hf : ∀ i ∈ s, f i ≠ ∞) (r : ℝ) : + ∏ i in s, f i ^ r = (∏ i in s, f i) ^ r := by + induction s using Finset.induction + case empty => simp + case insert i s hi ih => + have h2f : ∀ i ∈ s, f i ≠ ∞ := fun i hi ↦ hf i <| mem_insert_of_mem hi + rw [prod_insert hi, prod_insert hi, ih h2f, ← mul_rpow_of_ne_top <| hf i <| mem_insert_self ..] + apply prod_lt_top h2f |>.ne + +theorem prod_rpow_of_nonneg {ι} {s : Finset ι} {f : ι → ℝ≥0∞} {r : ℝ} (hr : 0 ≤ r) : + ∏ i in s, f i ^ r = (∏ i in s, f i) ^ r := by + induction s using Finset.induction + case empty => simp + case insert i s hi ih => simp_rw [prod_insert hi, ih, ← mul_rpow_of_nonneg _ _ hr] + theorem inv_rpow (x : ℝ≥0∞) (y : ℝ) : x⁻¹ ^ y = (x ^ y)⁻¹ := by rcases eq_or_ne y 0 with (rfl | hy); · simp only [rpow_zero, inv_one] replace hy := hy.lt_or_lt diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index f5d91bd870bef..0b5f1c153d95c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -255,6 +255,22 @@ theorem ofReal_cpow_of_nonpos {x : ℝ} (hx : x ≤ 0) (y : ℂ) : arg_ofReal_of_nonneg (neg_nonneg.2 hx), ofReal_zero, zero_mul, add_zero] #align complex.of_real_cpow_of_nonpos Complex.ofReal_cpow_of_nonpos +lemma cpow_ofReal (x : ℂ) (y : ℝ) : + x ^ (y : ℂ) = ↑(abs x ^ y) * (Real.cos (arg x * y) + Real.sin (arg x * y) * I) := by + rcases eq_or_ne x 0 with rfl | hx + · simp [ofReal_cpow le_rfl] + · rw [cpow_def_of_ne_zero hx, exp_eq_exp_re_mul_sin_add_cos, mul_comm (log x)] + norm_cast + rw [ofReal_mul_re, ofReal_mul_im, log_re, log_im, mul_comm y, mul_comm y, Real.exp_mul, + Real.exp_log] + rwa [abs.pos_iff] + +lemma cpow_ofReal_re (x : ℂ) (y : ℝ) : (x ^ (y : ℂ)).re = (abs x) ^ y * Real.cos (arg x * y) := by + rw [cpow_ofReal]; generalize arg x * y = z; simp [Real.cos] + +lemma cpow_ofReal_im (x : ℂ) (y : ℝ) : (x ^ y).im = (abs x) ^ y * Real.sin (arg x * y) := by + rw [cpow_ofReal]; generalize arg x * y = z; simp [Real.sin] + theorem abs_cpow_of_ne_zero {z : ℂ} (hz : z ≠ 0) (w : ℂ) : abs (z ^ w) = abs z ^ w.re / Real.exp (arg z * im w) := by rw [cpow_def_of_ne_zero hz, abs_exp, mul_re, log_re, log_im, Real.exp_sub, @@ -271,16 +287,15 @@ theorem abs_cpow_of_imp {z w : ℂ} (h : z = 0 → w.re = 0 → w = 0) : #align complex.abs_cpow_of_imp Complex.abs_cpow_of_imp theorem abs_cpow_le (z w : ℂ) : abs (z ^ w) ≤ abs z ^ w.re / Real.exp (arg z * im w) := by - rcases ne_or_eq z 0 with (hz | rfl) <;> [exact (abs_cpow_of_ne_zero hz w).le; rw [map_zero]] - rcases eq_or_ne w 0 with (rfl | hw); · simp - rw [zero_cpow hw, map_zero] - exact div_nonneg (Real.rpow_nonneg_of_nonneg le_rfl _) (Real.exp_pos _).le + by_cases h : z = 0 → w.re = 0 → w = 0 + · exact (abs_cpow_of_imp h).le + · push_neg at h + simp [h] #align complex.abs_cpow_le Complex.abs_cpow_le @[simp] theorem abs_cpow_real (x : ℂ) (y : ℝ) : abs (x ^ (y : ℂ)) = Complex.abs x ^ y := by - rcases eq_or_ne x 0 with (rfl | hx) <;> [rcases eq_or_ne y 0 with (rfl | hy); skip] <;> - simp [*, abs_cpow_of_ne_zero] + rw [abs_cpow_of_imp] <;> simp #align complex.abs_cpow_real Complex.abs_cpow_real @[simp] @@ -295,10 +310,7 @@ theorem abs_cpow_eq_rpow_re_of_pos {x : ℝ} (hx : 0 < x) (y : ℂ) : abs (x ^ y theorem abs_cpow_eq_rpow_re_of_nonneg {x : ℝ} (hx : 0 ≤ x) {y : ℂ} (hy : re y ≠ 0) : abs (x ^ y) = x ^ re y := by - rcases hx.eq_or_lt with (rfl | hlt) - · rw [ofReal_zero, zero_cpow, map_zero, Real.zero_rpow hy] - exact ne_of_apply_ne re hy - · exact abs_cpow_eq_rpow_re_of_pos hlt y + rw [abs_cpow_of_imp] <;> simp [*, arg_ofReal_of_nonneg, _root_.abs_of_nonneg] #align complex.abs_cpow_eq_rpow_re_of_nonneg Complex.abs_cpow_eq_rpow_re_of_nonneg end Complex @@ -768,6 +780,36 @@ theorem exists_rat_pow_btwn {α : Type*} [LinearOrderedField α] [Archimedean α end Real +namespace Complex + +lemma cpow_inv_two_re (x : ℂ) : (x ^ (2⁻¹ : ℂ)).re = sqrt ((abs x + x.re) / 2) := by + rw [← ofReal_ofNat, ← ofReal_inv, cpow_ofReal_re, ← div_eq_mul_inv, ← one_div, + ← Real.sqrt_eq_rpow, cos_half, ← sqrt_mul, ← mul_div_assoc, mul_add, mul_one, abs_mul_cos_arg] + exacts [abs.nonneg _, (neg_pi_lt_arg _).le, arg_le_pi _] + +lemma cpow_inv_two_im_eq_sqrt {x : ℂ} (hx : 0 ≤ x.im) : + (x ^ (2⁻¹ : ℂ)).im = sqrt ((abs x - x.re) / 2) := by + rw [← ofReal_ofNat, ← ofReal_inv, cpow_ofReal_im, ← div_eq_mul_inv, ← one_div, + ← Real.sqrt_eq_rpow, sin_half_eq_sqrt, ← sqrt_mul (abs.nonneg _), ← mul_div_assoc, mul_sub, + mul_one, abs_mul_cos_arg] + · rwa [arg_nonneg_iff] + · linarith [pi_pos, arg_le_pi x] + +lemma cpow_inv_two_im_eq_neg_sqrt {x : ℂ} (hx : x.im < 0) : + (x ^ (2⁻¹ : ℂ)).im = -sqrt ((abs x - x.re) / 2) := by + rw [← ofReal_ofNat, ← ofReal_inv, cpow_ofReal_im, ← div_eq_mul_inv, ← one_div, + ← Real.sqrt_eq_rpow, sin_half_eq_neg_sqrt, mul_neg, ← sqrt_mul (abs.nonneg _), + ← mul_div_assoc, mul_sub, mul_one, abs_mul_cos_arg] + · linarith [pi_pos, neg_pi_lt_arg x] + · exact (arg_neg_iff.2 hx).le + +lemma abs_cpow_inv_two_im (x : ℂ) : |(x ^ (2⁻¹ : ℂ)).im| = sqrt ((abs x - x.re) / 2) := by + rw [← ofReal_ofNat, ← ofReal_inv, cpow_ofReal_im, ← div_eq_mul_inv, ← one_div, + ← Real.sqrt_eq_rpow, _root_.abs_mul, _root_.abs_of_nonneg (sqrt_nonneg _), abs_sin_half, + ← sqrt_mul (abs.nonneg _), ← mul_div_assoc, mul_sub, mul_one, abs_mul_cos_arg] + +end Complex + section Tactics /-! diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index 462671ce53e0a..247c8a7cce4e7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -503,6 +503,23 @@ theorem cos_eq_sqrt_one_sub_sin_sq {x : ℝ} (hl : -(π / 2) ≤ x) (hu : x ≤ rw [← abs_cos_eq_sqrt_one_sub_sin_sq, abs_of_nonneg (cos_nonneg_of_mem_Icc ⟨hl, hu⟩)] #align real.cos_eq_sqrt_one_sub_sin_sq Real.cos_eq_sqrt_one_sub_sin_sq +lemma cos_half {x : ℝ} (hl : -π ≤ x) (hr : x ≤ π) : cos (x / 2) = sqrt ((1 + cos x) / 2) := by + have : 0 ≤ cos (x / 2) := cos_nonneg_of_mem_Icc <| by constructor <;> linarith + rw [← sqrt_sq this, cos_sq, add_div, two_mul, add_halves] + +lemma abs_sin_half (x : ℝ) : |sin (x / 2)| = sqrt ((1 - cos x) / 2) := by + rw [← sqrt_sq_eq_abs, sin_sq_eq_half_sub, two_mul, add_halves, sub_div] + +lemma sin_half_eq_sqrt {x : ℝ} (hl : 0 ≤ x) (hr : x ≤ 2 * π) : + sin (x / 2) = sqrt ((1 - cos x) / 2) := by + rw [← abs_sin_half, abs_of_nonneg] + apply sin_nonneg_of_nonneg_of_le_pi <;> linarith + +lemma sin_half_eq_neg_sqrt {x : ℝ} (hl : -(2 * π) ≤ x) (hr : x ≤ 0) : + sin (x / 2) = -sqrt ((1 - cos x) / 2) := by + rw [← abs_sin_half, abs_of_nonpos, neg_neg] + apply sin_nonpos_of_nonnpos_of_neg_pi_le <;> linarith + theorem sin_eq_zero_iff_of_lt_of_lt {x : ℝ} (hx₁ : -π < x) (hx₂ : x < π) : sin x = 0 ↔ x = 0 := ⟨fun h => by contrapose! h @@ -732,24 +749,12 @@ theorem sqrtTwoAddSeries_monotone_left {x y : ℝ} (h : x ≤ y) : theorem cos_pi_over_two_pow : ∀ n : ℕ, cos (π / 2 ^ (n + 1)) = sqrtTwoAddSeries 0 n / 2 | 0 => by simp | n + 1 => by - have : (2 : ℝ) ≠ 0 := two_ne_zero - rw [eq_div_iff_mul_eq this, eq_comm, sqrtTwoAddSeries, sqrt_eq_iff_sq_eq, mul_pow, cos_sq, - ← mul_div_assoc, pow_succ, mul_div_mul_left _ _ this, cos_pi_over_two_pow _, add_mul] - · congr - · norm_num - · rw [mul_comm, sq, mul_assoc, ← mul_div_assoc, mul_div_cancel_left _ this, ← mul_div_assoc, - mul_div_cancel_left _ this] - · exact add_nonneg two_pos.le (sqrtTwoAddSeries_zero_nonneg _) - refine le_of_lt <| mul_pos (cos_pos_of_mem_Ioo ⟨?_, ?_⟩) two_pos - · trans (0 : ℝ) - · rw [neg_lt_zero] - exact pi_div_two_pos - · exact div_pos pi_pos <| pow_pos two_pos _ - apply div_lt_div' (le_refl π) _ pi_pos two_pos - refine' lt_of_le_of_lt (le_of_eq (pow_one _).symm) _ - apply pow_lt_pow - · norm_num - · exact Nat.succ_lt_succ n.succ_pos + have A : (1 : ℝ) < 2 ^ (n + 1) := one_lt_pow one_lt_two n.succ_ne_zero + have B : π / 2 ^ (n + 1) < π := div_lt_self pi_pos A + have C : 0 < π / 2 ^ (n + 1) := by positivity + rw [pow_succ', div_mul_eq_div_div, cos_half, cos_pi_over_two_pow n, sqrtTwoAddSeries, + add_div_eq_mul_add_div, one_mul, ← div_mul_eq_div_div, sqrt_div, sqrt_mul_self] <;> + linarith [sqrtTwoAddSeries_nonneg le_rfl n] #align real.cos_pi_over_two_pow Real.cos_pi_over_two_pow theorem sin_sq_pi_over_two_pow (n : ℕ) : @@ -774,14 +779,9 @@ theorem sin_pi_over_two_pow_succ (n : ℕ) : · congr <;> norm_num · rw [sub_nonneg] exact (sqrtTwoAddSeries_lt_two _).le - refine le_of_lt <| mul_pos (sin_pos_of_pos_of_lt_pi ?_ ?_) two_pos - · exact div_pos pi_pos <| pow_pos two_pos _ - refine' lt_of_lt_of_le _ (le_of_eq (div_one _)) - rw [div_lt_div_left pi_pos (pow_pos two_pos _) one_pos] - refine' lt_of_le_of_lt (le_of_eq (pow_zero 2).symm) _ - apply pow_lt_pow - · norm_num - · apply Nat.succ_pos + refine mul_nonneg (sin_nonneg_of_nonneg_of_le_pi ?_ ?_) zero_le_two + · positivity + · exact div_le_self pi_pos.le <| one_le_pow_of_one_le one_le_two _ #align real.sin_pi_over_two_pow_succ Real.sin_pi_over_two_pow_succ @[simp] @@ -864,33 +864,19 @@ theorem cos_pi_div_three : cos (π / 3) = 1 / 2 := by linarith [cos_pi] #align real.cos_pi_div_three Real.cos_pi_div_three -/-- The square of the cosine of `π / 6` is `3 / 4` (this is sometimes more convenient than the -result for cosine itself). -/ -theorem sq_cos_pi_div_six : cos (π / 6) ^ 2 = 3 / 4 := by - have h1 : cos (π / 6) ^ 2 = 1 / 2 + 1 / 2 / 2 := by - convert cos_sq (π / 6) using 3 - have h2 : 2 * (π / 6) = π / 3 := by linarith - rw [h2, cos_pi_div_three] - rw [← sub_eq_zero] at h1 ⊢ - convert h1 using 1 - ring -#align real.sq_cos_pi_div_six Real.sq_cos_pi_div_six - /-- The cosine of `π / 6` is `√3 / 2`. -/ @[simp] theorem cos_pi_div_six : cos (π / 6) = sqrt 3 / 2 := by - suffices sqrt 3 = cos (π / 6) * 2 by - field_simp [(by norm_num : 0 ≠ 2), ← this] - rw [sqrt_eq_iff_sq_eq] - · have h1 := (mul_right_inj' (by norm_num : (4 : ℝ) ≠ 0)).mpr sq_cos_pi_div_six - rw [← sub_eq_zero] at h1 ⊢ - convert h1 using 1 - ring - · norm_num - · have : 0 < cos (π / 6) := by apply cos_pos_of_mem_Ioo; constructor <;> linarith [pi_pos] - linarith + rw [show (6 : ℝ) = 3 * 2 by norm_num, div_mul_eq_div_div, cos_half, cos_pi_div_three, one_add_div, + ← div_mul_eq_div_div, two_add_one_eq_three, sqrt_div, sqrt_mul_self] <;> linarith [pi_pos] #align real.cos_pi_div_six Real.cos_pi_div_six +/-- The square of the cosine of `π / 6` is `3 / 4` (this is sometimes more convenient than the +result for cosine itself). -/ +theorem sq_cos_pi_div_six : cos (π / 6) ^ 2 = 3 / 4 := by + rw [cos_pi_div_six, div_pow, sq_sqrt] <;> norm_num +#align real.sq_cos_pi_div_six Real.sq_cos_pi_div_six + /-- The sine of `π / 6` is `1 / 2`. -/ @[simp] theorem sin_pi_div_six : sin (π / 6) = 1 / 2 := by diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean index 01c2e7a722cc8..3f5231c886843 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean @@ -103,7 +103,8 @@ theorem U_complex_cos (n : ℕ) : (U ℂ n).eval (cos θ) * sin θ = sin ((n + 1 end Complex --- ### Real versions +/-! ### Real versions -/ + section Real open Real diff --git a/Mathlib/CategoryTheory/Abelian/Homology.lean b/Mathlib/CategoryTheory/Abelian/Homology.lean index d75b716963237..7a24eae161591 100644 --- a/Mathlib/CategoryTheory/Abelian/Homology.lean +++ b/Mathlib/CategoryTheory/Abelian/Homology.lean @@ -12,19 +12,23 @@ import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Images /-! -The object `homology f g w`, where `w : f ≫ g = 0`, can be identified with either a -cokernel or a kernel. The isomorphism with a cokernel is `homologyIsoCokernelLift`, which +The object `homology' f g w`, where `w : f ≫ g = 0`, can be identified with either a +cokernel or a kernel. The isomorphism with a cokernel is `homology'IsoCokernelLift`, which was obtained elsewhere. In the case of an abelian category, this file shows the isomorphism with a kernel as well. -We use these isomorphisms to obtain the analogous api for `homology`: -- `homology.ι` is the map from `homology f g w` into the cokernel of `f`. -- `homology.π'` is the map from `kernel g` to `homology f g w`. -- `homology.desc'` constructs a morphism from `homology f g w`, when it is viewed as a cokernel. -- `homology.lift` constructs a morphism to `homology f g w`, when it is viewed as a kernel. +We use these isomorphisms to obtain the analogous api for `homology'`: +- `homology'.ι` is the map from `homology' f g w` into the cokernel of `f`. +- `homology'.π'` is the map from `kernel g` to `homology' f g w`. +- `homology'.desc'` constructs a morphism from `homology' f g w`, when it is viewed as a cokernel. +- `homology'.lift` constructs a morphism to `homology' f g w`, when it is viewed as a kernel. - Various small lemmas are proved as well, mimicking the API for (co)kernels. With these definitions and lemmas, the isomorphisms between homology and a (co)kernel need not be used directly. + +Note: As part of the homology refactor, it is planned to remove the definitions in this file, +because it can be replaced by the content of `Algebra.Homology.ShortComplex.Homology`. + -/ @@ -103,98 +107,98 @@ instance (w : f ≫ g = 0) : IsIso (homologyCToK f g w) := end CategoryTheory.Abelian /-- The homology associated to `f` and `g` is isomorphic to a kernel. -/ -def homologyIsoKernelDesc : homology f g w ≅ kernel (cokernel.desc f g w) := - homologyIsoCokernelLift _ _ _ ≪≫ asIso (CategoryTheory.Abelian.homologyCToK _ _ _) -#align homology_iso_kernel_desc homologyIsoKernelDesc +def homology'IsoKernelDesc : homology' f g w ≅ kernel (cokernel.desc f g w) := + homology'IsoCokernelLift _ _ _ ≪≫ asIso (CategoryTheory.Abelian.homologyCToK _ _ _) +#align homology_iso_kernel_desc homology'IsoKernelDesc -namespace homology +namespace homology' -- `homology.π` is taken /-- The canonical map from the kernel of `g` to the homology of `f` and `g`. -/ -def π' : kernel g ⟶ homology f g w := - cokernel.π _ ≫ (homologyIsoCokernelLift _ _ _).inv -#align homology.π' homology.π' +def π' : kernel g ⟶ homology' f g w := + cokernel.π _ ≫ (homology'IsoCokernelLift _ _ _).inv +#align homology.π' homology'.π' /-- The canonical map from the homology of `f` and `g` to the cokernel of `f`. -/ -def ι : homology f g w ⟶ cokernel f := - (homologyIsoKernelDesc _ _ _).hom ≫ kernel.ι _ -#align homology.ι homology.ι +def ι : homology' f g w ⟶ cokernel f := + (homology'IsoKernelDesc _ _ _).hom ≫ kernel.ι _ +#align homology.ι homology'.ι /-- Obtain a morphism from the homology, given a morphism from the kernel. -/ -def desc' {W : A} (e : kernel g ⟶ W) (he : kernel.lift g f w ≫ e = 0) : homology f g w ⟶ W := - (homologyIsoCokernelLift _ _ _).hom ≫ cokernel.desc _ e he -#align homology.desc' homology.desc' +def desc' {W : A} (e : kernel g ⟶ W) (he : kernel.lift g f w ≫ e = 0) : homology' f g w ⟶ W := + (homology'IsoCokernelLift _ _ _).hom ≫ cokernel.desc _ e he +#align homology.desc' homology'.desc' /-- Obtain a moprhism to the homology, given a morphism to the kernel. -/ -def lift {W : A} (e : W ⟶ cokernel f) (he : e ≫ cokernel.desc f g w = 0) : W ⟶ homology f g w := - kernel.lift _ e he ≫ (homologyIsoKernelDesc _ _ _).inv -#align homology.lift homology.lift +def lift {W : A} (e : W ⟶ cokernel f) (he : e ≫ cokernel.desc f g w = 0) : W ⟶ homology' f g w := + kernel.lift _ e he ≫ (homology'IsoKernelDesc _ _ _).inv +#align homology.lift homology'.lift @[reassoc (attr := simp)] theorem π'_desc' {W : A} (e : kernel g ⟶ W) (he : kernel.lift g f w ≫ e = 0) : π' f g w ≫ desc' f g w e he = e := by dsimp [π', desc'] simp -#align homology.π'_desc' homology.π'_desc' +#align homology.π'_desc' homology'.π'_desc' @[reassoc (attr := simp)] theorem lift_ι {W : A} (e : W ⟶ cokernel f) (he : e ≫ cokernel.desc f g w = 0) : lift f g w e he ≫ ι _ _ _ = e := by dsimp [ι, lift] simp -#align homology.lift_ι homology.lift_ι +#align homology.lift_ι homology'.lift_ι @[reassoc (attr := simp)] theorem condition_π' : kernel.lift g f w ≫ π' f g w = 0 := by dsimp [π'] simp -#align homology.condition_π' homology.condition_π' +#align homology.condition_π' homology'.condition_π' @[reassoc (attr := simp)] theorem condition_ι : ι f g w ≫ cokernel.desc f g w = 0 := by dsimp [ι] simp -#align homology.condition_ι homology.condition_ι +#align homology.condition_ι homology'.condition_ι @[ext] -theorem hom_from_ext {W : A} (a b : homology f g w ⟶ W) +theorem hom_from_ext {W : A} (a b : homology' f g w ⟶ W) (h : π' f g w ≫ a = π' f g w ≫ b) : a = b := by dsimp [π'] at h - apply_fun fun e => (homologyIsoCokernelLift f g w).inv ≫ e + apply_fun fun e => (homology'IsoCokernelLift f g w).inv ≫ e swap · intro i j hh - apply_fun fun e => (homologyIsoCokernelLift f g w).hom ≫ e at hh + apply_fun fun e => (homology'IsoCokernelLift f g w).hom ≫ e at hh simpa using hh simp only [Category.assoc] at h exact coequalizer.hom_ext h -#align homology.hom_from_ext homology.hom_from_ext +#align homology.hom_from_ext homology'.hom_from_ext @[ext] -theorem hom_to_ext {W : A} (a b : W ⟶ homology f g w) (h : a ≫ ι f g w = b ≫ ι f g w) : a = b := by +theorem hom_to_ext {W : A} (a b : W ⟶ homology' f g w) (h : a ≫ ι f g w = b ≫ ι f g w) : a = b := by dsimp [ι] at h - apply_fun fun e => e ≫ (homologyIsoKernelDesc f g w).hom + apply_fun fun e => e ≫ (homology'IsoKernelDesc f g w).hom swap · intro i j hh - apply_fun fun e => e ≫ (homologyIsoKernelDesc f g w).inv at hh + apply_fun fun e => e ≫ (homology'IsoKernelDesc f g w).inv at hh simpa using hh simp only [← Category.assoc] at h exact equalizer.hom_ext h -#align homology.hom_to_ext homology.hom_to_ext +#align homology.hom_to_ext homology'.hom_to_ext @[reassoc (attr := simp)] theorem π'_ι : π' f g w ≫ ι f g w = kernel.ι _ ≫ cokernel.π _ := by - dsimp [π', ι, homologyIsoKernelDesc] + dsimp [π', ι, homology'IsoKernelDesc] simp -#align homology.π'_ι homology.π'_ι +#align homology.π'_ι homology'.π'_ι @[reassoc (attr := simp)] theorem π'_eq_π : (kernelSubobjectIso _).hom ≫ π' f g w = π _ _ _ := by - dsimp [π', homologyIsoCokernelLift] + dsimp [π', homology'IsoCokernelLift] simp only [← Category.assoc] rw [Iso.comp_inv_eq] - dsimp [π, homologyIsoCokernelImageToKernel'] + dsimp [π, homology'IsoCokernelImageToKernel'] simp -#align homology.π'_eq_π homology.π'_eq_π +#align homology.π'_eq_π homology'.π'_eq_π section @@ -222,56 +226,56 @@ theorem π'_map (α β h) : π' _ _ _ ≫ map w w' α β h = simp rw [this] simp only [Category.assoc] - dsimp [π', homologyIsoCokernelLift] + dsimp [π', homology'IsoCokernelLift] simp only [cokernelIsoOfEq_inv_comp_desc, cokernel.π_desc_assoc] congr 1 · congr exact h.symm · rw [Iso.inv_comp_eq, ← Category.assoc, Iso.eq_comp_inv] - dsimp [homologyIsoCokernelImageToKernel'] + dsimp [homology'IsoCokernelImageToKernel'] simp -#align homology.π'_map homology.π'_map +#align homology.π'_map homology'.π'_map -- Porting note: need to fill in f,g,f',g' in the next few results or time out theorem map_eq_desc'_lift_left (α β h) : map w w' α β h = - homology.desc' f g _ (homology.lift f' g' _ (kernel.ι _ ≫ β.left ≫ cokernel.π _) + homology'.desc' f g _ (homology'.lift f' g' _ (kernel.ι _ ≫ β.left ≫ cokernel.π _) (by simp)) (by ext simp only [← h, Category.assoc, zero_comp, lift_ι, kernel.lift_ι_assoc] erw [← reassoc_of% α.w] simp) := by - apply homology.hom_from_ext + apply homology'.hom_from_ext simp only [π'_map, π'_desc'] dsimp [π', lift] rw [Iso.eq_comp_inv] - dsimp [homologyIsoKernelDesc] + dsimp [homology'IsoKernelDesc] ext simp [h] -#align homology.map_eq_desc'_lift_left homology.map_eq_desc'_lift_left +#align homology.map_eq_desc'_lift_left homology'.map_eq_desc'_lift_left theorem map_eq_lift_desc'_left (α β h) : map w w' α β h = - homology.lift f' g' _ - (homology.desc' f g _ (kernel.ι _ ≫ β.left ≫ cokernel.π _) + homology'.lift f' g' _ + (homology'.desc' f g _ (kernel.ι _ ≫ β.left ≫ cokernel.π _) (by simp only [kernel.lift_ι_assoc, ← h] erw [← reassoc_of% α.w] simp)) (by -- Porting note: used to be ext - apply homology.hom_from_ext + apply homology'.hom_from_ext simp) := by rw [map_eq_desc'_lift_left] -- Porting note: once was known as ext - apply homology.hom_to_ext - apply homology.hom_from_ext + apply homology'.hom_to_ext + apply homology'.hom_from_ext simp -#align homology.map_eq_lift_desc'_left homology.map_eq_lift_desc'_left +#align homology.map_eq_lift_desc'_left homology'.map_eq_lift_desc'_left theorem map_eq_desc'_lift_right (α β h) : map w w' α β h = - homology.desc' f g _ (homology.lift f' g' _ (kernel.ι _ ≫ α.right ≫ cokernel.π _) + homology'.desc' f g _ (homology'.lift f' g' _ (kernel.ι _ ≫ α.right ≫ cokernel.π _) (by simp [h])) (by ext @@ -281,26 +285,26 @@ theorem map_eq_desc'_lift_right (α β h) : rw [map_eq_desc'_lift_left] ext simp [h] -#align homology.map_eq_desc'_lift_right homology.map_eq_desc'_lift_right +#align homology.map_eq_desc'_lift_right homology'.map_eq_desc'_lift_right theorem map_eq_lift_desc'_right (α β h) : map w w' α β h = - homology.lift f' g' _ - (homology.desc' f g _ (kernel.ι _ ≫ α.right ≫ cokernel.π _) + homology'.lift f' g' _ + (homology'.desc' f g _ (kernel.ι _ ≫ α.right ≫ cokernel.π _) (by simp only [kernel.lift_ι_assoc] erw [← reassoc_of% α.w] simp)) (by -- Porting note: once was known as ext - apply homology.hom_from_ext + apply homology'.hom_from_ext simp [h]) := by rw [map_eq_desc'_lift_right] -- Porting note: once was known as ext - apply homology.hom_to_ext - apply homology.hom_from_ext + apply homology'.hom_to_ext + apply homology'.hom_from_ext simp -#align homology.map_eq_lift_desc'_right homology.map_eq_lift_desc'_right +#align homology.map_eq_lift_desc'_right homology'.map_eq_lift_desc'_right @[reassoc (attr := simp)] theorem map_ι (α β h) : @@ -308,14 +312,14 @@ theorem map_ι (α β h) : ι f g w ≫ cokernel.map f f' α.left β.left (by simp [h, β.w.symm]) := by rw [map_eq_lift_desc'_left, lift_ι] -- Porting note: once was known as ext - apply homology.hom_from_ext + apply homology'.hom_from_ext simp only [← Category.assoc] rw [π'_ι, π'_desc', Category.assoc, Category.assoc, cokernel.π_desc] -#align homology.map_ι homology.map_ι +#align homology.map_ι homology'.map_ι end -end homology +end homology' namespace CategoryTheory.Functor @@ -323,8 +327,8 @@ variable {ι : Type*} {c : ComplexShape ι} {B : Type*} [Category B] [Abelian B] [Functor.Additive F] [PreservesFiniteLimits F] [PreservesFiniteColimits F] /-- When `F` is an exact additive functor, `F(Hᵢ(X)) ≅ Hᵢ(F(X))` for `X` a complex. -/ -noncomputable def homologyIso (C : HomologicalComplex A c) (j : ι) : - F.obj (C.homology j) ≅ ((F.mapHomologicalComplex c).obj C).homology j := +noncomputable def homology'Iso (C : HomologicalComplex A c) (j : ι) : + F.obj (C.homology' j) ≅ ((F.mapHomologicalComplex c).obj C).homology' j := (PreservesCokernel.iso F _).trans (cokernel.mapIso _ _ ((F.mapIso (imageSubobjectIso _)).trans @@ -337,19 +341,19 @@ noncomputable def homologyIso (C : HomologicalComplex A c) (j : ι) : simp only [Category.assoc, imageToKernel_arrow] erw [kernelSubobject_arrow', imageSubobject_arrow'] simp [← F.map_comp])) -#align category_theory.functor.homology_iso CategoryTheory.Functor.homologyIso +#align category_theory.functor.homology_iso CategoryTheory.Functor.homology'Iso /-- If `F` is an exact additive functor, then `F` commutes with `Hᵢ` (up to natural isomorphism). -/ -noncomputable def homologyFunctorIso (i : ι) : - homologyFunctor A c i ⋙ F ≅ F.mapHomologicalComplex c ⋙ homologyFunctor B c i := - NatIso.ofComponents (fun X => homologyIso F X i) (by +noncomputable def homology'FunctorIso (i : ι) : + homology'Functor A c i ⋙ F ≅ F.mapHomologicalComplex c ⋙ homology'Functor B c i := + NatIso.ofComponents (fun X => homology'Iso F X i) (by intro X Y f dsimp rw [← Iso.inv_comp_eq, ← Category.assoc, ← Iso.eq_comp_inv] refine' coequalizer.hom_ext _ - dsimp [homologyIso] + dsimp [homology'Iso] simp only [PreservesCokernel.iso_inv] - dsimp [homology.map] + dsimp [homology'.map] simp only [← Category.assoc, cokernel.π_desc] simp only [Category.assoc, cokernelComparison_map_desc, cokernel.π_desc] simp only [π_comp_cokernelComparison, ← F.map_comp] @@ -362,6 +366,6 @@ noncomputable def homologyFunctorIso (i : ι) : rotate_right; simp rw [← kernel_map_comp_kernelSubobjectIso_inv] any_goals simp) -#align category_theory.functor.homology_functor_iso CategoryTheory.Functor.homologyFunctorIso +#align category_theory.functor.homology_functor_iso CategoryTheory.Functor.homology'FunctorIso end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean index 1cf6bd7ccde08..03e2361395341 100644 --- a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean @@ -329,7 +329,7 @@ variable {C : Type u} [Category.{v} C] [Abelian C] /-- If `X` is a cochain complex of injective objects and we have a quasi-isomorphism `f : Y[0] ⟶ X`, then `X` is an injective resolution of `Y`. -/ def HomologicalComplex.Hom.fromSingle₀InjectiveResolution (X : CochainComplex C ℕ) (Y : C) - (f : (CochainComplex.single₀ C).obj Y ⟶ X) [QuasiIso f] (H : ∀ n, Injective (X.X n)) : + (f : (CochainComplex.single₀ C).obj Y ⟶ X) [QuasiIso' f] (H : ∀ n, Injective (X.X n)) : InjectiveResolution Y where cocomplex := X ι := f diff --git a/Mathlib/CategoryTheory/Abelian/LeftDerived.lean b/Mathlib/CategoryTheory/Abelian/LeftDerived.lean index 6ba9646fbadb8..593a3c57989e8 100644 --- a/Mathlib/CategoryTheory/Abelian/LeftDerived.lean +++ b/Mathlib/CategoryTheory/Abelian/LeftDerived.lean @@ -69,7 +69,7 @@ theorem exact_of_map_projectiveResolution (P : ProjectiveResolution X) def leftDerivedZeroToSelfApp [EnoughProjectives C] {X : C} (P : ProjectiveResolution X) : (F.leftDerived 0).obj X ⟶ F.obj X := (leftDerivedObjIso F 0 P).hom ≫ - homology.desc' _ _ _ (kernel.ι _ ≫ F.map (P.π.f 0)) + homology'.desc' _ _ _ (kernel.ι _ ≫ F.map (P.π.f 0)) (by rw [kernel.lift_ι_assoc, HomologicalComplex.dTo_eq _ (by simp : (ComplexShape.down ℕ).Rel 1 0), @@ -85,7 +85,7 @@ def leftDerivedZeroToSelfAppInv [EnoughProjectives C] [PreservesFiniteColimits F have := isIso_cokernel_desc_of_exact_of_epi _ _ (exact_of_map_projectiveResolution F P) refine' (asIso (cokernel.desc _ _ (exact_of_map_projectiveResolution F P).w)).inv ≫ - _ ≫ (homologyIsoCokernelLift _ _ _).inv ≫ (leftDerivedObjIso F 0 P).inv + _ ≫ (homology'IsoCokernelLift _ _ _).inv ≫ (leftDerivedObjIso F 0 P).inv refine' cokernel.map _ _ (𝟙 _) (kernel.lift _ (𝟙 _) (by simp)) _ ext -- Porting note: this used to just be `simp` @@ -103,11 +103,11 @@ theorem leftDerivedZeroToSelfApp_comp_inv [EnoughProjectives C] [PreservesFinite convert Category.comp_id (leftDerivedObjIso F 0 P).hom rw [← Category.assoc, ← Category.assoc, Iso.comp_inv_eq] -- Porting note: broken ext - apply homology.hom_from_ext + apply homology'.hom_from_ext simp only [← Category.assoc] - erw [homology.π'_desc', Category.assoc, Category.assoc, ← + erw [homology'.π'_desc', Category.assoc, Category.assoc, ← Category.assoc (F.map _), Abelian.cokernel.desc.inv _ _ (exact_of_map_projectiveResolution F P), - cokernel.π_desc, homology.π', Category.comp_id, Category.assoc (cokernel.π _), Iso.inv_hom_id, + cokernel.π_desc, homology'.π', Category.comp_id, Category.assoc (cokernel.π _), Iso.inv_hom_id, Category.comp_id, ← Category.assoc] -- Porting note: restructured proof to avoid `convert` conv_rhs => rw [← Category.id_comp (cokernel.π _)] @@ -141,8 +141,8 @@ theorem leftDerivedZeroToSelfAppInv_comp [EnoughProjectives C] [PreservesFiniteC -- Porting note: working around 'motive is not type correct' simp only [Category.comp_id] ext - simp only [cokernel.π_desc_assoc, Category.assoc, cokernel.π_desc, homology.desc'] - rw [← Category.assoc, ← Category.assoc (homologyIsoCokernelLift _ _ _).inv, Iso.inv_hom_id] + simp only [cokernel.π_desc_assoc, Category.assoc, cokernel.π_desc, homology'.desc'] + rw [← Category.assoc, ← Category.assoc (homology'IsoCokernelLift _ _ _).inv, Iso.inv_hom_id] simp only [Category.assoc, cokernel.π_desc, kernel.lift_ι_assoc, Category.id_comp] #align category_theory.abelian.functor.left_derived_zero_to_self_app_inv_comp CategoryTheory.Abelian.Functor.leftDerivedZeroToSelfAppInv_comp @@ -166,12 +166,12 @@ theorem leftDerived_zero_to_self_natural [EnoughProjectives C] {X : C} {Y : C} ( rw [Functor.leftDerived_map_eq F 0 f (ProjectiveResolution.lift f P Q) (by simp), Category.assoc, Category.assoc, ← Category.assoc _ (F.leftDerivedObjIso 0 Q).hom, Iso.inv_hom_id, Category.id_comp, Category.assoc, whisker_eq] - dsimp only [homologyFunctor_map] + dsimp only [homology'Functor_map] -- Porting note: broken ext - apply homology.hom_from_ext + apply homology'.hom_from_ext simp only [HomologicalComplex.Hom.sqTo_right, mapHomologicalComplex_map_f, - homology.π'_map_assoc, homology.π'_desc', kernel.lift_ι_assoc, Category.assoc, - homology.π'_desc'_assoc, ← map_comp, + homology'.π'_map_assoc, homology'.π'_desc', kernel.lift_ι_assoc, Category.assoc, + homology'.π'_desc'_assoc, ← map_comp, show (ProjectiveResolution.lift f P Q).f 0 ≫ _ = _ ≫ f from HomologicalComplex.congr_hom (ProjectiveResolution.lift_commutes f P Q) 0] #align category_theory.abelian.functor.left_derived_zero_to_self_natural CategoryTheory.Abelian.Functor.leftDerived_zero_to_self_natural diff --git a/Mathlib/CategoryTheory/Abelian/Projective.lean b/Mathlib/CategoryTheory/Abelian/Projective.lean index 65ef9c70ae635..f8c33b67b2e79 100644 --- a/Mathlib/CategoryTheory/Abelian/Projective.lean +++ b/Mathlib/CategoryTheory/Abelian/Projective.lean @@ -165,7 +165,7 @@ variable {C : Type u} [Category.{v} C] [Abelian C] then `X` is a projective resolution of `Y.` -/ def toSingle₀ProjectiveResolution {X : ChainComplex C ℕ} {Y : C} -- porting note: autoporter incorrectly went for `X.pt` at the end there - (f : X ⟶ (ChainComplex.single₀ C).obj Y) [QuasiIso f] (H : ∀ n, Projective (X.X n)) : + (f : X ⟶ (ChainComplex.single₀ C).obj Y) [QuasiIso' f] (H : ∀ n, Projective (X.X n)) : ProjectiveResolution Y where complex := X π := f diff --git a/Mathlib/CategoryTheory/Abelian/Refinements.lean b/Mathlib/CategoryTheory/Abelian/Refinements.lean new file mode 100644 index 0000000000000..8418752824196 --- /dev/null +++ b/Mathlib/CategoryTheory/Abelian/Refinements.lean @@ -0,0 +1,116 @@ +/- +Copyright (c) 2023 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Homology.ShortComplex.Exact + +/-! +# Refinements + +In order to prove injectivity/surjectivity/exactness properties for diagrams +in the category of abelian groups, we often need to do diagram chases. +Some of these can be carried out in more general abelian categories: +for example, a morphism `X ⟶ Y` in an abelian category `C` is a +monomorphism if and only if for all `A : C`, the induced map +`(A ⟶ X) → (A ⟶ Y)` of abelian groups is a monomorphism, i.e. injective. +Alternatively, the yoneda presheaf functor which sends `X` to the +presheaf of maps `A ⟶ X` for all `A : C` preserves and reflects +monomorphisms. + +However, if `p : X ⟶ Y` is an epimorphism in `C` and `A : C`, +`(A ⟶ X) → (A ⟶ Y)` may fail to be surjective (unless `p` is a split +epimorphism). + +In this file, the basic result is `epi_iff_surjective_up_to_refinements` +which states that `f : X ⟶ Y` is a morphism in an abelian category, +then it is an epimorphism if and only if for all `y : A ⟶ Y`, +there exists an epimorphism `π : A' ⟶ A` and `x : A' ⟶ X` such +that `π ≫ y = x ≫ f`. In order words, if we allow a precomposition +with an epimorphism, we may lift a morphism to `Y` to a morphism to `X`. +Following unpublished notes by George Bergman, we shall say that the +precomposition by an epimorphism `π ≫ y` is a refinement of `y`. Then, +we get that an epimorphism is a morphism that is "surjective up to refinements". +(This result is similar to the fact that a morphism of sheaves on +a topological space or a site is epi iff sections can be lifted +locally. Then, arguing "up to refinements" is very similar to +arguing locally for a Grothendieck topology (TODO: show that it +corresponds to arguing for the canonical topology on the abelian +category `C` by showing that a morphism in `C` is an epi iff +the corresponding morphisms of sheaves for the canonical +topology is an epi, and that the criteria +`epi_iff_surjective_up_to_refinements` could be deduced from +this equivalence.) + +Similarly, it is possible to show that a short complex in an abelian +category is exact if and only if it is exact up to refinements +(see `ShortComplex.exact_iff_exact_up_to_refinements`). + +As it is outlined in the documentation of the file +`CategoryTheory.Abelian.Pseudoelements`, the Freyd-Mitchell +embedding theorem implies the existence of a faithful and exact functor `ι` +from an abelian category `C` to the category of abelian groups. If we +define a pseudo-element of `X : C` to be an element in `ι.obj X`, one +may do diagram chases in any abelian category using these pseudo-elements. +However, using this approach would require proving this embedding theorem! +Currently, mathlib contains a weaker notion of pseudo-elements +`CategoryTheory.Abelian.Pseudoelements`. Some theorems can be obtained +using this notion, but there is the issue that for this notion +of pseudo-elements a morphism `X ⟶ Y` in `C` is not determined by +its action on pseudo-elements (see also `Counterexamples/Pseudoelement`). +On the contrary, the approach consisting of working up to refinements +does not require the introduction of other types: we only need to work +with morphisms `A ⟶ X` in `C` which we may consider as being +"sort of elements of `X`". One may carry diagram-chasing by tracking +these morphisms and sometimes introducing an auxiliary epimorphism `A' ⟶ A`. + +## References +* George Bergman, A note on abelian categories – translating element-chasing proofs, +and exact embedding in abelian groups (1974) +http://math.berkeley.edu/~gbergman/papers/unpub/elem-chase.pdf + +-/ + +namespace CategoryTheory + +open Category Limits + +variable {C : Type _} [Category C] [Abelian C] {X Y : C} (S : ShortComplex C) + {S₁ S₂ : ShortComplex C} + +lemma epi_iff_surjective_up_to_refinements (f : X ⟶ Y) : + Epi f ↔ ∀ ⦃A : C⦄ (y : A ⟶ Y), + ∃ (A' : C) (π : A' ⟶ A) (_ : Epi π) (x : A' ⟶ X), π ≫ y = x ≫ f := by + constructor + · intro _ A a + exact ⟨pullback a f, pullback.fst, inferInstance, pullback.snd, pullback.condition⟩ + · intro hf + obtain ⟨A, π, hπ, a', fac⟩ := hf (𝟙 Y) + rw [comp_id] at fac + exact epi_of_epi_fac fac.symm + +lemma surjective_up_to_refinements_of_epi (f : X ⟶ Y) [Epi f] {A : C} (y : A ⟶ Y) : + ∃ (A' : C) (π : A' ⟶ A) (_ : Epi π) (x : A' ⟶ X), π ≫ y = x ≫ f := + (epi_iff_surjective_up_to_refinements f).1 inferInstance y + +lemma ShortComplex.exact_iff_exact_up_to_refinements : + S.Exact ↔ ∀ ⦃A : C⦄ (x₂ : A ⟶ S.X₂) (_ : x₂ ≫ S.g = 0), + ∃ (A' : C) (π : A' ⟶ A) (_ : Epi π) (x₁ : A' ⟶ S.X₁), π ≫ x₂ = x₁ ≫ S.f := by + rw [S.exact_iff_epi_toCycles, epi_iff_surjective_up_to_refinements] + constructor + · intro hS A a ha + obtain ⟨A', π, hπ, x₁, fac⟩ := hS (S.liftCycles a ha) + exact ⟨A', π, hπ, x₁, by simpa only [assoc, liftCycles_i, toCycles_i] using fac =≫ S.iCycles⟩ + · intro hS A a + obtain ⟨A', π, hπ, x₁, fac⟩ := hS (a ≫ S.iCycles) (by simp) + exact ⟨A', π, hπ, x₁, by simp only [← cancel_mono S.iCycles, assoc, toCycles_i, fac]⟩ + +variable {S} + +lemma ShortComplex.Exact.exact_up_to_refinements + (hS : S.Exact) {A : C} (x₂ : A ⟶ S.X₂) (hx₂ : x₂ ≫ S.g = 0) : + ∃ (A' : C) (π : A' ⟶ A) (_ : Epi π) (x₁ : A' ⟶ S.X₁), π ≫ x₂ = x₁ ≫ S.f := by + rw [ShortComplex.exact_iff_exact_up_to_refinements] at hS + exact hS x₂ hx₂ + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Abelian/RightDerived.lean b/Mathlib/CategoryTheory/Abelian/RightDerived.lean index 31f6ac000e864..870ad3bd266c8 100644 --- a/Mathlib/CategoryTheory/Abelian/RightDerived.lean +++ b/Mathlib/CategoryTheory/Abelian/RightDerived.lean @@ -64,7 +64,7 @@ variable [Abelian C] [HasInjectiveResolutions C] [Abelian D] /-- The right derived functors of an additive functor. -/ def Functor.rightDerived (F : C ⥤ D) [F.Additive] (n : ℕ) : C ⥤ D := - injectiveResolutions C ⋙ F.mapHomotopyCategory _ ⋙ HomotopyCategory.homologyFunctor D _ n + injectiveResolutions C ⋙ F.mapHomotopyCategory _ ⋙ HomotopyCategory.homology'Functor D _ n #align category_theory.functor.right_derived CategoryTheory.Functor.rightDerived /-- We can compute a right derived functor using a chosen injective resolution. -/ @@ -72,11 +72,11 @@ def Functor.rightDerived (F : C ⥤ D) [F.Additive] (n : ℕ) : C ⥤ D := def Functor.rightDerivedObjIso (F : C ⥤ D) [F.Additive] (n : ℕ) {X : C} (P : InjectiveResolution X) : (F.rightDerived n).obj X ≅ - (homologyFunctor D _ n).obj ((F.mapHomologicalComplex _).obj P.cocomplex) := - (HomotopyCategory.homologyFunctor D _ n).mapIso + (homology'Functor D _ n).obj ((F.mapHomologicalComplex _).obj P.cocomplex) := + (HomotopyCategory.homology'Functor D _ n).mapIso (HomotopyCategory.isoOfHomotopyEquiv (F.mapHomotopyEquiv (InjectiveResolution.homotopyEquiv _ P))) ≪≫ - (HomotopyCategory.homologyFactors D _ n).app _ + (HomotopyCategory.homology'Factors D _ n).app _ #align category_theory.functor.right_derived_obj_iso CategoryTheory.Functor.rightDerivedObjIso /-- The 0-th derived functor of `F` on an injective object `X` is just `F.obj X`. -/ @@ -84,7 +84,7 @@ def Functor.rightDerivedObjIso (F : C ⥤ D) [F.Additive] (n : ℕ) {X : C} def Functor.rightDerivedObjInjectiveZero (F : C ⥤ D) [F.Additive] (X : C) [Injective X] : (F.rightDerived 0).obj X ≅ F.obj X := F.rightDerivedObjIso 0 (InjectiveResolution.self X) ≪≫ - (homologyFunctor _ _ _).mapIso ((CochainComplex.single₀MapHomologicalComplex F).app X) ≪≫ + (homology'Functor _ _ _).mapIso ((CochainComplex.single₀MapHomologicalComplex F).app X) ≪≫ (CochainComplex.homologyFunctor0Single₀ D).app (F.obj X) #align category_theory.functor.right_derived_obj_injective_zero CategoryTheory.Functor.rightDerivedObjInjectiveZero @@ -95,8 +95,8 @@ open ZeroObject def Functor.rightDerivedObjInjectiveSucc (F : C ⥤ D) [F.Additive] (n : ℕ) (X : C) [Injective X] : (F.rightDerived (n + 1)).obj X ≅ 0 := F.rightDerivedObjIso (n + 1) (InjectiveResolution.self X) ≪≫ - (homologyFunctor _ _ _).mapIso ((CochainComplex.single₀MapHomologicalComplex F).app X) ≪≫ - (CochainComplex.homologyFunctorSuccSingle₀ D n).app (F.obj X) ≪≫ (Functor.zero_obj _).isoZero + (homology'Functor _ _ _).mapIso ((CochainComplex.single₀MapHomologicalComplex F).app X) ≪≫ + (CochainComplex.homology'FunctorSuccSingle₀ D n).app (F.obj X) ≪≫ (Functor.zero_obj _).isoZero #align category_theory.functor.right_derived_obj_injective_succ CategoryTheory.Functor.rightDerivedObjInjectiveSucc /-- We can compute a right derived functor on a morphism using a descent of that morphism @@ -107,12 +107,12 @@ theorem Functor.rightDerived_map_eq (F : C ⥤ D) [F.Additive] (n : ℕ) {X Y : (w : Q.ι ≫ g = (CochainComplex.single₀ C).map f ≫ P.ι) : (F.rightDerived n).map f = (F.rightDerivedObjIso n Q).hom ≫ - (homologyFunctor D _ n).map ((F.mapHomologicalComplex _).map g) ≫ + (homology'Functor D _ n).map ((F.mapHomologicalComplex _).map g) ≫ (F.rightDerivedObjIso n P).inv := by dsimp only [Functor.rightDerived, Functor.rightDerivedObjIso] dsimp simp only [Category.comp_id, Category.id_comp] - rw [← homologyFunctor_map, HomotopyCategory.homologyFunctor_map_factors] + rw [← homology'Functor_map, HomotopyCategory.homology'Functor_map_factors] simp only [← Functor.map_comp] congr 1 apply HomotopyCategory.eq_of_homotopy @@ -129,7 +129,7 @@ theorem Functor.rightDerived_map_eq (F : C ⥤ D) [F.Additive] (n : ℕ) {X Y : def NatTrans.rightDerived {F G : C ⥤ D} [F.Additive] [G.Additive] (α : F ⟶ G) (n : ℕ) : F.rightDerived n ⟶ G.rightDerived n := whiskerLeft (injectiveResolutions C) - (whiskerRight (NatTrans.mapHomotopyCategory α _) (HomotopyCategory.homologyFunctor D _ n)) + (whiskerRight (NatTrans.mapHomotopyCategory α _) (HomotopyCategory.homology'Functor D _ n)) #align category_theory.nat_trans.right_derived CategoryTheory.NatTrans.rightDerived @[simp] @@ -153,12 +153,12 @@ theorem NatTrans.rightDerived_eq {F G : C ⥤ D} [F.Additive] [G.Additive] (α : (P : InjectiveResolution X) : (NatTrans.rightDerived α n).app X = (F.rightDerivedObjIso n P).hom ≫ - (homologyFunctor D _ n).map ((NatTrans.mapHomologicalComplex α _).app P.cocomplex) ≫ + (homology'Functor D _ n).map ((NatTrans.mapHomologicalComplex α _).app P.cocomplex) ≫ (G.rightDerivedObjIso n P).inv := by symm dsimp [NatTrans.rightDerived, Functor.rightDerivedObjIso] simp only [Category.comp_id, Category.id_comp] - rw [← homologyFunctor_map, HomotopyCategory.homologyFunctor_map_factors] + rw [← homology'Functor_map, HomotopyCategory.homology'Functor_map_factors] simp only [← Functor.map_comp] congr 1 apply HomotopyCategory.eq_of_homotopy @@ -210,7 +210,7 @@ theorem exact_of_map_injectiveResolution (P : InjectiveResolution X) [PreservesF def rightDerivedZeroToSelfApp [EnoughInjectives C] [PreservesFiniteLimits F] {X : C} (P : InjectiveResolution X) : (F.rightDerived 0).obj X ⟶ F.obj X := (rightDerivedObjIso F 0 P).hom ≫ - (homologyIsoKernelDesc _ _ _).hom ≫ + (homology'IsoKernelDesc _ _ _).hom ≫ kernel.map _ (((F.mapHomologicalComplex (ComplexShape.up ℕ)).obj P.cocomplex).dFrom 0) (cokernel.desc _ (𝟙 _) (by simp)) (𝟙 _) (by @@ -228,7 +228,7 @@ def rightDerivedZeroToSelfApp [EnoughInjectives C] [PreservesFiniteLimits F] {X /-- Given `P : InjectiveResolution X`, a morphism `F.obj X ⟶ (F.rightDerived 0).obj X`. -/ def rightDerivedZeroToSelfAppInv [EnoughInjectives C] {X : C} (P : InjectiveResolution X) : F.obj X ⟶ (F.rightDerived 0).obj X := - homology.lift _ _ _ (F.map (P.ι.f 0) ≫ cokernel.π _) + homology'.lift _ _ _ (F.map (P.ι.f 0) ≫ cokernel.π _) (by have : (ComplexShape.up ℕ).Rel 0 1 := rfl rw [Category.assoc, cokernel.π_desc, HomologicalComplex.dFrom_eq _ this, @@ -244,15 +244,15 @@ theorem rightDerivedZeroToSelfApp_comp_inv [EnoughInjectives C] [PreservesFinite rw [← Category.assoc, Iso.comp_inv_eq, Category.id_comp, Category.assoc, Category.assoc, ← Iso.eq_inv_comp, Iso.inv_hom_id] -- Porting note: broken ext - apply homology.hom_to_ext - apply homology.hom_from_ext - rw [Category.assoc, Category.assoc, homology.lift_ι, Category.id_comp] - erw [homology.π'_ι] -- Porting note: had to insist + apply homology'.hom_to_ext + apply homology'.hom_from_ext + rw [Category.assoc, Category.assoc, homology'.lift_ι, Category.id_comp] + erw [homology'.π'_ι] -- Porting note: had to insist rw [Category.assoc, ← Category.assoc _ _ (cokernel.π _), Abelian.kernel.lift.inv, ← Category.assoc, ← Category.assoc _ (kernel.ι _), Limits.kernel.lift_ι, Category.assoc, Category.assoc, ← - Category.assoc (homologyIsoKernelDesc _ _ _).hom _ _, ← homology.ι, ← Category.assoc] - erw [homology.π'_ι] -- Porting note: had to insist + Category.assoc (homology'IsoKernelDesc _ _ _).hom _ _, ← homology'.ι, ← Category.assoc] + erw [homology'.π'_ι] -- Porting note: had to insist rw [Category.assoc, ← Category.assoc (cokernel.π _)] erw [cokernel.π_desc] -- Porting note: had to insist rw [whisker_eq] @@ -273,9 +273,9 @@ theorem rightDerivedZeroToSelfAppInv_comp [EnoughInjectives C] [PreservesFiniteL · rw [Category.id_comp] ext simp only [Limits.kernel.lift_ι_assoc, - Category.assoc, Limits.kernel.lift_ι, homology.lift] + Category.assoc, Limits.kernel.lift_ι, homology'.lift] rw [← Category.assoc, ← Category.assoc, - Category.assoc _ _ (homologyIsoKernelDesc _ _ _).hom] + Category.assoc _ _ (homology'IsoKernelDesc _ _ _).hom] simp -- Porting note: this used to be an instance in ML3 · apply isIso_kernel_lift_of_exact_of_mono _ _ (exact_of_map_injectiveResolution F P) @@ -303,13 +303,13 @@ theorem rightDerivedZeroToSelf_natural [EnoughInjectives C] {X : C} {Y : C} (f : Category.assoc, Category.assoc, Category.assoc, Category.assoc, Iso.inv_hom_id, Category.comp_id, ← Category.assoc (F.rightDerivedObjIso 0 P).inv, Iso.inv_hom_id, Category.id_comp] - dsimp only [homologyFunctor_map] + dsimp only [homology'Functor_map] -- Porting note: broken ext - apply homology.hom_to_ext - rw [Category.assoc, homology.lift_ι, Category.assoc] - erw [homology.map_ι] -- Porting note: need to insist - rw [←Category.assoc (homology.lift _ _ _ _ _) _ _] - erw [homology.lift_ι] -- Porting note: need to insist + apply homology'.hom_to_ext + rw [Category.assoc, homology'.lift_ι, Category.assoc] + erw [homology'.map_ι] -- Porting note: need to insist + rw [←Category.assoc (homology'.lift _ _ _ _ _) _ _] + erw [homology'.lift_ι] -- Porting note: need to insist rw [Category.assoc] erw [cokernel.π_desc] -- Porting note: need to insist rw [← Category.assoc, ← Functor.map_comp, ← Category.assoc, diff --git a/Mathlib/CategoryTheory/Action.lean b/Mathlib/CategoryTheory/Action.lean index bf3c1852f85c5..99cccd566b98e 100644 --- a/Mathlib/CategoryTheory/Action.lean +++ b/Mathlib/CategoryTheory/Action.lean @@ -113,12 +113,12 @@ variable {X} (x : X) /-- The stabilizer of a point is isomorphic to the endomorphism monoid at the corresponding point. In fact they are definitionally equivalent. -/ -def stabilizerIsoEnd : Stabilizer.submonoid M x ≃* @End (ActionCategory M X) _ x := +def stabilizerIsoEnd : stabilizerSubmonoid M x ≃* @End (ActionCategory M X) _ x := MulEquiv.refl _ #align category_theory.action_category.stabilizer_iso_End CategoryTheory.ActionCategory.stabilizerIsoEnd @[simp] -theorem stabilizerIsoEnd_apply (f : Stabilizer.submonoid M x) : +theorem stabilizerIsoEnd_apply (f : stabilizerSubmonoid M x) : (stabilizerIsoEnd M x) f = f := rfl #align category_theory.action_category.stabilizer_iso_End_apply CategoryTheory.ActionCategory.stabilizerIsoEnd_apply diff --git a/Mathlib/CategoryTheory/Adhesive.lean b/Mathlib/CategoryTheory/Adhesive.lean index 3594818727808..4546551b9dc5f 100644 --- a/Mathlib/CategoryTheory/Adhesive.lean +++ b/Mathlib/CategoryTheory/Adhesive.lean @@ -26,12 +26,8 @@ import Mathlib.CategoryTheory.Limits.Shapes.KernelPair monomorphisms are stable under pushouts. - `CategoryTheory.Adhesive.toRegularMonoCategory`: Monomorphisms in adhesive categories are regular (this implies that adhesive categories are balanced). - -## TODO - -Show that the following are adhesive: -- functor categories into adhesive categories -- the categories of sheaves over a site +- `CategoryTHeory.adhesive_functor`: The category `C ⥤ D` is adhesive if `D` + has all pullbacks and all pushouts and is adhesive ## References - https://ncatlab.org/nlab/show/adhesive+category @@ -273,4 +269,71 @@ noncomputable instance (priority := 100) Adhesive.toRegularMonoCategory [Adhesiv example [Adhesive C] : Balanced C := inferInstance +section functor + +universe v'' u'' + +variable {D : Type u''} [Category.{v''} D] + +instance adhesive_functor [Adhesive C] [HasPullbacks C] [HasPushouts C] : + Adhesive (D ⥤ C) := by + constructor + intros W X Y Z f g h i hf H + rw [IsPushout.isVanKampen_iff] + apply isVanKampenColimit_of_evaluation + intro x + refine (IsVanKampenColimit.precompose_isIso_iff (diagramIsoSpan _).inv).mp ?_ + refine IsVanKampenColimit.of_iso ?_ (PushoutCocone.isoMk _).symm + refine (IsPushout.isVanKampen_iff (H.map ((evaluation _ _).obj x))).mp ?_ + apply Adhesive.van_kampen + +theorem adhesive_of_preserves_and_reflects (F : C ⥤ D) [Adhesive D] + [H₁ : ∀ {X Y S : C} (f : X ⟶ S) (g : Y ⟶ S) [Mono f], HasPullback f g] + [H₂ : ∀ {X Y S : C} (f : S ⟶ X) (g : S ⟶ Y) [Mono f], HasPushout f g] + [PreservesLimitsOfShape WalkingCospan F] + [ReflectsLimitsOfShape WalkingCospan F] + [PreservesColimitsOfShape WalkingSpan F] + [ReflectsColimitsOfShape WalkingSpan F] : + Adhesive C := by + apply Adhesive.mk (hasPullback_of_mono_left := H₁) (hasPushout_of_mono_left := H₂) + intros W X Y Z f g h i hf H + rw [IsPushout.isVanKampen_iff] + refine IsVanKampenColimit.of_mapCocone F ?_ + refine (IsVanKampenColimit.precompose_isIso_iff (diagramIsoSpan _).inv).mp ?_ + refine IsVanKampenColimit.of_iso ?_ (PushoutCocone.isoMk _).symm + refine (IsPushout.isVanKampen_iff (H.map F)).mp ?_ + apply Adhesive.van_kampen + +theorem adhesive_of_preserves_and_reflects_isomorphism (F : C ⥤ D) + [Adhesive D] [HasPullbacks C] [HasPushouts C] + [PreservesLimitsOfShape WalkingCospan F] + [PreservesColimitsOfShape WalkingSpan F] + [ReflectsIsomorphisms F] : + Adhesive C := by + haveI : ReflectsLimitsOfShape WalkingCospan F := + reflectsLimitsOfShapeOfReflectsIsomorphisms + haveI : ReflectsColimitsOfShape WalkingSpan F := + reflectsColimitsOfShapeOfReflectsIsomorphisms + exact adhesive_of_preserves_and_reflects F + +theorem adhesive_of_reflective [HasPullbacks D] [Adhesive C] [HasPullbacks C] [HasPushouts C] + [H₂ : ∀ {X Y S : D} (f : S ⟶ X) (g : S ⟶ Y) [Mono f], HasPushout f g] + {Gl : C ⥤ D} {Gr : D ⥤ C} (adj : Gl ⊣ Gr) [Full Gr] [Faithful Gr] + [PreservesLimitsOfShape WalkingCospan Gl] : + Adhesive D := by + have := adj.leftAdjointPreservesColimits + have := adj.rightAdjointPreservesLimits + apply Adhesive.mk (hasPushout_of_mono_left := H₂) + intro W X Y Z f g h i _ H + have := Adhesive.van_kampen (IsPushout.of_hasPushout (Gr.map f) (Gr.map g)) + rw [IsPushout.isVanKampen_iff] at this ⊢ + refine (IsVanKampenColimit.precompose_isIso_iff + (isoWhiskerLeft _ (asIso adj.counit) ≪≫ Functor.rightUnitor _).hom).mp ?_ + refine ((this.precompose_isIso (spanCompIso _ _ _).hom).map_reflective adj).of_iso + (IsColimit.uniqueUpToIso ?_ ?_) + · exact isColimitOfPreserves Gl ((IsColimit.precomposeHomEquiv _ _).symm $ pushoutIsPushout _ _) + · exact (IsColimit.precomposeHomEquiv _ _).symm H.isColimit + +end functor + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Category/Basic.lean b/Mathlib/CategoryTheory/Category/Basic.lean index 56697c5fb4c4d..af3eb3731f038 100644 --- a/Mathlib/CategoryTheory/Category/Basic.lean +++ b/Mathlib/CategoryTheory/Category/Basic.lean @@ -117,6 +117,7 @@ use in auto-params. macro (name := aesop_cat) "aesop_cat" c:Aesop.tactic_clause* : tactic => `(tactic| aesop $c* (options := { introsTransparency? := some .default, terminal := true }) + (simp_options := { decide := true }) (rule_sets [$(Lean.mkIdent `CategoryTheory):ident])) /-- diff --git a/Mathlib/CategoryTheory/Comma.lean b/Mathlib/CategoryTheory/Comma.lean index 489d3e03f7277..17f9ed66f228f 100644 --- a/Mathlib/CategoryTheory/Comma.lean +++ b/Mathlib/CategoryTheory/Comma.lean @@ -167,9 +167,8 @@ theorem eqToHom_left (X Y : Comma L R) (H : X = Y) : #align category_theory.comma.eq_to_hom_left CategoryTheory.Comma.eqToHom_left @[simp] -theorem eqToHom_right (X Y : Comma L R) (H : X = Y) : CommaMorphism.right (eqToHom H) = eqToHom (by - cases H - rfl) := by +theorem eqToHom_right (X Y : Comma L R) (H : X = Y) : + CommaMorphism.right (eqToHom H) = eqToHom (by cases H; rfl) := by cases H rfl #align category_theory.comma.eq_to_hom_right CategoryTheory.Comma.eqToHom_right diff --git a/Mathlib/CategoryTheory/ComposableArrows.lean b/Mathlib/CategoryTheory/ComposableArrows.lean new file mode 100644 index 0000000000000..4f6502d31d172 --- /dev/null +++ b/Mathlib/CategoryTheory/ComposableArrows.lean @@ -0,0 +1,441 @@ +/- +Copyright (c) 2023 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Category.Preorder +import Mathlib.CategoryTheory.EqToHom +import Mathlib.CategoryTheory.Functor.Const +import Mathlib.Tactic.FinCases +import Mathlib.Tactic.Linarith + +/-! +# Composable arrows + +If `C` is a category, the type of `n`-simplices in the nerve of `C` identifies +to the type of functors `Fin (n + 1) ⥤ C`, which can be thought as families of `n` composable +arrows in `C`. In this file, we introduce and study this category `ComposableArrows C n` +of `n` composable arrows in `C`. + +If `F : ComposableArrows C n`, we define `F.left` as the leftmost object, `F.right` as the +rightmost object, and `F.hom : F.left ⟶ F.right` is the canonical map. + +The most significant definition in this file is the constructor +`F.precomp f : ComposableArrows C (n + 1)` for `F : ComposableArrows C n` and `f : X ⟶ F.left`: +"it shifts `F` towards the right and inserts `f` on the left". This `precomp` has +good definitional properties. + +In the namespace `CategoryTheory.ComposableArrows`, we provide constructors +like `mk₁ f`, `mk₂ f g`, `mk₃ f g h` for `ComposableArrows C n` for small `n`. + +TODO (@joelriou): +* define various constructors for objects, morphisms, isomorphisms in `ComposableArrows C n` +* redefine `Arrow C` as `ComposableArrow C 1`? +* construct some elements in `ComposableArrows m (Fin (n + 1))` for small `n` +the precomposition with which shall induce funtors +`ComposableArrows C n ⥤ ComposableArrows C m` which correspond to simplicial operations +(specifically faces) with good definitional properties (this might be necessary for +up to `n = 7` in order to formalize spectral sequences following Verdier) + +-/ + +namespace CategoryTheory + +open Category + +variable (C : Type*) [Category C] + +/-- `ComposableArrows C n` is the type of functors `Fin (n + 1) ⥤ C`. -/ +abbrev ComposableArrows (n : ℕ) := Fin (n + 1) ⥤ C + +namespace ComposableArrows + +variable {C} {n m : ℕ} +variable (F G : ComposableArrows C n) + +/-- The `i`th object (with `i : ℕ` such that `i ≤ n`) of `F : ComposableArrows C n`. -/ +@[simp] +abbrev obj' (i : ℕ) (hi : i ≤ n := by linarith) : C := F.obj ⟨i, by linarith⟩ + +/-- The map `F.obj' i ⟶ F.obj' j` when `F : ComposableArrows C n`, and `i` and `j` +are natural numbers such that `i ≤ j ≤ n`. -/ +@[simp] +abbrev map' (i j : ℕ) (hij : i ≤ j := by linarith) (hjn : j ≤ n := by linarith) : + F.obj ⟨i, by linarith⟩ ⟶ F.obj ⟨j, by linarith⟩ := F.map (homOfLE (by + simp only [Fin.mk_le_mk] + linarith)) + +lemma map'_self (i : ℕ) (hi : i ≤ n := by linarith) : + F.map' i i = 𝟙 _ := F.map_id _ + +lemma map'_comp (i j k : ℕ) (hij : i ≤ j := by linarith) + (hjk : j ≤ k := by linarith) (hk : k ≤ n := by linarith) : + F.map' i k = F.map' i j ≫ F.map' j k := + F.map_comp _ _ + +/-- The leftmost object of `F : ComposableArrows C n`. -/ +abbrev left := obj' F 0 + +/-- The rightmost object of `F : ComposableArrows C n`. -/ +abbrev right := obj' F n + +/-- The canonical map `F.left ⟶ F.right` for `F : ComposableArrows C n`. -/ +abbrev hom : F.left ⟶ F.right := map' F 0 n + +variable {F G} + +/-- The map `F.obj' i ⟶ G.obj' i` induced on `i`th objects by a morphism `F ⟶ G` +in `ComposableArrows C n` when `i` is a natural number such that `i ≤ n`. -/ +@[simp] +abbrev app' (φ : F ⟶ G) (i : ℕ) (hi : i ≤ n := by linarith) : + F.obj' i ⟶ G.obj' i := φ.app _ + +/-- Constructor for `ComposableArrows C 0`. -/ +@[simps!] +def mk₀ (X : C) : ComposableArrows C 0 := (Functor.const (Fin 1)).obj X + +namespace Mk₁ + +variable (X₀ X₁ : C) + +/-- The map which sends `0 : Fin 2` to `X₀` and `1` to `X₁`. -/ +@[simp] +def obj : Fin 2 → C + | ⟨0, _⟩ => X₀ + | ⟨1, _⟩ => X₁ + +variable {X₀ X₁} (f : X₀ ⟶ X₁) + +/-- The obvious map `obj X₀ X₁ i ⟶ obj X₀ X₁ j` whenever `i j : Fin 2` satisfy `i ≤ j`. -/ +@[simp] +def map : ∀ (i j : Fin 2) (_ : i ≤ j), obj X₀ X₁ i ⟶ obj X₀ X₁ j + | ⟨0, _⟩, ⟨0, _⟩, _ => 𝟙 _ + | ⟨0, _⟩, ⟨1, _⟩, _ => f + | ⟨1, _⟩, ⟨1, _⟩, _ => 𝟙 _ + +lemma map_id (i : Fin 2) : map f i i (by simp) = 𝟙 _ := + match i with + | 0 => rfl + | 1 => rfl + +lemma map_comp {i j k : Fin 2} (hij : i ≤ j) (hjk : j ≤ k) : + map f i k (hij.trans hjk) = map f i j hij ≫ map f j k hjk := + match i with + | 0 => + match j with + | 0 => by rw [map_id, id_comp] + | 1 => by + obtain rfl : k = 1 := k.eq_one_of_neq_zero (by rintro rfl; simp at hjk) + rw [map_id, comp_id] + | 1 => by + obtain rfl := j.eq_one_of_neq_zero (by rintro rfl; simp at hij) + obtain rfl := k.eq_one_of_neq_zero (by rintro rfl; simp at hjk) + rw [map_id, id_comp] + +end Mk₁ + +/-- Constructor for `ComposableArrows C 1`. -/ +@[simps] +def mk₁ {X₀ X₁ : C} (f : X₀ ⟶ X₁) : ComposableArrows C 1 where + obj := Mk₁.obj X₀ X₁ + map g := Mk₁.map f _ _ (leOfHom g) + map_id := Mk₁.map_id f + map_comp g g' := Mk₁.map_comp f (leOfHom g) (leOfHom g') + +/-- Constructor for morphisms `F ⟶ G` in `ComposableArrows C n` which takes as inputs +a family of morphisms `F.obj i ⟶ G.obj i` and the naturality condition only for the +maps in `Fin (n + 1)` given by inequalities of the form `i ≤ i + 1`. -/ +@[simps] +def homMk {F G : ComposableArrows C n} (app : ∀ i, F.obj i ⟶ G.obj i) + (w : ∀ (i : ℕ) (hi : i < n), F.map' i (i + 1) ≫ app _ = app _ ≫ G.map' i (i + 1)) : + F ⟶ G where + app := app + naturality := by + suffices ∀ (k i j : ℕ) (hj : i + k = j) (hj' : j ≤ n), + F.map' i j ≫ app _ = app _ ≫ G.map' i j by + rintro ⟨i, hi⟩ ⟨j, hj⟩ hij + have hij' := leOfHom hij + simp only [Fin.mk_le_mk] at hij' + obtain ⟨k, hk⟩ := Nat.le.dest hij' + exact this k i j hk (by linarith) + intro k + induction' k with k hk + · intro i j hj hj' + simp only [Nat.zero_eq, add_zero] at hj + obtain rfl := hj + rw [F.map'_self i, G.map'_self i, id_comp, comp_id] + · intro i j hj hj' + rw [Nat.succ_eq_add_one, ← add_assoc] at hj + subst hj + rw [F.map'_comp i (i + k) (i + k + 1), G.map'_comp i (i + k) (i + k + 1), assoc, + w (i + k) (by linarith), reassoc_of% (hk i (i + k) rfl (by linarith))] + +/-- Constructor for isomorphisms `F ≅ G` in `ComposableArrows C n` which takes as inputs +a family of isomorphisms `F.obj i ≅ G.obj i` and the naturality condition only for the +maps in `Fin (n + 1)` given by inequalities of the form `i ≤ i + 1`. -/ +@[simps] +def isoMk {F G : ComposableArrows C n} (app : ∀ i, F.obj i ≅ G.obj i) + (w : ∀ (i : ℕ) (hi : i < n), + F.map' i (i + 1) ≫ (app _).hom = (app _).hom ≫ G.map' i (i + 1)) : + F ≅ G where + hom := homMk (fun i => (app i).hom) w + inv := homMk (fun i => (app i).inv) (fun i hi => by + dsimp only + rw [← cancel_epi ((app _).hom), ← reassoc_of% (w i hi), Iso.hom_inv_id, comp_id, + Iso.hom_inv_id_assoc]) + +lemma ext {F G : ComposableArrows C n} (h : ∀ i, F.obj i = G.obj i) + (w : ∀ (i : ℕ) (hi : i < n), F.map' i (i + 1) = + eqToHom (h _) ≫ G.map' i (i + 1) ≫ eqToHom (h _).symm) : F = G := + Functor.ext_of_iso + (isoMk (fun i => eqToIso (h i)) (fun i hi => by simp [w i hi])) h (fun i => rfl) + +/-- Constructor for morphisms in `ComposableArrows C 0`. -/ +@[simps!] +def homMk₀ {F G : ComposableArrows C 0} (f : F.obj' 0 ⟶ G.obj' 0) : F ⟶ G := + homMk (fun i => match i with + | ⟨0, _⟩ => f) (fun i hi => by simp at hi) + +@[ext] +lemma hom_ext₀ {F G : ComposableArrows C 0} {φ φ' : F ⟶ G} + (h : app' φ 0 = app' φ' 0) : + φ = φ' := by + ext i + fin_cases i + exact h + +/-- Constructor for isomorphisms in `ComposableArrows C 0`. -/ +@[simps!] +def isoMk₀ {F G : ComposableArrows C 0} (e : F.obj' 0 ≅ G.obj' 0) : F ≅ G where + hom := homMk₀ e.hom + inv := homMk₀ e.inv + +lemma ext₀ {F G : ComposableArrows C 0} (h : F.obj' 0 = G.obj 0) : F = G := + ext (fun i => match i with + | ⟨0, _⟩ => h) (fun i hi => by simp at hi) + +lemma mk₀_surjective (F : ComposableArrows C 0) : ∃ (X : C), F = mk₀ X := + ⟨F.obj' 0, ext₀ rfl⟩ + +/-- Constructor for morphisms in `ComposableArrows C 1`. -/ +@[simps!] +def homMk₁ {F G : ComposableArrows C 1} + (left : F.obj' 0 ⟶ G.obj' 0) (right : F.obj' 1 ⟶ G.obj' 1) + (w : F.map' 0 1 ≫ right = left ≫ G.map' 0 1 := by aesop_cat) : + F ⟶ G := + homMk (fun i => match i with + | ⟨0, _⟩ => left + | ⟨1, _⟩ => right) (by + intro i hi + obtain rfl : i = 0 := by simpa using hi + exact w) + +@[ext] +lemma hom_ext₁ {F G : ComposableArrows C 1} {φ φ' : F ⟶ G} + (h₀ : app' φ 0 = app' φ' 0) (h₁ : app' φ 1 = app' φ' 1) : + φ = φ' := by + ext i + match i with + | 0 => exact h₀ + | 1 => exact h₁ + +/-- Constructor for isomorphisms in `ComposableArrows C 1`. -/ +@[simps!] +def isoMk₁ {F G : ComposableArrows C 1} + (left : F.obj' 0 ≅ G.obj' 0) (right : F.obj' 1 ≅ G.obj' 1) + (w : F.map' 0 1 ≫ right.hom = left.hom ≫ G.map' 0 1 := by aesop_cat) : + F ≅ G where + hom := homMk₁ left.hom right.hom w + inv := homMk₁ left.inv right.inv (by + rw [← cancel_mono right.hom, assoc, assoc, w, right.inv_hom_id, left.inv_hom_id_assoc] + apply comp_id) + +lemma map'_eq_hom₁ (F : ComposableArrows C 1) : F.map' 0 1 = F.hom := rfl + +lemma ext₁ {F G : ComposableArrows C 1} + (left : F.left = G.left) (right : F.right = G.right) + (w : F.hom = eqToHom left ≫ G.hom ≫ eqToHom right.symm) : F = G := + Functor.ext_of_iso (isoMk₁ (eqToIso left) (eqToIso right) (by simp [map'_eq_hom₁, w])) + (fun i => by fin_cases i <;> assumption) + (fun i => by fin_cases i <;> rfl) + +lemma mk₁_surjective (X : ComposableArrows C 1) : ∃ (X₀ X₁ : C) (f : X₀ ⟶ X₁), X = mk₁ f := + ⟨_, _, X.map' 0 1, ext₁ rfl rfl (by simp)⟩ + +variable (F) + +namespace Precomp + +variable (X : C) + +/-- The map `Fin (n + 1 + 1) → C` which "shifts" `F.obj'` to the right and inserts `X` in +the zeroth position. -/ +def obj : Fin (n + 1 + 1) → C + | ⟨0, _⟩ => X + | ⟨i + 1, hi⟩ => F.obj' i + +@[simp] +lemma obj_zero : obj F X 0 = X := rfl + +@[simp] +lemma obj_one : obj F X 1 = F.obj' 0 := rfl + +@[simp] +lemma obj_succ (i : ℕ) (hi : i + 1 < n + 1 + 1) : obj F X ⟨i + 1, hi⟩ = F.obj' i := rfl + +variable {X} (f : X ⟶ F.left) + +/-- Auxiliary definition for the action on maps of the functor `F.precomp f`. +It sends `0 ≤ 1` to `f` and `i + 1 ≤ j + 1` to `F.map' i j`. -/ +def map : ∀ (i j : Fin (n + 1 + 1)) (_ : i ≤ j), obj F X i ⟶ obj F X j + | ⟨0, _⟩, ⟨0, _⟩, _ => 𝟙 X + | ⟨0, _⟩, ⟨1, _⟩, _ => f + | ⟨0, _⟩, ⟨j + 2, hj⟩, _ => f ≫ F.map' 0 (j + 1) + | ⟨i + 1, hi⟩, ⟨j + 1, hj⟩, hij => F.map' i j (by simpa using hij) + +@[simp] +lemma map_zero_zero : map F f 0 0 (by simp) = 𝟙 X := rfl + +@[simp] +lemma map_one_one : map F f 1 1 (by simp) = F.map (𝟙 _) := rfl + +@[simp] +lemma map_zero_one : map F f 0 1 (by simp) = f := rfl + +@[simp] +lemma map_zero_one' : map F f 0 ⟨0 + 1, by simp⟩ (by simp) = f := rfl + +@[simp] +lemma map_zero_succ_succ (j : ℕ) (hj : j + 2 < n + 1 + 1) : + map F f 0 ⟨j + 2, hj⟩ (by simp) = f ≫ F.map' 0 (j+1) := rfl + +@[simp] +lemma map_succ_succ (i j : ℕ) (hi : i + 1 < n + 1 + 1) (hj : j + 1 < n + 1 + 1) + (hij : i + 1 ≤ j + 1) : + map F f ⟨i + 1, hi⟩ ⟨j + 1, hj⟩ hij = F.map' i j := rfl + +@[simp] +lemma map_one_succ (j : ℕ) (hj : j + 1 < n + 1 + 1) : + map F f 1 ⟨j + 1, hj⟩ (by simp [Fin.le_def]) = F.map' 0 j := rfl + +lemma map_id (i : Fin (n + 1 + 1)) : map F f i i (by simp) = 𝟙 _ := by + obtain ⟨i, hi⟩ := i + cases i + · rfl + · apply F.map_id + +lemma map_comp {i j k : Fin (n + 1 + 1)} (hij : i ≤ j) (hjk : j ≤ k) : + map F f i k (hij.trans hjk) = map F f i j hij ≫ map F f j k hjk := by + obtain ⟨i, hi⟩ := i + obtain ⟨j, hj⟩ := j + obtain ⟨k, hk⟩ := k + cases i + · obtain _ | _ | j := j + · dsimp + rw [id_comp] + · obtain _ | _ | k := k + · simp at hjk + · dsimp + rw [F.map_id, comp_id] + · rfl + · obtain _ | _ | k := k + · simp [Fin.ext_iff] at hjk + · simp [Fin.le_def, Nat.succ_eq_add_one] at hjk + · dsimp + rw [assoc, ← F.map_comp, homOfLE_comp] + · obtain _ | j := j + · simp [Fin.ext_iff] at hij + · obtain _ | k := k + · simp [Fin.ext_iff] at hjk + · dsimp + rw [← F.map_comp, homOfLE_comp] + +end Precomp + +/-- "Precomposition" of `F : ComposableArrows C n` by a morphism `f : X ⟶ F.left`. -/ +@[simps] +def precomp {X : C} (f : X ⟶ F.left) : ComposableArrows C (n + 1) where + obj := Precomp.obj F X + map g := Precomp.map F f _ _ (leOfHom g) + map_id := Precomp.map_id F f + map_comp g g' := (Precomp.map_comp F f (leOfHom g) (leOfHom g')) + +/-- Constructor for `ComposableArrows C 2`. -/ +@[simp] +def mk₂ {X₀ X₁ X₂ : C} (f : X₀ ⟶ X₁) (g : X₁ ⟶ X₂) : ComposableArrows C 2 := + (mk₁ g).precomp f + +/-- Constructor for `ComposableArrows C 3`. -/ +@[simp] +def mk₃ {X₀ X₁ X₂ X₃ : C} (f : X₀ ⟶ X₁) (g : X₁ ⟶ X₂) (h : X₂ ⟶ X₃) : ComposableArrows C 3 := + (mk₂ g h).precomp f + +section + +variable {X₀ X₁ X₂ X₃ X₄ : C} (f : X₀ ⟶ X₁) (g : X₁ ⟶ X₂) (h : X₂ ⟶ X₃) (i : X₃ ⟶ X₄) + +/-! These examples are meant to test the good definitional properties of `precomp`, +and that `dsimp` can see through. -/ + +example : map' (mk₂ f g) 0 1 = f := by dsimp +example : map' (mk₂ f g) 1 2 = g := by dsimp +example : map' (mk₂ f g) 0 2 = f ≫ g := by dsimp +example : (mk₂ f g).hom = f ≫ g := by dsimp +example : map' (mk₂ f g) 0 0 = 𝟙 _ := by dsimp +example : map' (mk₂ f g) 1 1 = 𝟙 _ := by dsimp +example : map' (mk₂ f g) 2 2 = 𝟙 _ := by dsimp + +example : map' (mk₃ f g h) 0 1 = f := by dsimp +example : map' (mk₃ f g h) 1 2 = g := by dsimp +example : map' (mk₃ f g h) 2 3 = h := by dsimp +example : map' (mk₃ f g h) 0 3 = f ≫ g ≫ h := by dsimp +example : (mk₃ f g h).hom = f ≫ g ≫ h := by dsimp +example : map' (mk₃ f g h) 0 2 = f ≫ g := by dsimp +example : map' (mk₃ f g h) 1 3 = g ≫ h := by dsimp + +end + +/-- The map `ComposableArrows C m → ComposableArrows C n` obtained by precomposition with +a functor `Fin (n + 1) ⥤ Fin (m + 1)`. -/ +@[simps!] +def whiskerLeft (F : ComposableArrows C m) (Φ : Fin (n + 1) ⥤ Fin (m + 1)) : + ComposableArrows C n := Φ ⋙ F + +/-- The functor `ComposableArrows C m ⥤ ComposableArrows C n` obtained by precomposition with +a functor `Fin (n + 1) ⥤ Fin (m + 1)`. -/ +@[simps!] +def whiskerLeftFunctor (Φ : Fin (n + 1) ⥤ Fin (m + 1)) : + ComposableArrows C m ⥤ ComposableArrows C n where + obj F := F.whiskerLeft Φ + map f := CategoryTheory.whiskerLeft Φ f + +/-- The functor `Fin n ⥤ Fin (n + 1)` which sends `i` to `i.succ`. -/ +@[simps] +def _root_.Fin.succFunctor (n : ℕ) : Fin n ⥤ Fin (n + 1) where + obj i := i.succ + map {i j} hij := homOfLE (Fin.succ_le_succ_iff.2 (leOfHom hij)) + +/-- The functor `ComposableArrows C (n + 1) ⥤ ComposableArrows C n` which forgets +the first arrow. -/ +@[simps!] +def δ₀Functor : ComposableArrows C (n + 1) ⥤ ComposableArrows C n := + whiskerLeftFunctor (Fin.succFunctor (n + 1)) + +/-- The `ComposableArrows C n` obtained by forgetting the first arrow. -/ +abbrev δ₀ (F : ComposableArrows C (n + 1)) := δ₀Functor.obj F + +@[simp] +lemma precomp_δ₀ {X : C} (f : X ⟶ F.left) : (F.precomp f).δ₀ = F := rfl + +end ComposableArrows + +variable {C} + +/-- The functor `ComposableArrows C n ⥤ ComposableArrows D n` obtained by postcomposition +with a functor `C ⥤ D`. -/ +@[simps!] +def Functor.mapComposableArrows {D : Type*} [Category D] (G : C ⥤ D) (n : ℕ) : + ComposableArrows C n ⥤ ComposableArrows D n := + (whiskeringRight _ _ _).obj G + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Conj.lean b/Mathlib/CategoryTheory/Conj.lean index a4e5be93c0e68..6974e7a01a0a1 100644 --- a/Mathlib/CategoryTheory/Conj.lean +++ b/Mathlib/CategoryTheory/Conj.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Algebra.Hom.Equiv.Units.Basic +import Mathlib.Algebra.Group.Units.Equiv import Mathlib.CategoryTheory.Endomorphism #align_import category_theory.conj from "leanprover-community/mathlib"@"32253a1a1071173b33dc7d6a218cf722c6feb514" diff --git a/Mathlib/CategoryTheory/DifferentialObject.lean b/Mathlib/CategoryTheory/DifferentialObject.lean index 440bd01778d6e..1e2da407c290b 100644 --- a/Mathlib/CategoryTheory/DifferentialObject.lean +++ b/Mathlib/CategoryTheory/DifferentialObject.lean @@ -200,8 +200,8 @@ def mapDifferentialObject (F : C ⥤ D) slice_lhs 2 3 => rw [← Functor.comp_map F (shiftFunctor D (1 : S)), ← η.naturality f.f] slice_lhs 1 2 => rw [Functor.comp_map, ← F.map_comp, f.comm, F.map_comp] rw [Category.assoc] } - map_id := by intros; ext; simp - map_comp := by intros; ext; simp + map_id := by intros; ext; simp [autoParam] + map_comp := by intros; ext; simp [autoParam] #align category_theory.functor.map_differential_object CategoryTheory.Functor.mapDifferentialObject end Functor @@ -219,12 +219,10 @@ variable [(shiftFunctor C (1 : S)).PreservesZeroMorphisms] open scoped ZeroObject -instance hasZeroObject : HasZeroObject (DifferentialObject S C) := by - -- Porting note(https://github.com/leanprover-community/mathlib4/issues/4998): added `aesop_cat` - -- Porting note: added `simp only [eq_iff_true_of_subsingleton]` - refine' ⟨⟨⟨0, 0, by aesop_cat⟩, fun X => ⟨⟨⟨⟨0, by aesop_cat⟩⟩, fun f => _⟩⟩, - fun X => ⟨⟨⟨⟨0, by aesop_cat⟩⟩, fun f => _⟩⟩⟩⟩ <;> ext <;> - simp only [eq_iff_true_of_subsingleton] +instance hasZeroObject : HasZeroObject (DifferentialObject S C) where + zero := ⟨{ obj := 0, d := 0 }, + { unique_to := fun X => ⟨⟨⟨{ f := 0 }⟩, fun f => by ext⟩⟩, + unique_from := fun X => ⟨⟨⟨{ f := 0 }⟩, fun f => by ext⟩⟩ }⟩ #align category_theory.differential_object.has_zero_object CategoryTheory.DifferentialObject.hasZeroObject end DifferentialObject diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index 1b2b79c8817d4..e75472b543473 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -273,7 +273,7 @@ theorem costructuredArrow_yoneda_equivalence_naturality {F₁ F₂ : Cᵒᵖ ⥤ congr ext _ f simpa using congr_fun (α.naturality f.op).symm (unop X).snd - · aesop + · simp [autoParam] #align category_theory.category_of_elements.costructured_arrow_yoneda_equivalence_naturality CategoryTheory.CategoryOfElements.costructuredArrow_yoneda_equivalence_naturality end CategoryOfElements diff --git a/Mathlib/CategoryTheory/Endomorphism.lean b/Mathlib/CategoryTheory/Endomorphism.lean index 5e769eaa7691e..fb1b90f8dbc23 100644 --- a/Mathlib/CategoryTheory/Endomorphism.lean +++ b/Mathlib/CategoryTheory/Endomorphism.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Scott Morrison, Simon Hudon -/ -import Mathlib.Algebra.Hom.Equiv.Basic +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.Units import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.Opposites diff --git a/Mathlib/CategoryTheory/EqToHom.lean b/Mathlib/CategoryTheory/EqToHom.lean index 9b69f0040493e..30258f19525ea 100644 --- a/Mathlib/CategoryTheory/EqToHom.lean +++ b/Mathlib/CategoryTheory/EqToHom.lean @@ -214,6 +214,12 @@ theorem ext {F G : C ⥤ D} (h_obj : ∀ X, F.obj X = G.obj X) simpa using h_map X Y f #align category_theory.functor.ext CategoryTheory.Functor.ext +lemma ext_of_iso {F G : C ⥤ D} (e : F ≅ G) (hobj : ∀ X, F.obj X = G.obj X) + (happ : ∀ X, e.hom.app X = eqToHom (hobj X)) : F = G := + Functor.ext hobj (fun X Y f => by + rw [← cancel_mono (e.hom.app Y), e.hom.naturality f, happ, happ, Category.assoc, + Category.assoc, eqToHom_trans, eqToHom_refl, Category.comp_id]) + /-- Two morphisms are conjugate via eqToHom if and only if they are heterogeneously equal. -/ theorem conj_eqToHom_iff_heq {W X Y Z : C} (f : W ⟶ X) (g : Y ⟶ Z) (h : W = Y) (h' : X = Z) : f = eqToHom h ≫ g ≫ eqToHom h'.symm ↔ HEq f g := by diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index f465e2f4b0e54..078771d545d26 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -8,6 +8,8 @@ import Mathlib.CategoryTheory.Limits.Shapes.StrictInitial import Mathlib.CategoryTheory.Limits.Shapes.Types import Mathlib.Topology.Category.TopCat.Limits.Pullbacks import Mathlib.CategoryTheory.Limits.FunctorCategory +import Mathlib.CategoryTheory.Limits.Constructions.FiniteProductsOfBinaryProducts +import Mathlib.CategoryTheory.Limits.VanKampen #align_import category_theory.extensive from "leanprover-community/mathlib"@"178a32653e369dce2da68dc6b2694e385d484ef1" @@ -16,9 +18,6 @@ import Mathlib.CategoryTheory.Limits.FunctorCategory # Extensive categories ## Main definitions -- `CategoryTheory.IsVanKampenColimit`: A (colimit) cocone over a diagram `F : J ⥤ C` is van - Kampen if for every cocone `c'` over the pullback of the diagram `F' : J ⥤ C'`, - `c'` is colimiting iff `c'` is the pullback of `c`. - `CategoryTheory.FinitaryExtensive`: A category is (finitary) extensive if it has finite coproducts, and binary coproducts are van Kampen. @@ -30,11 +29,15 @@ import Mathlib.CategoryTheory.Limits.FunctorCategory - `CategoryTheory.BinaryCofan.isPullback_initial_to_of_isVanKampen`: In extensive categories, sums are disjoint, i.e. the pullback of `X ⟶ X ⨿ Y` and `Y ⟶ X ⨿ Y` is the initial object. - `CategoryTheory.types.finitaryExtensive`: The category of types is extensive. - +- `CategoryTheory.FinitaryExtensive_TopCat`: + The category `Top` is extensive. +- `CategoryTheory.FinitaryExtensive_functor`: The category `C ⥤ D` is extensive if `D` + has all pullbacks and is extensive. +- `CategoryTheory.FinitaryExtensive.isVanKampen_finiteCoproducts`: Finite coproducts in a + finitary extensive category are van Kampen. ## TODO Show that the following are finitary extensive: -- the categories of sheaves over a site - `Scheme` - `AffineScheme` (`CommRingᵒᵖ`) @@ -49,89 +52,52 @@ open CategoryTheory.Limits namespace CategoryTheory -universe v' u' v u +universe v' u' v u v'' u'' variable {J : Type v'} [Category.{u'} J] {C : Type u} [Category.{v} C] - -/-- A natural transformation is equifibered if every commutative square of the following form is -a pullback. -``` -F(X) → F(Y) - ↓ ↓ -G(X) → G(Y) -``` --/ -def NatTrans.Equifibered {F G : J ⥤ C} (α : F ⟶ G) : Prop := - ∀ ⦃i j : J⦄ (f : i ⟶ j), IsPullback (F.map f) (α.app i) (α.app j) (G.map f) -#align category_theory.nat_trans.equifibered CategoryTheory.NatTrans.Equifibered - -theorem NatTrans.equifibered_of_isIso {F G : J ⥤ C} (α : F ⟶ G) [IsIso α] : Equifibered α := - fun _ _ f => IsPullback.of_vert_isIso ⟨NatTrans.naturality _ f⟩ -#align category_theory.nat_trans.equifibered_of_is_iso CategoryTheory.NatTrans.equifibered_of_isIso - -theorem NatTrans.Equifibered.comp {F G H : J ⥤ C} {α : F ⟶ G} {β : G ⟶ H} (hα : Equifibered α) - (hβ : Equifibered β) : Equifibered (α ≫ β) := - fun _ _ f => (hα f).paste_vert (hβ f) -#align category_theory.nat_trans.equifibered.comp CategoryTheory.NatTrans.Equifibered.comp - -/-- A (colimit) cocone over a diagram `F : J ⥤ C` is universal if it is stable under pullbacks. -/ -def IsUniversalColimit {F : J ⥤ C} (c : Cocone F) : Prop := - ∀ ⦃F' : J ⥤ C⦄ (c' : Cocone F') (α : F' ⟶ F) (f : c'.pt ⟶ c.pt) - (_ : α ≫ c.ι = c'.ι ≫ (Functor.const J).map f) (_ : NatTrans.Equifibered α), - (∀ j : J, IsPullback (c'.ι.app j) (α.app j) f (c.ι.app j)) → Nonempty (IsColimit c') -#align category_theory.is_universal_colimit CategoryTheory.IsUniversalColimit - -/-- A (colimit) cocone over a diagram `F : J ⥤ C` is van Kampen if for every cocone `c'` over the -pullback of the diagram `F' : J ⥤ C'`, `c'` is colimiting iff `c'` is the pullback of `c`. - -TODO: Show that this is iff the functor `C ⥤ Catᵒᵖ` sending `x` to `C/x` preserves it. -TODO: Show that this is iff the inclusion functor `C ⥤ Span(C)` preserves it. --/ -def IsVanKampenColimit {F : J ⥤ C} (c : Cocone F) : Prop := - ∀ ⦃F' : J ⥤ C⦄ (c' : Cocone F') (α : F' ⟶ F) (f : c'.pt ⟶ c.pt) - (_ : α ≫ c.ι = c'.ι ≫ (Functor.const J).map f) (_ : NatTrans.Equifibered α), - Nonempty (IsColimit c') ↔ ∀ j : J, IsPullback (c'.ι.app j) (α.app j) f (c.ι.app j) -#align category_theory.is_van_kampen_colimit CategoryTheory.IsVanKampenColimit - -theorem IsVanKampenColimit.isUniversal {F : J ⥤ C} {c : Cocone F} (H : IsVanKampenColimit c) : - IsUniversalColimit c := - fun _ c' α f h hα => (H c' α f h hα).mpr -#align category_theory.is_van_kampen_colimit.is_universal CategoryTheory.IsVanKampenColimit.isUniversal - -/-- A van Kampen colimit is a colimit. -/ -noncomputable def IsVanKampenColimit.isColimit {F : J ⥤ C} {c : Cocone F} - (h : IsVanKampenColimit c) : IsColimit c := by - refine' ((h c (𝟙 F) (𝟙 c.pt : _) (by rw [Functor.map_id, Category.comp_id, Category.id_comp]) - (NatTrans.equifibered_of_isIso _)).mpr fun j => _).some - haveI : IsIso (𝟙 c.pt) := inferInstance - exact IsPullback.of_vert_isIso ⟨by erw [NatTrans.id_app, Category.comp_id, Category.id_comp]⟩ -#align category_theory.is_van_kampen_colimit.is_colimit CategoryTheory.IsVanKampenColimit.isColimit - -theorem IsInitial.isVanKampenColimit [HasStrictInitialObjects C] {X : C} (h : IsInitial X) : - IsVanKampenColimit (asEmptyCocone X) := by - intro F' c' α f hf hα - have : F' = Functor.empty C := by apply Functor.hext <;> rintro ⟨⟨⟩⟩ - subst this - haveI := h.isIso_to f - refine' ⟨by rintro _ ⟨⟨⟩⟩, - fun _ => ⟨IsColimit.ofIsoColimit h (Cocones.ext (asIso f).symm <| by rintro ⟨⟨⟩⟩)⟩⟩ -#align category_theory.is_initial.is_van_kampen_colimit CategoryTheory.IsInitial.isVanKampenColimit +variable {D : Type u''} [Category.{v''} D] section Extensive variable {X Y : C} -/-- A category is (finitary) extensive if it has finite coproducts, -and binary coproducts are van Kampen. +/-- A category has pullback of inclusions if it has all pullbacks along coproduct injections. -/ +class HasPullbacksOfInclusions (C : Type u) [Category.{v} C] [HasBinaryCoproducts C] : Prop where + [hasPullbackInl : ∀ {X Y Z : C} (f : Z ⟶ X ⨿ Y), HasPullback coprod.inl f] + +attribute [instance] HasPullbacksOfInclusions.hasPullbackInl + +/-- +A functor preserves pullback of inclusions if it preserves all pullbacks along coproduct injections. +-/ +class PreservesPullbacksOfInclusions {C : Type*} [Category C] {D : Type*} [Category D] + (F : C ⥤ D) [HasBinaryCoproducts C] where + [preservesPullbackInl : ∀ {X Y Z : C} (f : Z ⟶ X ⨿ Y), PreservesLimit (cospan coprod.inl f) F] + +attribute [instance] PreservesPullbacksOfInclusions.preservesPullbackInl + +/-- A category is (finitary) pre-extensive if it has finite coproducts, +and binary coproducts are universal. -/ +class FinitaryPreExtensive (C : Type u) [Category.{v} C] : Prop where + [hasFiniteCoproducts : HasFiniteCoproducts C] + [HasPullbacksOfInclusions : HasPullbacksOfInclusions C] + /-- In a finitary extensive category, all coproducts are van Kampen-/ + universal' : ∀ {X Y : C} (c : BinaryCofan X Y), IsColimit c → IsUniversalColimit c + +attribute [instance] FinitaryPreExtensive.hasFiniteCoproducts +attribute [instance] FinitaryPreExtensive.HasPullbacksOfInclusions -TODO: Show that this is iff all finite coproducts are van Kampen. -/ +/-- A category is (finitary) extensive if it has finite coproducts, +and binary coproducts are van Kampen. -/ class FinitaryExtensive (C : Type u) [Category.{v} C] : Prop where [hasFiniteCoproducts : HasFiniteCoproducts C] + [HasPullbacksOfInclusions : HasPullbacksOfInclusions C] /-- In a finitary extensive category, all coproducts are van Kampen-/ van_kampen' : ∀ {X Y : C} (c : BinaryCofan X Y), IsColimit c → IsVanKampenColimit c #align category_theory.finitary_extensive CategoryTheory.FinitaryExtensive attribute [instance] FinitaryExtensive.hasFiniteCoproducts +attribute [instance] FinitaryExtensive.HasPullbacksOfInclusions theorem FinitaryExtensive.vanKampen [FinitaryExtensive C] {F : Discrete WalkingPair ⥤ C} (c : Cocone F) (hc : IsColimit c) : IsVanKampenColimit c := by @@ -146,97 +112,63 @@ theorem FinitaryExtensive.vanKampen [FinitaryExtensive C] {F : Discrete WalkingP exact FinitaryExtensive.van_kampen' c hc #align category_theory.finitary_extensive.van_kampen CategoryTheory.FinitaryExtensive.vanKampen -theorem mapPair_equifibered {F F' : Discrete WalkingPair ⥤ C} (α : F ⟶ F') : - NatTrans.Equifibered α := by - rintro ⟨⟨⟩⟩ ⟨j⟩ ⟨⟨rfl : _ = j⟩⟩ - all_goals - dsimp; simp only [Discrete.functor_map_id] - exact IsPullback.of_horiz_isIso ⟨by simp only [Category.comp_id, Category.id_comp]⟩ -#align category_theory.map_pair_equifibered CategoryTheory.mapPair_equifibered - -theorem BinaryCofan.isVanKampen_iff (c : BinaryCofan X Y) : - IsVanKampenColimit c ↔ - ∀ {X' Y' : C} (c' : BinaryCofan X' Y') (αX : X' ⟶ X) (αY : Y' ⟶ Y) (f : c'.pt ⟶ c.pt) - (_ : αX ≫ c.inl = c'.inl ≫ f) (_ : αY ≫ c.inr = c'.inr ≫ f), - Nonempty (IsColimit c') ↔ IsPullback c'.inl αX f c.inl ∧ IsPullback c'.inr αY f c.inr := by - constructor - · introv H hαX hαY - rw [H c' (mapPair αX αY) f (by ext ⟨⟨⟩⟩ <;> dsimp <;> assumption) (mapPair_equifibered _)] - constructor - · intro H - exact ⟨H _, H _⟩ - · rintro H ⟨⟨⟩⟩ - exacts [H.1, H.2] - · introv H F' hα h - let X' := F'.obj ⟨WalkingPair.left⟩ - let Y' := F'.obj ⟨WalkingPair.right⟩ - have : F' = pair X' Y' := by - apply Functor.hext - · rintro ⟨⟨⟩⟩ <;> rfl - · rintro ⟨⟨⟩⟩ ⟨j⟩ ⟨⟨rfl : _ = j⟩⟩ <;> simp - clear_value X' Y' - subst this - change BinaryCofan X' Y' at c' - rw [H c' _ _ _ (NatTrans.congr_app hα ⟨WalkingPair.left⟩) - (NatTrans.congr_app hα ⟨WalkingPair.right⟩)] - constructor - · rintro H ⟨⟨⟩⟩ - exacts [H.1, H.2] - · intro H - exact ⟨H _, H _⟩ -#align category_theory.binary_cofan.is_van_kampen_iff CategoryTheory.BinaryCofan.isVanKampen_iff - -theorem BinaryCofan.isVanKampen_mk {X Y : C} (c : BinaryCofan X Y) - (cofans : ∀ X Y : C, BinaryCofan X Y) (colimits : ∀ X Y, IsColimit (cofans X Y)) - (cones : ∀ {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z), PullbackCone f g) - (limits : ∀ {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z), IsLimit (cones f g)) - (h₁ : ∀ {X' Y' : C} (αX : X' ⟶ X) (αY : Y' ⟶ Y) (f : (cofans X' Y').pt ⟶ c.pt) - (_ : αX ≫ c.inl = (cofans X' Y').inl ≫ f) (_ : αY ≫ c.inr = (cofans X' Y').inr ≫ f), - IsPullback (cofans X' Y').inl αX f c.inl ∧ IsPullback (cofans X' Y').inr αY f c.inr) - (h₂ : ∀ {Z : C} (f : Z ⟶ c.pt), - IsColimit (BinaryCofan.mk (cones f c.inl).fst (cones f c.inr).fst)) : - IsVanKampenColimit c := by - rw [BinaryCofan.isVanKampen_iff] - introv hX hY - constructor - · rintro ⟨h⟩ - let e := h.coconePointUniqueUpToIso (colimits _ _) - obtain ⟨hl, hr⟩ := h₁ αX αY (e.inv ≫ f) (by simp [hX]) (by simp [hY]) - constructor - · rw [← Category.id_comp αX, ← Iso.hom_inv_id_assoc e f] - haveI : IsIso (𝟙 X') := inferInstance - have : c'.inl ≫ e.hom = 𝟙 X' ≫ (cofans X' Y').inl := by - dsimp - simp - exact (IsPullback.of_vert_isIso ⟨this⟩).paste_vert hl - · rw [← Category.id_comp αY, ← Iso.hom_inv_id_assoc e f] - haveI : IsIso (𝟙 Y') := inferInstance - have : c'.inr ≫ e.hom = 𝟙 Y' ≫ (cofans X' Y').inr := by - dsimp - simp - exact (IsPullback.of_vert_isIso ⟨this⟩).paste_vert hr - · rintro ⟨H₁, H₂⟩ - refine' ⟨IsColimit.ofIsoColimit _ <| (isoBinaryCofanMk _).symm⟩ - let e₁ : X' ≅ _ := H₁.isLimit.conePointUniqueUpToIso (limits _ _) - let e₂ : Y' ≅ _ := H₂.isLimit.conePointUniqueUpToIso (limits _ _) - have he₁ : c'.inl = e₁.hom ≫ (cones f c.inl).fst := by simp - have he₂ : c'.inr = e₂.hom ≫ (cones f c.inr).fst := by simp - rw [he₁, he₂] - apply BinaryCofan.isColimitCompRightIso (BinaryCofan.mk _ _) - apply BinaryCofan.isColimitCompLeftIso (BinaryCofan.mk _ _) - exact h₂ f -#align category_theory.binary_cofan.is_van_kampen_mk CategoryTheory.BinaryCofan.isVanKampen_mk - -theorem BinaryCofan.mono_inr_of_isVanKampen [HasInitial C] {X Y : C} {c : BinaryCofan X Y} - (h : IsVanKampenColimit c) : Mono c.inr := by - refine' PullbackCone.mono_of_isLimitMkIdId _ (IsPullback.isLimit _) - refine' (h (BinaryCofan.mk (initial.to Y) (𝟙 Y)) (mapPair (initial.to X) (𝟙 Y)) c.inr _ - (mapPair_equifibered _)).mp ⟨_⟩ ⟨WalkingPair.right⟩ - · ext ⟨⟨⟩⟩ <;> dsimp; simp - · exact ((BinaryCofan.isColimit_iff_isIso_inr initialIsInitial _).mpr (by - dsimp - infer_instance)).some -#align category_theory.binary_cofan.mono_inr_of_is_van_kampen CategoryTheory.BinaryCofan.mono_inr_of_isVanKampen +namespace HasPullbacksOfInclusions + +instance (priority := 100) [HasBinaryCoproducts C] [HasPullbacks C] : + HasPullbacksOfInclusions C := ⟨⟩ + +variable [HasBinaryCoproducts C] [HasPullbacksOfInclusions C] {X Y Z : C} (f : Z ⟶ X ⨿ Y) + +instance preservesPullbackInl' : + HasPullback f coprod.inl := + hasPullback_symmetry _ _ + +instance hasPullbackInr' : + HasPullback f coprod.inr := by + have : IsPullback (𝟙 _) (f ≫ (coprod.braiding X Y).hom) f (coprod.braiding Y X).hom := + IsPullback.of_horiz_isIso ⟨by simp⟩ + have := (IsPullback.of_hasPullback (f ≫ (coprod.braiding X Y).hom) coprod.inl).paste_horiz this + simp only [coprod.braiding_hom, Category.comp_id, colimit.ι_desc, BinaryCofan.mk_pt, + BinaryCofan.ι_app_left, BinaryCofan.mk_inl] at this + exact ⟨⟨⟨_, this.isLimit⟩⟩⟩ + +instance hasPullbackInr : + HasPullback coprod.inr f := + hasPullback_symmetry _ _ + +end HasPullbacksOfInclusions + +namespace PreservesPullbacksOfInclusions + +variable {D : Type*} [Category D] [HasBinaryCoproducts C] (F : C ⥤ D) + +noncomputable +instance (priority := 100) [PreservesLimitsOfShape WalkingCospan F] : + PreservesPullbacksOfInclusions F := ⟨⟩ + +variable [PreservesPullbacksOfInclusions F] {X Y Z : C} (f : Z ⟶ X ⨿ Y) + +noncomputable +instance preservesPullbackInl' : + PreservesLimit (cospan f coprod.inl) F := + preservesPullbackSymmetry _ _ _ + +noncomputable +instance preservesPullbackInr' : + PreservesLimit (cospan f coprod.inr) F := by + apply preservesLimitOfIsoDiagram (K₁ := cospan (f ≫ (coprod.braiding X Y).hom) coprod.inl) + apply cospanExt (Iso.refl _) (Iso.refl _) (coprod.braiding X Y).symm <;> simp + +noncomputable +instance preservesPullbackInr : + PreservesLimit (cospan coprod.inr f) F := + preservesPullbackSymmetry _ _ _ + +end PreservesPullbacksOfInclusions + +instance (priority := 100) FinitaryExtensive.toFinitaryPreExtensive [FinitaryExtensive C] : + FinitaryPreExtensive C := + ⟨fun c hc ↦ (FinitaryExtensive.van_kampen' c hc).isUniversal⟩ theorem FinitaryExtensive.mono_inr_of_isColimit [FinitaryExtensive C] {c : BinaryCofan X Y} (hc : IsColimit c) : Mono c.inr := @@ -254,49 +186,25 @@ instance [FinitaryExtensive C] (X Y : C) : Mono (coprod.inl : X ⟶ X ⨿ Y) := instance [FinitaryExtensive C] (X Y : C) : Mono (coprod.inr : Y ⟶ X ⨿ Y) := (FinitaryExtensive.mono_inr_of_isColimit (coprodIsCoprod X Y) : _) -theorem BinaryCofan.isPullback_initial_to_of_isVanKampen [HasInitial C] {c : BinaryCofan X Y} - (h : IsVanKampenColimit c) : IsPullback (initial.to _) (initial.to _) c.inl c.inr := by - refine' ((h (BinaryCofan.mk (initial.to Y) (𝟙 Y)) (mapPair (initial.to X) (𝟙 Y)) c.inr _ - (mapPair_equifibered _)).mp ⟨_⟩ ⟨WalkingPair.left⟩).flip - · ext ⟨⟨⟩⟩ <;> dsimp; simp - · exact ((BinaryCofan.isColimit_iff_isIso_inr initialIsInitial _).mpr (by - dsimp - infer_instance)).some -#align category_theory.binary_cofan.is_pullback_initial_to_of_is_van_kampen CategoryTheory.BinaryCofan.isPullback_initial_to_of_isVanKampen - theorem FinitaryExtensive.isPullback_initial_to_binaryCofan [FinitaryExtensive C] {c : BinaryCofan X Y} (hc : IsColimit c) : IsPullback (initial.to _) (initial.to _) c.inl c.inr := BinaryCofan.isPullback_initial_to_of_isVanKampen (FinitaryExtensive.vanKampen c hc) #align category_theory.finitary_extensive.is_pullback_initial_to_binary_cofan CategoryTheory.FinitaryExtensive.isPullback_initial_to_binaryCofan -theorem hasStrictInitial_of_isUniversal [HasInitial C] - (H : IsUniversalColimit (BinaryCofan.mk (𝟙 (⊥_ C)) (𝟙 (⊥_ C)))) : HasStrictInitialObjects C := - hasStrictInitialObjects_of_initial_is_strict - (by - intro A f - suffices IsColimit (BinaryCofan.mk (𝟙 A) (𝟙 A)) by - obtain ⟨l, h₁, h₂⟩ := Limits.BinaryCofan.IsColimit.desc' this (f ≫ initial.to A) (𝟙 A) - rcases (Category.id_comp _).symm.trans h₂ with rfl - exact ⟨⟨_, ((Category.id_comp _).symm.trans h₁).symm, initialIsInitial.hom_ext _ _⟩⟩ - refine' (H (BinaryCofan.mk (𝟙 _) (𝟙 _)) (mapPair f f) f (by ext ⟨⟨⟩⟩ <;> dsimp <;> simp) - (mapPair_equifibered _) _).some - rintro ⟨⟨⟩⟩ <;> dsimp <;> - exact IsPullback.of_horiz_isIso ⟨(Category.id_comp _).trans (Category.comp_id _).symm⟩) -#align category_theory.has_strict_initial_of_is_universal CategoryTheory.hasStrictInitial_of_isUniversal - -instance (priority := 100) hasStrictInitialObjects_of_finitaryExtensive [FinitaryExtensive C] : - HasStrictInitialObjects C := - hasStrictInitial_of_isUniversal (FinitaryExtensive.vanKampen _ +instance (priority := 100) hasStrictInitialObjects_of_finitaryPreExtensive + [FinitaryPreExtensive C] : HasStrictInitialObjects C := + hasStrictInitial_of_isUniversal (FinitaryPreExtensive.universal' _ ((BinaryCofan.isColimit_iff_isIso_inr initialIsInitial _).mpr (by dsimp - infer_instance)).some).isUniversal -#align category_theory.has_strict_initial_objects_of_finitary_extensive CategoryTheory.hasStrictInitialObjects_of_finitaryExtensive + infer_instance)).some) +#align category_theory.has_strict_initial_objects_of_finitary_extensive CategoryTheory.hasStrictInitialObjects_of_finitaryPreExtensive theorem finitaryExtensive_iff_of_isTerminal (C : Type u) [Category.{v} C] [HasFiniteCoproducts C] + [HasPullbacksOfInclusions C] (T : C) (HT : IsTerminal T) (c₀ : BinaryCofan T T) (hc₀ : IsColimit c₀) : FinitaryExtensive C ↔ IsVanKampenColimit c₀ := by - refine' ⟨fun H => H.2 c₀ hc₀, fun H => _⟩ + refine' ⟨fun H => H.van_kampen' c₀ hc₀, fun H => _⟩ constructor simp_rw [BinaryCofan.isVanKampen_iff] at H ⊢ intro X Y c hc X' Y' c' αX αY f hX hY @@ -458,7 +366,7 @@ noncomputable def finitaryExtensiveTopCatAux (Z : TopCat.{u}) set_option linter.uppercaseLean3 false in #align category_theory.finitary_extensive_Top_aux CategoryTheory.finitaryExtensiveTopCatAux -instance : FinitaryExtensive TopCat.{u} := by +instance finitaryExtensive_TopCat : FinitaryExtensive TopCat.{u} := by rw [finitaryExtensive_iff_of_isTerminal TopCat.{u} _ TopCat.isTerminalPUnit _ (TopCat.binaryCofanIsColimit _ _)] apply BinaryCofan.isVanKampen_mk _ _ (fun X Y => TopCat.binaryCofanIsColimit X Y) _ @@ -506,71 +414,63 @@ end TopCat section Functor -universe v'' u'' - -variable {D : Type u''} [Category.{v''} D] - -theorem NatTrans.Equifibered.whiskerRight {F G : J ⥤ C} {α : F ⟶ G} (hα : Equifibered α) - (H : C ⥤ D) [PreservesLimitsOfShape WalkingCospan H] : Equifibered (whiskerRight α H) := - fun _ _ f => (hα f).map H -#align category_theory.nat_trans.equifibered.whisker_right CategoryTheory.NatTrans.Equifibered.whiskerRight - -theorem IsVanKampenColimit.of_iso {F : J ⥤ C} {c c' : Cocone F} (H : IsVanKampenColimit c) - (e : c ≅ c') : IsVanKampenColimit c' := by - intro F' c'' α f h hα - have : c'.ι ≫ (Functor.const J).map e.inv.hom = c.ι := by - ext j - exact e.inv.2 j - rw [H c'' α (f ≫ e.inv.1) (by rw [Functor.map_comp, ← reassoc_of% h, this]) hα] - apply forall_congr' - intro j - conv_lhs => rw [← Category.comp_id (α.app j)] - haveI : IsIso e.inv.hom := Functor.map_isIso (Cocones.forget _) e.inv - exact (IsPullback.of_vert_isIso ⟨by simp⟩).paste_vert_iff (NatTrans.congr_app h j).symm -#align category_theory.is_van_kampen_colimit.of_iso CategoryTheory.IsVanKampenColimit.of_iso - -theorem IsVanKampenColimit.of_map {D : Type*} [Category D] (G : C ⥤ D) {F : J ⥤ C} {c : Cocone F} - [PreservesLimitsOfShape WalkingCospan G] [ReflectsLimitsOfShape WalkingCospan G] - [PreservesColimitsOfShape J G] [ReflectsColimitsOfShape J G] - (H : IsVanKampenColimit (G.mapCocone c)) : IsVanKampenColimit c := by - intro F' c' α f h hα - refine' (Iff.trans _ (H (G.mapCocone c') (whiskerRight α G) (G.map f) - (by ext j; simpa using G.congr_map (NatTrans.congr_app h j)) - (hα.whiskerRight G))).trans (forall_congr' fun j => _) - · exact ⟨fun h => ⟨isColimitOfPreserves G h.some⟩, fun h => ⟨isColimitOfReflects G h.some⟩⟩ - · exact IsPullback.map_iff G (NatTrans.congr_app h.symm j) -#align category_theory.is_van_kampen_colimit.of_map CategoryTheory.IsVanKampenColimit.of_map - -theorem isVanKampenColimit_of_evaluation [HasPullbacks D] [HasColimitsOfShape J D] (F : J ⥤ C ⥤ D) - (c : Cocone F) (hc : ∀ x : C, IsVanKampenColimit (((evaluation C D).obj x).mapCocone c)) : - IsVanKampenColimit c := by - intro F' c' α f e hα - have := fun x => hc x (((evaluation C D).obj x).mapCocone c') (whiskerRight α _) - (((evaluation C D).obj x).map f) - (by - ext y - dsimp - exact NatTrans.congr_app (NatTrans.congr_app e y) x) - (hα.whiskerRight _) +theorem finitaryExtensive_of_reflective + [HasFiniteCoproducts D] [HasPullbacksOfInclusions D] [FinitaryExtensive C] + {Gl : C ⥤ D} {Gr : D ⥤ C} (adj : Gl ⊣ Gr) [Full Gr] [Faithful Gr] + [∀ X Y (f : X ⟶ Gl.obj Y), HasPullback (Gr.map f) (adj.unit.app Y)] + [∀ X Y (f : X ⟶ Gl.obj Y), PreservesLimit (cospan (Gr.map f) (adj.unit.app Y)) Gl] + [PreservesPullbacksOfInclusions Gl] : + FinitaryExtensive D := by + have : PreservesColimitsOfSize Gl := adj.leftAdjointPreservesColimits constructor - · rintro ⟨hc'⟩ j - refine' ⟨⟨(NatTrans.congr_app e j).symm⟩, ⟨evaluationJointlyReflectsLimits _ _⟩⟩ - refine' fun x => (isLimitMapConePullbackConeEquiv _ _).symm _ - exact ((this x).mp ⟨PreservesColimit.preserves hc'⟩ _).isLimit - · exact fun H => ⟨evaluationJointlyReflectsColimits _ fun x => - ((this x).mpr fun j => (H j).map ((evaluation C D).obj x)).some⟩ -#align category_theory.is_van_kampen_colimit_of_evaluation CategoryTheory.isVanKampenColimit_of_evaluation - -instance [HasPullbacks C] [FinitaryExtensive C] : FinitaryExtensive (D ⥤ C) := + intros X Y c hc + apply (IsVanKampenColimit.precompose_isIso_iff + (isoWhiskerLeft _ (asIso adj.counit) ≪≫ Functor.rightUnitor _).hom).mp + have : ∀ (Z : C) (i : Discrete WalkingPair) (f : Z ⟶ (colimit.cocone (pair X Y ⋙ Gr)).pt), + PreservesLimit (cospan f ((colimit.cocone (pair X Y ⋙ Gr)).ι.app i)) Gl := by + have : pair X Y ⋙ Gr = pair (Gr.obj X) (Gr.obj Y) := by + apply Functor.hext + · rintro ⟨⟨⟩⟩ <;> rfl + · rintro ⟨⟨⟩⟩ ⟨j⟩ ⟨⟨rfl : _ = j⟩⟩ <;> simp + rw [this] + rintro Z ⟨_|_⟩ f <;> dsimp <;> infer_instance + refine ((FinitaryExtensive.vanKampen _ (colimit.isColimit $ pair X Y ⋙ _)).map_reflective + adj).of_iso (IsColimit.uniqueUpToIso ?_ ?_) + · exact isColimitOfPreserves Gl (colimit.isColimit _) + · exact (IsColimit.precomposeHomEquiv _ _).symm hc + +instance finitaryExtensive_functor [HasPullbacks C] [FinitaryExtensive C] : + FinitaryExtensive (D ⥤ C) := haveI : HasFiniteCoproducts (D ⥤ C) := ⟨fun _ => Limits.functorCategoryHasColimitsOfShape⟩ ⟨fun c hc => isVanKampenColimit_of_evaluation _ c fun _ => FinitaryExtensive.vanKampen _ <| PreservesColimit.preserves hc⟩ +noncomputable +instance {C} [Category C] {D} [Category D] (F : C ⥤ D) + {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) [IsIso f] : PreservesLimit (cospan f g) F := + have := hasPullback_of_left_iso f g + preservesLimitOfPreservesLimitCone (IsPullback.of_hasPullback f g).isLimit + ((isLimitMapConePullbackConeEquiv _ pullback.condition).symm + (IsPullback.of_vert_isIso ⟨by simp only [← F.map_comp, pullback.condition]⟩).isLimit) + +noncomputable +instance {C} [Category C] {D} [Category D] (F : C ⥤ D) + {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) [IsIso g] : PreservesLimit (cospan f g) F := + preservesPullbackSymmetry _ _ _ + theorem finitaryExtensive_of_preserves_and_reflects (F : C ⥤ D) [FinitaryExtensive D] - [HasFiniteCoproducts C] [PreservesLimitsOfShape WalkingCospan F] + [HasFiniteCoproducts C] [HasPullbacksOfInclusions C] + [PreservesPullbacksOfInclusions F] [ReflectsLimitsOfShape WalkingCospan F] [PreservesColimitsOfShape (Discrete WalkingPair) F] - [ReflectsColimitsOfShape (Discrete WalkingPair) F] : FinitaryExtensive C := - ⟨fun _ hc => (FinitaryExtensive.vanKampen _ (isColimitOfPreserves F hc)).of_map F⟩ + [ReflectsColimitsOfShape (Discrete WalkingPair) F] : FinitaryExtensive C := by + constructor + intros X Y c hc + refine IsVanKampenColimit.of_iso ?_ (hc.uniqueUpToIso (coprodIsCoprod X Y)).symm + have (i : Discrete WalkingPair) (Z : C) (f : Z ⟶ X ⨿ Y) : + PreservesLimit (cospan f ((BinaryCofan.mk coprod.inl coprod.inr).ι.app i)) F := by + rcases i with ⟨_|_⟩ <;> dsimp <;> infer_instance + refine (FinitaryExtensive.vanKampen _ + (isColimitOfPreserves F (coprodIsCoprod X Y))).of_mapCocone F #align category_theory.finitary_extensive_of_preserves_and_reflects CategoryTheory.finitaryExtensive_of_preserves_and_reflects theorem finitaryExtensive_of_preserves_and_reflects_isomorphism (F : C ⥤ D) [FinitaryExtensive D] @@ -585,6 +485,141 @@ theorem finitaryExtensive_of_preserves_and_reflects_isomorphism (F : C ⥤ D) [F end Functor +section FiniteCoproducts + +theorem FinitaryPreExtensive.isUniversal_finiteCoproducts_Fin [FinitaryPreExtensive C] {n : ℕ} + {F : Discrete (Fin n) ⥤ C} {c : Cocone F} (hc : IsColimit c) : IsUniversalColimit c := by + let f : Fin n → C := F.obj ∘ Discrete.mk + have : F = Discrete.functor f := + Functor.hext (fun _ ↦ rfl) (by rintro ⟨i⟩ ⟨j⟩ ⟨⟨rfl : i = j⟩⟩; simp) + clear_value f + subst this + induction' n with n IH + · exact (isVanKampenColimit_of_isEmpty _ hc).isUniversal + · apply IsUniversalColimit.of_iso _ + ((extendCofanIsColimit f (coproductIsCoproduct _) (coprodIsCoprod _ _)).uniqueUpToIso hc) + apply @isUniversalColimit_extendCofan _ _ _ _ _ _ _ _ ?_ + · apply IH + exact coproductIsCoproduct _ + · apply FinitaryPreExtensive.universal' + exact coprodIsCoprod _ _ + · dsimp + infer_instance + +theorem FinitaryPreExtensive.isUniversal_finiteCoproducts [FinitaryPreExtensive C] {ι : Type*} + [Finite ι] {F : Discrete ι ⥤ C} {c : Cocone F} (hc : IsColimit c) : IsUniversalColimit c := by + obtain ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin ι + apply (IsUniversalColimit.whiskerEquivalence_iff (Discrete.equivalence e).symm).mp + apply FinitaryPreExtensive.isUniversal_finiteCoproducts_Fin + exact (IsColimit.whiskerEquivalenceEquiv (Discrete.equivalence e).symm) hc + +theorem FinitaryExtensive.isVanKampen_finiteCoproducts_Fin [FinitaryExtensive C] {n : ℕ} + {F : Discrete (Fin n) ⥤ C} {c : Cocone F} (hc : IsColimit c) : IsVanKampenColimit c := by + let f : Fin n → C := F.obj ∘ Discrete.mk + have : F = Discrete.functor f := + Functor.hext (fun _ ↦ rfl) (by rintro ⟨i⟩ ⟨j⟩ ⟨⟨rfl : i = j⟩⟩; simp) + clear_value f + subst this + induction' n with n IH + · exact isVanKampenColimit_of_isEmpty _ hc + · apply IsVanKampenColimit.of_iso _ + ((extendCofanIsColimit f (coproductIsCoproduct _) (coprodIsCoprod _ _)).uniqueUpToIso hc) + apply @isVanKampenColimit_extendCofan _ _ _ _ _ _ _ _ ?_ + · apply IH + exact coproductIsCoproduct _ + · apply FinitaryExtensive.van_kampen' + exact coprodIsCoprod _ _ + · dsimp + infer_instance + +theorem FinitaryExtensive.isVanKampen_finiteCoproducts [FinitaryExtensive C] {ι : Type*} + [Finite ι] {F : Discrete ι ⥤ C} {c : Cocone F} (hc : IsColimit c) : IsVanKampenColimit c := by + obtain ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin ι + apply (IsVanKampenColimit.whiskerEquivalence_iff (Discrete.equivalence e).symm).mp + apply FinitaryExtensive.isVanKampen_finiteCoproducts_Fin + exact (IsColimit.whiskerEquivalenceEquiv (Discrete.equivalence e).symm) hc + +lemma FinitaryPreExtensive.hasPullbacks_of_is_coproduct [FinitaryPreExtensive C] {ι : Type*} + [Finite ι] {F : Discrete ι ⥤ C} {c : Cocone F} (hc : IsColimit c) (i : Discrete ι) {X : C} + (g : X ⟶ _) : HasPullback g (c.ι.app i) := by + classical + let f : ι → C := F.obj ∘ Discrete.mk + have : F = Discrete.functor f := + Functor.hext (fun i ↦ rfl) (by rintro ⟨i⟩ ⟨j⟩ ⟨⟨rfl : i = j⟩⟩; simp) + clear_value f + subst this + change Cofan f at c + obtain ⟨i⟩ := i + let e : ∐ f ≅ f i ⨿ (∐ fun j : ({i}ᶜ : Set ι) ↦ f j) := + { hom := Sigma.desc (fun j ↦ if h : j = i then eqToHom (congr_arg f h) ≫ coprod.inl else + Sigma.ι (fun j : ({i}ᶜ : Set ι) ↦ f j) ⟨j, h⟩ ≫ coprod.inr) + inv := coprod.desc (Sigma.ι f i) (Sigma.desc <| fun j ↦ Sigma.ι f j) + hom_inv_id := by aesop_cat + inv_hom_id := by + ext j + · simp + · simp only [coprod.desc_comp, colimit.ι_desc, Cofan.mk_pt, Cofan.mk_ι_app, + eqToHom_refl, Category.id_comp, dite_true, BinaryCofan.mk_pt, BinaryCofan.ι_app_right, + BinaryCofan.mk_inr, colimit.ι_desc_assoc, Discrete.functor_obj, Category.comp_id] + exact dif_neg j.prop } + let e' : c.pt ≅ f i ⨿ (∐ fun j : ({i}ᶜ : Set ι) ↦ f j) := + hc.coconePointUniqueUpToIso (getColimitCocone _).2 ≪≫ e + have : coprod.inl ≫ e'.inv = c.ι.app ⟨i⟩ + · simp only [Iso.trans_inv, coprod.desc_comp, colimit.ι_desc, BinaryCofan.mk_pt, + BinaryCofan.ι_app_left, BinaryCofan.mk_inl] + exact colimit.comp_coconePointUniqueUpToIso_inv _ _ + clear_value e' + rw [← this] + have : IsPullback (𝟙 _) (g ≫ e'.hom) g e'.inv := IsPullback.of_horiz_isIso ⟨by simp⟩ + exact ⟨⟨⟨_, ((IsPullback.of_hasPullback (g ≫ e'.hom) coprod.inl).paste_horiz this).isLimit⟩⟩⟩ + +lemma FinitaryExtensive.mono_ι [FinitaryExtensive C] {ι : Type*} [Finite ι] {F : Discrete ι ⥤ C} + {c : Cocone F} (hc : IsColimit c) (i : Discrete ι) : + Mono (c.ι.app i) := + mono_of_cofan_isVanKampen (isVanKampen_finiteCoproducts hc) _ + +instance [FinitaryExtensive C] {ι : Type*} [Finite ι] (X : ι → C) (i : ι) : + Mono (Sigma.ι X i) := + FinitaryExtensive.mono_ι (coproductIsCoproduct _) ⟨i⟩ + +lemma FinitaryExtensive.isPullback_initial_to [FinitaryExtensive C] + {ι : Type*} [Finite ι] {F : Discrete ι ⥤ C} + {c : Cocone F} (hc : IsColimit c) (i j : Discrete ι) (e : i ≠ j) : + IsPullback (initial.to _) (initial.to _) (c.ι.app i) (c.ι.app j) := + isPullback_initial_to_of_cofan_isVanKampen (isVanKampen_finiteCoproducts hc) i j e + +lemma FinitaryExtensive.isPullback_initial_to_sigma_ι [FinitaryExtensive C] {ι : Type*} [Finite ι] + (X : ι → C) (i j : ι) (e : i ≠ j) : + IsPullback (initial.to _) (initial.to _) (Sigma.ι X i) (Sigma.ι X j) := + FinitaryExtensive.isPullback_initial_to (coproductIsCoproduct _) ⟨i⟩ ⟨j⟩ + (ne_of_apply_ne Discrete.as e) + +instance FinitaryPreExtensive.hasPullbacks_of_inclusions [FinitaryPreExtensive C] {X Z : C} + {α : Type*} (f : X ⟶ Z) {Y : (a : α) → C} (i : (a : α) → Y a ⟶ Z) [Finite α] + [hi : IsIso (Sigma.desc i)] (a : α) : HasPullback f (i a) := by + apply FinitaryPreExtensive.hasPullbacks_of_is_coproduct (c := Cofan.mk Z i) + exact @IsColimit.ofPointIso (t := Cofan.mk Z i) (P := _) hi + +lemma FinitaryPreExtensive.sigma_desc_iso [FinitaryPreExtensive C] {α : Type} [Finite α] {X : C} + {Z : α → C} (π : (a : α) → Z a ⟶ X) {Y : C} (f : Y ⟶ X) (hπ : IsIso (Sigma.desc π)) : + IsIso (Sigma.desc ((fun _ ↦ pullback.fst) : (a : α) → pullback f (π a) ⟶ _)) := by + suffices IsColimit (Cofan.mk _ ((fun _ ↦ pullback.fst) : (a : α) → pullback f (π a) ⟶ _)) by + change IsIso (this.coconePointUniqueUpToIso (getColimitCocone _).2).inv + infer_instance + let : IsColimit (Cofan.mk X π) + · refine @IsColimit.ofPointIso (t := Cofan.mk X π) (P := coproductIsCoproduct Z) ?_ + convert hπ + simp [coproductIsCoproduct] + refine (FinitaryPreExtensive.isUniversal_finiteCoproducts this + (Cofan.mk _ ((fun _ ↦ pullback.fst) : (a : α) → pullback f (π a) ⟶ _)) + (Discrete.natTrans <| fun i ↦ pullback.snd) f ?_ + (NatTrans.equifibered_of_discrete _) ?_).some + · ext + simp [pullback.condition] + · exact fun j ↦ IsPullback.of_hasPullback f (π j.as) + +end FiniteCoproducts + end Extensive end CategoryTheory diff --git a/Mathlib/CategoryTheory/Functor/Flat.lean b/Mathlib/CategoryTheory/Functor/Flat.lean index 655d7b3b20828..57d41d0658911 100644 --- a/Mathlib/CategoryTheory/Functor/Flat.lean +++ b/Mathlib/CategoryTheory/Functor/Flat.lean @@ -184,8 +184,6 @@ theorem fac (x : J) : lift F hc s ≫ (F.mapCone c).π.app x = s.π.app x := by simp [lift, ← Functor.map_comp] #align category_theory.preserves_finite_limits_of_flat.fac CategoryTheory.PreservesFiniteLimitsOfFlat.fac -attribute [local simp] eqToHom_map - theorem uniq {K : J ⥤ C} {c : Cone K} (hc : IsLimit c) (s : Cone (K ⋙ F)) (f₁ f₂ : s.pt ⟶ F.obj c.pt) (h₁ : ∀ j : J, f₁ ≫ (F.mapCone c).π.app j = s.π.app j) (h₂ : ∀ j : J, f₂ ≫ (F.mapCone c).π.app j = s.π.app j) : f₁ = f₂ := by @@ -208,7 +206,7 @@ theorem uniq {K : J ⥤ C} {c : Cone K} (hc : IsLimit c) (s : Cone (K ⋙ F)) intro j injection c₀.π.naturality (BiconeHom.left j) with _ e₁ injection c₀.π.naturality (BiconeHom.right j) with _ e₂ - simpa (config := {zeta := false}) using e₁.symm.trans e₂ + convert e₁.symm.trans e₂ <;> simp have : c.extend g₁.right = c.extend g₂.right := by unfold Cone.extend congr 1 diff --git a/Mathlib/CategoryTheory/Functor/LeftDerived.lean b/Mathlib/CategoryTheory/Functor/LeftDerived.lean index 877e846cb5ac1..a4d13b15e4ce6 100644 --- a/Mathlib/CategoryTheory/Functor/LeftDerived.lean +++ b/Mathlib/CategoryTheory/Functor/LeftDerived.lean @@ -60,7 +60,7 @@ variable [Preadditive D] [HasEqualizers D] [HasCokernels D] [HasImages D] [HasIm /-- The left derived functors of an additive functor. -/ def Functor.leftDerived (F : C ⥤ D) [F.Additive] (n : ℕ) : C ⥤ D := - projectiveResolutions C ⋙ F.mapHomotopyCategory _ ⋙ HomotopyCategory.homologyFunctor D _ n + projectiveResolutions C ⋙ F.mapHomotopyCategory _ ⋙ HomotopyCategory.homology'Functor D _ n #align category_theory.functor.left_derived CategoryTheory.Functor.leftDerived -- TODO the left derived functors are additive (and linear when `F` is linear) @@ -69,11 +69,11 @@ def Functor.leftDerived (F : C ⥤ D) [F.Additive] (n : ℕ) : C ⥤ D := def Functor.leftDerivedObjIso (F : C ⥤ D) [F.Additive] (n : ℕ) {X : C} (P : ProjectiveResolution X) : (F.leftDerived n).obj X ≅ - (homologyFunctor D _ n).obj ((F.mapHomologicalComplex _).obj P.complex) := - (HomotopyCategory.homologyFunctor D _ n).mapIso + (homology'Functor D _ n).obj ((F.mapHomologicalComplex _).obj P.complex) := + (HomotopyCategory.homology'Functor D _ n).mapIso (HomotopyCategory.isoOfHomotopyEquiv (F.mapHomotopyEquiv (ProjectiveResolution.homotopyEquiv _ P))) ≪≫ - (HomotopyCategory.homologyFactors D _ n).app _ + (HomotopyCategory.homology'Factors D _ n).app _ #align category_theory.functor.left_derived_obj_iso CategoryTheory.Functor.leftDerivedObjIso section @@ -85,8 +85,8 @@ variable [HasZeroObject D] def Functor.leftDerivedObjProjectiveZero (F : C ⥤ D) [F.Additive] (X : C) [Projective X] : (F.leftDerived 0).obj X ≅ F.obj X := F.leftDerivedObjIso 0 (ProjectiveResolution.self X) ≪≫ - (homologyFunctor _ _ _).mapIso ((ChainComplex.single₀MapHomologicalComplex F).app X) ≪≫ - (ChainComplex.homologyFunctor0Single₀ D).app (F.obj X) + (homology'Functor _ _ _).mapIso ((ChainComplex.single₀MapHomologicalComplex F).app X) ≪≫ + (ChainComplex.homology'Functor0Single₀ D).app (F.obj X) #align category_theory.functor.left_derived_obj_projective_zero CategoryTheory.Functor.leftDerivedObjProjectiveZero open ZeroObject @@ -96,8 +96,8 @@ open ZeroObject def Functor.leftDerivedObjProjectiveSucc (F : C ⥤ D) [F.Additive] (n : ℕ) (X : C) [Projective X] : (F.leftDerived (n + 1)).obj X ≅ 0 := F.leftDerivedObjIso (n + 1) (ProjectiveResolution.self X) ≪≫ - (homologyFunctor _ _ _).mapIso ((ChainComplex.single₀MapHomologicalComplex F).app X) ≪≫ - (ChainComplex.homologyFunctorSuccSingle₀ D n).app (F.obj X) ≪≫ (Functor.zero_obj _).isoZero + (homology'Functor _ _ _).mapIso ((ChainComplex.single₀MapHomologicalComplex F).app X) ≪≫ + (ChainComplex.homology'FunctorSuccSingle₀ D n).app (F.obj X) ≪≫ (Functor.zero_obj _).isoZero #align category_theory.functor.left_derived_obj_projective_succ CategoryTheory.Functor.leftDerivedObjProjectiveSucc end @@ -110,11 +110,11 @@ theorem Functor.leftDerived_map_eq (F : C ⥤ D) [F.Additive] (n : ℕ) {X Y : C (w : g ≫ Q.π = P.π ≫ (ChainComplex.single₀ C).map f) : (F.leftDerived n).map f = (F.leftDerivedObjIso n P).hom ≫ - (homologyFunctor D _ n).map ((F.mapHomologicalComplex _).map g) ≫ + (homology'Functor D _ n).map ((F.mapHomologicalComplex _).map g) ≫ (F.leftDerivedObjIso n Q).inv := by dsimp only [Functor.leftDerived, Functor.leftDerivedObjIso] dsimp; simp only [Category.comp_id, Category.id_comp] - rw [← homologyFunctor_map, HomotopyCategory.homologyFunctor_map_factors] + rw [← homology'Functor_map, HomotopyCategory.homology'Functor_map_factors] simp only [← Functor.map_comp] congr 1 apply HomotopyCategory.eq_of_homotopy @@ -129,7 +129,7 @@ theorem Functor.leftDerived_map_eq (F : C ⥤ D) [F.Additive] (n : ℕ) {X Y : C def NatTrans.leftDerived {F G : C ⥤ D} [F.Additive] [G.Additive] (α : F ⟶ G) (n : ℕ) : F.leftDerived n ⟶ G.leftDerived n := whiskerLeft (projectiveResolutions C) - (whiskerRight (NatTrans.mapHomotopyCategory α _) (HomotopyCategory.homologyFunctor D _ n)) + (whiskerRight (NatTrans.mapHomotopyCategory α _) (HomotopyCategory.homology'Functor D _ n)) #align category_theory.nat_trans.left_derived CategoryTheory.NatTrans.leftDerived @[simp] @@ -154,12 +154,12 @@ theorem NatTrans.leftDerived_eq {F G : C ⥤ D} [F.Additive] [G.Additive] (α : (P : ProjectiveResolution X) : (NatTrans.leftDerived α n).app X = (F.leftDerivedObjIso n P).hom ≫ - (homologyFunctor D _ n).map ((NatTrans.mapHomologicalComplex α _).app P.complex) ≫ + (homology'Functor D _ n).map ((NatTrans.mapHomologicalComplex α _).app P.complex) ≫ (G.leftDerivedObjIso n P).inv := by symm dsimp [NatTrans.leftDerived, Functor.leftDerivedObjIso] simp only [Category.comp_id, Category.id_comp] - rw [← homologyFunctor_map, HomotopyCategory.homologyFunctor_map_factors] + rw [← homology'Functor_map, HomotopyCategory.homology'Functor_map_factors] simp only [← Functor.map_comp] congr 1 apply HomotopyCategory.eq_of_homotopy diff --git a/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean b/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean index f2a2357610d3e..46a945db9e513 100644 --- a/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean @@ -688,7 +688,7 @@ theorem full_empty : full ∅ = (⊥ : Subgroupoid C) := by @[simp] theorem full_univ : full Set.univ = (⊤ : Subgroupoid C) := by ext - simp only [mem_full_iff, mem_univ, mem_top] + simp only [mem_full_iff, mem_univ, and_self, mem_top] #align category_theory.subgroupoid.full_univ CategoryTheory.Subgroupoid.full_univ theorem full_mono {D E : Set C} (h : D ≤ E) : full D ≤ full E := by diff --git a/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean b/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean index 61cc401d9b4e5..ffeda75bc2923 100644 --- a/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean +++ b/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean @@ -6,8 +6,8 @@ Authors: Rémi Bottinelli import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.PathCategory import Mathlib.Algebra.Group.Defs -import Mathlib.Algebra.Hom.Group.Defs -import Mathlib.Algebra.Hom.Equiv.Basic +import Mathlib.Algebra.Group.Hom.Defs +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Combinatorics.Quiver.Path import Mathlib.Combinatorics.Quiver.ConnectedComponent diff --git a/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean b/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean index 545ea9debd3a9..e15dcdf2a94eb 100644 --- a/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean +++ b/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean @@ -80,7 +80,6 @@ instance functor_category_isIdempotentComplete [IsIdempotentComplete C] : use Y, i, e constructor · ext j - apply equalizer.hom_ext dsimp rw [assoc, equalizer.lift_ι, ← equalizer.condition, id_comp, comp_id] · ext j diff --git a/Mathlib/CategoryTheory/Limits/FintypeCat.lean b/Mathlib/CategoryTheory/Limits/FintypeCat.lean new file mode 100644 index 0000000000000..2d72f827f8dfa --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/FintypeCat.lean @@ -0,0 +1,78 @@ +/- +Copyright (c) 2023 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.CategoryTheory.FintypeCat +import Mathlib.CategoryTheory.Limits.Shapes.FiniteLimits +import Mathlib.CategoryTheory.Limits.Types +import Mathlib.CategoryTheory.Limits.Creates +import Mathlib.CategoryTheory.Limits.Preserves.Finite +import Mathlib.Data.FunLike.Fintype + +/-! +# (Co)limits in the category of finite types + +We show that finite (co)limits exist in `FintypeCat` and that they are preserved by the natural +inclusion `FintypeCat.incl`. +-/ + +open CategoryTheory Limits Functor + +universe u + +namespace CategoryTheory.Limits.FintypeCat + +instance {J : Type} [SmallCategory J] [FinCategory J] (K : J ⥤ FintypeCat.{u}) (j : J) : + Finite ((K ⋙ FintypeCat.incl.{u}).obj j) := by + simp only [comp_obj, FintypeCat.incl_obj] + infer_instance + +/-- Any functor from a finite category to Types that only involves finite objects, +has a finite limit. -/ +noncomputable instance finiteLimitOfFiniteDiagram {J : Type} [SmallCategory J] [FinCategory J] + (K : J ⥤ Type*) [∀ j, Finite (K.obj j)] : Fintype (limit K) := by + have : Fintype (sections K) := Fintype.ofFinite (sections K) + exact Fintype.ofEquiv (sections K) (Types.limitEquivSections K).symm + +noncomputable instance inclusionCreatesFiniteLimits {J : Type} [SmallCategory J] [FinCategory J] : + CreatesLimitsOfShape J FintypeCat.incl.{u} where + CreatesLimit {K} := createsLimitOfFullyFaithfulOfIso + (FintypeCat.of <| limit <| K ⋙ FintypeCat.incl) (Iso.refl _) + +instance {J : Type} [SmallCategory J] [FinCategory J] : HasLimitsOfShape J FintypeCat.{u} where + has_limit F := hasLimit_of_created F FintypeCat.incl + +instance hasFiniteLimits : HasFiniteLimits FintypeCat.{u} where + out _ := inferInstance + +noncomputable instance inclusionPreservesFiniteLimits : + PreservesFiniteLimits FintypeCat.incl.{u} where + preservesFiniteLimits _ := + preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape FintypeCat.incl + +/-- Any functor from a finite category to Types that only involves finite objects, +has a finite colimit. -/ +noncomputable instance finiteColimitOfFiniteDiagram {J : Type} [SmallCategory J] [FinCategory J] + (K : J ⥤ Type*) [∀ j, Finite (K.obj j)] : Fintype (colimit K) := by + have : Finite (Types.Quot K) := Quot.finite (Types.Quot.Rel K) + have : Fintype (Types.Quot K) := Fintype.ofFinite (Types.Quot K) + exact Fintype.ofEquiv (Types.Quot K) (Types.colimitEquivQuot K).symm + +noncomputable instance inclusionCreatesFiniteColimits {J : Type} [SmallCategory J] [FinCategory J] : + CreatesColimitsOfShape J FintypeCat.incl.{u} where + CreatesColimit {K} := createsColimitOfFullyFaithfulOfIso + (FintypeCat.of <| colimit <| K ⋙ FintypeCat.incl) (Iso.refl _) + +instance {J : Type} [SmallCategory J] [FinCategory J] : HasColimitsOfShape J FintypeCat.{u} where + has_colimit F := hasColimit_of_created F FintypeCat.incl + +instance hasFiniteColimits : HasFiniteColimits FintypeCat.{u} where + out _ := inferInstance + +noncomputable instance inclusionPreservesFiniteColimits : + PreservesFiniteColimits FintypeCat.incl.{u} where + preservesFiniteColimits _ := + preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape FintypeCat.incl + +end CategoryTheory.Limits.FintypeCat diff --git a/Mathlib/CategoryTheory/Limits/Fubini.lean b/Mathlib/CategoryTheory/Limits/Fubini.lean index 50ec514a1b28e..6bfd5e58686e6 100644 --- a/Mathlib/CategoryTheory/Limits/Fubini.lean +++ b/Mathlib/CategoryTheory/Limits/Fubini.lean @@ -10,7 +10,7 @@ import Mathlib.CategoryTheory.Functor.Currying #align_import category_theory.limits.fubini from "leanprover-community/mathlib"@"59382264386afdbaf1727e617f5fdda511992eb9" /-! -# A Fubini theorem for categorical limits +# A Fubini theorem for categorical (co)limits We prove that $lim_{J × K} G = lim_J (lim_K G(j, -))$ for a functor `G : J × K ⥤ C`, when all the appropriate limits exist. @@ -32,9 +32,7 @@ For convenience, we also provide `limitIsoLimitCurryCompLim G : limit G ≅ limit ((curry.obj G) ⋙ lim)` in terms of the uncurried functor. -## Future work - -The dual statement. +All statements have their counterpart for colimits. -/ @@ -54,13 +52,26 @@ variable (F : J ⥤ K ⥤ C) /-- A structure carrying a diagram of cones over the functors `F.obj j`. -/ structure DiagramOfCones where + /-- For each object, a cone. -/ obj : ∀ j : J, Cone (F.obj j) + /-- For each map, a map of cones. -/ map : ∀ {j j' : J} (f : j ⟶ j'), (Cones.postcompose (F.map f)).obj (obj j) ⟶ obj j' id : ∀ j : J, (map (𝟙 j)).hom = 𝟙 _ := by aesop_cat comp : ∀ {j₁ j₂ j₃ : J} (f : j₁ ⟶ j₂) (g : j₂ ⟶ j₃), (map (f ≫ g)).hom = (map f).hom ≫ (map g).hom := by aesop_cat #align category_theory.limits.diagram_of_cones CategoryTheory.Limits.DiagramOfCones +/-- A structure carrying a diagram of cocones over the functors `F.obj j`. +-/ +structure DiagramOfCocones where + /-- For each object, a cocone. -/ + obj : ∀ j : J, Cocone (F.obj j) + /-- For each map, a map of cocones. -/ + map : ∀ {j j' : J} (f : j ⟶ j'), (obj j) ⟶ (Cocones.precompose (F.map f)).obj (obj j') + id : ∀ j : J, (map (𝟙 j)).hom = 𝟙 _ := by aesop_cat + comp : ∀ {j₁ j₂ j₃ : J} (f : j₁ ⟶ j₂) (g : j₂ ⟶ j₃), + (map (f ≫ g)).hom = (map f).hom ≫ (map g).hom := by aesop_cat + variable {F} /-- Extract the functor `J ⥤ C` consisting of the cone points and the maps between them, @@ -74,6 +85,16 @@ def DiagramOfCones.conePoints (D : DiagramOfCones F) : J ⥤ C where map_comp f g := D.comp f g #align category_theory.limits.diagram_of_cones.cone_points CategoryTheory.Limits.DiagramOfCones.conePoints +/-- Extract the functor `J ⥤ C` consisting of the cocone points and the maps between them, +from a `DiagramOfCocones`. +-/ +@[simps] +def DiagramOfCocones.coconePoints (D : DiagramOfCocones F) : J ⥤ C where + obj j := (D.obj j).pt + map f := (D.map f).hom + map_id j := D.id j + map_comp f g := D.comp f g + /-- Given a diagram `D` of limit cones over the `F.obj j`, and a cone over `uncurry.obj F`, we can construct a cone over the diagram consisting of the cone points from `D`. -/ @@ -109,6 +130,40 @@ def coneOfConeUncurry {D : DiagramOfCones F} (Q : ∀ j, IsLimit (D.obj j)) exact this) } #align category_theory.limits.cone_of_cone_uncurry CategoryTheory.Limits.coneOfConeUncurry +/-- Given a diagram `D` of colimit cocones over the `F.obj j`, and a cocone over `uncurry.obj F`, +we can construct a cocone over the diagram consisting of the cocone points from `D`. +-/ +@[simps] +def coconeOfCoconeUncurry {D : DiagramOfCocones F} (Q : ∀ j, IsColimit (D.obj j)) + (c : Cocone (uncurry.obj F)) : Cocone D.coconePoints where + pt := c.pt + ι := + { app := fun j => + (Q j).desc + { pt := c.pt + ι := + { app := fun k => c.ι.app (j, k) + naturality := fun k k' f => by + dsimp; simp only [Category.comp_id] + conv_lhs => + arg 1; equals (F.map (𝟙 _)).app _ ≫ (F.obj j).map f => + simp; + conv_lhs => arg 1; rw [← uncurry_obj_map F ((𝟙 j,f) : (j,k) ⟶ (j,k'))] + rw [c.w] } } + naturality := fun j j' f => + (Q j).hom_ext + (by + dsimp + intro k + simp only [Limits.CoconeMorphism.w_assoc, Limits.Cocones.precompose_obj_ι, + Limits.IsColimit.fac_assoc, Limits.IsColimit.fac, NatTrans.comp_app, Category.comp_id, + Category.assoc] + have := @NatTrans.naturality _ _ _ _ _ _ c.ι (j, k) (j', k) (f, 𝟙 k) + dsimp at this + simp only [Category.id_comp, Category.comp_id, CategoryTheory.Functor.map_id, + NatTrans.id_app] at this + exact this) } + /-- `coneOfConeUncurry Q c` is a limit cone when `c` is a limit cone. -/ def coneOfConeUncurryIsLimit {D : DiagramOfCones F} (Q : ∀ j, IsLimit (D.obj j)) @@ -149,6 +204,44 @@ def coneOfConeUncurryIsLimit {D : DiagramOfCones F} (Q : ∀ j, IsLimit (D.obj j simp #align category_theory.limits.cone_of_cone_uncurry_is_limit CategoryTheory.Limits.coneOfConeUncurryIsLimit +/-- `coconeOfCoconeUncurry Q c` is a colimit cocone when `c` is a colimit cocone. +-/ +def coconeOfCoconeUncurryIsColimit {D : DiagramOfCocones F} (Q : ∀ j, IsColimit (D.obj j)) + {c : Cocone (uncurry.obj F)} (P : IsColimit c) : IsColimit (coconeOfCoconeUncurry Q c) where + desc s := + P.desc + { pt := s.pt + ι := + { app := fun p => (D.obj p.1).ι.app p.2 ≫ s.ι.app p.1 + naturality := fun p p' f => by + dsimp; simp only [Category.id_comp, Category.assoc] + rcases p with ⟨j, k⟩ + rcases p' with ⟨j', k'⟩ + rcases f with ⟨fj, fk⟩ + dsimp + slice_lhs 2 3 => rw [(D.obj j').ι.naturality] + simp only [Functor.const_obj_map, Category.id_comp, Category.assoc] + have w := (D.map fj).w k + dsimp at w + slice_lhs 1 2 => rw [← w] + have n := s.ι.naturality fj + dsimp at n + simp only [Category.comp_id] at n + rw [← n] + simp } } + fac s j := by + apply (Q j).hom_ext + intro k + simp + uniq s m w := by + refine' P.uniq + { pt := s.pt + ι := _ } m _ + rintro ⟨j, k⟩ + dsimp + rw [← w j] + simp + section variable (F) @@ -214,6 +307,65 @@ end section +variable (F) + +variable [HasColimitsOfShape K C] + +/-- Given a functor `F : J ⥤ K ⥤ C`, with all needed colimits, +we can construct a diagram consisting of the colimit cocone over each functor `F.obj j`, +and the universal cocone morphisms between these. +-/ +@[simps] +noncomputable def DiagramOfCocones.mkOfHasColimits : DiagramOfCocones F where + obj j := colimit.cocone (F.obj j) + map f := { hom := colim.map (F.map f) } + +-- Satisfying the inhabited linter. +noncomputable instance diagramOfCoconesInhabited : Inhabited (DiagramOfCocones F) := + ⟨DiagramOfCocones.mkOfHasColimits F⟩ + +@[simp] +theorem DiagramOfCocones.mkOfHasColimits_coconePoints : + (DiagramOfCocones.mkOfHasColimits F).coconePoints = F ⋙ colim := + rfl + +variable [HasColimit (uncurry.obj F)] + +variable [HasColimit (F ⋙ colim)] + +/-- The Fubini theorem for a functor `F : J ⥤ K ⥤ C`, +showing that the colimit of `uncurry.obj F` can be computed as +the colimit of the colimits of the functors `F.obj j`. +-/ +noncomputable def colimitUncurryIsoColimitCompColim : + colimit (uncurry.obj F) ≅ colimit (F ⋙ colim) := by + let c := colimit.cocone (uncurry.obj F) + let P : IsColimit c := colimit.isColimit _ + let G := DiagramOfCocones.mkOfHasColimits F + let Q : ∀ j, IsColimit (G.obj j) := fun j => colimit.isColimit _ + have Q' := coconeOfCoconeUncurryIsColimit Q P + have Q'' := colimit.isColimit (F ⋙ colim) + exact IsColimit.coconePointUniqueUpToIso Q' Q'' + +@[simp, reassoc] +theorem colimitUncurryIsoColimitCompColim_ι_ι_inv {j} {k} : + colimit.ι (F.obj j) k ≫ colimit.ι (F ⋙ colim) j ≫ (colimitUncurryIsoColimitCompColim F).inv = + colimit.ι (uncurry.obj F) (j, k) := by + dsimp [colimitUncurryIsoColimitCompColim, IsColimit.coconePointUniqueUpToIso, + IsColimit.uniqueUpToIso] + simp + +@[simp, reassoc] +theorem colimitUncurryIsoColimitCompColim_ι_hom {j} {k} : + colimit.ι _ (j, k) ≫ (colimitUncurryIsoColimitCompColim F).hom = + (colimit.ι _ k ≫ colimit.ι (F ⋙ colim) j : _ ⟶ (colimit (F ⋙ colim))) := by + rw [← cancel_mono (colimitUncurryIsoColimitCompColim F).inv] + simp + +end + +section + variable (F) [HasLimitsOfShape J C] [HasLimitsOfShape K C] -- With only moderate effort these could be derived if needed: @@ -251,6 +403,39 @@ end section +variable (F) [HasColimitsOfShape J C] [HasColimitsOfShape K C] + +variable [HasColimitsOfShape (J × K) C] [HasColimitsOfShape (K × J) C] + +/-- The colimit of `F.flip ⋙ colim` is isomorphic to the colimit of `F ⋙ colim`. -/ +noncomputable def colimitFlipCompColimIsoColimitCompColim : + colimit (F.flip ⋙ colim) ≅ colimit (F ⋙ colim) := + (colimitUncurryIsoColimitCompColim _).symm ≪≫ + HasColimit.isoOfNatIso (uncurryObjFlip _) ≪≫ + HasColimit.isoOfEquivalence (Prod.braiding _ _) + (NatIso.ofComponents fun _ => by rfl) ≪≫ + colimitUncurryIsoColimitCompColim _ + +@[simp, reassoc] +theorem colimitFlipCompColimIsoColimitCompColim_ι_ι_hom (j) (k) : + colimit.ι (F.flip.obj k) j ≫ colimit.ι (F.flip ⋙ colim) k ≫ + (colimitFlipCompColimIsoColimitCompColim F).hom = + (colimit.ι _ k ≫ colimit.ι (F ⋙ colim) j : _ ⟶ colimit (F⋙ colim)) := by + dsimp [colimitFlipCompColimIsoColimitCompColim] + slice_lhs 1 3 => simp only [] + simp + +@[simp, reassoc] +theorem colimitFlipCompColimIsoColimitCompColim_ι_ι_inv (k) (j) : + colimit.ι (F.obj j) k ≫ colimit.ι (F ⋙ colim) j ≫ + (colimitFlipCompColimIsoColimitCompColim F).inv = + (colimit.ι _ j ≫ colimit.ι (F.flip ⋙ colim) k : _ ⟶ colimit (F.flip ⋙ colim)) := by + dsimp [colimitFlipCompColimIsoColimitCompColim] + slice_lhs 1 3 => simp only [] + simp + +end + variable (G : J × K ⥤ C) section @@ -292,6 +477,41 @@ end section +variable [HasColimitsOfShape K C] + +variable [HasColimit G] + +variable [HasColimit (curry.obj G ⋙ colim)] + +/-- The Fubini theorem for a functor `G : J × K ⥤ C`, +showing that the colimit of `G` can be computed as +the colimit of the colimits of the functors `G.obj (j, _)`. +-/ +noncomputable def colimitIsoColimitCurryCompColim : colimit G ≅ colimit (curry.obj G ⋙ colim) := by + have i : G ≅ uncurry.obj ((@curry J _ K _ C _).obj G) := currying.symm.unitIso.app G + haveI : Limits.HasColimit (uncurry.obj ((@curry J _ K _ C _).obj G)) := hasColimitOfIso i.symm + trans colimit (uncurry.obj ((@curry J _ K _ C _).obj G)) + apply HasColimit.isoOfNatIso i + exact colimitUncurryIsoColimitCompColim ((@curry J _ K _ C _).obj G) + +@[simp, reassoc] +theorem colimitIsoColimitCurryCompColim_ι_ι_inv {j} {k} : + colimit.ι ((curry.obj G).obj j) k ≫ colimit.ι (curry.obj G ⋙ colim) j ≫ + (colimitIsoColimitCurryCompColim G).inv = colimit.ι _ (j, k) := by + simp [colimitIsoColimitCurryCompColim, Trans.simple, HasColimit.isoOfNatIso, + colimitUncurryIsoColimitCompColim] + +@[simp, reassoc] +theorem colimitIsoColimitCurryCompColim_ι_hom {j} {k} : + colimit.ι _ (j, k) ≫ (colimitIsoColimitCurryCompColim G).hom = + (colimit.ι (_) k ≫ colimit.ι (curry.obj G ⋙ colim) j : _ ⟶ colimit (_ ⋙ colim)) := by + rw [← cancel_mono (colimitIsoColimitCurryCompColim G).inv] + simp + +end + +section + variable [HasLimits C] -- Certainly one could weaken the hypotheses here. @@ -342,6 +562,48 @@ theorem limitCurrySwapCompLimIsoLimitCurryCompLim_inv_π_π {j} {k} : end +section + +variable [HasColimits C] + +open CategoryTheory.prod + +/-- A variant of the Fubini theorem for a functor `G : J × K ⥤ C`, +showing that $\colim_k \colim_j G(j,k) ≅ \colim_j \colim_k G(j,k)$. +-/ +noncomputable def colimitCurrySwapCompColimIsoColimitCurryCompColim : + colimit (curry.obj (Prod.swap K J ⋙ G) ⋙ colim) ≅ colimit (curry.obj G ⋙ colim) := + calc + colimit (curry.obj (Prod.swap K J ⋙ G) ⋙ colim) ≅ colimit (Prod.swap K J ⋙ G) := + (colimitIsoColimitCurryCompColim _).symm + _ ≅ colimit G := (HasColimit.isoOfEquivalence (Prod.braiding K J) (Iso.refl _)) + _ ≅ colimit (curry.obj G ⋙ colim) := colimitIsoColimitCurryCompColim _ + +@[simp] +theorem colimitCurrySwapCompColimIsoColimitCurryCompColim_ι_ι_hom {j} {k} : + colimit.ι _ j ≫ colimit.ι (curry.obj (Prod.swap K J ⋙ G) ⋙ colim) k ≫ + (colimitCurrySwapCompColimIsoColimitCurryCompColim G).hom = + (colimit.ι _ k ≫ colimit.ι (curry.obj G ⋙ colim) j : _ ⟶ colimit (curry.obj G⋙ colim)) := by + dsimp [colimitCurrySwapCompColimIsoColimitCurryCompColim] + slice_lhs 1 3 => simp only [] + simp + +@[simp] +theorem colimitCurrySwapCompColimIsoColimitCurryCompColim_ι_ι_inv {j} {k} : + colimit.ι _ k ≫ colimit.ι (curry.obj G ⋙ colim) j ≫ + (colimitCurrySwapCompColimIsoColimitCurryCompColim G).inv = + (colimit.ι _ j ≫ + colimit.ι (curry.obj _ ⋙ colim) k : + _ ⟶ colimit (curry.obj (Prod.swap K J ⋙ G) ⋙ colim)) := by + dsimp [colimitCurrySwapCompColimIsoColimitCurryCompColim] + slice_lhs 1 3 => simp only [] + simp only [colimitIsoColimitCurryCompColim_ι_ι_inv, HasColimit.isoOfEquivalence_inv_π, + Functor.id_obj, Functor.comp_obj, Prod.braiding_inverse_obj, Prod.braiding_functor_obj, + Prod.braiding_counitIso_inv_app, Prod.swap_obj, Iso.refl_hom, NatTrans.id_app, Category.id_comp, + Category.assoc, colimitIsoColimitCurryCompColim_ι_hom, curry_obj_obj_obj] + erw [CategoryTheory.Bifunctor.map_id] + simp + end end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean b/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean index e647a6fbb14c0..86c7716f92308 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean @@ -312,7 +312,7 @@ def preservesColimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [Preser /-- Transfer preservation of colimits of shape along a natural isomorphism in the functor. -/ def preservesColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfShape J F] : PreservesColimitsOfShape J G where - preservesColimit {K} := preservesColimitOfNatIso K h + preservesColimit {K} := preservesColimitOfNatIso K h #align category_theory.limits.preserves_colimits_of_shape_of_nat_iso CategoryTheory.Limits.preservesColimitsOfShapeOfNatIso /-- Transfer preservation of colimits along a natural isomorphism in the functor. -/ diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean b/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean index 58891ed32398e..acca9b381583c 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean @@ -183,4 +183,6 @@ for `Fintype J` -/ class PreservesFiniteCoproducts (F : C ⥤ D) where preserves : ∀ (J : Type) [Fintype J], PreservesColimitsOfShape (Discrete J) F +attribute [instance] PreservesFiniteCoproducts.preserves + end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index 5188a501193e7..d302afdcb0d28 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -2029,7 +2029,7 @@ theorem biprod.symmetry (P Q : C) : (biprod.braiding P Q).hom ≫ (biprod.braidi /-- The associator isomorphism which associates a binary biproduct. -/ @[simps] -def biprod.associator (P Q R : C) : (P ⊞ Q) ⊞ R ≅ P ⊞ (Q ⊞ R) where +def biprod.associator (P Q R : C) : (P ⊞ Q) ⊞ R ≅ P ⊞ (Q ⊞ R) where hom := biprod.lift (biprod.fst ≫ biprod.fst) (biprod.lift (biprod.fst ≫ biprod.snd) biprod.snd) inv := biprod.lift (biprod.lift biprod.fst (biprod.snd ≫ biprod.fst)) (biprod.snd ≫ biprod.snd) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean index 4d5f9d4f7166c..1c90554a71e83 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean @@ -709,7 +709,7 @@ theorem cokernel.π_desc {W : C} (k : Y ⟶ W) (h : f ≫ k = 0) : -- porting note: added to ease the port of `Abelian.Exact` @[reassoc (attr := simp)] lemma colimit_ι_zero_cokernel_desc {C : Type*} [Category C] - [HasZeroMorphisms C] {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) (h : f ≫ g = 0) [HasCokernel f]: + [HasZeroMorphisms C] {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) (h : f ≫ g = 0) [HasCokernel f] : colimit.ι (parallelPair f 0) WalkingParallelPair.zero ≫ cokernel.desc f g h = 0 := by rw [(colimit.w (parallelPair f 0) WalkingParallelPairHom.left).symm] aesop_cat diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean index 7b36b40597842..778c9842ac457 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean @@ -2393,7 +2393,7 @@ noncomputable def pullbackAssoc [HasPullback ((pullback.snd : Z₁ ⟶ X₂) ≫ @[reassoc (attr := simp)] theorem pullbackAssoc_inv_fst_fst [HasPullback ((pullback.snd : Z₁ ⟶ X₂) ≫ f₃) f₄] - [HasPullback f₁ ((pullback.fst : Z₂ ⟶ X₂) ≫ f₂)]: + [HasPullback f₁ ((pullback.fst : Z₂ ⟶ X₂) ≫ f₂)] : (pullbackAssoc f₁ f₂ f₃ f₄).inv ≫ pullback.fst ≫ pullback.fst = pullback.fst := by trans l₁' ≫ pullback.fst rw [← Category.assoc] @@ -2584,7 +2584,7 @@ theorem hasPushout_assoc_symm [HasPushout g₁ (g₂ ≫ f₃)] : HasPushout (g /-- The canonical isomorphism `(X₁ ⨿[Z₁] X₂) ⨿[Z₂] X₃ ≅ X₁ ⨿[Z₁] (X₂ ⨿[Z₂] X₃)`. -/ noncomputable def pushoutAssoc [HasPushout (g₃ ≫ (pushout.inr : X₂ ⟶ Y₁)) g₄] - [HasPushout g₁ (g₂ ≫ (pushout.inl : X₂ ⟶ Y₂))]: + [HasPushout g₁ (g₂ ≫ (pushout.inl : X₂ ⟶ Y₂))] : pushout (g₃ ≫ pushout.inr : _ ⟶ pushout g₁ g₂) g₄ ≅ pushout g₁ (g₂ ≫ pushout.inl : _ ⟶ pushout g₃ g₄) := (pushoutPushoutLeftIsPushout g₁ g₂ g₃ g₄).coconePointUniqueUpToIso diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean b/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean index 0f65102d9ae49..5cf300e987d05 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean @@ -606,12 +606,15 @@ def limitOfDiagramInitial {X : J} (tX : IsInitial X) (F : J ⥤ C) : dsimp; simp #align category_theory.limits.limit_of_diagram_initial CategoryTheory.Limits.limitOfDiagramInitial +instance hasLimit_of_domain_hasInitial [HasInitial J] {F : J ⥤ C} : HasLimit F := + HasLimit.mk { cone := _, isLimit := limitOfDiagramInitial (initialIsInitial) F } + -- See note [dsimp, simp] -- This is reducible to allow usage of lemmas about `cone_point_unique_up_to_iso`. /-- For a functor `F : J ⥤ C`, if `J` has an initial object then the image of it is isomorphic to the limit of `F`. -/ @[reducible] -def limitOfInitial (F : J ⥤ C) [HasInitial J] [HasLimit F] : limit F ≅ F.obj (⊥_ J) := +def limitOfInitial (F : J ⥤ C) [HasInitial J] : limit F ≅ F.obj (⊥_ J) := IsLimit.conePointUniqueUpToIso (limit.isLimit _) (limitOfDiagramInitial initialIsInitial F) #align category_theory.limits.limit_of_initial CategoryTheory.Limits.limitOfInitial @@ -638,12 +641,16 @@ def limitOfDiagramTerminal {X : J} (hX : IsTerminal X) (F : J ⥤ C) lift S := S.π.app _ #align category_theory.limits.limit_of_diagram_terminal CategoryTheory.Limits.limitOfDiagramTerminal +instance hasLimit_of_domain_hasTerminal [HasTerminal J] {F : J ⥤ C} + [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : HasLimit F := + HasLimit.mk { cone := _, isLimit := limitOfDiagramTerminal (terminalIsTerminal) F } + -- This is reducible to allow usage of lemmas about `cone_point_unique_up_to_iso`. /-- For a functor `F : J ⥤ C`, if `J` has a terminal object and all the morphisms in the diagram are isomorphisms, then the image of the terminal object is isomorphic to the limit of `F`. -/ @[reducible] -def limitOfTerminal (F : J ⥤ C) [HasTerminal J] [HasLimit F] - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : limit F ≅ F.obj (⊤_ J) := +def limitOfTerminal (F : J ⥤ C) [HasTerminal J] [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : + limit F ≅ F.obj (⊤_ J) := IsLimit.conePointUniqueUpToIso (limit.isLimit _) (limitOfDiagramTerminal terminalIsTerminal F) #align category_theory.limits.limit_of_terminal CategoryTheory.Limits.limitOfTerminal @@ -670,11 +677,14 @@ def colimitOfDiagramTerminal {X : J} (tX : IsTerminal X) (F : J ⥤ C) : simp #align category_theory.limits.colimit_of_diagram_terminal CategoryTheory.Limits.colimitOfDiagramTerminal +instance hasColimit_of_domain_hasTerminal [HasTerminal J] {F : J ⥤ C} : HasColimit F := + HasColimit.mk { cocone := _, isColimit := colimitOfDiagramTerminal (terminalIsTerminal) F } + -- This is reducible to allow usage of lemmas about `cocone_point_unique_up_to_iso`. /-- For a functor `F : J ⥤ C`, if `J` has a terminal object then the image of it is isomorphic to the colimit of `F`. -/ @[reducible] -def colimitOfTerminal (F : J ⥤ C) [HasTerminal J] [HasColimit F] : colimit F ≅ F.obj (⊤_ J) := +def colimitOfTerminal (F : J ⥤ C) [HasTerminal J] : colimit F ≅ F.obj (⊤_ J) := IsColimit.coconePointUniqueUpToIso (colimit.isColimit _) (colimitOfDiagramTerminal terminalIsTerminal F) #align category_theory.limits.colimit_of_terminal CategoryTheory.Limits.colimitOfTerminal @@ -702,12 +712,16 @@ def colimitOfDiagramInitial {X : J} (hX : IsInitial X) (F : J ⥤ C) desc S := S.ι.app _ #align category_theory.limits.colimit_of_diagram_initial CategoryTheory.Limits.colimitOfDiagramInitial +instance hasColimit_of_domain_hasInitial [HasInitial J] {F : J ⥤ C} + [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : HasColimit F := + HasColimit.mk { cocone := _, isColimit := colimitOfDiagramInitial (initialIsInitial) F } + -- This is reducible to allow usage of lemmas about `cocone_point_unique_up_to_iso`. /-- For a functor `F : J ⥤ C`, if `J` has an initial object and all the morphisms in the diagram are isomorphisms, then the image of the initial object is isomorphic to the colimit of `F`. -/ @[reducible] -def colimitOfInitial (F : J ⥤ C) [HasInitial J] [HasColimit F] - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : colimit F ≅ F.obj (⊥_ J) := +def colimitOfInitial (F : J ⥤ C) [HasInitial J] [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : + colimit F ≅ F.obj (⊥_ J) := IsColimit.coconePointUniqueUpToIso (colimit.isColimit _) (colimitOfDiagramInitial initialIsInitial _) #align category_theory.limits.colimit_of_initial CategoryTheory.Limits.colimitOfInitial @@ -719,7 +733,7 @@ theorem isIso_π_of_isInitial {j : J} (I : IsInitial j) (F : J ⥤ C) [HasLimit ⟨⟨limit.lift _ (coneOfDiagramInitial I F), ⟨by ext; simp, by simp⟩⟩⟩ #align category_theory.limits.is_iso_π_of_is_initial CategoryTheory.Limits.isIso_π_of_isInitial -instance isIso_π_initial [HasInitial J] (F : J ⥤ C) [HasLimit F] : IsIso (limit.π F (⊥_ J)) := +instance isIso_π_initial [HasInitial J] (F : J ⥤ C) : IsIso (limit.π F (⊥_ J)) := isIso_π_of_isInitial initialIsInitial F #align category_theory.limits.is_iso_π_initial CategoryTheory.Limits.isIso_π_initial @@ -728,8 +742,8 @@ theorem isIso_π_of_isTerminal {j : J} (I : IsTerminal j) (F : J ⥤ C) [HasLimi ⟨⟨limit.lift _ (coneOfDiagramTerminal I F), by ext; simp, by simp⟩⟩ #align category_theory.limits.is_iso_π_of_is_terminal CategoryTheory.Limits.isIso_π_of_isTerminal -instance isIso_π_terminal [HasTerminal J] (F : J ⥤ C) [HasLimit F] - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : IsIso (limit.π F (⊤_ J)) := +instance isIso_π_terminal [HasTerminal J] (F : J ⥤ C) [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : + IsIso (limit.π F (⊤_ J)) := isIso_π_of_isTerminal terminalIsTerminal F #align category_theory.limits.is_iso_π_terminal CategoryTheory.Limits.isIso_π_terminal @@ -740,7 +754,7 @@ theorem isIso_ι_of_isTerminal {j : J} (I : IsTerminal j) (F : J ⥤ C) [HasColi ⟨⟨colimit.desc _ (coconeOfDiagramTerminal I F), ⟨by simp, by ext; simp⟩⟩⟩ #align category_theory.limits.is_iso_ι_of_is_terminal CategoryTheory.Limits.isIso_ι_of_isTerminal -instance isIso_ι_terminal [HasTerminal J] (F : J ⥤ C) [HasColimit F] : IsIso (colimit.ι F (⊤_ J)) := +instance isIso_ι_terminal [HasTerminal J] (F : J ⥤ C) : IsIso (colimit.ι F (⊤_ J)) := isIso_ι_of_isTerminal terminalIsTerminal F #align category_theory.limits.is_iso_ι_terminal CategoryTheory.Limits.isIso_ι_terminal @@ -755,8 +769,8 @@ theorem isIso_ι_of_isInitial {j : J} (I : IsInitial j) (F : J ⥤ C) [HasColimi ⟩⟩ #align category_theory.limits.is_iso_ι_of_is_initial CategoryTheory.Limits.isIso_ι_of_isInitial -instance isIso_ι_initial [HasInitial J] (F : J ⥤ C) [HasColimit F] - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : IsIso (colimit.ι F (⊥_ J)) := +instance isIso_ι_initial [HasInitial J] (F : J ⥤ C) [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : + IsIso (colimit.ι F (⊥_ J)) := isIso_ι_of_isInitial initialIsInitial F #align category_theory.limits.is_iso_ι_initial CategoryTheory.Limits.isIso_ι_initial diff --git a/Mathlib/CategoryTheory/Limits/VanKampen.lean b/Mathlib/CategoryTheory/Limits/VanKampen.lean new file mode 100644 index 0000000000000..b95b00a45ef28 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/VanKampen.lean @@ -0,0 +1,783 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.CategoryTheory.Limits.Shapes.CommSq +import Mathlib.CategoryTheory.Limits.Shapes.StrictInitial +import Mathlib.CategoryTheory.Limits.Shapes.Types +import Mathlib.Topology.Category.TopCat.Limits.Pullbacks +import Mathlib.CategoryTheory.Limits.FunctorCategory +import Mathlib.CategoryTheory.Limits.Constructions.FiniteProductsOfBinaryProducts +import Mathlib.CategoryTheory.Adjunction.FullyFaithful + +#align_import category_theory.extensive from "leanprover-community/mathlib"@"178a32653e369dce2da68dc6b2694e385d484ef1" + +/-! + +# Universal colimits and van Kampen colimits + +## Main definitions +- `CategoryTheory.IsUniversalColimit`: A (colimit) cocone over a diagram `F : J ⥤ C` is universal + if it is stable under pullbacks. +- `CategoryTheory.IsVanKampenColimit`: A (colimit) cocone over a diagram `F : J ⥤ C` is van + Kampen if for every cocone `c'` over the pullback of the diagram `F' : J ⥤ C'`, + `c'` is colimiting iff `c'` is the pullback of `c`. + +## References +- https://ncatlab.org/nlab/show/van+Kampen+colimit +- [Stephen Lack and Paweł Sobociński, Adhesive Categories][adhesive2004] + +-/ + + +open CategoryTheory.Limits + +namespace CategoryTheory + +universe v' u' v u + +variable {J : Type v'} [Category.{u'} J] {C : Type u} [Category.{v} C] +variable {K : Type*} [Category K] {D : Type*} [Category D] + +section NatTrans + +/-- A natural transformation is equifibered if every commutative square of the following form is +a pullback. +``` +F(X) → F(Y) + ↓ ↓ +G(X) → G(Y) +``` +-/ +def NatTrans.Equifibered {F G : J ⥤ C} (α : F ⟶ G) : Prop := + ∀ ⦃i j : J⦄ (f : i ⟶ j), IsPullback (F.map f) (α.app i) (α.app j) (G.map f) +#align category_theory.nat_trans.equifibered CategoryTheory.NatTrans.Equifibered + +theorem NatTrans.equifibered_of_isIso {F G : J ⥤ C} (α : F ⟶ G) [IsIso α] : Equifibered α := + fun _ _ f => IsPullback.of_vert_isIso ⟨NatTrans.naturality _ f⟩ +#align category_theory.nat_trans.equifibered_of_is_iso CategoryTheory.NatTrans.equifibered_of_isIso + +theorem NatTrans.Equifibered.comp {F G H : J ⥤ C} {α : F ⟶ G} {β : G ⟶ H} (hα : Equifibered α) + (hβ : Equifibered β) : Equifibered (α ≫ β) := + fun _ _ f => (hα f).paste_vert (hβ f) +#align category_theory.nat_trans.equifibered.comp CategoryTheory.NatTrans.Equifibered.comp + +theorem NatTrans.Equifibered.whiskerRight {F G : J ⥤ C} {α : F ⟶ G} (hα : Equifibered α) + (H : C ⥤ D) [∀ (i j : J) (f : j ⟶ i), PreservesLimit (cospan (α.app i) (G.map f)) H] : + Equifibered (whiskerRight α H) := + fun _ _ f => (hα f).map H +#align category_theory.nat_trans.equifibered.whisker_right CategoryTheory.NatTrans.Equifibered.whiskerRight + +theorem NatTrans.Equifibered.whiskerLeft {K : Type*} [Category K] {F G : J ⥤ C} {α : F ⟶ G} + (hα : Equifibered α) (H : K ⥤ J) : Equifibered (whiskerLeft H α) := + fun _ _ f => hα (H.map f) + +theorem mapPair_equifibered {F F' : Discrete WalkingPair ⥤ C} (α : F ⟶ F') : + NatTrans.Equifibered α := by + rintro ⟨⟨⟩⟩ ⟨j⟩ ⟨⟨rfl : _ = j⟩⟩ + all_goals + dsimp; simp only [Discrete.functor_map_id] + exact IsPullback.of_horiz_isIso ⟨by simp only [Category.comp_id, Category.id_comp]⟩ +#align category_theory.map_pair_equifibered CategoryTheory.mapPair_equifibered + +theorem NatTrans.equifibered_of_discrete {ι : Type*} {F G : Discrete ι ⥤ C} + (α : F ⟶ G) : NatTrans.Equifibered α := by + rintro ⟨i⟩ ⟨j⟩ ⟨⟨rfl : i = j⟩⟩ + simp only [Discrete.functor_map_id] + exact IsPullback.of_horiz_isIso ⟨by rw [Category.id_comp, Category.comp_id]⟩ + +end NatTrans + +/-- A (colimit) cocone over a diagram `F : J ⥤ C` is universal if it is stable under pullbacks. -/ +def IsUniversalColimit {F : J ⥤ C} (c : Cocone F) : Prop := + ∀ ⦃F' : J ⥤ C⦄ (c' : Cocone F') (α : F' ⟶ F) (f : c'.pt ⟶ c.pt) + (_ : α ≫ c.ι = c'.ι ≫ (Functor.const J).map f) (_ : NatTrans.Equifibered α), + (∀ j : J, IsPullback (c'.ι.app j) (α.app j) f (c.ι.app j)) → Nonempty (IsColimit c') +#align category_theory.is_universal_colimit CategoryTheory.IsUniversalColimit + +/-- A (colimit) cocone over a diagram `F : J ⥤ C` is van Kampen if for every cocone `c'` over the +pullback of the diagram `F' : J ⥤ C'`, `c'` is colimiting iff `c'` is the pullback of `c`. + +TODO: Show that this is iff the functor `C ⥤ Catᵒᵖ` sending `x` to `C/x` preserves it. +TODO: Show that this is iff the inclusion functor `C ⥤ Span(C)` preserves it. +-/ +def IsVanKampenColimit {F : J ⥤ C} (c : Cocone F) : Prop := + ∀ ⦃F' : J ⥤ C⦄ (c' : Cocone F') (α : F' ⟶ F) (f : c'.pt ⟶ c.pt) + (_ : α ≫ c.ι = c'.ι ≫ (Functor.const J).map f) (_ : NatTrans.Equifibered α), + Nonempty (IsColimit c') ↔ ∀ j : J, IsPullback (c'.ι.app j) (α.app j) f (c.ι.app j) +#align category_theory.is_van_kampen_colimit CategoryTheory.IsVanKampenColimit + +theorem IsVanKampenColimit.isUniversal {F : J ⥤ C} {c : Cocone F} (H : IsVanKampenColimit c) : + IsUniversalColimit c := + fun _ c' α f h hα => (H c' α f h hα).mpr +#align category_theory.is_van_kampen_colimit.is_universal CategoryTheory.IsVanKampenColimit.isUniversal + +/-- A van Kampen colimit is a colimit. -/ +noncomputable def IsVanKampenColimit.isColimit {F : J ⥤ C} {c : Cocone F} + (h : IsVanKampenColimit c) : IsColimit c := by + refine' ((h c (𝟙 F) (𝟙 c.pt : _) (by rw [Functor.map_id, Category.comp_id, Category.id_comp]) + (NatTrans.equifibered_of_isIso _)).mpr fun j => _).some + haveI : IsIso (𝟙 c.pt) := inferInstance + exact IsPullback.of_vert_isIso ⟨by erw [NatTrans.id_app, Category.comp_id, Category.id_comp]⟩ +#align category_theory.is_van_kampen_colimit.is_colimit CategoryTheory.IsVanKampenColimit.isColimit + +theorem IsInitial.isVanKampenColimit [HasStrictInitialObjects C] {X : C} (h : IsInitial X) : + IsVanKampenColimit (asEmptyCocone X) := by + intro F' c' α f hf hα + have : F' = Functor.empty C := by apply Functor.hext <;> rintro ⟨⟨⟩⟩ + subst this + haveI := h.isIso_to f + refine' ⟨by rintro _ ⟨⟨⟩⟩, + fun _ => ⟨IsColimit.ofIsoColimit h (Cocones.ext (asIso f).symm <| by rintro ⟨⟨⟩⟩)⟩⟩ +#align category_theory.is_initial.is_van_kampen_colimit CategoryTheory.IsInitial.isVanKampenColimit + +section Functor + +theorem IsUniversalColimit.of_iso {F : J ⥤ C} {c c' : Cocone F} (hc : IsUniversalColimit c) + (e : c ≅ c') : IsUniversalColimit c' := by + intro F' c'' α f h hα H + have : c'.ι ≫ (Functor.const J).map e.inv.hom = c.ι := by + ext j + exact e.inv.2 j + apply hc c'' α (f ≫ e.inv.1) (by rw [Functor.map_comp, ← reassoc_of% h, this]) hα + intro j + rw [← Category.comp_id (α.app j)] + have : IsIso e.inv.hom := Functor.map_isIso (Cocones.forget _) e.inv + exact (H j).paste_vert (IsPullback.of_vert_isIso ⟨by simp⟩) + +theorem IsVanKampenColimit.of_iso {F : J ⥤ C} {c c' : Cocone F} (H : IsVanKampenColimit c) + (e : c ≅ c') : IsVanKampenColimit c' := by + intro F' c'' α f h hα + have : c'.ι ≫ (Functor.const J).map e.inv.hom = c.ι := by + ext j + exact e.inv.2 j + rw [H c'' α (f ≫ e.inv.1) (by rw [Functor.map_comp, ← reassoc_of% h, this]) hα] + apply forall_congr' + intro j + conv_lhs => rw [← Category.comp_id (α.app j)] + haveI : IsIso e.inv.hom := Functor.map_isIso (Cocones.forget _) e.inv + exact (IsPullback.of_vert_isIso ⟨by simp⟩).paste_vert_iff (NatTrans.congr_app h j).symm +#align category_theory.is_van_kampen_colimit.of_iso CategoryTheory.IsVanKampenColimit.of_iso + +theorem IsVanKampenColimit.precompose_isIso {F G : J ⥤ C} (α : F ⟶ G) [IsIso α] + {c : Cocone G} (hc : IsVanKampenColimit c) : + IsVanKampenColimit ((Cocones.precompose α).obj c) := by + intros F' c' α' f e hα + refine (hc c' (α' ≫ α) f ((Category.assoc _ _ _).trans e) + (hα.comp (NatTrans.equifibered_of_isIso _))).trans ?_ + apply forall_congr' + intro j + simp only [Functor.const_obj_obj, NatTrans.comp_app, + Cocones.precompose_obj_pt, Cocones.precompose_obj_ι] + have : IsPullback (α.app j ≫ c.ι.app j) (α.app j) (𝟙 _) (c.ι.app j) := + IsPullback.of_vert_isIso ⟨Category.comp_id _⟩ + rw [← IsPullback.paste_vert_iff this _, Category.comp_id] + exact (congr_app e j).symm + +theorem IsUniversalColimit.precompose_isIso {F G : J ⥤ C} (α : F ⟶ G) [IsIso α] + {c : Cocone G} (hc : IsUniversalColimit c) : + IsUniversalColimit ((Cocones.precompose α).obj c) := by + intros F' c' α' f e hα H + apply (hc c' (α' ≫ α) f ((Category.assoc _ _ _).trans e) + (hα.comp (NatTrans.equifibered_of_isIso _))) + intro j + simp only [Functor.const_obj_obj, NatTrans.comp_app, + Cocones.precompose_obj_pt, Cocones.precompose_obj_ι] + rw [← Category.comp_id f] + exact (H j).paste_vert (IsPullback.of_vert_isIso ⟨Category.comp_id _⟩) + +theorem IsVanKampenColimit.precompose_isIso_iff {F G : J ⥤ C} (α : F ⟶ G) [IsIso α] + {c : Cocone G} : IsVanKampenColimit ((Cocones.precompose α).obj c) ↔ IsVanKampenColimit c := + ⟨fun hc ↦ IsVanKampenColimit.of_iso (IsVanKampenColimit.precompose_isIso (inv α) hc) + (Cocones.ext (Iso.refl _) (by simp)), + IsVanKampenColimit.precompose_isIso α⟩ + +theorem IsUniversalColimit.of_mapCocone (G : C ⥤ D) {F : J ⥤ C} {c : Cocone F} + [PreservesLimitsOfShape WalkingCospan G] [ReflectsColimitsOfShape J G] + (hc : IsUniversalColimit (G.mapCocone c)) : IsUniversalColimit c := + fun F' c' α f h hα H ↦ + ⟨ReflectsColimit.reflects (hc (G.mapCocone c') (whiskerRight α G) (G.map f) + (by ext j; simpa using G.congr_map (NatTrans.congr_app h j)) + (hα.whiskerRight G) (fun j ↦ (H j).map G)).some⟩ + +theorem IsVanKampenColimit.of_mapCocone (G : C ⥤ D) {F : J ⥤ C} {c : Cocone F} + [∀ (i j : J) (X : C) (f : X ⟶ F.obj j) (g : i ⟶ j), PreservesLimit (cospan f (F.map g)) G] + [∀ (i : J) (X : C) (f : X ⟶ c.pt), PreservesLimit (cospan f (c.ι.app i)) G] + [ReflectsLimitsOfShape WalkingCospan G] + [PreservesColimitsOfShape J G] + [ReflectsColimitsOfShape J G] + (H : IsVanKampenColimit (G.mapCocone c)) : IsVanKampenColimit c := by + intro F' c' α f h hα + refine' (Iff.trans _ (H (G.mapCocone c') (whiskerRight α G) (G.map f) + (by ext j; simpa using G.congr_map (NatTrans.congr_app h j)) + (hα.whiskerRight G))).trans (forall_congr' fun j => _) + · exact ⟨fun h => ⟨isColimitOfPreserves G h.some⟩, fun h => ⟨isColimitOfReflects G h.some⟩⟩ + · exact IsPullback.map_iff G (NatTrans.congr_app h.symm j) +#align category_theory.is_van_kampen_colimit.of_map CategoryTheory.IsVanKampenColimit.of_mapCocone + +theorem IsVanKampenColimit.mapCocone_iff (G : C ⥤ D) {F : J ⥤ C} {c : Cocone F} + [IsEquivalence G] : IsVanKampenColimit (G.mapCocone c) ↔ IsVanKampenColimit c := + ⟨IsVanKampenColimit.of_mapCocone G, fun hc ↦ by + let e : F ⋙ G ⋙ Functor.inv G ≅ F := NatIso.hcomp (Iso.refl F) G.asEquivalence.unitIso.symm + apply IsVanKampenColimit.of_mapCocone G.inv + apply (IsVanKampenColimit.precompose_isIso_iff e.inv).mp + refine hc.of_iso (Cocones.ext (G.asEquivalence.unitIso.app c.pt) ?_) + simp [Functor.asEquivalence]⟩ + +theorem IsUniversalColimit.whiskerEquivalence {K : Type*} [Category K] (e : J ≌ K) + {F : K ⥤ C} {c : Cocone F} (hc : IsUniversalColimit c) : + IsUniversalColimit (c.whisker e.functor) := by + intro F' c' α f e' hα H + convert hc (c'.whisker e.inverse) (whiskerLeft e.inverse α ≫ (e.invFunIdAssoc F).hom) f ?_ + ((hα.whiskerLeft _).comp (NatTrans.equifibered_of_isIso _)) ?_ using 1 + · exact (IsColimit.whiskerEquivalenceEquiv e.symm).nonempty_congr + · convert congr_arg (whiskerLeft e.inverse) e' + ext + simp + · intro k + rw [← Category.comp_id f] + refine (H (e.inverse.obj k)).paste_vert ?_ + have : IsIso (𝟙 (Cocone.whisker e.functor c).pt) := inferInstance + exact IsPullback.of_vert_isIso ⟨by simp⟩ + +theorem IsUniversalColimit.whiskerEquivalence_iff {K : Type*} [Category K] (e : J ≌ K) + {F : K ⥤ C} {c : Cocone F} : + IsUniversalColimit (c.whisker e.functor) ↔ IsUniversalColimit c := + ⟨fun hc ↦ ((hc.whiskerEquivalence e.symm).precompose_isIso (e.invFunIdAssoc F).inv).of_iso + (Cocones.ext (Iso.refl _) (by simp)), IsUniversalColimit.whiskerEquivalence e⟩ + +theorem IsVanKampenColimit.whiskerEquivalence {K : Type*} [Category K] (e : J ≌ K) + {F : K ⥤ C} {c : Cocone F} (hc : IsVanKampenColimit c) : + IsVanKampenColimit (c.whisker e.functor) := by + intro F' c' α f e' hα + convert hc (c'.whisker e.inverse) (whiskerLeft e.inverse α ≫ (e.invFunIdAssoc F).hom) f ?_ + ((hα.whiskerLeft _).comp (NatTrans.equifibered_of_isIso _)) using 1 + · exact (IsColimit.whiskerEquivalenceEquiv e.symm).nonempty_congr + · simp only [Functor.const_obj_obj, Functor.comp_obj, Cocone.whisker_pt, Cocone.whisker_ι, + whiskerLeft_app, NatTrans.comp_app, Equivalence.invFunIdAssoc_hom_app, Functor.id_obj] + constructor + · intro H k + rw [← Category.comp_id f] + refine (H (e.inverse.obj k)).paste_vert ?_ + have : IsIso (𝟙 (Cocone.whisker e.functor c).pt) := inferInstance + exact IsPullback.of_vert_isIso ⟨by simp⟩ + · intro H j + have : α.app j = F'.map (e.unit.app _) ≫ α.app _ ≫ F.map (e.counit.app (e.functor.obj j)) + · simp [← Functor.map_comp] + rw [← Category.id_comp f, this] + refine IsPullback.paste_vert ?_ (H (e.functor.obj j)) + exact IsPullback.of_vert_isIso ⟨by simp⟩ + · ext k + simpa using congr_app e' (e.inverse.obj k) + +theorem IsVanKampenColimit.whiskerEquivalence_iff {K : Type*} [Category K] (e : J ≌ K) + {F : K ⥤ C} {c : Cocone F} : + IsVanKampenColimit (c.whisker e.functor) ↔ IsVanKampenColimit c := + ⟨fun hc ↦ ((hc.whiskerEquivalence e.symm).precompose_isIso (e.invFunIdAssoc F).inv).of_iso + (Cocones.ext (Iso.refl _) (by simp)), IsVanKampenColimit.whiskerEquivalence e⟩ + +theorem isVanKampenColimit_of_evaluation [HasPullbacks D] [HasColimitsOfShape J D] (F : J ⥤ C ⥤ D) + (c : Cocone F) (hc : ∀ x : C, IsVanKampenColimit (((evaluation C D).obj x).mapCocone c)) : + IsVanKampenColimit c := by + intro F' c' α f e hα + have := fun x => hc x (((evaluation C D).obj x).mapCocone c') (whiskerRight α _) + (((evaluation C D).obj x).map f) + (by + ext y + dsimp + exact NatTrans.congr_app (NatTrans.congr_app e y) x) + (hα.whiskerRight _) + constructor + · rintro ⟨hc'⟩ j + refine' ⟨⟨(NatTrans.congr_app e j).symm⟩, ⟨evaluationJointlyReflectsLimits _ _⟩⟩ + refine' fun x => (isLimitMapConePullbackConeEquiv _ _).symm _ + exact ((this x).mp ⟨PreservesColimit.preserves hc'⟩ _).isLimit + · exact fun H => ⟨evaluationJointlyReflectsColimits _ fun x => + ((this x).mpr fun j => (H j).map ((evaluation C D).obj x)).some⟩ +#align category_theory.is_van_kampen_colimit_of_evaluation CategoryTheory.isVanKampenColimit_of_evaluation + +end Functor + +section reflective + +theorem IsUniversalColimit.map_reflective + {Gl : C ⥤ D} {Gr : D ⥤ C} (adj : Gl ⊣ Gr) [Full Gr] [Faithful Gr] + {F : J ⥤ D} {c : Cocone (F ⋙ Gr)} + (H : IsUniversalColimit c) + [∀ X (f : X ⟶ Gl.obj c.pt), HasPullback (Gr.map f) (adj.unit.app c.pt)] + [∀ X (f : X ⟶ Gl.obj c.pt), PreservesLimit (cospan (Gr.map f) (adj.unit.app c.pt)) Gl] : + IsUniversalColimit (Gl.mapCocone c) := by + have := adj.rightAdjointPreservesLimits + have : PreservesColimitsOfSize.{u', v'} Gl := adj.leftAdjointPreservesColimits + intros F' c' α f h hα hc' + have : HasPullback (Gl.map (Gr.map f)) (Gl.map (adj.unit.app c.pt)) := + ⟨⟨_, isLimitPullbackConeMapOfIsLimit _ pullback.condition + (IsPullback.of_hasPullback _ _).isLimit⟩⟩ + let α' := α ≫ (Functor.associator _ _ _).hom ≫ whiskerLeft F adj.counit ≫ F.rightUnitor.hom + have hα' : NatTrans.Equifibered α' := hα.comp (NatTrans.equifibered_of_isIso _) + have hadj : ∀ X, Gl.map (adj.unit.app X) = inv (adj.counit.app _) + · intro X + apply IsIso.eq_inv_of_inv_hom_id + exact adj.left_triangle_components + haveI : ∀ X, IsIso (Gl.map (adj.unit.app X)) := by + simp_rw [hadj] + infer_instance + have hα'' : ∀ j, Gl.map (Gr.map $ α'.app j) = adj.counit.app _ ≫ α.app j + · intro j + rw [← cancel_mono (adj.counit.app $ F.obj j)] + dsimp + simp only [Category.comp_id, Adjunction.counit_naturality_assoc, Category.id_comp, + Adjunction.counit_naturality, Category.assoc, Functor.map_comp] + have hc'' : ∀ j, α.app j ≫ Gl.map (c.ι.app j) = c'.ι.app j ≫ f := NatTrans.congr_app h + let β := isoWhiskerLeft F' (asIso adj.counit) ≪≫ F'.rightUnitor + let c'' : Cocone (F' ⋙ Gr) + · refine + { pt := pullback (Gr.map f) (adj.unit.app _) + ι := { app := λ j ↦ pullback.lift (Gr.map $ c'.ι.app j) (Gr.map (α'.app j) ≫ c.ι.app j) ?_ + naturality := ?_ } } + · rw [← Gr.map_comp, ← hc''] + erw [← adj.unit_naturality] + rw [Gl.map_comp, hα''] + dsimp + simp only [Category.assoc, Functor.map_comp, adj.right_triangle_components_assoc] + · intros i j g + dsimp + ext + all_goals simp only [Category.comp_id, Category.id_comp, Category.assoc, + ← Functor.map_comp, pullback.lift_fst, pullback.lift_snd, ← Functor.map_comp_assoc] + · congr 1 + exact c'.w _ + · rw [α.naturality_assoc] + dsimp + rw [adj.counit_naturality, ← Category.assoc, Gr.map_comp_assoc] + congr 1 + exact c.w _ + let cf : (Cocones.precompose β.hom).obj c' ⟶ Gl.mapCocone c'' + · refine { hom := pullback.lift ?_ f ?_ ≫ (PreservesPullback.iso _ _ _).inv, w := ?_ } + exact (inv $ adj.counit.app c'.pt) + · rw [IsIso.inv_comp_eq, ← adj.counit_naturality_assoc f, ← cancel_mono (adj.counit.app $ + Gl.obj c.pt), Category.assoc, Category.assoc, adj.left_triangle_components] + erw [Category.comp_id] + rfl + · intro j + rw [← Category.assoc, Iso.comp_inv_eq] + ext + all_goals simp only [PreservesPullback.iso_hom_fst, PreservesPullback.iso_hom_snd, + pullback.lift_fst, pullback.lift_snd, Category.assoc, + Functor.mapCocone_ι_app, ← Gl.map_comp] + · rw [IsIso.comp_inv_eq, adj.counit_naturality] + dsimp + rw [Category.comp_id] + · rw [Gl.map_comp, hα'', Category.assoc, hc''] + dsimp + rw [Category.comp_id, Category.assoc] + have : cf.hom ≫ (PreservesPullback.iso _ _ _).hom ≫ pullback.fst ≫ adj.counit.app _ = 𝟙 _ + · simp only [IsIso.inv_hom_id, Iso.inv_hom_id_assoc, Category.assoc, pullback.lift_fst_assoc] + have : IsIso cf + · apply @Cocones.cocone_iso_of_hom_iso (i := ?_) + rw [← IsIso.eq_comp_inv] at this + rw [this] + infer_instance + have ⟨Hc''⟩ := H c'' (whiskerRight α' Gr) pullback.snd ?_ (hα'.whiskerRight Gr) ?_ + · exact ⟨IsColimit.precomposeHomEquiv β c' $ + (isColimitOfPreserves Gl Hc'').ofIsoColimit (asIso cf).symm⟩ + · ext j + dsimp + simp only [Category.comp_id, Category.id_comp, Category.assoc, + Functor.map_comp, pullback.lift_snd] + · intro j + apply IsPullback.of_right _ _ (IsPullback.of_hasPullback _ _) + · dsimp + simp only [Category.comp_id, Category.id_comp, Category.assoc, Functor.map_comp, + pullback.lift_fst] + rw [← Category.comp_id (Gr.map f)] + refine ((hc' j).map Gr).paste_vert (IsPullback.of_vert_isIso ⟨?_⟩) + rw [← adj.unit_naturality, Category.comp_id, ← Category.assoc, + ← Category.id_comp (Gr.map ((Gl.mapCocone c).ι.app j))] + congr 1 + rw [← cancel_mono (Gr.map (adj.counit.app (F.obj j)))] + dsimp + simp only [Category.comp_id, Adjunction.right_triangle_components, Category.id_comp, + Category.assoc] + · dsimp + simp only [Category.comp_id, Category.id_comp, Category.assoc, Functor.map_comp, + pullback.lift_snd] + +theorem IsVanKampenColimit.map_reflective [HasColimitsOfShape J C] + {Gl : C ⥤ D} {Gr : D ⥤ C} (adj : Gl ⊣ Gr) [Full Gr] [Faithful Gr] + {F : J ⥤ D} {c : Cocone (F ⋙ Gr)} (H : IsVanKampenColimit c) + [∀ X (f : X ⟶ Gl.obj c.pt), HasPullback (Gr.map f) (adj.unit.app c.pt)] + [∀ X (f : X ⟶ Gl.obj c.pt), PreservesLimit (cospan (Gr.map f) (adj.unit.app c.pt)) Gl] + [∀ X i (f : X ⟶ c.pt), PreservesLimit (cospan f (c.ι.app i)) Gl] : + IsVanKampenColimit (Gl.mapCocone c) := by + have := adj.rightAdjointPreservesLimits + have : PreservesColimitsOfSize.{u', v'} Gl := adj.leftAdjointPreservesColimits + intro F' c' α f h hα + refine ⟨?_, H.isUniversal.map_reflective adj c' α f h hα⟩ + intro ⟨hc'⟩ j + let α' := α ≫ (Functor.associator _ _ _).hom ≫ whiskerLeft F adj.counit ≫ F.rightUnitor.hom + have hα' : NatTrans.Equifibered α' := hα.comp (NatTrans.equifibered_of_isIso _) + have hα'' : ∀ j, Gl.map (Gr.map $ α'.app j) = adj.counit.app _ ≫ α.app j + · intro j + rw [← cancel_mono (adj.counit.app $ F.obj j)] + dsimp + simp only [Category.comp_id, Adjunction.counit_naturality_assoc, Category.id_comp, + Adjunction.counit_naturality, Category.assoc, Functor.map_comp] + let β := isoWhiskerLeft F' (asIso adj.counit) ≪≫ F'.rightUnitor + let hl := (IsColimit.precomposeHomEquiv β c').symm hc' + let hr := isColimitOfPreserves Gl (colimit.isColimit $ F' ⋙ Gr) + have : α.app j = β.inv.app _ ≫ Gl.map (Gr.map $ α'.app j) + · rw [hα''] + simp + rw [this] + have : f = (hl.coconePointUniqueUpToIso hr).hom ≫ + Gl.map (colimit.desc _ ⟨_, whiskerRight α' Gr ≫ c.2⟩) + · symm + convert @IsColimit.coconePointUniqueUpToIso_hom_desc _ _ _ _ ((F' ⋙ Gr) ⋙ Gl) + (Gl.mapCocone ⟨_, (whiskerRight α' Gr ≫ c.2 : _)⟩) _ _ hl hr using 2 + · apply hr.hom_ext + intro j + rw [hr.fac, Functor.mapCocone_ι_app, ← Gl.map_comp, colimit.cocone_ι, colimit.ι_desc] + rfl + · clear_value α' + apply hl.hom_ext + intro j + rw [hl.fac] + dsimp + simp only [Category.comp_id, hα'', Category.assoc, Gl.map_comp] + congr 1 + exact (NatTrans.congr_app h j).symm + rw [this] + have := ((H (colimit.cocone $ F' ⋙ Gr) (whiskerRight α' Gr) + (colimit.desc _ ⟨_, whiskerRight α' Gr ≫ c.2⟩) ?_ (hα'.whiskerRight Gr)).mp + ⟨(getColimitCocone $ F' ⋙ Gr).2⟩ j).map Gl + convert IsPullback.paste_vert _ this + refine IsPullback.of_vert_isIso ⟨?_⟩ + rw [← IsIso.inv_comp_eq, ← Category.assoc, NatIso.inv_inv_app] + exact IsColimit.comp_coconePointUniqueUpToIso_hom hl hr _ + · clear_value α' + ext j + simp + +end reflective + +section Initial + +theorem hasStrictInitial_of_isUniversal [HasInitial C] + (H : IsUniversalColimit (BinaryCofan.mk (𝟙 (⊥_ C)) (𝟙 (⊥_ C)))) : HasStrictInitialObjects C := + hasStrictInitialObjects_of_initial_is_strict + (by + intro A f + suffices IsColimit (BinaryCofan.mk (𝟙 A) (𝟙 A)) by + obtain ⟨l, h₁, h₂⟩ := Limits.BinaryCofan.IsColimit.desc' this (f ≫ initial.to A) (𝟙 A) + rcases(Category.id_comp _).symm.trans h₂ with rfl + exact ⟨⟨_, ((Category.id_comp _).symm.trans h₁).symm, initialIsInitial.hom_ext _ _⟩⟩ + refine' (H (BinaryCofan.mk (𝟙 _) (𝟙 _)) (mapPair f f) f (by ext ⟨⟨⟩⟩ <;> dsimp <;> simp) + (mapPair_equifibered _) _).some + rintro ⟨⟨⟩⟩ <;> dsimp <;> + exact IsPullback.of_horiz_isIso ⟨(Category.id_comp _).trans (Category.comp_id _).symm⟩) +#align category_theory.has_strict_initial_of_is_universal CategoryTheory.hasStrictInitial_of_isUniversal + +theorem isVanKampenColimit_of_isEmpty [HasStrictInitialObjects C] [IsEmpty J] {F : J ⥤ C} + (c : Cocone F) (hc : IsColimit c) : IsVanKampenColimit c := by + have : IsInitial c.pt + · have := (IsColimit.precomposeInvEquiv (Functor.uniqueFromEmpty _) _).symm + (hc.whiskerEquivalence (equivalenceOfIsEmpty (Discrete PEmpty.{1}) J)) + exact IsColimit.ofIsoColimit this (Cocones.ext (Iso.refl c.pt) (fun {X} ↦ isEmptyElim X)) + replace this := IsInitial.isVanKampenColimit this + apply (IsVanKampenColimit.whiskerEquivalence_iff + (equivalenceOfIsEmpty (Discrete PEmpty.{1}) J)).mp + exact (this.precompose_isIso (Functor.uniqueFromEmpty + ((equivalenceOfIsEmpty (Discrete PEmpty.{1}) J).functor ⋙ F)).hom).of_iso + (Cocones.ext (Iso.refl _) (by simp)) + +end Initial + +section BinaryCoproduct + +variable {X Y : C} + +theorem BinaryCofan.isVanKampen_iff (c : BinaryCofan X Y) : + IsVanKampenColimit c ↔ + ∀ {X' Y' : C} (c' : BinaryCofan X' Y') (αX : X' ⟶ X) (αY : Y' ⟶ Y) (f : c'.pt ⟶ c.pt) + (_ : αX ≫ c.inl = c'.inl ≫ f) (_ : αY ≫ c.inr = c'.inr ≫ f), + Nonempty (IsColimit c') ↔ IsPullback c'.inl αX f c.inl ∧ IsPullback c'.inr αY f c.inr := by + constructor + · introv H hαX hαY + rw [H c' (mapPair αX αY) f (by ext ⟨⟨⟩⟩ <;> dsimp <;> assumption) (mapPair_equifibered _)] + constructor + · intro H + exact ⟨H _, H _⟩ + · rintro H ⟨⟨⟩⟩ + exacts [H.1, H.2] + · introv H F' hα h + let X' := F'.obj ⟨WalkingPair.left⟩ + let Y' := F'.obj ⟨WalkingPair.right⟩ + have : F' = pair X' Y' := by + apply Functor.hext + · rintro ⟨⟨⟩⟩ <;> rfl + · rintro ⟨⟨⟩⟩ ⟨j⟩ ⟨⟨rfl : _ = j⟩⟩ <;> simp + clear_value X' Y' + subst this + change BinaryCofan X' Y' at c' + rw [H c' _ _ _ (NatTrans.congr_app hα ⟨WalkingPair.left⟩) + (NatTrans.congr_app hα ⟨WalkingPair.right⟩)] + constructor + · rintro H ⟨⟨⟩⟩ + exacts [H.1, H.2] + · intro H + exact ⟨H _, H _⟩ +#align category_theory.binary_cofan.is_van_kampen_iff CategoryTheory.BinaryCofan.isVanKampen_iff + +theorem BinaryCofan.isVanKampen_mk {X Y : C} (c : BinaryCofan X Y) + (cofans : ∀ X Y : C, BinaryCofan X Y) (colimits : ∀ X Y, IsColimit (cofans X Y)) + (cones : ∀ {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z), PullbackCone f g) + (limits : ∀ {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z), IsLimit (cones f g)) + (h₁ : ∀ {X' Y' : C} (αX : X' ⟶ X) (αY : Y' ⟶ Y) (f : (cofans X' Y').pt ⟶ c.pt) + (_ : αX ≫ c.inl = (cofans X' Y').inl ≫ f) (_ : αY ≫ c.inr = (cofans X' Y').inr ≫ f), + IsPullback (cofans X' Y').inl αX f c.inl ∧ IsPullback (cofans X' Y').inr αY f c.inr) + (h₂ : ∀ {Z : C} (f : Z ⟶ c.pt), + IsColimit (BinaryCofan.mk (cones f c.inl).fst (cones f c.inr).fst)) : + IsVanKampenColimit c := by + rw [BinaryCofan.isVanKampen_iff] + introv hX hY + constructor + · rintro ⟨h⟩ + let e := h.coconePointUniqueUpToIso (colimits _ _) + obtain ⟨hl, hr⟩ := h₁ αX αY (e.inv ≫ f) (by simp [hX]) (by simp [hY]) + constructor + · rw [← Category.id_comp αX, ← Iso.hom_inv_id_assoc e f] + haveI : IsIso (𝟙 X') := inferInstance + have : c'.inl ≫ e.hom = 𝟙 X' ≫ (cofans X' Y').inl := by + dsimp + simp + exact (IsPullback.of_vert_isIso ⟨this⟩).paste_vert hl + · rw [← Category.id_comp αY, ← Iso.hom_inv_id_assoc e f] + haveI : IsIso (𝟙 Y') := inferInstance + have : c'.inr ≫ e.hom = 𝟙 Y' ≫ (cofans X' Y').inr := by + dsimp + simp + exact (IsPullback.of_vert_isIso ⟨this⟩).paste_vert hr + · rintro ⟨H₁, H₂⟩ + refine' ⟨IsColimit.ofIsoColimit _ <| (isoBinaryCofanMk _).symm⟩ + let e₁ : X' ≅ _ := H₁.isLimit.conePointUniqueUpToIso (limits _ _) + let e₂ : Y' ≅ _ := H₂.isLimit.conePointUniqueUpToIso (limits _ _) + have he₁ : c'.inl = e₁.hom ≫ (cones f c.inl).fst := by simp + have he₂ : c'.inr = e₂.hom ≫ (cones f c.inr).fst := by simp + rw [he₁, he₂] + apply BinaryCofan.isColimitCompRightIso (BinaryCofan.mk _ _) + apply BinaryCofan.isColimitCompLeftIso (BinaryCofan.mk _ _) + exact h₂ f +#align category_theory.binary_cofan.is_van_kampen_mk CategoryTheory.BinaryCofan.isVanKampen_mk + +theorem BinaryCofan.mono_inr_of_isVanKampen [HasInitial C] {X Y : C} {c : BinaryCofan X Y} + (h : IsVanKampenColimit c) : Mono c.inr := by + refine' PullbackCone.mono_of_isLimitMkIdId _ (IsPullback.isLimit _) + refine' (h (BinaryCofan.mk (initial.to Y) (𝟙 Y)) (mapPair (initial.to X) (𝟙 Y)) c.inr _ + (mapPair_equifibered _)).mp ⟨_⟩ ⟨WalkingPair.right⟩ + · ext ⟨⟨⟩⟩ <;> dsimp; simp + · exact ((BinaryCofan.isColimit_iff_isIso_inr initialIsInitial _).mpr (by + dsimp + infer_instance)).some +#align category_theory.binary_cofan.mono_inr_of_is_van_kampen CategoryTheory.BinaryCofan.mono_inr_of_isVanKampen + +theorem BinaryCofan.isPullback_initial_to_of_isVanKampen [HasInitial C] {c : BinaryCofan X Y} + (h : IsVanKampenColimit c) : IsPullback (initial.to _) (initial.to _) c.inl c.inr := by + refine' ((h (BinaryCofan.mk (initial.to Y) (𝟙 Y)) (mapPair (initial.to X) (𝟙 Y)) c.inr _ + (mapPair_equifibered _)).mp ⟨_⟩ ⟨WalkingPair.left⟩).flip + · ext ⟨⟨⟩⟩ <;> dsimp; simp + · exact ((BinaryCofan.isColimit_iff_isIso_inr initialIsInitial _).mpr (by + dsimp + infer_instance)).some +#align category_theory.binary_cofan.is_pullback_initial_to_of_is_van_kampen CategoryTheory.BinaryCofan.isPullback_initial_to_of_isVanKampen + +end BinaryCoproduct + +section FiniteCoproducts + +theorem isUniversalColimit_extendCofan {n : ℕ} (f : Fin (n + 1) → C) + {c₁ : Cofan fun i : Fin n ↦ f i.succ} {c₂ : BinaryCofan (f 0) c₁.pt} + (t₁ : IsUniversalColimit c₁) (t₂ : IsUniversalColimit c₂) + [∀ {Z} (i : Z ⟶ c₂.pt), HasPullback c₂.inr i] : + IsUniversalColimit (extendCofan c₁ c₂) := by + intro F c α i e hα H + let F' : Fin (n + 1) → C := F.obj ∘ Discrete.mk + have : F = Discrete.functor F' + · apply Functor.hext + · exact fun i ↦ rfl + · rintro ⟨i⟩ ⟨j⟩ ⟨⟨rfl : i = j⟩⟩ + simp + have t₁' := @t₁ (Discrete.functor (fun j ↦ F.obj ⟨j.succ⟩)) + (Cofan.mk (pullback c₂.inr i) <| fun j ↦ pullback.lift (α.app _ ≫ c₁.inj _) (c.ι.app _) ?_) + (Discrete.natTrans <| fun i ↦ α.app _) pullback.fst ?_ (NatTrans.equifibered_of_discrete _) ?_ + rotate_left + · simpa only [Functor.const_obj_obj, pair_obj_right, Discrete.functor_obj, Category.assoc, + extendCofan_pt, Functor.const_obj_obj, NatTrans.comp_app, extendCofan_ι_app, + Fin.cases_succ, Functor.const_map_app] using congr_app e ⟨j.succ⟩ + · ext j + dsimp + simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Cofan.inj] + · intro j + simp only [pair_obj_right, Functor.const_obj_obj, Discrete.functor_obj, id_eq, + extendCofan_pt, eq_mpr_eq_cast, Cofan.mk_pt, Cofan.mk_ι_app, Discrete.natTrans_app] + refine IsPullback.of_right ?_ ?_ (IsPullback.of_hasPullback (BinaryCofan.inr c₂) i).flip + · simp only [Functor.const_obj_obj, pair_obj_right, limit.lift_π, + PullbackCone.mk_pt, PullbackCone.mk_π_app] + exact H _ + · simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Cofan.inj] + obtain ⟨H₁⟩ := t₁' + have t₂' := @t₂ (pair (F.obj ⟨0⟩) (pullback c₂.inr i)) (BinaryCofan.mk (c.ι.app ⟨0⟩) pullback.snd) + (mapPair (α.app _) pullback.fst) i ?_ (mapPair_equifibered _) ?_ + rotate_left + · ext ⟨⟨⟩⟩ + · simpa [mapPair] using congr_app e ⟨0⟩ + · simpa using pullback.condition + · rintro ⟨⟨⟩⟩ + · simp only [pair_obj_right, Functor.const_obj_obj, pair_obj_left, BinaryCofan.mk_pt, + BinaryCofan.ι_app_left, BinaryCofan.mk_inl, mapPair_left] + exact H ⟨0⟩ + · simp only [pair_obj_right, Functor.const_obj_obj, BinaryCofan.mk_pt, BinaryCofan.ι_app_right, + BinaryCofan.mk_inr, mapPair_right] + exact (IsPullback.of_hasPullback (BinaryCofan.inr c₂) i).flip + obtain ⟨H₂⟩ := t₂' + clear_value F' + subst this + refine ⟨IsColimit.ofIsoColimit (extendCofanIsColimit + (fun i ↦ (Discrete.functor F').obj ⟨i⟩) H₁ H₂) <| Cocones.ext (Iso.refl _) ?_⟩ + dsimp + rintro ⟨j⟩ + simp only [Discrete.functor_obj, limit.lift_π, PullbackCone.mk_pt, + PullbackCone.mk_π_app, Category.comp_id] + induction' j using Fin.inductionOn + · simp only [Fin.cases_zero] + · simp only [Fin.cases_succ] + +theorem isVanKampenColimit_extendCofan {n : ℕ} (f : Fin (n + 1) → C) + {c₁ : Cofan fun i : Fin n ↦ f i.succ} {c₂ : BinaryCofan (f 0) c₁.pt} + (t₁ : IsVanKampenColimit c₁) (t₂ : IsVanKampenColimit c₂) + [∀ {Z} (i : Z ⟶ c₂.pt), HasPullback c₂.inr i] + [HasFiniteCoproducts C] : + IsVanKampenColimit (extendCofan c₁ c₂) := by + intro F c α i e hα + refine ⟨?_, isUniversalColimit_extendCofan f t₁.isUniversal t₂.isUniversal c α i e hα⟩ + intro ⟨Hc⟩ ⟨j⟩ + have t₂' := (@t₂ (pair (F.obj ⟨0⟩) (∐ fun (j : Fin n) ↦ F.obj ⟨j.succ⟩)) + (BinaryCofan.mk (P := c.pt) (c.ι.app _) (Sigma.desc <| fun b ↦ c.ι.app _)) + (mapPair (α.app _) (Sigma.desc <| fun b ↦ α.app _ ≫ c₁.inj _)) i ?_ + (mapPair_equifibered _)).mp ⟨?_⟩ + rotate_left + · ext ⟨⟨⟩⟩ + · simpa only [pair_obj_left, Functor.const_obj_obj, pair_obj_right, Discrete.functor_obj, + NatTrans.comp_app, mapPair_left, BinaryCofan.ι_app_left, BinaryCofan.mk_pt, + BinaryCofan.mk_inl, Functor.const_map_app, extendCofan_pt, + extendCofan_ι_app, Fin.cases_zero] using congr_app e ⟨0⟩ + · dsimp + ext j + simpa only [colimit.ι_desc_assoc, Discrete.functor_obj, Cofan.mk_pt, Cofan.mk_ι_app, + Category.assoc, extendCofan_pt, Functor.const_obj_obj, NatTrans.comp_app, extendCofan_ι_app, + Fin.cases_succ, Functor.const_map_app] using congr_app e ⟨j.succ⟩ + · let F' : Fin (n + 1) → C := F.obj ∘ Discrete.mk + have : F = Discrete.functor F' + · apply Functor.hext + · exact fun i ↦ rfl + · rintro ⟨i⟩ ⟨j⟩ ⟨⟨rfl : i = j⟩⟩ + simp + clear_value F' + subst this + apply BinaryCofan.IsColimit.mk _ (fun {T} f₁ f₂ ↦ Hc.desc (Cofan.mk T (Fin.cases f₁ + (fun i ↦ Sigma.ι (fun (j : Fin n) ↦ (Discrete.functor F').obj ⟨j.succ⟩) _ ≫ f₂)))) + · intro T f₁ f₂ + simp only [Discrete.functor_obj, pair_obj_left, BinaryCofan.mk_pt, Functor.const_obj_obj, + BinaryCofan.ι_app_left, BinaryCofan.mk_inl, IsColimit.fac, Cofan.mk_pt, Cofan.mk_ι_app, + Fin.cases_zero] + · intro T f₁ f₂ + simp only [Discrete.functor_obj, pair_obj_right, BinaryCofan.mk_pt, Functor.const_obj_obj, + BinaryCofan.ι_app_right, BinaryCofan.mk_inr] + ext j + simp only [colimit.ι_desc_assoc, Discrete.functor_obj, Cofan.mk_pt, + Cofan.mk_ι_app, IsColimit.fac, Fin.cases_succ] + · intro T f₁ f₂ f₃ m₁ m₂ + simp at m₁ m₂ ⊢ + refine Hc.uniq (Cofan.mk T (Fin.cases f₁ + (fun i ↦ Sigma.ι (fun (j : Fin n) ↦ (Discrete.functor F').obj ⟨j.succ⟩) _ ≫ f₂))) _ ?_ + intro ⟨j⟩ + simp only [Discrete.functor_obj, Cofan.mk_pt, Functor.const_obj_obj, Cofan.mk_ι_app] + induction' j using Fin.inductionOn with j _ + · simp only [Fin.cases_zero, m₁] + · simp only [← m₂, colimit.ι_desc_assoc, Discrete.functor_obj, + Cofan.mk_pt, Cofan.mk_ι_app, Fin.cases_succ] + induction' j using Fin.inductionOn with j _ + · exact t₂' ⟨WalkingPair.left⟩ + · have t₁' := (@t₁ (Discrete.functor (fun j ↦ F.obj ⟨j.succ⟩)) (Cofan.mk _ _) (Discrete.natTrans + <| fun i ↦ α.app _) (Sigma.desc (fun j ↦ α.app _ ≫ c₁.inj _)) ?_ + (NatTrans.equifibered_of_discrete _)).mp ⟨coproductIsCoproduct _⟩ ⟨j⟩ + rotate_left + · ext ⟨j⟩ + dsimp + erw [colimit.ι_desc] -- Why? + rfl + simpa [Functor.const_obj_obj, Discrete.functor_obj, extendCofan_pt, extendCofan_ι_app, + Fin.cases_succ, BinaryCofan.mk_pt, colimit.cocone_x, Cofan.mk_pt, Cofan.mk_ι_app, + BinaryCofan.ι_app_right, BinaryCofan.mk_inr, colimit.ι_desc, + Discrete.natTrans_app] using t₁'.paste_horiz (t₂' ⟨WalkingPair.right⟩) + +theorem isPullback_of_cofan_isVanKampen [HasInitial C] {ι : Type*} {X : ι → C} + {c : Cofan X} (hc : IsVanKampenColimit c) (i j : ι) [DecidableEq ι] : + IsPullback (P := (if j = i then X i else ⊥_ C)) + (if h : j = i then eqToHom (if_pos h) else eqToHom (if_neg h) ≫ initial.to (X i)) + (if h : j = i then eqToHom ((if_pos h).trans (congr_arg X h.symm)) + else eqToHom (if_neg h) ≫ initial.to (X j)) + (Cofan.inj c i) (Cofan.inj c j) := by + refine (hc (Cofan.mk (X i) (f := fun k ↦ if k = i then X i else ⊥_ C) + (fun k ↦ if h : k = i then (eqToHom $ if_pos h) else (eqToHom $ if_neg h) ≫ initial.to _)) + (Discrete.natTrans (fun k ↦ if h : k.1 = i then (eqToHom $ (if_pos h).trans + (congr_arg X h.symm)) else (eqToHom $ if_neg h) ≫ initial.to _)) + (c.inj i) ?_ (NatTrans.equifibered_of_discrete _)).mp ⟨?_⟩ ⟨j⟩ + · ext ⟨k⟩ + simp only [Discrete.functor_obj, Functor.const_obj_obj, NatTrans.comp_app, + Discrete.natTrans_app, Cofan.mk_pt, Cofan.mk_ι_app, Functor.const_map_app] + split + · subst ‹k = i›; rfl + · simp + · refine mkCofanColimit _ (fun t ↦ (eqToHom (if_pos rfl).symm) ≫ t.inj i) ?_ ?_ + · intro t j + simp only [Cofan.mk_pt, cofan_mk_inj] + split + · subst ‹j = i›; simp + · rw [Category.assoc, ← IsIso.eq_inv_comp] + exact initialIsInitial.hom_ext _ _ + · intro t m hm + simp [← hm i] + +theorem isPullback_initial_to_of_cofan_isVanKampen [HasInitial C] {ι : Type*} {F : Discrete ι ⥤ C} + {c : Cocone F} (hc : IsVanKampenColimit c) (i j : Discrete ι) (hi : i ≠ j) : + IsPullback (initial.to _) (initial.to _) (c.ι.app i) (c.ι.app j) := by + classical + let f : ι → C := F.obj ∘ Discrete.mk + have : F = Discrete.functor f := + Functor.hext (fun i ↦ rfl) (by rintro ⟨i⟩ ⟨j⟩ ⟨⟨rfl : i = j⟩⟩; simp) + clear_value f + subst this + convert isPullback_of_cofan_isVanKampen hc i.as j.as + exact (if_neg (mt (Discrete.ext _ _) hi.symm)).symm + +theorem mono_of_cofan_isVanKampen [HasInitial C] {ι : Type*} {F : Discrete ι ⥤ C} + {c : Cocone F} (hc : IsVanKampenColimit c) (i : Discrete ι) : Mono (c.ι.app i) := by + classical + let f : ι → C := F.obj ∘ Discrete.mk + have : F = Discrete.functor f := + Functor.hext (fun i ↦ rfl) (by rintro ⟨i⟩ ⟨j⟩ ⟨⟨rfl : i = j⟩⟩; simp) + clear_value f + subst this + refine' PullbackCone.mono_of_isLimitMkIdId _ (IsPullback.isLimit _) + nth_rw 1 [← Category.id_comp (c.ι.app i)] + convert IsPullback.paste_vert _ (isPullback_of_cofan_isVanKampen hc i.as i.as) + swap + · exact (eqToHom (if_pos rfl).symm) + · simp + · exact IsPullback.of_vert_isIso ⟨by simp⟩ + +end FiniteCoproducts + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean b/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean new file mode 100644 index 0000000000000..e1bb1ce977f88 --- /dev/null +++ b/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean @@ -0,0 +1,338 @@ +/- +Copyright (c) 2023 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Localization.Opposite + +/-! +# Calculus of fractions + +Following the definitions by [Gabriel and Zisman][gabriel-zisman-1967], +given a morphism property `W : MorphismProperty C` on a category `C`, +we introduce the class `W.HasLeftCalculusOfFractions`. The main +result (TODO) is that if `L : C ⥤ D` is a localization functor for `W`, +then for any morphism `L.obj X ⟶ L.obj Y` in `D`, there exists an auxiliary +object `Y' : C` and morphisms `g : X ⟶ Y'` and `s : Y ⟶ Y'`, with `W s`, such +that the given morphism is a sort of fraction `g / s`, or more precisely of +the form `L.map g ≫ (Localization.isoOfHom L W s hs).inv`. + +## References + +* [P. Gabriel, M. Zisman, *Calculus of fractions and homotopy theory*][gabriel-zisman-1967] + +-/ + +namespace CategoryTheory + +open Category + +namespace MorphismProperty + +variable {C D : Type _} [Category C] [Category D] + +/-- A left fraction from `X : C` to `Y : C` for `W : MorphismProperty C` consists of the +datum of an object `Y' : C` and maps `f : X ⟶ Y'` and `s : Y ⟶ Y'` such that `W s`. -/ +structure LeftFraction (W : MorphismProperty C) (X Y : C) where + /-- the auxiliary object of a left fraction -/ + {Y' : C} + /-- the numerator of a left fraction -/ + f : X ⟶ Y' + /-- the denominator of a left fraction -/ + s : Y ⟶ Y' + /-- the condition that the denominator belongs to the given morphism property -/ + hs : W s + +namespace LeftFraction + +variable (W : MorphismProperty C) {X Y : C} + +/-- The left fraction from `X` to `Y` given by a morphism `f : X ⟶ Y`. -/ +@[simps] +def ofHom (f : X ⟶ Y) [W.ContainsIdentities] : + W.LeftFraction X Y := mk f (𝟙 Y) (W.id_mem Y) + +variable {W} + +/-- The left fraction from `X` to `Y` given by a morphism `s : Y ⟶ X` such that `W s`. -/ +@[simps] +def ofInv (s : Y ⟶ X) (hs : W s) : + W.LeftFraction X Y := mk (𝟙 X) s hs + +/-- If `φ : W.LeftFraction X Y` and `L` is a functor which inverts `W`, this is the +induced morphism `L.obj X ⟶ L.obj Y` -/ +noncomputable def map (φ : W.LeftFraction X Y) (L : C ⥤ D) (hL : W.IsInvertedBy L) : + L.obj X ⟶ L.obj Y := + have := hL _ φ.hs + L.map φ.f ≫ inv (L.map φ.s) + +@[reassoc (attr := simp)] +lemma map_comp_map_s (φ : W.LeftFraction X Y) (L : C ⥤ D) (hL : W.IsInvertedBy L) : + φ.map L hL ≫ L.map φ.s = L.map φ.f := by + letI := hL _ φ.hs + simp [map] + +variable (W) + +lemma map_ofHom (f : X ⟶ Y) (L : C ⥤ D) (hL : W.IsInvertedBy L) [W.ContainsIdentities] : + (ofHom W f).map L hL = L.map f := by + simp [map] + +@[reassoc (attr := simp)] +lemma map_ofInv_hom_id (s : Y ⟶ X) (hs : W s) (L : C ⥤ D) (hL : W.IsInvertedBy L) : + (ofInv s hs).map L hL ≫ L.map s = 𝟙 _ := by + letI := hL _ hs + simp [map] + +@[reassoc (attr := simp)] +lemma map_hom_ofInv_id (s : Y ⟶ X) (hs : W s) (L : C ⥤ D) (hL : W.IsInvertedBy L) : + L.map s ≫ (ofInv s hs).map L hL = 𝟙 _ := by + letI := hL _ hs + simp [map] + +variable {W} + +lemma cases (α : W.LeftFraction X Y) : + ∃ (Y' : C) (f : X ⟶ Y') (s : Y ⟶ Y') (hs : W s), α = LeftFraction.mk f s hs := + ⟨_, _, _, _, rfl⟩ + +end LeftFraction + +/-- A right fraction from `X : C` to `Y : C` for `W : MorphismProperty C` consists of the +datum of an object `X' : C` and maps `s : X' ⟶ X` and `f : X' ⟶ Y` such that `W s`. -/ +structure RightFraction (W : MorphismProperty C) (X Y : C) where + /-- the auxiliary object of a right fraction -/ + {X' : C} + /-- the denominator of a right fraction -/ + s : X' ⟶ X + /-- the condition that the denominator belongs to the given morphism property -/ + hs : W s + /-- the numerator of a right fraction -/ + f : X' ⟶ Y + +namespace RightFraction + +variable (W : MorphismProperty C) + +variable {X Y : C} + +/-- The right fraction from `X` to `Y` given by a morphism `f : X ⟶ Y`. -/ +@[simps] +def ofHom (f : X ⟶ Y) [W.ContainsIdentities] : + W.RightFraction X Y := mk (𝟙 X) (W.id_mem X) f + +variable {W} + +/-- The right fraction from `X` to `Y` given by a morphism `s : Y ⟶ X` such that `W s`. -/ +@[simps] +def ofInv (s : Y ⟶ X) (hs : W s) : + W.RightFraction X Y := mk s hs (𝟙 Y) + +/-- If `φ : W.RightFraction X Y` and `L` is a functor which inverts `W`, this is the +induced morphism `L.obj X ⟶ L.obj Y` -/ +noncomputable def map (φ : W.RightFraction X Y) (L : C ⥤ D) (hL : W.IsInvertedBy L) : + L.obj X ⟶ L.obj Y := + have := hL _ φ.hs + inv (L.map φ.s) ≫ L.map φ.f + +@[reassoc (attr := simp)] +lemma map_s_comp_map (φ : W.RightFraction X Y) (L : C ⥤ D) (hL : W.IsInvertedBy L) : + L.map φ.s ≫ φ.map L hL = L.map φ.f := by + letI := hL _ φ.hs + simp [map] + +variable (W) + +@[simp] +lemma map_ofHom (f : X ⟶ Y) (L : C ⥤ D) (hL : W.IsInvertedBy L) [W.ContainsIdentities] : + (ofHom W f).map L hL = L.map f := by + simp [map] + +@[reassoc (attr := simp)] +lemma map_ofInv_hom_id (s : Y ⟶ X) (hs : W s) (L : C ⥤ D) (hL : W.IsInvertedBy L) : + (ofInv s hs).map L hL ≫ L.map s = 𝟙 _ := by + letI := hL _ hs + simp [map] + +@[reassoc (attr := simp)] +lemma map_hom_ofInv_id (s : Y ⟶ X) (hs : W s) (L : C ⥤ D) (hL : W.IsInvertedBy L) : + L.map s ≫ (ofInv s hs).map L hL = 𝟙 _ := by + letI := hL _ hs + simp [map] + +variable {W} + +lemma cases (α : W.RightFraction X Y) : + ∃ (X' : C) (s : X' ⟶ X) (hs : W s) (f : X' ⟶ Y) , α = RightFraction.mk s hs f := + ⟨_, _, _, _, rfl⟩ + +end RightFraction + +variable (W : MorphismProperty C) + +/-- A multiplicative morphism property `W` has left calculus of fractions if +any right fraction can be turned into a left fraction and that two morphisms +that can be equalized by precomposition with a morphism in `W` can also +be equalized by postcomposition with a morphism in `W`. -/ +class HasLeftCalculusOfFractions extends W.IsMultiplicative : Prop where + exists_leftFraction ⦃X Y : C⦄ (φ : W.RightFraction X Y) : + ∃ (ψ : W.LeftFraction X Y), φ.f ≫ ψ.s = φ.s ≫ ψ.f + ext : ∀ ⦃X' X Y : C⦄ (f₁ f₂ : X ⟶ Y) (s : X' ⟶ X) (_ : W s) + (_ : s ≫ f₁ = s ≫ f₂), ∃ (Y' : C) (t : Y ⟶ Y') (_ : W t), f₁ ≫ t = f₂ ≫ t + +/-- A multiplicative morphism property `W` has right calculus of fractions if +any left fraction can be turned into a right fraction and that two morphisms +that can be equalized by postcomposition with a morphism in `W` can also +be equalized by precomposition with a morphism in `W`. -/ +class HasRightCalculusOfFractions extends W.IsMultiplicative : Prop where + exists_rightFraction ⦃X Y : C⦄ (φ : W.LeftFraction X Y) : + ∃ (ψ : W.RightFraction X Y), ψ.s ≫ φ.f = ψ.f ≫ φ.s + ext : ∀ ⦃X Y Y' : C⦄ (f₁ f₂ : X ⟶ Y) (s : Y ⟶ Y') (_ : W s) + (_ : f₁ ≫ s = f₂ ≫ s), ∃ (X' : C) (t : X' ⟶ X) (_ : W t), t ≫ f₁ = t ≫ f₂ + +variable {W} + +lemma RightFraction.exists_leftFraction [W.HasLeftCalculusOfFractions] {X Y : C} + (φ : W.RightFraction X Y) : ∃ (ψ : W.LeftFraction X Y), φ.f ≫ ψ.s = φ.s ≫ ψ.f := + HasLeftCalculusOfFractions.exists_leftFraction φ + +/-- A choice of a left fraction deduced from a right fraction for a morphism property `W` +when `W` has left calculus of fractions. -/ +noncomputable def RightFraction.leftFraction [W.HasLeftCalculusOfFractions] {X Y : C} + (φ : W.RightFraction X Y) : W.LeftFraction X Y := + φ.exists_leftFraction.choose + +@[reassoc] +lemma RightFraction.leftFraction_fac [W.HasLeftCalculusOfFractions] {X Y : C} + (φ : W.RightFraction X Y) : φ.f ≫ φ.leftFraction.s = φ.s ≫ φ.leftFraction.f := + φ.exists_leftFraction.choose_spec + +lemma LeftFraction.exists_rightFraction [W.HasRightCalculusOfFractions] {X Y : C} + (φ : W.LeftFraction X Y) : ∃ (ψ : W.RightFraction X Y), ψ.s ≫ φ.f = ψ.f ≫ φ.s := + HasRightCalculusOfFractions.exists_rightFraction φ + +/-- A choice of a right fraction deduced from a left fraction for a morphism property `W` +when `W` has right calculus of fractions. -/ +noncomputable def LeftFraction.rightFraction [W.HasRightCalculusOfFractions] {X Y : C} + (φ : W.LeftFraction X Y) : W.RightFraction X Y := + φ.exists_rightFraction.choose + +@[reassoc] +lemma LeftFraction.rightFraction_fac [W.HasRightCalculusOfFractions] {X Y : C} + (φ : W.LeftFraction X Y) : φ.rightFraction.s ≫ φ.f = φ.rightFraction.f ≫ φ.s := + φ.exists_rightFraction.choose_spec + +/-- The equivalence relation on left fractions for a morphism property `W`. -/ +def LeftFractionRel {X Y : C} (z₁ z₂ : W.LeftFraction X Y) : Prop := + ∃ (Z : C) (t₁ : z₁.Y' ⟶ Z) (t₂ : z₂.Y' ⟶ Z) (_ : z₁.s ≫ t₁ = z₂.s ≫ t₂) + (_ : z₁.f ≫ t₁ = z₂.f ≫ t₂), W (z₁.s ≫ t₁) + +namespace LeftFractionRel + +lemma refl {X Y : C} (z : W.LeftFraction X Y) : LeftFractionRel z z := + ⟨z.Y', 𝟙 _, 𝟙 _, rfl, rfl, by simpa only [Category.comp_id] using z.hs⟩ + +lemma symm {X Y : C} {z₁ z₂ : W.LeftFraction X Y} (h : LeftFractionRel z₁ z₂) : + LeftFractionRel z₂ z₁ := by + obtain ⟨Z, t₁, t₂, hst, hft, ht⟩ := h + exact ⟨Z, t₂, t₁, hst.symm, hft.symm, by simpa only [← hst] using ht⟩ + +lemma trans {X Y : C} {z₁ z₂ z₃ : W.LeftFraction X Y} + [HasLeftCalculusOfFractions W] + (h₁₂ : LeftFractionRel z₁ z₂) (h₂₃ : LeftFractionRel z₂ z₃) : + LeftFractionRel z₁ z₃ := by + obtain ⟨Z₄, t₁, t₂, hst, hft, ht⟩ := h₁₂ + obtain ⟨Z₅, u₂, u₃, hsu, hfu, hu⟩ := h₂₃ + obtain ⟨⟨v₄, v₅, hv₅⟩, fac⟩ := HasLeftCalculusOfFractions.exists_leftFraction + (RightFraction.mk (z₁.s ≫ t₁) ht (z₃.s ≫ u₃)) + simp only [Category.assoc] at fac + have eq : z₂.s ≫ u₂ ≫ v₅ = z₂.s ≫ t₂ ≫ v₄ := by + simpa only [← reassoc_of% hsu, reassoc_of% hst] using fac + obtain ⟨Z₇, w, hw, fac'⟩ := HasLeftCalculusOfFractions.ext _ _ _ z₂.hs eq + simp only [Category.assoc] at fac' + refine' ⟨Z₇, t₁ ≫ v₄ ≫ w, u₃ ≫ v₅ ≫ w, _, _, _⟩ + · rw [reassoc_of% fac] + · rw [reassoc_of% hft, ← fac', reassoc_of% hfu] + · rw [← reassoc_of% fac, ← reassoc_of% hsu, ← Category.assoc] + exact W.comp_mem _ _ hu (W.comp_mem _ _ hv₅ hw) + +end LeftFractionRel + +section + +variable [W.HasLeftCalculusOfFractions] (W) + +lemma equivalenceLeftFractionRel (X Y : C) : + @_root_.Equivalence (W.LeftFraction X Y) LeftFractionRel where + refl := LeftFractionRel.refl + symm := LeftFractionRel.symm + trans := LeftFractionRel.trans + +variable {W} + +namespace LeftFraction + +/-- Auxiliary definition for the composition of left fractions. -/ +@[simp] +def comp₀ {X Y Z : C} (z₁ : W.LeftFraction X Y) (z₂ : W.LeftFraction Y Z) + (z₃ : W.LeftFraction z₁.Y' z₂.Y') : + W.LeftFraction X Z := + mk (z₁.f ≫ z₃.f) (z₂.s ≫ z₃.s) (W.comp_mem _ _ z₂.hs z₃.hs) + +/-- The equivalence class of `z₁.comp₀ z₂ z₃` does not depend on the choice of `z₃` provided +they satisfy the compatibility `z₂.f ≫ z₃.s = z₁.s ≫ z₃.f`. -/ +lemma comp₀_rel {X Y Z : C} (z₁ : W.LeftFraction X Y) (z₂ : W.LeftFraction Y Z) + (z₃ z₃' : W.LeftFraction z₁.Y' z₂.Y') (h₃ : z₂.f ≫ z₃.s = z₁.s ≫ z₃.f) + (h₃' : z₂.f ≫ z₃'.s = z₁.s ≫ z₃'.f) : + LeftFractionRel (z₁.comp₀ z₂ z₃) (z₁.comp₀ z₂ z₃') := by + obtain ⟨z₄, fac⟩ := HasLeftCalculusOfFractions.exists_leftFraction + (RightFraction.mk z₃.s z₃.hs z₃'.s) + dsimp at fac + have eq : z₁.s ≫ z₃.f ≫ z₄.f = z₁.s ≫ z₃'.f ≫ z₄.s := by + rw [← reassoc_of% h₃, ← reassoc_of% h₃', fac] + obtain ⟨Y, t, ht, fac'⟩ := HasLeftCalculusOfFractions.ext _ _ _ z₁.hs eq + simp only [assoc] at fac' + refine' ⟨Y, z₄.f ≫ t, z₄.s ≫ t, _, _, _⟩ + · simp only [comp₀, assoc, reassoc_of% fac] + · simp only [comp₀, assoc, fac'] + · simp only [comp₀, assoc, ← reassoc_of% fac] + exact W.comp_mem _ _ z₂.hs (W.comp_mem _ _ z₃'.hs (W.comp_mem _ _ z₄.hs ht)) + +variable (W) + +/-- The morphisms in the constructed localized category for a morphism property `W` +that has left calculus of fractions are equivalence classes of left fractions. -/ +def Localization.Hom (X Y : C) := + Quot (LeftFractionRel : W.LeftFraction X Y → W.LeftFraction X Y → Prop) + +variable {W} + +/-- The morphism in the constructed localized category that is induced by a left fraction. -/ +def Localization.Hom.mk {X Y : C} (z : W.LeftFraction X Y) : Localization.Hom W X Y := + Quot.mk _ z + +lemma Localization.Hom.mk_surjective {X Y : C} (f : Localization.Hom W X Y) : + ∃ (z : W.LeftFraction X Y), f = mk z := by + obtain ⟨z⟩ := f + exact ⟨z, rfl⟩ + +/-- Auxiliary definition towards the definition of the composition of morphisms +in the constructed localized category for a morphism property that has +left calculus of fractions. -/ +noncomputable def comp {X Y Z : C} (z₁ : W.LeftFraction X Y) (z₂ : W.LeftFraction Y Z) : + Localization.Hom W X Z := + Localization.Hom.mk (z₁.comp₀ z₂ (RightFraction.mk z₁.s z₁.hs z₂.f).leftFraction) + +lemma comp_eq {X Y Z : C} (z₁ : W.LeftFraction X Y) (z₂ : W.LeftFraction Y Z) + (z₃ : W.LeftFraction z₁.Y' z₂.Y') (h₃ : z₂.f ≫ z₃.s = z₁.s ≫ z₃.f) : + z₁.comp z₂ = Localization.Hom.mk (z₁.comp₀ z₂ z₃) := + Quot.sound (LeftFraction.comp₀_rel _ _ _ _ + (RightFraction.leftFraction_fac (RightFraction.mk z₁.s z₁.hs z₂.f)) h₃) + +end LeftFraction + +end + +end MorphismProperty + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 42ee4e89ccf75..27e1eceffdeab 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -407,10 +407,10 @@ theorem inv_hom_id_tensor {V W X Y Z : C} (f : V ≅ W) (g : X ⟶ Y) (h : Y ⟶ #align category_theory.monoidal_category.inv_hom_id_tensor CategoryTheory.MonoidalCategory.inv_hom_id_tensor @[reassoc (attr := simp)] -theorem tensorHom_inv_id {V W X Y Z : C} (f : V ≅ W) (g : X ⟶ Y) (h : Y ⟶ Z) : +theorem tensor_hom_inv_id {V W X Y Z : C} (f : V ≅ W) (g : X ⟶ Y) (h : Y ⟶ Z) : (g ⊗ f.hom) ≫ (h ⊗ f.inv) = (g ⊗ 𝟙 V) ≫ (h ⊗ 𝟙 V) := by rw [← tensor_comp, f.hom_inv_id, comp_tensor_id] -#align category_theory.monoidal_category.tensor_hom_inv_id CategoryTheory.MonoidalCategory.tensorHom_inv_id +#align category_theory.monoidal_category.tensor_hom_inv_id CategoryTheory.MonoidalCategory.tensor_hom_inv_id @[reassoc (attr := simp)] theorem tensor_inv_hom_id {V W X Y Z : C} (f : V ≅ W) (g : X ⟶ Y) (h : Y ⟶ Z) : @@ -431,10 +431,10 @@ theorem inv_hom_id_tensor' {V W X Y Z : C} (f : V ⟶ W) [IsIso f] (g : X ⟶ Y) #align category_theory.monoidal_category.inv_hom_id_tensor' CategoryTheory.MonoidalCategory.inv_hom_id_tensor' @[reassoc (attr := simp)] -theorem tensorHom_inv_id' {V W X Y Z : C} (f : V ⟶ W) [IsIso f] (g : X ⟶ Y) (h : Y ⟶ Z) : +theorem tensor_hom_inv_id' {V W X Y Z : C} (f : V ⟶ W) [IsIso f] (g : X ⟶ Y) (h : Y ⟶ Z) : (g ⊗ f) ≫ (h ⊗ inv f) = (g ⊗ 𝟙 V) ≫ (h ⊗ 𝟙 V) := by rw [← tensor_comp, IsIso.hom_inv_id, comp_tensor_id] -#align category_theory.monoidal_category.tensor_hom_inv_id' CategoryTheory.MonoidalCategory.tensorHom_inv_id' +#align category_theory.monoidal_category.tensor_hom_inv_id' CategoryTheory.MonoidalCategory.tensor_hom_inv_id' @[reassoc (attr := simp)] theorem tensor_inv_hom_id' {V W X Y Z : C} (f : V ⟶ W) [IsIso f] (g : X ⟶ Y) (h : Y ⟶ Z) : diff --git a/Mathlib/CategoryTheory/Monoidal/Discrete.lean b/Mathlib/CategoryTheory/Monoidal/Discrete.lean index ab6f6147ccc33..641d322d3126d 100644 --- a/Mathlib/CategoryTheory/Monoidal/Discrete.lean +++ b/Mathlib/CategoryTheory/Monoidal/Discrete.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Mathlib.Algebra.Hom.Group.Defs +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.CategoryTheory.DiscreteCategory import Mathlib.CategoryTheory.Monoidal.NaturalTransformation diff --git a/Mathlib/CategoryTheory/Monoidal/Skeleton.lean b/Mathlib/CategoryTheory/Monoidal/Skeleton.lean index 20cbd2cc196ee..35cdf9bea613d 100644 --- a/Mathlib/CategoryTheory/Monoidal/Skeleton.lean +++ b/Mathlib/CategoryTheory/Monoidal/Skeleton.lean @@ -67,7 +67,7 @@ noncomputable instance instBraidedCategory [BraidedCategory C] : BraidedCategory braidedCategoryOfFullyFaithful (Monoidal.fromTransported (skeletonEquivalence C).symm) /-- -The skeleton of a braided monoidal category can be viewed as a commutative monoid, where the +The skeleton of a braided monoidal category can be viewed as a commutative monoid, where the multiplication is given by the tensor product, and satisfies the monoid axioms since it is a skeleton. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean index 9e9b5cdf1f969..5f67c79f369ab 100644 --- a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean @@ -67,8 +67,8 @@ When `P` is a monoidal predicate, the full subcategory for `P` inherits the mono -/ instance fullMonoidalSubcategory : MonoidalCategory (FullSubcategory P) := Monoidal.induced (fullSubcategoryInclusion P) - { μIsoSymm := fun X Y => eqToIso rfl - εIsoSymm := eqToIso rfl } + { μIso := fun X Y => eqToIso rfl + εIso := eqToIso rfl } #align category_theory.monoidal_category.full_monoidal_subcategory CategoryTheory.MonoidalCategory.fullMonoidalSubcategory /-- The forgetful monoidal functor from a full monoidal subcategory into the original category diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index c88a195690064..5c674a6ffa602 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -44,33 +44,33 @@ variable {D : Type u₂} [Category.{v₂} D] definitions of `⊗`, `𝟙_`, `▷`, `◁` that are preserved by `F`. -/ structure InducingFunctorData [MonoidalCategoryStruct D] (F : D ⥤ C) where - /-- Analogous to the reversed version of `CategoryTheory.LaxMonoidalFunctor.μIso` -/ - μIsoSymm : ∀ X Y, - F.obj (X ⊗ Y) ≅ F.obj X ⊗ F.obj Y + /-- Analogous to `CategoryTheory.LaxMonoidalFunctor.μIso` -/ + μIso : ∀ X Y, + F.obj X ⊗ F.obj Y ≅ F.obj (X ⊗ Y) whiskerLeft_eq : ∀ (X : D) {Y₁ Y₂ : D} (f : Y₁ ⟶ Y₂), - F.map (X ◁ f) = (μIsoSymm _ _).hom ≫ (F.obj X ◁ F.map f) ≫ (μIsoSymm _ _).inv := + F.map (X ◁ f) = (μIso _ _).inv ≫ (F.obj X ◁ F.map f) ≫ (μIso _ _).hom := by aesop_cat whiskerRight_eq : ∀ {X₁ X₂ : D} (f : X₁ ⟶ X₂) (Y : D), - F.map (f ▷ Y) = (μIsoSymm _ _).hom ≫ (F.map f ▷ F.obj Y) ≫ (μIsoSymm _ _).inv := + F.map (f ▷ Y) = (μIso _ _).inv ≫ (F.map f ▷ F.obj Y) ≫ (μIso _ _).hom := by aesop_cat tensorHom_eq : ∀ {X₁ Y₁ X₂ Y₂ : D} (f : X₁ ⟶ Y₁) (g : X₂ ⟶ Y₂), - F.map (f ⊗ g) = (μIsoSymm _ _).hom ≫ (F.map f ⊗ F.map g) ≫ (μIsoSymm _ _).inv := + F.map (f ⊗ g) = (μIso _ _).inv ≫ (F.map f ⊗ F.map g) ≫ (μIso _ _).hom := by aesop_cat - /-- Analogous to the reversed version of `CategoryTheory.LaxMonoidalFunctor.εIso` -/ - εIsoSymm : F.obj (𝟙_ _) ≅ 𝟙_ _ + /-- Analogous to `CategoryTheory.LaxMonoidalFunctor.εIso` -/ + εIso : 𝟙_ _ ≅ F.obj (𝟙_ _) associator_eq : ∀ X Y Z : D, F.map (α_ X Y Z).hom = - ((μIsoSymm _ _ ≪≫ (μIsoSymm _ _ ⊗ .refl _)) + (((μIso _ _).symm ≪≫ ((μIso _ _).symm ⊗ .refl _)) ≪≫ α_ (F.obj X) (F.obj Y) (F.obj Z) - ≪≫ ((.refl _ ⊗ (μIsoSymm _ _).symm) ≪≫ (μIsoSymm _ _).symm)).hom := + ≪≫ ((.refl _ ⊗ μIso _ _) ≪≫ μIso _ _)).hom := by aesop_cat leftUnitor_eq : ∀ X : D, F.map (λ_ X).hom = - ((μIsoSymm _ _ ≪≫ (εIsoSymm ⊗ .refl _)) ≪≫ λ_ (F.obj X)).hom := + (((μIso _ _).symm ≪≫ (εIso.symm ⊗ .refl _)) ≪≫ λ_ (F.obj X)).hom := by aesop_cat rightUnitor_eq : ∀ X : D, F.map (ρ_ X).hom = - ((μIsoSymm _ _ ≪≫ (.refl _ ⊗ εIsoSymm)) ≪≫ ρ_ (F.obj X)).hom := + (((μIso _ _).symm ≪≫ (.refl _ ⊗ εIso.symm)) ≪≫ ρ_ (F.obj X)).hom := by aesop_cat -- these are theorems so don't need docstrings (std4#217) @@ -95,53 +95,43 @@ abbrev induced [MonoidalCategoryStruct D] (F : D ⥤ C) [Faithful F] MonoidalCategory.{v₂} D where tensorHom_def {X₁ Y₁ X₂ Y₂} f g := F.map_injective <| by rw [fData.tensorHom_eq, Functor.map_comp, fData.whiskerRight_eq, fData.whiskerLeft_eq] - simp only [tensorHom_def, assoc, Iso.inv_hom_id_assoc] + simp only [tensorHom_def, assoc, Iso.hom_inv_id_assoc] tensor_id X₁ X₂ := F.map_injective <| by cases fData; aesop_cat tensor_comp {X₁ Y₁ Z₁ X₂ Y₂ Z₂} f₁ f₂ g₁ g₂ := F.map_injective <| by cases fData; aesop_cat whiskerLeft_id X Y := F.map_injective <| by simp [fData.whiskerLeft_eq] id_whiskerRight X Y := F.map_injective <| by simp [fData.whiskerRight_eq] triangle X Y := F.map_injective <| by cases fData; aesop_cat pentagon W X Y Z := F.map_injective <| by - have := MonoidalCategory.pentagon (F.obj W) (F.obj X) (F.obj Y) (F.obj Z) simp only [Functor.map_comp, fData.tensorHom_eq, fData.associator_eq, Iso.trans_assoc, - Iso.trans_hom, tensorIso_hom, Iso.refl_hom, Iso.symm_hom, Functor.map_id, comp_tensor_id, - associator_conjugation, tensor_id, assoc, id_tensor_comp, Iso.inv_hom_id_assoc, - tensor_inv_hom_id_assoc, id_comp, inv_hom_id_tensor_assoc, id_tensor_comp_tensor_id_assoc, - Iso.cancel_iso_hom_left] - congr 1 - simp only [←assoc] - congr 2 - simp only [assoc, ←tensor_comp, id_comp, Iso.inv_hom_id, tensor_id] - congr 1 - conv_rhs => rw [←tensor_id_comp_id_tensor] - simp only [assoc] - congr 1 - rw [Iso.inv_comp_eq] - conv_lhs => rw [←id_comp (𝟙 (F.obj W)), tensor_comp] - slice_lhs 0 2 => rw [this] - rw [assoc] - congr 1 - rw [←associator_naturality, tensor_id] + Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, Functor.map_id, comp_tensor_id, + associator_conjugation, tensor_id, assoc, id_tensor_comp, Iso.hom_inv_id_assoc, + tensor_hom_inv_id_assoc, id_comp, hom_inv_id_tensor_assoc, Iso.inv_hom_id_assoc, + id_tensor_comp_tensor_id_assoc, Iso.cancel_iso_inv_left] + slice_lhs 6 8 => + rw [← id_tensor_comp, hom_inv_id_tensor, tensor_id, comp_id, + tensor_id] + simp only [comp_id, assoc, pentagon_assoc, Iso.inv_hom_id_assoc, + ← associator_naturality_assoc, tensor_id, tensor_id_comp_id_tensor_assoc] leftUnitor_naturality {X Y : D} f := F.map_injective <| by have := leftUnitor_naturality (F.map f) simp only [Functor.map_comp, fData.tensorHom_eq, Functor.map_id, fData.leftUnitor_eq, - Iso.trans_assoc, Iso.trans_hom, tensorIso_hom, Iso.refl_hom, assoc, Iso.inv_hom_id_assoc, - id_tensor_comp_tensor_id_assoc, Iso.cancel_iso_hom_left] + Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, assoc, + Iso.hom_inv_id_assoc, id_tensor_comp_tensor_id_assoc, Iso.cancel_iso_inv_left] rw [←this, ←assoc, ←tensor_comp, id_comp, comp_id] rightUnitor_naturality {X Y : D} f := F.map_injective <| by have := rightUnitor_naturality (F.map f) simp only [Functor.map_comp, fData.tensorHom_eq, Functor.map_id, fData.rightUnitor_eq, - Iso.trans_assoc, Iso.trans_hom, tensorIso_hom, Iso.refl_hom, assoc, Iso.inv_hom_id_assoc, - tensor_id_comp_id_tensor_assoc, Iso.cancel_iso_hom_left] + Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, assoc, + Iso.hom_inv_id_assoc, tensor_id_comp_id_tensor_assoc, Iso.cancel_iso_inv_left] rw [←this, ←assoc, ←tensor_comp, id_comp, comp_id] associator_naturality {X₁ X₂ X₃ Y₁ Y₂ Y₃} f₁ f₂ f₃ := F.map_injective <| by have := associator_naturality (F.map f₁) (F.map f₂) (F.map f₃) simp [fData.associator_eq, fData.tensorHom_eq] - simp_rw [←assoc, ←tensor_comp, assoc, Iso.inv_hom_id, ←assoc] + simp_rw [←assoc, ←tensor_comp, assoc, Iso.hom_inv_id, ←assoc] congr 1 conv_rhs => rw [←comp_id (F.map f₁), ←id_comp (F.map f₁)] simp only [tensor_comp] - simp only [tensor_id, comp_id, assoc, tensor_inv_hom_id_assoc, id_comp] + simp only [tensor_id, comp_id, assoc, tensor_hom_inv_id_assoc, id_comp] slice_rhs 2 3 => rw [←this] simp only [← assoc, Iso.inv_hom_id, comp_id] congr 2 @@ -158,8 +148,8 @@ def fromInduced [MonoidalCategoryStruct D] (F : D ⥤ C) [Faithful F] MonoidalFunctor D C := letI := induced F fData { toFunctor := F - ε := fData.εIsoSymm.inv - μ := fun X Y => (fData.μIsoSymm X Y).inv + ε := fData.εIso.hom + μ := fun X Y => (fData.μIso X Y).hom μ_natural := by cases fData; aesop_cat associativity := by cases fData; aesop_cat left_unitality := by cases fData; aesop_cat @@ -191,8 +181,8 @@ def transportStruct (e : C ≌ D) : MonoidalCategoryStruct.{v₂} D where def transport (e : C ≌ D) : MonoidalCategory.{v₂} D := letI : MonoidalCategoryStruct.{v₂} D := transportStruct e induced e.inverse - { μIsoSymm := fun X Y => (e.unitIso.app _).symm - εIsoSymm := (e.unitIso.app _).symm } + { μIso := fun X Y => e.unitIso.app _ + εIso := e.unitIso.app _ } #align category_theory.monoidal.transport CategoryTheory.Monoidal.transport /-- A type synonym for `D`, which will carry the transported monoidal structure. -/ diff --git a/Mathlib/CategoryTheory/PEmpty.lean b/Mathlib/CategoryTheory/PEmpty.lean index 191e450550f61..7239f79887454 100644 --- a/Mathlib/CategoryTheory/PEmpty.lean +++ b/Mathlib/CategoryTheory/PEmpty.lean @@ -13,29 +13,42 @@ import Mathlib.CategoryTheory.DiscreteCategory Defines a category structure on `PEmpty`, and the unique functor `PEmpty ⥤ C` for any category `C`. -/ -universe w v u +universe w v v' u u' -- morphism levels before object levels. See note [CategoryTheory universes]. namespace CategoryTheory -namespace Functor +variable (C : Type u) [Category.{v} C] (D : Type u') [Category.{v'} D] + +instance (α : Type*) [IsEmpty α] : IsEmpty (Discrete α) := Function.isEmpty Discrete.as + +/-- The (unique) functor from an empty category. -/ +def functorOfIsEmpty [IsEmpty C] : C ⥤ D where + obj := isEmptyElim + map := fun {X} ↦ isEmptyElim X + map_id := fun {X} ↦ isEmptyElim X + map_comp := fun {X} ↦ isEmptyElim X + +variable {C D} -variable (C : Type u) [Category.{v} C] +/-- Any two functors out of an empty category are isomorphic. -/ +def Functor.isEmptyExt [IsEmpty C] (F G : C ⥤ D) : F ≅ G := + NatIso.ofComponents isEmptyElim (fun {X} ↦ isEmptyElim X) + +variable (C D) + +/-- The equivalence between two empty categories. -/ +def equivalenceOfIsEmpty [IsEmpty C] [IsEmpty D] : C ≌ D where + functor := functorOfIsEmpty C D + inverse := functorOfIsEmpty D C + unitIso := Functor.isEmptyExt _ _ + counitIso := Functor.isEmptyExt _ _ + functor_unitIso_comp := isEmptyElim /-- Equivalence between two empty categories. -/ -def emptyEquivalence : Discrete.{w} PEmpty ≌ Discrete.{v} PEmpty where - functor := - { obj := PEmpty.elim ∘ Discrete.as - map := fun {X} _ _ => X.as.elim } - inverse := - { obj := PEmpty.elim ∘ Discrete.as - map := fun {X} _ _ => X.as.elim } - unitIso := - { hom := { app := fun X => X.as.elim } - inv := { app := fun X => X.as.elim } } - counitIso := - { hom := { app := fun X => X.as.elim } - inv := { app := fun X => X.as.elim } } -#align category_theory.functor.empty_equivalence CategoryTheory.Functor.emptyEquivalence +def emptyEquivalence : Discrete.{w} PEmpty ≌ Discrete.{v} PEmpty := equivalenceOfIsEmpty _ _ +#align category_theory.functor.empty_equivalence CategoryTheory.emptyEquivalence + +namespace Functor /-- The canonical functor out of the empty category. -/ def empty : Discrete.{w} PEmpty ⥤ C := diff --git a/Mathlib/CategoryTheory/Preadditive/Basic.lean b/Mathlib/CategoryTheory/Preadditive/Basic.lean index a20e083b6f963..bf7b43315c2e7 100644 --- a/Mathlib/CategoryTheory/Preadditive/Basic.lean +++ b/Mathlib/CategoryTheory/Preadditive/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel, Jakob von Raumer -/ import Mathlib.Algebra.BigOperators.Basic -import Mathlib.Algebra.Hom.Group.Defs +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Algebra.Module.Basic import Mathlib.CategoryTheory.Endomorphism import Mathlib.CategoryTheory.Limits.Shapes.Kernels diff --git a/Mathlib/CategoryTheory/Shift/Induced.lean b/Mathlib/CategoryTheory/Shift/Induced.lean index 42471d5d38930..500790854268b 100644 --- a/Mathlib/CategoryTheory/Shift/Induced.lean +++ b/Mathlib/CategoryTheory/Shift/Induced.lean @@ -118,7 +118,7 @@ noncomputable def induced : HasShift D A := add_zero_hom_app := fun n => by have := hF.2 suffices (Induced.add F s i hF n 0).hom = - eqToHom (by rw [add_zero]; rfl) ≫ whiskerLeft (s n) (Induced.zero F s i hF).inv by + eqToHom (by rw [add_zero]; rfl) ≫ whiskerLeft (s n) (Induced.zero F s i hF).inv by intro X simpa using NatTrans.congr_app this X apply ((whiskeringLeft C D D).obj F).map_injective diff --git a/Mathlib/CategoryTheory/Sites/Coherent.lean b/Mathlib/CategoryTheory/Sites/Coherent.lean index e5e4853ffde40..a4cb8b2c1809d 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent.lean @@ -63,7 +63,7 @@ def coherentCoverage [Precoherent C] : Coverage C where obtain ⟨β,_,X₂,π₂,h,i,ι,hh⟩ := Precoherent.pullback f α X₁ π₁ hS refine ⟨Presieve.ofArrows X₂ π₂, ⟨β, inferInstance, X₂, π₂, rfl, h⟩, ?_⟩ rintro _ _ ⟨b⟩ - refine ⟨(X₁ (i b)), ι _, π₁ _, ⟨_⟩, hh _⟩ + exact ⟨(X₁ (i b)), ι _, π₁ _, ⟨_⟩, hh _⟩ /-- The coherent Grothendieck topology on a precoherent category `C`. diff --git a/Mathlib/CategoryTheory/Sites/Coverage.lean b/Mathlib/CategoryTheory/Sites/Coverage.lean index 8ce411ebb7c5c..034dadcd169d8 100644 --- a/Mathlib/CategoryTheory/Sites/Coverage.lean +++ b/Mathlib/CategoryTheory/Sites/Coverage.lean @@ -385,6 +385,22 @@ theorem isSheaf_coverage (K : Coverage C) (P : Cᵒᵖ ⥤ Type w) : apply hx simp +/-- +A presheaf is a sheaf for the Grothendieck topology generated by a union of coverages iff it is a +sheaf for the Grothendieck topology generated by each coverage separately. +-/ +theorem isSheaf_sup (K L : Coverage C) (P : Cᵒᵖ ⥤ Type w) : + (Presieve.IsSheaf ((K ⊔ L).toGrothendieck C)) P ↔ + (Presieve.IsSheaf (K.toGrothendieck C)) P ∧ (Presieve.IsSheaf (L.toGrothendieck C)) P := by + refine ⟨fun h ↦ ⟨Presieve.isSheaf_of_le _ ((gi C).gc.monotone_l le_sup_left) h, + Presieve.isSheaf_of_le _ ((gi C).gc.monotone_l le_sup_right) h⟩, fun h ↦ ?_⟩ + rw [isSheaf_coverage, isSheaf_coverage] at h + rw [isSheaf_coverage] + intro X R hR + cases' hR with hR hR + · exact h.1 R hR + · exact h.2 R hR + end Presieve end CategoryTheory diff --git a/Mathlib/CategoryTheory/Sites/EffectiveEpimorphic.lean b/Mathlib/CategoryTheory/Sites/EffectiveEpimorphic.lean index 05edddb89c604..8a8ab55e0e7aa 100644 --- a/Mathlib/CategoryTheory/Sites/EffectiveEpimorphic.lean +++ b/Mathlib/CategoryTheory/Sites/EffectiveEpimorphic.lean @@ -58,7 +58,7 @@ def Sieve.generateSingleton {X Y : C} (f : Y ⟶ X) : Sieve X where arrows Z := { g | ∃ (e : Z ⟶ Y), e ≫ f = g } downward_closed := by rintro W Z g ⟨e,rfl⟩ q - refine ⟨q ≫ e, by simp⟩ + exact ⟨q ≫ e, by simp⟩ lemma Sieve.generateSingleton_eq {X Y : C} (f : Y ⟶ X) : Sieve.generate (Presieve.singleton f) = Sieve.generateSingleton f := by @@ -226,7 +226,7 @@ def Sieve.generateFamily {B : C} {α : Type*} (X : α → C) (π : (a : α) → arrows Y := { f | ∃ (a : α) (g : Y ⟶ X a), g ≫ π a = f } downward_closed := by rintro Y₁ Y₂ g₁ ⟨a,q,rfl⟩ e - refine ⟨a, e ≫ q, by simp⟩ + exact ⟨a, e ≫ q, by simp⟩ lemma Sieve.generateFamily_eq {B : C} {α : Type*} (X : α → C) (π : (a : α) → (X a ⟶ B)) : Sieve.generate (Presieve.ofArrows X π) = Sieve.generateFamily X π := by @@ -235,7 +235,7 @@ lemma Sieve.generateFamily_eq {B : C} {α : Type*} (X : α → C) (π : (a : α) · rintro ⟨W, g, f, ⟨a⟩, rfl⟩ exact ⟨a, g, rfl⟩ · rintro ⟨a, g, rfl⟩ - refine ⟨_, g, π a, ⟨a⟩, rfl⟩ + exact ⟨_, g, π a, ⟨a⟩, rfl⟩ /-- This structure encodes the data required for a family of morphisms to be effective epimorphic. @@ -540,6 +540,14 @@ instance {B : C} {α : Type*} (X : α → C) (π : (a : α) → (X a ⟶ B)) [Ha [IsIso (Sigma.desc π)] : EffectiveEpiFamily X π := ⟨⟨EffectiveEpiFamilyStruct_of_isIso_desc X π⟩⟩ +/-- The identity is an effective epi. -/ +def EffectiveEpiStructId {X : C} : EffectiveEpiStruct (𝟙 X) where + desc e _ := e + fac _ _ := by simp only [Category.id_comp] + uniq _ _ _ h := by simp only [Category.id_comp] at h; exact h + +instance {X : C} : EffectiveEpi (𝟙 X) := ⟨⟨EffectiveEpiStructId⟩⟩ + end instances section Epi diff --git a/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean b/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean index 67cf9086dec69..44731334c89a7 100644 --- a/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean +++ b/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean @@ -5,6 +5,7 @@ Authors: Bhavik Mehta -/ import Mathlib.CategoryTheory.Sites.IsSheafFor import Mathlib.CategoryTheory.Limits.Shapes.Types +import Mathlib.Tactic.ApplyFun #align_import category_theory.sites.sheaf_of_types from "leanprover-community/mathlib"@"70fd9563a21e7b963887c9360bd29b2393e6225a" @@ -185,7 +186,8 @@ namespace Presieve variable [R.hasPullbacks] -/-- The rightmost object of the fork diagram of https://stacks.math.columbia.edu/tag/00VM, which +/-- +The rightmost object of the fork diagram of https://stacks.math.columbia.edu/tag/00VM, which contains the data used to check a family of elements for a presieve is compatible. -/ @[simp] def SecondObj : Type max v u := @@ -259,6 +261,104 @@ theorem sheaf_condition : R.IsSheafFor P ↔ Nonempty (IsLimit (Fork.ofι _ (w P simp [forkMap] #align category_theory.equalizer.presieve.sheaf_condition CategoryTheory.Equalizer.Presieve.sheaf_condition +namespace Arrows + +open Presieve + +variable {B : C} {I : Type} (X : I → C) (π : (i : I) → X i ⟶ B) [UnivLE.{w, max v u}] + [(ofArrows X π).hasPullbacks] +-- TODO: allow `I : Type w`  + +/-- +The middle object of the fork diagram of . +The difference between this and `Equalizer.FirstObj P (ofArrows X π)` arrises if the family of +arrows `π` contains duplicates. The `Presieve.ofArrows` doesn't see those. +-/ +def FirstObj : Type max v u := ∏ (fun i ↦ P.obj (op (X i))) + +@[ext] +lemma FirstObj.ext (z₁ z₂ : FirstObj P X) (h : ∀ i, (Pi.π _ i : FirstObj P X ⟶ _) z₁ = + (Pi.π _ i : FirstObj P X ⟶ _) z₂) : z₁ = z₂ := by + apply Limits.Types.limit_ext + rintro ⟨i⟩ + exact h i + +/-- +The rightmost object of the fork diagram of https://stacks.math.columbia.edu/tag/00VM. +The difference between this and `Equalizer.Presieve.SecondObj P (ofArrows X π)` arrises if the +family of arrows `π` contains duplicates. The `Presieve.ofArrows` doesn't see those. +-/ +def SecondObj : Type max v u := + ∏ (fun (ij : I × I) ↦ P.obj (op (pullback (π ij.1) (π ij.2)))) + +@[ext] +lemma SecondObj.ext (z₁ z₂ : SecondObj P X π) (h : ∀ ij, (Pi.π _ ij : SecondObj P X π ⟶ _) z₁ = + (Pi.π _ ij : SecondObj P X π ⟶ _) z₂) : z₁ = z₂ := by + apply Limits.Types.limit_ext + rintro ⟨i⟩ + exact h i + +/-- +The left morphism of the fork diagram. +-/ +def forkMap : P.obj (op B) ⟶ FirstObj P X := Pi.lift (fun i ↦ P.map (π i).op) + +/-- +The first of the two parallel morphisms of the fork diagram, induced by the first projection in +each pullback. +-/ +def firstMap : FirstObj P X ⟶ SecondObj P X π := Pi.lift fun _ => Pi.π _ _ ≫ P.map pullback.fst.op + +/-- +The second of the two parallel morphisms of the fork diagram, induced by the second projection in +each pullback. +-/ +def secondMap : FirstObj P X ⟶ SecondObj P X π := Pi.lift fun _ => Pi.π _ _ ≫ P.map pullback.snd.op + +theorem w : forkMap P X π ≫ firstMap P X π = forkMap P X π ≫ secondMap P X π := by + ext x ij + simp only [firstMap, secondMap, forkMap, types_comp_apply, Types.pi_lift_π_apply, + ← FunctorToTypes.map_comp_apply, ← op_comp, pullback.condition] + +/-- +The family of elements given by `x : FirstObj P S` is compatible iff `firstMap` and `secondMap` +map it to the same point. +-/ +theorem compatible_iff (x : FirstObj P X) : (Arrows.Compatible P π ((Types.productIso _).hom x)) ↔ + firstMap P X π x = secondMap P X π x := by + rw [Arrows.pullbackCompatible_iff] + constructor + · intro t + ext ij + simpa [firstMap, secondMap] using t ij.1 ij.2 + · intro t i j + apply_fun Pi.π (fun (ij : I × I) ↦ P.obj (op (pullback (π ij.1) (π ij.2)))) ⟨i, j⟩ at t + simpa [firstMap, secondMap] using t + +/-- +`P` is a sheaf for `Presieve.ofArrows X π`, iff the fork given by `w` is an equalizer. +See . +-/ +theorem sheaf_condition : (ofArrows X π).IsSheafFor P ↔ + Nonempty (IsLimit (Fork.ofι (forkMap P X π) (w P X π))) := by + rw [Types.type_equalizer_iff_unique, isSheafFor_arrows_iff] + erw [← Equiv.forall_congr_left (Types.productIso _).toEquiv.symm] + simp_rw [← compatible_iff, ← Iso.toEquiv_fun, Equiv.apply_symm_apply] + apply ball_congr + intro x _ + apply exists_unique_congr + intro t + erw [Equiv.eq_symm_apply] + constructor + · intro q + funext i + simpa [forkMap] using q i + · intro q i + rw [← q] + simp [forkMap] + +end Arrows + end Presieve end diff --git a/Mathlib/CategoryTheory/Sites/Grothendieck.lean b/Mathlib/CategoryTheory/Sites/Grothendieck.lean index 439908a692ca3..99d42586fb5d5 100644 --- a/Mathlib/CategoryTheory/Sites/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Sites/Grothendieck.lean @@ -116,13 +116,13 @@ theorem mem_sieves_iff_coe : S ∈ J.sieves X ↔ S ∈ J X := #align category_theory.grothendieck_topology.mem_sieves_iff_coe CategoryTheory.GrothendieckTopology.mem_sieves_iff_coe -/ --- Also known as the maximality axiom. +/-- Also known as the maximality axiom. -/ @[simp] theorem top_mem (X : C) : ⊤ ∈ J X := J.top_mem' X #align category_theory.grothendieck_topology.top_mem CategoryTheory.GrothendieckTopology.top_mem --- Also known as the stability axiom. +/-- Also known as the stability axiom. -/ @[simp] theorem pullback_stable (f : Y ⟶ X) (hS : S ∈ J X) : S.pullback f ∈ J Y := J.pullback_stable' f hS diff --git a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean index 975520d537cc1..8205369ab54e0 100644 --- a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean +++ b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean @@ -113,6 +113,9 @@ In special cases, this condition can be simplified, see `pullbackCompatible_iff` This is referred to as a "compatible family" in Definition C2.1.2 of [Elephant], and on nlab: https://ncatlab.org/nlab/show/sheaf#GeneralDefinitionInComponents + +For a more explicit version in the case where `R` is of the form `Presieve.ofArrows`, see +`CategoryTheory.Presieve.Arrows.Compatible`. -/ def FamilyOfElements.Compatible (x : FamilyOfElements P R) : Prop := ∀ ⦃Y₁ Y₂ Z⦄ (g₁ : Z ⟶ Y₁) (g₂ : Z ⟶ Y₂) ⦃f₁ : Y₁ ⟶ X⦄ ⦃f₂ : Y₂ ⟶ X⦄ (h₁ : R f₁) (h₂ : R f₂), @@ -130,6 +133,9 @@ This is the definition for a "matching" family given in [MM92], Chapter III, Sec Equation (5). Viewing the type `FamilyOfElements` as the middle object of the fork in https://stacks.math.columbia.edu/tag/00VM, this condition expresses that `pr₀* (x) = pr₁* (x)`, using the notation defined there. + +For a more explicit version in the case where `R` is of the form `Presieve.ofArrows`, see +`CategoryTheory.Presieve.Arrows.PullbackCompatible`. -/ def FamilyOfElements.PullbackCompatible (x : FamilyOfElements P R) [R.hasPullbacks] : Prop := ∀ ⦃Y₁ Y₂⦄ ⦃f₁ : Y₁ ⟶ X⦄ ⦃f₂ : Y₂ ⟶ X⦄ (h₁ : R f₁) (h₂ : R f₂), @@ -145,7 +151,7 @@ theorem pullbackCompatible_iff (x : FamilyOfElements P R) [R.hasPullbacks] : haveI := hasPullbacks.has_pullbacks hf₁ hf₂ apply pullback.condition · intro t Y₁ Y₂ Z g₁ g₂ f₁ f₂ hf₁ hf₂ comm - haveI := hasPullbacks.has_pullbacks hf₁ hf₂ + haveI := hasPullbacks.has_pullbacks hf₁ hf₂ rw [← pullback.lift_fst _ _ comm, op_comp, FunctorToTypes.map_comp_apply, t hf₁ hf₂, ← FunctorToTypes.map_comp_apply, ← op_comp, pullback.lift_snd] #align category_theory.presieve.pullback_compatible_iff CategoryTheory.Presieve.pullbackCompatible_iff @@ -700,4 +706,85 @@ theorem isSheafFor_subsieve (P : Cᵒᵖ ⥤ Type w) {S : Sieve X} {R : Presieve isSheafFor_subsieve_aux P h (by simpa using trans (𝟙 _)) fun Y f _ => (trans f).isSeparatedFor #align category_theory.presieve.is_sheaf_for_subsieve CategoryTheory.Presieve.isSheafFor_subsieve +section Arrows + +variable {B : C} {I : Type*} {X : I → C} (π : (i : I) → X i ⟶ B) (P) + +/-- +A more explicit version of `FamilyOfElements.Compatible` for a `Presieve.ofArrows`. +-/ +def Arrows.Compatible (x : (i : I) → P.obj (op (X i))) : Prop := + ∀ i j Z (gi : Z ⟶ X i) (gj : Z ⟶ X j), gi ≫ π i = gj ≫ π j → + P.map gi.op (x i) = P.map gj.op (x j) + +lemma FamilyOfElements.isAmalgamation_iff_ofArrows (x : FamilyOfElements P (ofArrows X π)) + (t : P.obj (op B)) : + x.IsAmalgamation t ↔ ∀ (i : I), P.map (π i).op t = x _ (ofArrows.mk i) := + ⟨fun h i ↦ h _ (ofArrows.mk i), fun h _ f ⟨i⟩ ↦ h i⟩ + +namespace Arrows.Compatible + +variable {x : (i : I) → P.obj (op (X i))} (hx : Compatible P π x) +variable {P π} + +theorem exists_familyOfElements : + ∃ (x' : FamilyOfElements P (ofArrows X π)), ∀ (i : I), x' _ (ofArrows.mk i) = x i := by + choose i h h' using @ofArrows_surj _ _ _ _ _ π + exact ⟨fun Y f hf ↦ P.map (eqToHom (h f hf).symm).op (x _), + fun j ↦ (hx _ j (X j) _ (𝟙 _) <| by rw [← h', id_comp]).trans <| by simp⟩ + +/-- +A `FamilyOfElements` associated to an explicit family of elements. +-/ +noncomputable +def familyOfElements : FamilyOfElements P (ofArrows X π) := + (exists_familyOfElements hx).choose + +@[simp] +theorem familyOfElements_ofArrows_mk (i : I) : + hx.familyOfElements _ (ofArrows.mk i) = x i := + (exists_familyOfElements hx).choose_spec _ + +theorem familyOfElements_compatible : hx.familyOfElements.Compatible := by + rintro Y₁ Y₂ Z g₁ g₂ f₁ f₂ ⟨i⟩ ⟨j⟩ hgf + simp [hx i j Z g₁ g₂ hgf] + +end Arrows.Compatible + +theorem isSheafFor_arrows_iff : (ofArrows X π).IsSheafFor P ↔ + (∀ (x : (i : I) → P.obj (op (X i))), Arrows.Compatible P π x → + ∃! t, ∀ i, P.map (π i).op t = x i) := by + refine ⟨fun h x hx ↦ ?_, fun h x hx ↦ ?_⟩ + · obtain ⟨t, ht₁, ht₂⟩ := h _ hx.familyOfElements_compatible + refine ⟨t, fun i ↦ ?_, fun t' ht' ↦ ht₂ _ fun _ _ ⟨i⟩ ↦ ?_⟩ + · rw [ht₁ _ (ofArrows.mk i), hx.familyOfElements_ofArrows_mk] + · rw [ht', hx.familyOfElements_ofArrows_mk] + · obtain ⟨t, hA, ht⟩ := h (fun i ↦ x (π i) (ofArrows.mk _)) + (fun i j Z gi gj ↦ hx gi gj (ofArrows.mk _) (ofArrows.mk _)) + exact ⟨t, fun Y f ⟨i⟩ ↦ hA i, fun y hy ↦ ht y (fun i ↦ hy (π i) (ofArrows.mk _))⟩ + +variable [(ofArrows X π).hasPullbacks] + +/-- +A more explicit version of `FamilyOfElements.PullbackCompatible` for a `Presieve.ofArrows`. +-/ +def Arrows.PullbackCompatible (x : (i : I) → P.obj (op (X i))) : Prop := + ∀ i j, P.map (pullback.fst (f := π i) (g := π j)).op (x i) = + P.map (pullback.snd (f := π i) (g := π j)).op (x j) + +theorem Arrows.pullbackCompatible_iff (x : (i : I) → P.obj (op (X i))) : + Compatible P π x ↔ PullbackCompatible P π x := by + refine ⟨fun t i j ↦ ?_, fun t i j Z gi gj comm ↦ ?_⟩ + · apply t + exact pullback.condition + · rw [← pullback.lift_fst _ _ comm, op_comp, FunctorToTypes.map_comp_apply, t i j, + ← FunctorToTypes.map_comp_apply, ← op_comp, pullback.lift_snd] + +theorem isSheafFor_arrows_iff_pullbacks : (ofArrows X π).IsSheafFor P ↔ + (∀ (x : (i : I) → P.obj (op (X i))), Arrows.PullbackCompatible P π x → + ∃! t, ∀ i, P.map (π i).op t = x i) := by + simp_rw [← Arrows.pullbackCompatible_iff, isSheafFor_arrows_iff] + +end Arrows + end Presieve diff --git a/Mathlib/CategoryTheory/Sites/LeftExact.lean b/Mathlib/CategoryTheory/Sites/LeftExact.lean index 7e3df8e2fa48d..274e20230b533 100644 --- a/Mathlib/CategoryTheory/Sites/LeftExact.lean +++ b/Mathlib/CategoryTheory/Sites/LeftExact.lean @@ -7,6 +7,7 @@ import Mathlib.CategoryTheory.Sites.Sheafification import Mathlib.CategoryTheory.Sites.Limits import Mathlib.CategoryTheory.Limits.FunctorCategory import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesFiniteLimit +import Mathlib.CategoryTheory.Adhesive #align_import category_theory.sites.left_exact from "leanprover-community/mathlib"@"59382264386afdbaf1727e617f5fdda511992eb9" @@ -18,7 +19,7 @@ In this file we show that sheafification commutes with finite limits. open CategoryTheory Limits Opposite -universe w v u +universe w' w v u -- porting note: was `C : Type max v u` which made most instances non automatically applicable -- it seems to me it is better to declare `C : Type u`: it works better, and it is more general @@ -237,17 +238,30 @@ variable [PreservesLimits (forget D)] variable [ReflectsIsomorphisms (forget D)] -variable (K : Type max v u) +variable (K : Type w') variable [SmallCategory K] [FinCategory K] [HasLimitsOfShape K D] instance preservesLimitsOfShape_presheafToSheaf : PreservesLimitsOfShape K (presheafToSheaf J D) := by + let e := (FinCategory.equivAsType K).symm.trans (AsSmall.equiv.{0, 0, max v u}) + haveI : HasLimitsOfShape (AsSmall.{max v u} (FinCategory.AsType K)) D := + Limits.hasLimitsOfShape_of_equivalence e + haveI : FinCategory (AsSmall.{max v u} (FinCategory.AsType K)) := by + constructor + · show Fintype (ULift _) + infer_instance + · intro j j' + show Fintype (ULift _) + infer_instance + refine @preservesLimitsOfShapeOfEquiv _ _ _ _ _ _ _ _ e.symm _ (show _ from ?_) constructor; intro F; constructor; intro S hS apply isLimitOfReflects (sheafToPresheaf J D) - have : ReflectsLimitsOfShape K (forget D) := reflectsLimitsOfShapeOfReflectsIsomorphisms + have : ReflectsLimitsOfShape (AsSmall.{max v u} (FinCategory.AsType K)) (forget D) := + reflectsLimitsOfShapeOfReflectsIsomorphisms -- porting note: the mathlib proof was by `apply is_limit_of_preserves (J.sheafification D) hS` - have : PreservesLimitsOfShape K (presheafToSheaf J D ⋙ sheafToPresheaf J D) := + have : PreservesLimitsOfShape (AsSmall.{max v u} (FinCategory.AsType K)) + (presheafToSheaf J D ⋙ sheafToPresheaf J D) := preservesLimitsOfShapeOfNatIso (J.sheafificationIsoPresheafToSheafCompSheafToPreasheaf D) exact isLimitOfPreserves (presheafToSheaf J D ⋙ sheafToPresheaf J D) hS @@ -257,4 +271,23 @@ instance preservesfiniteLimits_presheafToSheaf [HasFiniteLimits D] : intros infer_instance +instance [FinitaryExtensive D] [HasFiniteCoproducts D] [HasPullbacks D] : + FinitaryExtensive (Sheaf J D) := + finitaryExtensive_of_reflective (sheafificationAdjunction _ _) + +instance [Adhesive D] [HasPullbacks D] [HasPushouts D] : Adhesive (Sheaf J D) := + adhesive_of_reflective (sheafificationAdjunction _ _) + +instance SheafOfTypes.finitary_extensive {C : Type u} [SmallCategory C] + (J : GrothendieckTopology C) : FinitaryExtensive (Sheaf J (Type u)) := + inferInstance + +instance SheafOfTypes.adhesive {C : Type u} [SmallCategory C] (J : GrothendieckTopology C) : + Adhesive (Sheaf J (Type u)) := + inferInstance + +instance SheafOfTypes.balanced {C : Type u} [SmallCategory C] (J : GrothendieckTopology C) : + Balanced (Sheaf J (Type u)) := + inferInstance + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Sites/Limits.lean b/Mathlib/CategoryTheory/Sites/Limits.lean index e2c25780bd101..5e73643bc4a7b 100644 --- a/Mathlib/CategoryTheory/Sites/Limits.lean +++ b/Mathlib/CategoryTheory/Sites/Limits.lean @@ -5,6 +5,7 @@ Authors: Adam Topaz -/ import Mathlib.CategoryTheory.Limits.Creates import Mathlib.CategoryTheory.Sites.Sheafification +import Mathlib.CategoryTheory.Limits.Shapes.FiniteProducts #align_import category_theory.sites.limits from "leanprover-community/mathlib"@"95e83ced9542828815f53a1096a4d373c1b08a77" @@ -39,13 +40,13 @@ open Opposite section Limits -universe w v u z +universe w v u z z' variable {C : Type u} [Category.{v} C] {J : GrothendieckTopology C} variable {D : Type w} [Category.{max v u} D] -variable {K : Type z} [SmallCategory K] +variable {K : Type z} [Category.{z'} K] noncomputable section @@ -166,6 +167,9 @@ instance createsLimitsOfShape : CreatesLimitsOfShape K (sheafToPresheaf J D) whe instance : HasLimitsOfShape K (Sheaf J D) := hasLimitsOfShape_of_hasLimitsOfShape_createsLimitsOfShape (sheafToPresheaf J D) +instance [HasFiniteProducts D] : HasFiniteProducts (Sheaf J D) := + ⟨inferInstance⟩ + end instance createsLimits [HasLimits D] : CreatesLimits (sheafToPresheaf J D) := @@ -180,13 +184,13 @@ end Limits section Colimits -universe w v u +universe w v u z z' variable {C : Type u} [Category.{v} C] {J : GrothendieckTopology C} variable {D : Type w} [Category.{max v u} D] -variable {K : Type max v u} [SmallCategory K] +variable {K : Type z} [Category.{z'} K] -- Now we need a handful of instances to obtain sheafification... variable [ConcreteCategory.{max v u} D] @@ -245,6 +249,9 @@ instance [HasColimitsOfShape K D] : HasColimitsOfShape K (Sheaf J D) := ⟨fun _ => HasColimit.mk ⟨sheafifyCocone (colimit.cocone _), isColimitSheafifyCocone _ (colimit.isColimit _)⟩⟩ +instance [HasFiniteCoproducts D] : HasFiniteCoproducts (Sheaf J D) := + ⟨inferInstance⟩ + instance [HasColimits D] : HasColimits (Sheaf J D) := ⟨inferInstance⟩ diff --git a/Mathlib/CategoryTheory/Sites/RegularExtensive.lean b/Mathlib/CategoryTheory/Sites/RegularExtensive.lean index 4518e82aeb6f1..b90afd9dcfde2 100644 --- a/Mathlib/CategoryTheory/Sites/RegularExtensive.lean +++ b/Mathlib/CategoryTheory/Sites/RegularExtensive.lean @@ -5,6 +5,7 @@ Authors: Dagur Asgeirsson, Filippo A. E. Nuccio, Riccardo Brasca -/ import Mathlib.CategoryTheory.Preadditive.Projective import Mathlib.CategoryTheory.Sites.Coherent +import Mathlib.CategoryTheory.Extensive import Mathlib.CategoryTheory.Sites.EqualizerSheafCondition import Mathlib.Tactic.ApplyFun /-! @@ -19,8 +20,8 @@ covering sieves of this coverage are generated by presieves consisting of a sing epimorphism. The second one is called the *extensive* coverage and for that to exist, the category `C` must -satisfy a condition called `Extensive C`. This means `C` has finite coproducts and that those -are preserved by pullbacks. The covering sieves of this coverage are generated by presieves +satisfy a condition called `FinitaryPreExtensive C`. This means `C` has finite coproducts and that +those are preserved by pullbacks. The covering sieves of this coverage are generated by presieves consisting finitely many arrows that together induce an isomorphism from the coproduct to the target. @@ -67,39 +68,6 @@ class Preregular : Prop where exists_fac : ∀ {X Y Z : C} (f : X ⟶ Y) (g : Z ⟶ Y) [EffectiveEpi g], (∃ (W : C) (h : W ⟶ X) (_ : EffectiveEpi h) (i : W ⟶ Z), i ≫ g = h ≫ f) -/-- -Describes the property of having pullbacks of morphsims into a finite coproduct, where one -of the morphisms is an inclusion map into the coproduct (up to isomorphism). --/ -class HasPullbacksOfInclusions : Prop where - /-- For any morphism `f : X ⟶ Z`, where `Z` is the coproduct of `i : (a : α) → Y a ⟶ Z` with - `α` finite, the pullback of `f` and `i a` exists for every `a : α`. -/ - has_pullback : ∀ {X Z : C} {α : Type w} (f : X ⟶ Z) {Y : (a : α) → C} - (i : (a : α) → Y a ⟶ Z) [Fintype α] [HasCoproduct Y] [IsIso (Sigma.desc i)] (a : α), - HasPullback f (i a) - -instance [HasPullbacksOfInclusions C] {X Z : C} {α : Type w} (f : X ⟶ Z) {Y : (a : α) → C} - (i : (a : α) → Y a ⟶ Z) [Fintype α] [HasCoproduct Y] [IsIso (Sigma.desc i)] (a : α) : - HasPullback f (i a) := HasPullbacksOfInclusions.has_pullback f i a - -/-- -If `C` has pullbacks then it has the pullbacks relevant to `HasPullbacksOfInclusions`. --/ -instance (priority := 10) [HasPullbacks C] : - HasPullbacksOfInclusions C := ⟨fun _ _ _ => inferInstance⟩ - -/-- -A category is *extensive* if it has all finite coproducts and those coproducts are preserved -by pullbacks (we only require the relevant pullbacks to exist, via `HasPullbacksOfInclusions`). - -TODO: relate this to the class `FinitaryExtensive` --/ -class Extensive extends HasFiniteCoproducts C, HasPullbacksOfInclusions C : Prop where - /-- Pulling back an isomorphism from a coproduct yields an isomorphism. -/ - sigma_desc_iso : ∀ {α : Type} [Fintype α] {X : C} {Z : α → C} (π : (a : α) → Z a ⟶ X) - {Y : C} (f : Y ⟶ X) (_ : IsIso (Sigma.desc π)), - IsIso (Sigma.desc ((fun _ ↦ pullback.fst) : (a : α) → pullback f (π a) ⟶ _)) - /-- The regular coverage on a regular category `C`. -/ @@ -122,7 +90,7 @@ def regularCoverage [Preregular C] : Coverage C where /-- The extensive coverage on an extensive category `C` -/ -def extensiveCoverage [Extensive C] : Coverage C where +def extensiveCoverage [FinitaryPreExtensive C] : Coverage C where covering B := { S | ∃ (α : Type) (_ : Fintype α) (X : α → C) (π : (a : α) → (X a ⟶ B)), S = Presieve.ofArrows X π ∧ IsIso (Sigma.desc π) } pullback := by @@ -131,16 +99,16 @@ def extensiveCoverage [Extensive C] : Coverage C where let π' : (a : α) → Z' a ⟶ Y := fun a ↦ pullback.fst refine ⟨@Presieve.ofArrows C _ _ α Z' π', ⟨?_, ?_⟩⟩ · constructor - exact ⟨hα, Z', π', ⟨by simp only, Extensive.sigma_desc_iso (fun x => π x) f h_iso⟩⟩ + exact ⟨hα, Z', π', ⟨by simp only, FinitaryPreExtensive.sigma_desc_iso (fun x => π x) f h_iso⟩⟩ · intro W g hg rcases hg with ⟨a⟩ refine ⟨Z a, pullback.snd, π a, ?_, by rw [CategoryTheory.Limits.pullback.condition]⟩ rw [hS] - refine Presieve.ofArrows.mk a + exact Presieve.ofArrows.mk a /-- The union of the extensive and regular coverages generates the coherent topology on `C`. -/ -lemma extensive_regular_generate_coherent [Preregular C] [Extensive C] [Precoherent C] : +lemma extensive_regular_generate_coherent [Preregular C] [FinitaryPreExtensive C] [Precoherent C] : ((extensiveCoverage C) ⊔ (regularCoverage C)).toGrothendieck = (coherentTopology C) := by ext B S diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index c7dc5e3c0897e..6bd8768eae9a6 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -168,6 +168,12 @@ theorem ofArrows_bind {ι : Type*} (Z : ι → C) (g : ∀ i : ι, Z i ⟶ X) exact bind_comp _ (ofArrows.mk _) (ofArrows.mk _) #align category_theory.presieve.of_arrows_bind CategoryTheory.Presieve.ofArrows_bind +theorem ofArrows_surj {ι : Type*} {Y : ι → C} (f : ∀ i, Y i ⟶ X) {Z : C} (g : Z ⟶ X) + (hg : ofArrows Y f g) : ∃ (i : ι) (h : Y i = Z), + g = eqToHom h.symm ≫ f i := by + cases' hg with i + exact ⟨i, rfl, by simp only [eqToHom_refl, id_comp]⟩ + /-- Given a presieve on `F(X)`, we can define a presieve on `X` by taking the preimage via `F`. -/ def functorPullback (R : Presieve (F.obj X)) : Presieve X := fun _ f => R (F.map f) #align category_theory.presieve.functor_pullback CategoryTheory.Presieve.functorPullback @@ -191,6 +197,10 @@ class hasPullbacks (R : Presieve X) : Prop where instance (R : Presieve X) [HasPullbacks C] : R.hasPullbacks := ⟨fun _ _ ↦ inferInstance⟩ +instance {α : Type v₂} {X : α → C} {B : C} (π : (a : α) → X a ⟶ B) + [(Presieve.ofArrows X π).hasPullbacks] (a b : α) : HasPullback (π a) (π b) := + Presieve.hasPullbacks.has_pullbacks (Presieve.ofArrows.mk _) (Presieve.ofArrows.mk _) + section FunctorPushforward variable {E : Type u₃} [Category.{v₃} E] (G : D ⥤ E) diff --git a/Mathlib/CategoryTheory/Sites/Types.lean b/Mathlib/CategoryTheory/Sites/Types.lean index b7d89512addcb..8bda080f4e657 100644 --- a/Mathlib/CategoryTheory/Sites/Types.lean +++ b/Mathlib/CategoryTheory/Sites/Types.lean @@ -191,7 +191,7 @@ theorem typesGrothendieckTopology_eq_canonical : have : (fun _ => ULift.up true) = fun _ => ULift.up false := (hs PUnit fun _ => x).isSeparatedFor.ext fun β f hf => funext fun y => hsx.elim <| S.2 hf fun _ => y - simp at this + simp [Function.funext_iff] at this #align category_theory.types_grothendieck_topology_eq_canonical CategoryTheory.typesGrothendieckTopology_eq_canonical end CategoryTheory diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 8b7fe98e0e6e0..091c67c66427d 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -421,11 +421,11 @@ lemma yonedaEquiv_naturality' {X Y : Cᵒᵖ} {F : Cᵒᵖ ⥤ Type v₁} (f : y (g : X ⟶ Y) : F.map g (yonedaEquiv f) = yonedaEquiv (yoneda.map g.unop ≫ f) := yonedaEquiv_naturality _ _ -lemma yonedaEquiv_comp {X : C} {F G : Cᵒᵖ ⥤ Type v₁} (α : yoneda.obj X ⟶ F) (β : F ⟶ G) : +lemma yonedaEquiv_comp {X : C} {F G : Cᵒᵖ ⥤ Type v₁} (α : yoneda.obj X ⟶ F) (β : F ⟶ G) : yonedaEquiv (α ≫ β) = β.app _ (yonedaEquiv α) := rfl -lemma yonedaEquiv_comp' {X : Cᵒᵖ} {F G : Cᵒᵖ ⥤ Type v₁} (α : yoneda.obj (unop X) ⟶ F) (β : F ⟶ G) : +lemma yonedaEquiv_comp' {X : Cᵒᵖ} {F G : Cᵒᵖ ⥤ Type v₁} (α : yoneda.obj (unop X) ⟶ F) (β : F ⟶ G) : yonedaEquiv (α ≫ β) = β.app X (yonedaEquiv α) := rfl diff --git a/Mathlib/Combinatorics/Additive/Behrend.lean b/Mathlib/Combinatorics/Additive/Behrend.lean index c88d0cc6ea789..7c21bed69936c 100644 --- a/Mathlib/Combinatorics/Additive/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/Behrend.lean @@ -164,7 +164,7 @@ theorem map_injOn : {x : Fin n → ℕ | ∀ i, x i < d}.InjOn (map d) := by theorem map_le_of_mem_box (hx : x ∈ box n d) : map (2 * d - 1) x ≤ ∑ i : Fin n, (d - 1) * (2 * d - 1) ^ (i : ℕ) := - map_monotone (2 * d - 1) fun _ => Nat.le_pred_of_lt <| mem_box.1 hx _ + map_monotone (2 * d - 1) fun _ => Nat.le_sub_one_of_lt <| mem_box.1 hx _ #align behrend.map_le_of_mem_box Behrend.map_le_of_mem_box nonrec theorem addSalemSpencer_sphere : AddSalemSpencer (sphere n d k : Set (Fin n → ℕ)) := by @@ -196,7 +196,7 @@ theorem addSalemSpencer_image_sphere : theorem sum_sq_le_of_mem_box (hx : x ∈ box n d) : ∑ i : Fin n, x i ^ 2 ≤ n * (d - 1) ^ 2 := by rw [mem_box] at hx have : ∀ i, x i ^ 2 ≤ (d - 1) ^ 2 := fun i => - Nat.pow_le_pow_of_le_left (Nat.le_pred_of_lt (hx i)) _ + Nat.pow_le_pow_of_le_left (Nat.le_sub_one_of_lt (hx i)) _ exact (sum_le_card_nsmul univ _ _ fun i _ => this i).trans (by rw [card_fin, smul_eq_mul]) #align behrend.sum_sq_le_of_mem_box Behrend.sum_sq_le_of_mem_box @@ -213,7 +213,7 @@ theorem sum_lt : (∑ i : Fin n, d * (2 * d + 1) ^ (i : ℕ)) < (2 * d + 1) ^ n theorem card_sphere_le_rothNumberNat (n d k : ℕ) : (sphere n d k).card ≤ rothNumberNat ((2 * d - 1) ^ n) := by cases n - · dsimp; refine' (card_le_univ _).trans_eq _; simp + · dsimp; refine' (card_le_univ _).trans_eq _; rfl cases d · simp refine' addSalemSpencer_image_sphere.le_rothNumberNat _ _ (card_image_of_injOn _) diff --git a/Mathlib/Combinatorics/Additive/SalemSpencer.lean b/Mathlib/Combinatorics/Additive/SalemSpencer.lean index e76c674f69828..e6c3b105fb423 100644 --- a/Mathlib/Combinatorics/Additive/SalemSpencer.lean +++ b/Mathlib/Combinatorics/Additive/SalemSpencer.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies, Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ -import Mathlib.Algebra.Hom.Freiman +import Mathlib.Algebra.Group.Freiman import Mathlib.Analysis.Asymptotics.Asymptotics import Mathlib.Analysis.Convex.StrictConvexSpace diff --git a/Mathlib/Combinatorics/Catalan.lean b/Mathlib/Combinatorics/Catalan.lean index 8e1ffd2835c4d..4065aeea8e04b 100644 --- a/Mathlib/Combinatorics/Catalan.lean +++ b/Mathlib/Combinatorics/Catalan.lean @@ -50,7 +50,7 @@ open BigOperators open Finset -open Finset.Nat.antidiagonal (fst_le snd_le) +open Finset.antidiagonal (fst_le snd_le) /-- The recursive definition of the sequence of Catalan numbers: `catalan (n + 1) = ∑ i : Fin n.succ, catalan i * catalan (n - i)` -/ @@ -70,7 +70,7 @@ theorem catalan_succ (n : ℕ) : catalan (n + 1) = ∑ i : Fin n.succ, catalan i #align catalan_succ catalan_succ theorem catalan_succ' (n : ℕ) : - catalan (n + 1) = ∑ ij in Nat.antidiagonal n, catalan ij.1 * catalan ij.2 := by + catalan (n + 1) = ∑ ij in antidiagonal n, catalan ij.1 * catalan ij.2 := by rw [catalan_succ, Nat.sum_antidiagonal_eq_sum_range_succ (fun x y => catalan x * catalan y) n, sum_range] #align catalan_succ' catalan_succ' @@ -163,7 +163,7 @@ def pairwiseNode (a b : Finset (Tree Unit)) : Finset (Tree Unit) := def treesOfNumNodesEq : ℕ → Finset (Tree Unit) | 0 => {nil} | n + 1 => - (Finset.Nat.antidiagonal n).attach.biUnion fun ijh => + (antidiagonal n).attach.biUnion fun ijh => -- Porting note: `unusedHavesSuffices` linter is not happy with this. Commented out. -- have := Nat.lt_succ_of_le (fst_le ijh.2) -- have := Nat.lt_succ_of_le (snd_le ijh.2) @@ -181,7 +181,7 @@ theorem treesOfNumNodesEq_zero : treesOfNumNodesEq 0 = {nil} := by rw [treesOfNu theorem treesOfNumNodesEq_succ (n : ℕ) : treesOfNumNodesEq (n + 1) = - (Nat.antidiagonal n).biUnion fun ij => + (antidiagonal n).biUnion fun ij => pairwiseNode (treesOfNumNodesEq ij.1) (treesOfNumNodesEq ij.2) := by rw [treesOfNumNodesEq] ext diff --git a/Mathlib/Combinatorics/Composition.lean b/Mathlib/Combinatorics/Composition.lean index b9c33b84d30c6..abbd4cd22bdcc 100644 --- a/Mathlib/Combinatorics/Composition.lean +++ b/Mathlib/Combinatorics/Composition.lean @@ -149,27 +149,21 @@ theorem blocks_length : c.blocks.length = c.length := rfl #align composition.blocks_length Composition.blocks_length --- porting note: TODO, refactor to `List.get` -set_option linter.deprecated false in /-- The blocks of a composition, seen as a function on `Fin c.length`. When composing analytic functions using compositions, this is the main player. -/ -def blocksFun : Fin c.length → ℕ := fun i => nthLe c.blocks i i.2 +def blocksFun : Fin c.length → ℕ := c.blocks.get #align composition.blocks_fun Composition.blocksFun --- porting note: TODO, refactor to `List.get` -set_option linter.deprecated false in theorem ofFn_blocksFun : ofFn c.blocksFun = c.blocks := - ofFn_nthLe _ + ofFn_get _ #align composition.of_fn_blocks_fun Composition.ofFn_blocksFun theorem sum_blocksFun : ∑ i, c.blocksFun i = n := by conv_rhs => rw [← c.blocks_sum, ← ofFn_blocksFun, sum_ofFn] #align composition.sum_blocks_fun Composition.sum_blocksFun --- porting note: TODO, refactor to `List.get` -set_option linter.deprecated false in theorem blocksFun_mem_blocks (i : Fin c.length) : c.blocksFun i ∈ c.blocks := - nthLe_mem _ _ _ + get_mem _ _ _ #align composition.blocks_fun_mem_blocks Composition.blocksFun_mem_blocks @[simp] @@ -177,17 +171,13 @@ theorem one_le_blocks {i : ℕ} (h : i ∈ c.blocks) : 1 ≤ i := c.blocks_pos h #align composition.one_le_blocks Composition.one_le_blocks --- porting note: TODO, refactor to `List.get` -set_option linter.deprecated false in @[simp] -theorem one_le_blocks' {i : ℕ} (h : i < c.length) : 1 ≤ nthLe c.blocks i h := - c.one_le_blocks (nthLe_mem (blocks c) i h) +theorem one_le_blocks' {i : ℕ} (h : i < c.length) : 1 ≤ c.blocks.get ⟨i, h⟩ := + c.one_le_blocks (get_mem (blocks c) i h) #align composition.one_le_blocks' Composition.one_le_blocks' --- porting note: TODO, refactor to `List.get` -set_option linter.deprecated false in @[simp] -theorem blocks_pos' (i : ℕ) (h : i < c.length) : 0 < nthLe c.blocks i h := +theorem blocks_pos' (i : ℕ) (h : i < c.length) : 0 < c.blocks.get ⟨i, h⟩ := c.one_le_blocks' h #align composition.blocks_pos' Composition.blocks_pos' @@ -232,7 +222,7 @@ theorem sizeUpTo_le (i : ℕ) : c.sizeUpTo i ≤ n := by #align composition.size_up_to_le Composition.sizeUpTo_le theorem sizeUpTo_succ {i : ℕ} (h : i < c.length) : - c.sizeUpTo (i + 1) = c.sizeUpTo i + c.blocks.nthLe i h := by + c.sizeUpTo (i + 1) = c.sizeUpTo i + c.blocks.get ⟨i, h⟩ := by simp only [sizeUpTo] rw [sum_take_succ _ _ h] #align composition.size_up_to_succ Composition.sizeUpTo_succ @@ -350,7 +340,7 @@ theorem sizeUpTo_index_le (j : Fin n) : c.sizeUpTo (c.index j) ≤ j := by have i₁_succ : i₁.succ = i := Nat.succ_pred_eq_of_pos i_pos have := Nat.find_min (c.index_exists j.2) i₁_lt_i simp [lt_trans i₁_lt_i (c.index j).2, i₁_succ] at this - exact Nat.lt_le_antisymm H this + exact Nat.lt_le_asymm H this #align composition.size_up_to_index_le Composition.sizeUpTo_index_le /-- Mapping an element `j` of `Fin n` to the element in the block containing it, identified with @@ -412,9 +402,7 @@ theorem disjoint_range {i₁ i₂ : Fin c.length} (h : i₁ ≠ i₂) : theorem mem_range_embedding (j : Fin n) : j ∈ Set.range (c.embedding (c.index j)) := by have : c.embedding (c.index j) (c.invEmbedding j) ∈ Set.range (c.embedding (c.index j)) := Set.mem_range_self _ - -- porting note: previously `rwa` closed - rw [c.embedding_comp_inv j] at this - assumption + rwa [c.embedding_comp_inv j] at this #align composition.mem_range_embedding Composition.mem_range_embedding theorem mem_range_embedding_iff' {j : Fin n} {i : Fin c.length} : @@ -501,11 +489,9 @@ theorem ones_blocks (n : ℕ) : (ones n).blocks = replicate n (1 : ℕ) := rfl #align composition.ones_blocks Composition.ones_blocks --- porting note: TODO, refactor to `List.get` -set_option linter.deprecated false in @[simp] theorem ones_blocksFun (n : ℕ) (i : Fin (ones n).length) : (ones n).blocksFun i = 1 := by - simp only [blocksFun, ones, blocks, i.2, List.nthLe_replicate] + simp only [blocksFun, ones, blocks, i.2, List.get_replicate] #align composition.ones_blocks_fun Composition.ones_blocksFun @[simp] @@ -554,7 +540,7 @@ theorem eq_ones_iff_length {c : Composition n} : c = ones n ↔ c.length = n := rw [← ofFn_blocksFun, mem_ofFn c.blocksFun, Set.mem_range] at hi obtain ⟨j : Fin c.length, hj : c.blocksFun j = i⟩ := hi rw [← hj] at i_blocks - exact Finset.sum_lt_sum (fun i _ => by simp [blocksFun]) ⟨j, Finset.mem_univ _, i_blocks⟩ + exact Finset.sum_lt_sum (fun i _ => one_le_blocksFun c i) ⟨j, Finset.mem_univ _, i_blocks⟩ } _ = n := c.sum_blocksFun #align composition.eq_ones_iff_length Composition.eq_ones_iff_length @@ -719,39 +705,33 @@ theorem sum_take_map_length_splitWrtComposition (l : List α) (c : Composition l exact map_length_splitWrtComposition l c #align list.sum_take_map_length_split_wrt_composition List.sum_take_map_length_splitWrtComposition --- porting note: TODO, refactor to `List.get` -set_option linter.deprecated false in -theorem nthLe_splitWrtCompositionAux (l : List α) (ns : List ℕ) {i : ℕ} (hi) : - nthLe (l.splitWrtCompositionAux ns) i hi = +theorem get_splitWrtCompositionAux (l : List α) (ns : List ℕ) {i : ℕ} (hi) : + (l.splitWrtCompositionAux ns).get ⟨i, hi⟩ = (l.take (ns.take (i + 1)).sum).drop (ns.take i).sum := by induction' ns with n ns IH generalizing l i · cases hi cases' i with i - · rw [Nat.add_zero, List.take_zero, sum_nil, nthLe_zero]; dsimp - simp only [splitWrtCompositionAux_cons, head!, sum, foldl, zero_add] - · simp only [splitWrtCompositionAux_cons, take, sum_cons, - Nat.add_eq, add_zero, gt_iff_lt, nthLe_cons, IH]; dsimp - rw [Nat.succ_sub_succ_eq_sub, ←Nat.succ_eq_add_one,tsub_zero] - simp only [← drop_take, drop_drop] - rw [add_comm] -#align list.nth_le_split_wrt_composition_aux List.nthLe_splitWrtCompositionAux - --- porting note: TODO, refactor to `List.get` -set_option linter.deprecated false in + · rw [Nat.add_zero, List.take_zero, sum_nil] + simpa using get_zero hi + · simp only [splitWrtCompositionAux._eq_2, get_cons_succ, IH, take, + sum_cons, Nat.add_eq, add_zero, splitAt_eq_take_drop, ← drop_take, drop_drop] + rw [Nat.succ_eq_add_one, add_comm i 1, add_comm] +#align list.nth_le_split_wrt_composition_aux List.get_splitWrtCompositionAux + /-- The `i`-th sublist in the splitting of a list `l` along a composition `c`, is the slice of `l` between the indices `c.sizeUpTo i` and `c.sizeUpTo (i+1)`, i.e., the indices in the `i`-th block of the composition. -/ -theorem nthLe_splitWrtComposition (l : List α) (c : Composition n) {i : ℕ} +theorem get_splitWrtComposition' (l : List α) (c : Composition n) {i : ℕ} (hi : i < (l.splitWrtComposition c).length) : - nthLe (l.splitWrtComposition c) i hi = (l.take (c.sizeUpTo (i + 1))).drop (c.sizeUpTo i) := - nthLe_splitWrtCompositionAux _ _ _ -#align list.nth_le_split_wrt_composition List.nthLe_splitWrtComposition + (l.splitWrtComposition c).get ⟨i, hi⟩ = (l.take (c.sizeUpTo (i + 1))).drop (c.sizeUpTo i) := + get_splitWrtCompositionAux _ _ _ +#align list.nth_le_split_wrt_composition List.get_splitWrtComposition' --- porting note: restatement of `nthLe_splitWrtComposition` +-- porting note: restatement of `get_splitWrtComposition` theorem get_splitWrtComposition (l : List α) (c : Composition n) (i : Fin (l.splitWrtComposition c).length) : get (l.splitWrtComposition c) i = (l.take (c.sizeUpTo (i + 1))).drop (c.sizeUpTo i) := - nthLe_splitWrtComposition _ _ _ + get_splitWrtComposition' _ _ _ theorem join_splitWrtCompositionAux {ns : List ℕ} : ∀ {l : List α}, ns.sum = l.length → (l.splitWrtCompositionAux ns).join = l := by @@ -944,8 +924,6 @@ theorem blocks_length : c.blocks.length = c.length := length_ofFn _ #align composition_as_set.blocks_length CompositionAsSet.blocks_length --- porting note: TODO, refactor to `List.get` -set_option linter.deprecated false in theorem blocks_partial_sum {i : ℕ} (h : i < c.boundaries.card) : (c.blocks.take i).sum = c.boundary ⟨i, h⟩ := by induction' i with i IH @@ -955,9 +933,7 @@ theorem blocks_partial_sum {i : ℕ} (h : i < c.boundaries.card) : simp [blocks, Nat.lt_of_succ_lt_succ h] have B : i < c.boundaries.card := lt_of_lt_of_le A (by simp [blocks, length, Nat.sub_le]) rw [sum_take_succ _ _ A, IH B] - simp only [blocks, blocksFun, nthLe_ofFn'] - apply add_tsub_cancel_of_le - simp + simp [blocks, blocksFun, get_ofFn] #align composition_as_set.blocks_partial_sum CompositionAsSet.blocks_partial_sum theorem mem_boundaries_iff_exists_blocks_sum_take_eq {j : Fin (n + 1)} : diff --git a/Mathlib/Combinatorics/Hall/Finite.lean b/Mathlib/Combinatorics/Hall/Finite.lean index 6a3afeeec7bfb..c2bff7c6070dd 100644 --- a/Mathlib/Combinatorics/Hall/Finite.lean +++ b/Mathlib/Combinatorics/Hall/Finite.lean @@ -62,7 +62,7 @@ theorem hall_cond_of_erase {x : ι} (a : α) rw [← erase_biUnion] by_cases hb : a ∈ s'.biUnion fun x => t x · rw [card_erase_of_mem hb] - exact Nat.le_pred_of_lt ha' + exact Nat.le_sub_one_of_lt ha' · rw [erase_eq_of_not_mem hb] exact Nat.le_of_lt ha' · rw [nonempty_iff_ne_empty, not_not] at he diff --git a/Mathlib/Combinatorics/Optimization/ValuedCSP.lean b/Mathlib/Combinatorics/Optimization/ValuedCSP.lean index ec681e2065bcb..5f36f2258e048 100644 --- a/Mathlib/Combinatorics/Optimization/ValuedCSP.lean +++ b/Mathlib/Combinatorics/Optimization/ValuedCSP.lean @@ -19,7 +19,7 @@ General-Valued CSP subsumes Min-Cost-Hom (including 3-SAT for example) and Finit * `ValuedCsp.Term.evalSolution`: An evaluation of the VCSP term for given solution. * `ValuedCsp.Instance`: An instance of a VCSP problem over given template. * `ValuedCsp.Instance.evalSolution`: An evaluation of the VCSP instance for given solution. -* `ValuedCsp.Instance.OptimumSolution`: Is given solution a minimum of the VCSP instance? +* `ValuedCsp.Instance.IsOptimumSolution`: Is given solution a minimum of the VCSP instance? ## References * [D. A. Cohen, M. C. Cooper, P. Creed, P. G. Jeavons, S. Živný, diff --git a/Mathlib/Combinatorics/SetFamily/Compression/UV.lean b/Mathlib/Combinatorics/SetFamily/Compression/UV.lean index 321f14a1ca5fc..9313c3b949c97 100644 --- a/Mathlib/Combinatorics/SetFamily/Compression/UV.lean +++ b/Mathlib/Combinatorics/SetFamily/Compression/UV.lean @@ -143,9 +143,9 @@ theorem compress_injOn : Set.InjOn (compress u v) ↑(s.filter (compress u v · intro a ha b hb hab rw [mem_coe, mem_filter] at ha hb rw [compress] at ha hab - split_ifs at ha hab with has + split_ifs at ha hab with has · rw [compress] at hb hab - split_ifs at hb hab with hbs + split_ifs at hb hab with hbs · exact sup_sdiff_injOn u v has hbs hab · exact (hb.2 hb.1).elim · exact (ha.2 ha.1).elim diff --git a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean new file mode 100644 index 0000000000000..e6c3aa38e93d2 --- /dev/null +++ b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean @@ -0,0 +1,357 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.BigOperators.Order +import Mathlib.Algebra.Order.Pi +import Mathlib.Data.Finset.Sups +import Mathlib.Order.Birkhoff +import Mathlib.Order.Booleanisation +import Mathlib.Order.Sublattice +import Mathlib.Tactic.Ring + +/-! +# The four functions theorem and corollaries + +This file proves the four functions theorem. The statement is that if +`f₁ a * f₂ b ≤ f₃ (a ⊓ b) * f (a ⊔ b)` for all `a`, `b` in a finite distributive lattice, then +`(∑ x in s, f₁ x) * (∑ x in t, f₂ x) ≤ (∑ x in s ⊼ t, f₃ x) * (∑ x in s ⊻ t, f₄ x)` where +`s ⊼ t = {a ⊓ b | a ∈ s, b ∈ t}`, `s ⊻ t = {a ⊔ b | a ∈ s, b ∈ t}`. + +The proof uses Birkhoff's representation theorem to restrict to the case where the finite +distributive lattice is in fact a finite powerset algebra, namely `Finset α` for some finite `α`. +Then it proves this new statement by induction on the size of `α`. + +## Main declarations + +The two versions of the four functions theorem are +* `Finset.four_functions_theorem` for finite powerset algebras. +* `four_functions_theorem` for any finite distributive lattices. + +We deduce a number of corollaries: +* `Finset.le_card_infs_mul_card_sups`: Daykin inequality. `|s| |t| ≤ |s ⊼ t| |s ⊻ t|` +* `holley`: Holley inequality. +* `fkg`: Fortuin-Kastelyn-Ginibre inequality. +* `Finset.card_le_card_diffs`: Marica-Schönheim inequality. `|s| ≤ |{a \ b | a, b ∈ s}|` + +## TODO + +Prove that lattices in which `Finset.le_card_infs_mul_card_sups` holds are distributive. See +Daykin, *A lattice is distributive iff |A| |B| <= |A ∨ B| |A ∧ B|* + +Prove the Fishburn-Shepp inequality. + +Is `collapse` a construct generally useful for set family inductions? If so, we should move it to an +earlier file and give it a proper API. +-/ + +open Finset Fintype +open scoped BigOperators FinsetFamily + +variable {α β : Type*} + +section Finset +variable [DecidableEq α] [LinearOrderedCommSemiring β] [ExistsAddOfLE β] {𝒜 ℬ : Finset (Finset α)} + {a : α} {f f₁ f₂ f₃ f₄ g μ : Finset α → β} {s t u : Finset α} + +/-- The `n = 1` case of the Ahlswede-Daykin inequality. Note that we can't just expand everything +out and bound termwise since `c₀ * d₁` appears twice on the RHS of the assumptions while `c₁ * d₀` +does not appear. -/ +private lemma ineq {a₀ a₁ b₀ b₁ c₀ c₁ d₀ d₁ : β} + (ha₀ : 0 ≤ a₀) (ha₁ : 0 ≤ a₁) (hb₀ : 0 ≤ b₀) (hb₁ : 0 ≤ b₁) + (hc₀ : 0 ≤ c₀) (hc₁ : 0 ≤ c₁) (hd₀ : 0 ≤ d₀) (hd₁ : 0 ≤ d₁) + (h₀₀ : a₀ * b₀ ≤ c₀ * d₀) (h₁₀ : a₁ * b₀ ≤ c₀ * d₁) + (h₀₁ : a₀ * b₁ ≤ c₀ * d₁) (h₁₁ : a₁ * b₁ ≤ c₁ * d₁) : + (a₀ + a₁) * (b₀ + b₁) ≤ (c₀ + c₁) * (d₀ + d₁) := by + calc + _ = a₀ * b₀ + (a₀ * b₁ + a₁ * b₀) + a₁ * b₁ := by ring + _ ≤ c₀ * d₀ + (c₀ * d₁ + c₁ * d₀) + c₁ * d₁ := add_le_add_three h₀₀ ?_ h₁₁ + _ = (c₀ + c₁) * (d₀ + d₁) := by ring + obtain hcd | hcd := (mul_nonneg hc₀ hd₁).eq_or_gt + · rw [hcd] at h₀₁ h₁₀ + rw [h₀₁.antisymm, h₁₀.antisymm, add_zero] <;> positivity + refine' le_of_mul_le_mul_right _ hcd + calc (a₀ * b₁ + a₁ * b₀) * (c₀ * d₁) + = a₀ * b₁ * (c₀ * d₁) + c₀ * d₁ * (a₁ * b₀) := by ring + _ ≤ a₀ * b₁ * (a₁ * b₀) + c₀ * d₁ * (c₀ * d₁) := mul_add_mul_le_mul_add_mul h₀₁ h₁₀ + _ = a₀ * b₀ * (a₁ * b₁) + c₀ * d₁ * (c₀ * d₁) := by ring + _ ≤ c₀ * d₀ * (c₁ * d₁) + c₀ * d₁ * (c₀ * d₁) := + add_le_add_right (mul_le_mul h₀₀ h₁₁ (by positivity) $ by positivity) _ + _ = (c₀ * d₁ + c₁ * d₀) * (c₀ * d₁) := by ring + +private def collapse (𝒜 : Finset (Finset α)) (a : α) (f : Finset α → β) (s : Finset α) : β := + ∑ t in 𝒜.filter λ t ↦ t.erase a = s, f t + +private lemma erase_eq_iff (hs : a ∉ s) : t.erase a = s ↔ t = s ∨ t = insert a s := by + by_cases ht : a ∈ t <;> + · simp [ne_of_mem_of_not_mem', erase_eq_iff_eq_insert, *]; + aesop + +private lemma filter_collapse_eq (ha : a ∉ s) (𝒜 : Finset (Finset α)) : + (𝒜.filter λ t ↦ t.erase a = s) = + if s ∈ 𝒜 then + (if insert a s ∈ 𝒜 then {s, insert a s} else {s}) + else if + insert a s ∈ 𝒜 then {insert a s} else ∅ := by + ext t; split_ifs <;> simp [erase_eq_iff ha] <;> aesop + +lemma collapse_eq (ha : a ∉ s) (𝒜 : Finset (Finset α)) (f : Finset α → β) : + collapse 𝒜 a f s = (if s ∈ 𝒜 then f s else 0) + + if insert a s ∈ 𝒜 then f (insert a s) else 0 := by + rw [collapse, filter_collapse_eq ha] + split_ifs <;> simp [(ne_of_mem_of_not_mem' (mem_insert_self a s) ha).symm, *] + +lemma collapse_of_mem (ha : a ∉ s) (ht : t ∈ 𝒜) (hu : u ∈ 𝒜) (hts : t = s) + (hus : u = insert a s) : collapse 𝒜 a f s = f t + f u := by + subst hts; subst hus; simp_rw [collapse_eq ha, if_pos ht, if_pos hu] + +lemma le_collapse_of_mem (ha : a ∉ s) (hf : 0 ≤ f) (hts : t = s) (ht : t ∈ 𝒜) : + f t ≤ collapse 𝒜 a f s := by + subst hts + rw [collapse_eq ha, if_pos ht] + split_ifs + · exact le_add_of_nonneg_right $ hf _ + · rw [add_zero] + +lemma le_collapse_of_insert_mem (ha : a ∉ s) (hf : 0 ≤ f) (hts : t = insert a s) (ht : t ∈ 𝒜) : + f t ≤ collapse 𝒜 a f s := by + rw [collapse_eq ha, ←hts, if_pos ht] + split_ifs + · exact le_add_of_nonneg_left $ hf _ + · rw [zero_add] + +lemma collapse_nonneg (hf : 0 ≤ f) : 0 ≤ collapse 𝒜 a f := λ _s ↦ sum_nonneg $ λ _t _ ↦ hf _ + +lemma collapse_modular (hu : a ∉ u) (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ f₂) (h₃ : 0 ≤ f₃) (h₄ : 0 ≤ f₄) + (h : ∀ ⦃s⦄, s ⊆ insert a u → ∀ ⦃t⦄, t ⊆ insert a u → f₁ s * f₂ t ≤ f₃ (s ∩ t) * f₄ (s ∪ t)) + (𝒜 ℬ : Finset (Finset α)) : + ∀ ⦃s⦄, s ⊆ u → ∀ ⦃t⦄, t ⊆ u → collapse 𝒜 a f₁ s * collapse ℬ a f₂ t ≤ + collapse (𝒜 ⊼ ℬ) a f₃ (s ∩ t) * collapse (𝒜 ⊻ ℬ) a f₄ (s ∪ t) := by + rintro s hsu t htu + -- Gather a bunch of facts we'll need a lot + have := hsu.trans $ subset_insert a _ + have := htu.trans $ subset_insert a _ + have := insert_subset_insert a hsu + have := insert_subset_insert a htu + have has := not_mem_mono hsu hu + have hat := not_mem_mono htu hu + have : a ∉ s ∩ t := not_mem_mono ((inter_subset_left _ t).trans hsu) hu + have := not_mem_union.2 ⟨has, hat⟩ + rw [collapse_eq has] + split_ifs + · rw [collapse_eq hat] + split_ifs + · rw [collapse_of_mem ‹_› (inter_mem_infs ‹_› ‹_›) (inter_mem_infs ‹_› ‹_›) rfl + (insert_inter_distrib _ _ _).symm, collapse_of_mem ‹_› (union_mem_sups ‹_› ‹_›) + (union_mem_sups ‹_› ‹_›) rfl (insert_union_distrib _ _ _).symm] + refine' ineq (h₁ _) (h₁ _) (h₂ _) (h₂ _) (h₃ _) (h₃ _) (h₄ _) (h₄ _) (h ‹_› ‹_›) _ _ _ + · simpa [*] using h ‹insert a s ⊆ _› ‹t ⊆ _› + · simpa [*] using h ‹s ⊆ _› ‹insert a t ⊆ _› + · simpa [*] using h ‹insert a s ⊆ _› ‹insert a t ⊆ _› + · rw [add_zero, add_mul] + refine' (add_le_add (h ‹_› ‹_›) $ h ‹_› ‹_›).trans _ + rw [collapse_of_mem ‹_› (union_mem_sups ‹_› ‹_›) (union_mem_sups ‹_› ‹_›) rfl + (insert_union _ _ _), insert_inter_of_not_mem ‹_›, ←mul_add] + exact mul_le_mul_of_nonneg_right (le_collapse_of_mem ‹_› h₃ rfl $ inter_mem_infs ‹_› ‹_›) $ + add_nonneg (h₄ _) $ h₄ _ + · rw [zero_add, add_mul] + refine' (add_le_add (h ‹_› ‹_›) $ h ‹_› ‹_›).trans _ + rw [collapse_of_mem ‹_› (inter_mem_infs ‹_› ‹_›) (inter_mem_infs ‹_› ‹_›) + (inter_insert_of_not_mem ‹_›) (insert_inter_distrib _ _ _).symm, union_insert, + insert_union_distrib, ←add_mul] + exact mul_le_mul_of_nonneg_left (le_collapse_of_insert_mem ‹_› h₄ + (insert_union_distrib _ _ _).symm $ union_mem_sups ‹_› ‹_›) $ add_nonneg (h₃ _) $ h₃ _ + · rw [add_zero, mul_zero] + exact mul_nonneg (collapse_nonneg h₃ _) $ collapse_nonneg h₄ _ + · rw [add_zero, collapse_eq hat, mul_add] + split_ifs + · refine' (add_le_add (h ‹_› ‹_›) $ h ‹_› ‹_›).trans _ + rw [collapse_of_mem ‹_› (union_mem_sups ‹_› ‹_›) (union_mem_sups ‹_› ‹_›) rfl + (union_insert _ _ _), inter_insert_of_not_mem ‹_›, ←mul_add] + exact mul_le_mul_of_nonneg_right (le_collapse_of_mem ‹_› h₃ rfl $ inter_mem_infs ‹_› ‹_›) $ + add_nonneg (h₄ _) $ h₄ _ + · rw [mul_zero, add_zero] + exact (h ‹_› ‹_›).trans $ mul_le_mul (le_collapse_of_mem ‹_› h₃ rfl $ + inter_mem_infs ‹_› ‹_›) (le_collapse_of_mem ‹_› h₄ rfl $ union_mem_sups ‹_› ‹_›) + (h₄ _) $ collapse_nonneg h₃ _ + · rw [mul_zero, zero_add] + refine' (h ‹_› ‹_›).trans $ mul_le_mul _ (le_collapse_of_insert_mem ‹_› h₄ + (union_insert _ _ _) $ union_mem_sups ‹_› ‹_›) (h₄ _) $ collapse_nonneg h₃ _ + exact le_collapse_of_mem (not_mem_mono (inter_subset_left _ _) ‹_›) h₃ + (inter_insert_of_not_mem ‹_›) $ inter_mem_infs ‹_› ‹_› + · simp_rw [mul_zero, add_zero] + exact mul_nonneg (collapse_nonneg h₃ _) $ collapse_nonneg h₄ _ + · rw [zero_add, collapse_eq hat, mul_add] + split_ifs + · refine' (add_le_add (h ‹_› ‹_›) $ h ‹_› ‹_›).trans _ + rw [collapse_of_mem ‹_› (inter_mem_infs ‹_› ‹_›) (inter_mem_infs ‹_› ‹_›) + (insert_inter_of_not_mem ‹_›) (insert_inter_distrib _ _ _).symm, + insert_inter_of_not_mem ‹_›, ←insert_inter_distrib, insert_union, insert_union_distrib, + ←add_mul] + exact mul_le_mul_of_nonneg_left (le_collapse_of_insert_mem ‹_› h₄ + (insert_union_distrib _ _ _).symm $ union_mem_sups ‹_› ‹_›) $ add_nonneg (h₃ _) $ h₃ _ + · rw [mul_zero, add_zero] + refine' (h ‹_› ‹_›).trans $ mul_le_mul (le_collapse_of_mem ‹_› h₃ + (insert_inter_of_not_mem ‹_›) $ inter_mem_infs ‹_› ‹_›) (le_collapse_of_insert_mem ‹_› h₄ + (insert_union _ _ _) $ union_mem_sups ‹_› ‹_›) (h₄ _) $ collapse_nonneg h₃ _ + · rw [mul_zero, zero_add] + exact (h ‹_› ‹_›).trans $ mul_le_mul (le_collapse_of_insert_mem ‹_› h₃ + (insert_inter_distrib _ _ _).symm $ inter_mem_infs ‹_› ‹_›) (le_collapse_of_insert_mem ‹_› + h₄ (insert_union_distrib _ _ _).symm $ union_mem_sups ‹_› ‹_›) (h₄ _) $ collapse_nonneg h₃ _ + · simp_rw [mul_zero, add_zero] + exact mul_nonneg (collapse_nonneg h₃ _) $ collapse_nonneg h₄ _ + · simp_rw [add_zero, zero_mul] + exact mul_nonneg (collapse_nonneg h₃ _) $ collapse_nonneg h₄ _ + +lemma sum_collapse (h𝒜 : 𝒜 ⊆ (insert a u).powerset) (hu : a ∉ u) : + ∑ s in u.powerset, collapse 𝒜 a f s = ∑ s in 𝒜, f s := by + calc + _ = ∑ s in u.powerset ∩ 𝒜, f s + ∑ s in u.powerset.image (insert a) ∩ 𝒜, f s := ?_ + _ = ∑ s in u.powerset ∩ 𝒜, f s + ∑ s in ((insert a u).powerset \ u.powerset) ∩ 𝒜, f s := ?_ + _ = ∑ s in 𝒜, f s := ?_ + · rw [←sum_ite_mem, ←sum_ite_mem, sum_image, ←sum_add_distrib] + · exact sum_congr rfl λ s hs ↦ collapse_eq (not_mem_mono (mem_powerset.1 hs) hu) _ _ + · exact (insert_erase_invOn.2.injOn).mono λ s hs ↦ not_mem_mono (mem_powerset.1 hs) hu + · congr with s + simp only [mem_image, mem_powerset, mem_sdiff, subset_insert_iff] + refine' ⟨_, λ h ↦ ⟨_, h.1, _⟩⟩ + · rintro ⟨s, hs, rfl⟩ + exact ⟨subset_insert_iff.1 $ insert_subset_insert _ hs, λ h ↦ hu $ h $ mem_insert_self _ _⟩ + · rw [insert_erase (erase_ne_self.1 λ hs ↦ ?_)] + rw [hs] at h + exact h.2 h.1 + · rw [←sum_union (disjoint_sdiff_self_right.mono inf_le_left inf_le_left), ←inter_distrib_right, + union_sdiff_of_subset (powerset_mono.2 $ subset_insert _ _), inter_eq_right.2 h𝒜] + +/-- The **Four Functions Theorem** on a powerset algebra. See `four_functions_theorem` for the +finite distributive lattice generalisation. -/ +protected lemma Finset.four_functions_theorem (u : Finset α) + (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ f₂) (h₃ : 0 ≤ f₃) (h₄ : 0 ≤ f₄) + (h : ∀ ⦃s⦄, s ⊆ u → ∀ ⦃t⦄, t ⊆ u → f₁ s * f₂ t ≤ f₃ (s ∩ t) * f₄ (s ∪ t)) + {𝒜 ℬ : Finset (Finset α)} (h𝒜 : 𝒜 ⊆ u.powerset) (hℬ : ℬ ⊆ u.powerset) : + (∑ s in 𝒜, f₁ s) * ∑ s in ℬ, f₂ s ≤ (∑ s in 𝒜 ⊼ ℬ, f₃ s) * ∑ s in 𝒜 ⊻ ℬ, f₄ s := by + induction' u using Finset.induction with a u hu ih generalizing f₁ f₂ f₃ f₄ 𝒜 ℬ + · simp only [Finset.powerset_empty, Finset.subset_singleton_iff] at h𝒜 hℬ + obtain rfl | rfl := h𝒜 <;> obtain rfl | rfl := hℬ <;> simp; exact h subset_rfl subset_rfl + specialize ih (collapse_nonneg h₁) (collapse_nonneg h₂) (collapse_nonneg h₃) (collapse_nonneg h₄) + (collapse_modular hu h₁ h₂ h₃ h₄ h 𝒜 ℬ) Subset.rfl Subset.rfl + have : 𝒜 ⊼ ℬ ⊆ powerset (insert a u) := by simpa using infs_subset h𝒜 hℬ + have : 𝒜 ⊻ ℬ ⊆ powerset (insert a u) := by simpa using sups_subset h𝒜 hℬ + simpa (config := {decide := true}) only + [powerset_sups_powerset_self, powerset_infs_powerset_self, sum_collapse, *] using ih + +variable (f₁ f₂ f₃ f₄) [Fintype α] + +private lemma four_functions_theorem_aux (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ f₂) (h₃ : 0 ≤ f₃) (h₄ : 0 ≤ f₄) + (h : ∀ s t, f₁ s * f₂ t ≤ f₃ (s ∩ t) * f₄ (s ∪ t)) (𝒜 ℬ : Finset (Finset α)) : + (∑ s in 𝒜, f₁ s) * ∑ s in ℬ, f₂ s ≤ (∑ s in 𝒜 ⊼ ℬ, f₃ s) * ∑ s in 𝒜 ⊻ ℬ, f₄ s := by + refine' univ.four_functions_theorem h₁ h₂ h₃ h₄ _ _ _ <;> simp [h] + +end Finset + +section DistribLattice +variable [DistribLattice α] [DecidableEq α] [LinearOrderedCommSemiring β] [ExistsAddOfLE β] + (f f₁ f₂ f₃ f₄ g μ : α → β) + +open Function + +/-- The **Four Functions Theorem**, aka **Ahlswede-Daykin Inequality**. -/ +lemma four_functions_theorem (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ f₂) (h₃ : 0 ≤ f₃) (h₄ : 0 ≤ f₄) + (h : ∀ a b, f₁ a * f₂ b ≤ f₃ (a ⊓ b) * f₄ (a ⊔ b)) (s t : Finset α) : + (∑ a in s, f₁ a) * ∑ a in t, f₂ a ≤ (∑ a in s ⊼ t, f₃ a) * ∑ a in s ⊻ t, f₄ a := by + classical + set L : Sublattice α := ⟨latticeClosure (s ∪ t), isSublattice_latticeClosure.1, + isSublattice_latticeClosure.2⟩ + have : Finite L := (s.finite_toSet.union t.finite_toSet).latticeClosure.to_subtype + set s' : Finset L := s.preimage (↑) $ Subtype.coe_injective.injOn _ + set t' : Finset L := t.preimage (↑) $ Subtype.coe_injective.injOn _ + have hs' : s'.map ⟨L.subtype, Subtype.coe_injective⟩ = s := by + simp [map_eq_image, image_preimage, filter_eq_self] + exact λ a ha ↦ subset_latticeClosure $ Set.subset_union_left _ _ ha + have ht' : t'.map ⟨L.subtype, Subtype.coe_injective⟩ = t := by + simp [map_eq_image, image_preimage, filter_eq_self] + exact λ a ha ↦ subset_latticeClosure $ Set.subset_union_right _ _ ha + clear_value s' t' + obtain ⟨β, _, _, g, hg⟩ := exists_birkhoff_representation L + have := four_functions_theorem_aux (extend g (f₁ ∘ (↑)) 0) (extend g (f₂ ∘ (↑)) 0) + (extend g (f₃ ∘ (↑)) 0) (extend g (f₄ ∘ (↑)) 0) (extend_nonneg (λ _ ↦ h₁ _) le_rfl) + (extend_nonneg (λ _ ↦ h₂ _) le_rfl) (extend_nonneg (λ _ ↦ h₃ _) le_rfl) + (extend_nonneg (λ _ ↦ h₄ _) le_rfl) ?_ (s'.map ⟨g, hg⟩) (t'.map ⟨g, hg⟩) + simpa only [←hs', ←ht', ←map_sups, ←map_infs, sum_map, Embedding.coeFn_mk, hg.extend_apply] + using this + rintro s t + classical + obtain ⟨a, rfl⟩ | hs := em (∃ a, g a = s) + · obtain ⟨b, rfl⟩ | ht := em (∃ b, g b = t) + · simp_rw [←sup_eq_union, ←inf_eq_inter, ←map_sup, ←map_inf, hg.extend_apply] + exact h _ _ + · simpa [extend_apply' _ _ _ ht] using mul_nonneg + (extend_nonneg (λ a : L ↦ h₃ a) le_rfl _) (extend_nonneg (λ a : L ↦ h₄ a) le_rfl _) + · simpa [extend_apply' _ _ _ hs] using mul_nonneg + (extend_nonneg (λ a : L ↦ h₃ a) le_rfl _) (extend_nonneg (λ a : L ↦ h₄ a) le_rfl _) + +/-- An inequality of Daykin. Interestingly, any lattice in which this inequality holds is +distributive. -/ +lemma Finset.le_card_infs_mul_card_sups (s t : Finset α) : + s.card * t.card ≤ (s ⊼ t).card * (s ⊻ t).card := by + simpa using four_functions_theorem (1 : α → ℕ) 1 1 1 zero_le_one zero_le_one zero_le_one + zero_le_one (λ _ _ ↦ le_rfl) s t + +variable [Fintype α] + +/-- Special case of the **Four Functions Theorem** when `s = t = univ`. -/ +lemma four_functions_theorem_univ (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ f₂) (h₃ : 0 ≤ f₃) (h₄ : 0 ≤ f₄) + (h : ∀ a b, f₁ a * f₂ b ≤ f₃ (a ⊓ b) * f₄ (a ⊔ b)) : + (∑ a, f₁ a) * ∑ a, f₂ a ≤ (∑ a, f₃ a) * ∑ a, f₄ a := by + simpa using four_functions_theorem f₁ f₂ f₃ f₄ h₁ h₂ h₃ h₄ h univ univ + +/-- The **Holley Inequality**. -/ +lemma holley (hμ₀ : 0 ≤ μ) (hf : 0 ≤ f) (hg : 0 ≤ g) (hμ : Monotone μ) + (hfg : ∑ a, f a = ∑ a, g a) (h : ∀ a b, f a * g b ≤ f (a ⊓ b) * g (a ⊔ b)) : + ∑ a, μ a * f a ≤ ∑ a, μ a * g a := by + obtain rfl | hf := hf.eq_or_lt + · simp [eq_comm, Fintype.sum_eq_zero_iff_of_nonneg hg] at hfg + simp [hfg] + obtain rfl | hg := hg.eq_or_lt + · simp [Fintype.sum_eq_zero_iff_of_nonneg hf.le] at hfg + simp [hfg] + have' := four_functions_theorem g (μ * f) f (μ * g) hg.le (mul_nonneg hμ₀ hf.le) hf.le + (mul_nonneg hμ₀ hg.le) (λ a b ↦ _) univ univ + simpa [hfg, sum_pos hg] using this + · simp_rw [Pi.mul_apply, mul_left_comm _ (μ _), mul_comm (g _)] + rw [sup_comm, inf_comm] + exact mul_le_mul (hμ le_sup_left) (h _ _) (mul_nonneg (hf.le _) $ hg.le _) $ hμ₀ _ + +/-- The **Fortuin-Kastelyn-Ginibre Inequality**. -/ +lemma fkg (hμ₀ : 0 ≤ μ) (hf₀ : 0 ≤ f) (hg₀ : 0 ≤ g) (hf : Monotone f) (hg : Monotone g) + (hμ : ∀ a b, μ a * μ b ≤ μ (a ⊓ b) * μ (a ⊔ b)) : + (∑ a, μ a * f a) * ∑ a, μ a * g a ≤ (∑ a, μ a) * ∑ a, μ a * (f a * g a) := by + refine' four_functions_theorem_univ (μ * f) (μ * g) μ _ (mul_nonneg hμ₀ hf₀) (mul_nonneg hμ₀ hg₀) + hμ₀ (mul_nonneg hμ₀ $ mul_nonneg hf₀ hg₀) (λ a b ↦ _) + dsimp + rw [mul_mul_mul_comm, ←mul_assoc (μ (a ⊓ b))] + exact mul_le_mul (hμ _ _) (mul_le_mul (hf le_sup_left) (hg le_sup_right) (hg₀ _) $ hf₀ _) + (mul_nonneg (hf₀ _) $ hg₀ _) $ mul_nonneg (hμ₀ _) $ hμ₀ _ + +end DistribLattice + +open Booleanisation + +variable [DecidableEq α] [GeneralizedBooleanAlgebra α] + +/-- A slight generalisation of the **Marica-Schönheim Inequality**. -/ +lemma Finset.le_card_diffs_mul_card_diffs (s t : Finset α) : + s.card * t.card ≤ (s \\ t).card * (t \\ s).card := by + have : ∀ s t : Finset α, (s \\ t).map ⟨_, liftLatticeHom_injective⟩ = + s.map ⟨_, liftLatticeHom_injective⟩ \\ t.map ⟨_, liftLatticeHom_injective⟩ + · rintro s t + simp_rw [map_eq_image] + exact image_image₂_distrib λ a b ↦ rfl + simpa [←card_compls (_ ⊻ _), ←map_sup, ←map_inf, ←this] using + (s.map ⟨_, liftLatticeHom_injective⟩).le_card_infs_mul_card_sups + (t.map ⟨_, liftLatticeHom_injective⟩)ᶜˢ + +/-- The **Marica-Schönheim Inequality**. -/ +lemma Finset.card_le_card_diffs (s : Finset α) : s.card ≤ (s \\ s).card := + le_of_pow_le_pow 2 (zero_le _) two_pos $ by simpa [←sq] using s.le_card_diffs_mul_card_diffs s diff --git a/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean b/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean index bd302168245f8..0edfa6f513ebf 100644 --- a/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean +++ b/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean @@ -59,10 +59,11 @@ theorem IsLowerSet.le_card_inter_finset' (h𝒜 : IsLowerSet (𝒜 : Set (Finset induction' s using Finset.induction with a s hs ih generalizing 𝒜 ℬ · simp_rw [subset_empty, ← subset_singleton_iff', subset_singleton_iff] at h𝒜s hℬs obtain rfl | rfl := h𝒜s - · simp only [card_empty, empty_inter, mul_zero, zero_mul] + · simp only [card_empty, zero_mul, empty_inter, mul_zero, le_refl] obtain rfl | rfl := hℬs - · simp only [card_empty, inter_empty, mul_zero, zero_mul] - · simp only [card_empty, pow_zero, inter_singleton_of_mem, mem_singleton, card_singleton] + · simp only [card_empty, inter_empty, mul_zero, zero_mul, le_refl] + · simp only [card_empty, pow_zero, inter_singleton_of_mem, mem_singleton, card_singleton, + le_refl] rw [card_insert_of_not_mem hs, ← card_memberSubfamily_add_card_nonMemberSubfamily a 𝒜, ← card_memberSubfamily_add_card_nonMemberSubfamily a ℬ, add_mul, mul_add, mul_add, add_comm (_ * _), add_add_add_comm] diff --git a/Mathlib/Combinatorics/SetFamily/Shadow.lean b/Mathlib/Combinatorics/SetFamily/Shadow.lean index 2b84680792ac2..f87fb428d4fe9 100644 --- a/Mathlib/Combinatorics/SetFamily/Shadow.lean +++ b/Mathlib/Combinatorics/SetFamily/Shadow.lean @@ -71,6 +71,9 @@ theorem shadow_empty : ∂ (∅ : Finset (Finset α)) = ∅ := rfl #align finset.shadow_empty Finset.shadow_empty +@[simp] lemma shadow_iterate_empty (k : ℕ) : ∂^[k] (∅ : Finset (Finset α)) = ∅ := by + induction' k <;> simp [*, shadow_empty] + @[simp] theorem shadow_singleton_empty : ∂ ({∅} : Finset (Finset α)) = ∅ := rfl @@ -83,9 +86,9 @@ theorem shadow_monotone : Monotone (shadow : Finset (Finset α) → Finset (Fins sup_mono #align finset.shadow_monotone Finset.shadow_monotone -/-- `s` is in the shadow of `𝒜` iff there is a `t ∈ 𝒜` from which we can remove one element to -get `s`. -/ -theorem mem_shadow_iff : s ∈ ∂ 𝒜 ↔ ∃ t ∈ 𝒜, ∃ a ∈ t, erase t a = s := by +/-- `t` is in the shadow of `𝒜` iff there is a `s ∈ 𝒜` from which we can remove one element to +get `t`. -/ +lemma mem_shadow_iff : t ∈ ∂ 𝒜 ↔ ∃ s ∈ 𝒜, ∃ a ∈ s, erase s a = t := by simp only [shadow, mem_sup, mem_image] #align finset.mem_shadow_iff Finset.mem_shadow_iff @@ -93,24 +96,74 @@ theorem erase_mem_shadow (hs : s ∈ 𝒜) (ha : a ∈ s) : erase s a ∈ ∂ mem_shadow_iff.2 ⟨s, hs, a, ha, rfl⟩ #align finset.erase_mem_shadow Finset.erase_mem_shadow +/-- `t ∈ ∂𝒜` iff `t` is exactly one element less than something from `𝒜`. + +See also `Finset.mem_shadow_iff_exists_mem_card_add_one`. -/ +lemma mem_shadow_iff_exists_sdiff : t ∈ ∂ 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ (s \ t).card = 1 := by + simp_rw [mem_shadow_iff, ←covby_iff_card_sdiff_eq_one, covby_iff_exists_erase, eq_comm] + /-- `t` is in the shadow of `𝒜` iff we can add an element to it so that the resulting finset is in `𝒜`. -/ -theorem mem_shadow_iff_insert_mem : s ∈ ∂ 𝒜 ↔ ∃ (a : _) (_ : a ∉ s), insert a s ∈ 𝒜 := by - refine' mem_shadow_iff.trans ⟨_, _⟩ - · rintro ⟨s, hs, a, ha, rfl⟩ - refine' ⟨a, not_mem_erase a s, _⟩ - rwa [insert_erase ha] - · rintro ⟨a, ha, hs⟩ - exact ⟨insert a s, hs, a, mem_insert_self _ _, erase_insert ha⟩ +lemma mem_shadow_iff_insert_mem : t ∈ ∂ 𝒜 ↔ ∃ a, a ∉ t ∧ insert a t ∈ 𝒜 := by + simp_rw [mem_shadow_iff_exists_sdiff, ←covby_iff_card_sdiff_eq_one, covby_iff_exists_insert] + aesop #align finset.mem_shadow_iff_insert_mem Finset.mem_shadow_iff_insert_mem +/-- `s ∈ ∂ 𝒜` iff `s` is exactly one element less than something from `𝒜`. + +See also `Finset.mem_shadow_iff_exists_sdiff`. -/ +lemma mem_shadow_iff_exists_mem_card_add_one : + t ∈ ∂ 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ s.card = t.card + 1 := by + refine mem_shadow_iff_exists_sdiff.trans $ exists_congr fun t ↦ and_congr_right fun _ ↦ + and_congr_right fun hst ↦ ?_ + rw [card_sdiff hst, tsub_eq_iff_eq_add_of_le, add_comm] + exact card_mono hst +#align finset.mem_shadow_iff_exists_mem_card_add_one Finset.mem_shadow_iff_exists_mem_card_add_one + +lemma mem_shadow_iterate_iff_exists_card : + t ∈ ∂^[k] 𝒜 ↔ ∃ u : Finset α, u.card = k ∧ Disjoint t u ∧ t ∪ u ∈ 𝒜 := by + induction' k with k ih generalizing t + · simp + simp only [mem_shadow_iff_insert_mem, ih, Function.iterate_succ_apply', card_eq_succ] + aesop + +/-- `t ∈ ∂^k 𝒜` iff `t` is exactly `k` elements less than something from `𝒜`. + +See also `Finset.mem_shadow_iff_exists_mem_card_add`. -/ +lemma mem_shadow_iterate_iff_exists_sdiff : t ∈ ∂^[k] 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ (s \ t).card = k := by + rw [mem_shadow_iterate_iff_exists_card] + constructor + · rintro ⟨u, rfl, htu, hsuA⟩ + exact ⟨_, hsuA, subset_union_left _ _, by rw [union_sdiff_cancel_left htu]⟩ + · rintro ⟨s, hs, hts, rfl⟩ + refine ⟨s \ t, rfl, disjoint_sdiff, ?_⟩ + rwa [union_sdiff_self_eq_union, union_eq_right.2 hts] + +/-- `t ∈ ∂^k 𝒜` iff `t` is exactly `k` elements less than something in `𝒜`. + +See also `Finset.mem_shadow_iterate_iff_exists_sdiff`. -/ +lemma mem_shadow_iterate_iff_exists_mem_card_add : + t ∈ ∂^[k] 𝒜 ↔ ∃ s ∈ 𝒜, t ⊆ s ∧ s.card = t.card + k := by + refine mem_shadow_iterate_iff_exists_sdiff.trans $ exists_congr fun t ↦ and_congr_right fun _ ↦ + and_congr_right fun hst ↦ ?_ + rw [card_sdiff hst, tsub_eq_iff_eq_add_of_le, add_comm] + exact card_mono hst +#align finset.mem_shadow_iff_exists_mem_card_add Finset.mem_shadow_iterate_iff_exists_mem_card_add + /-- The shadow of a family of `r`-sets is a family of `r - 1`-sets. -/ -protected theorem Set.Sized.shadow (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : +protected theorem _root_.Set.Sized.shadow (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : (∂ 𝒜 : Set (Finset α)).Sized (r - 1) := by intro A h obtain ⟨A, hA, i, hi, rfl⟩ := mem_shadow_iff.1 h rw [card_erase_of_mem hi, h𝒜 hA] -#align finset.set.sized.shadow Finset.Set.Sized.shadow +#align finset.set.sized.shadow Set.Sized.shadow + +/-- The `k`-th shadow of a family of `r`-sets is a family of `r - k`-sets. -/ +lemma _root_.Set.Sized.shadow_iterate (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : + (∂^[k] 𝒜 : Set (Finset α)).Sized (r - k) := by + simp_rw [Set.Sized, mem_coe, mem_shadow_iterate_iff_exists_sdiff] + rintro t ⟨s, hs, hts, rfl⟩ + rw [card_sdiff hts, ←h𝒜 hs, Nat.sub_sub_self (card_le_of_subset hts)] theorem sized_shadow_iff (h : ∅ ∉ 𝒜) : (∂ 𝒜 : Set (Finset α)).Sized r ↔ (𝒜 : Set (Finset α)).Sized (r + 1) := by @@ -119,55 +172,12 @@ theorem sized_shadow_iff (h : ∅ ∉ 𝒜) : rw [← h𝒜 (erase_mem_shadow hs ha), card_erase_add_one ha] #align finset.sized_shadow_iff Finset.sized_shadow_iff -/-- `s ∈ ∂ 𝒜` iff `s` is exactly one element less than something from `𝒜` -/ -theorem mem_shadow_iff_exists_mem_card_add_one : - s ∈ ∂ 𝒜 ↔ ∃ t ∈ 𝒜, s ⊆ t ∧ t.card = s.card + 1 := by - refine' mem_shadow_iff_insert_mem.trans ⟨_, _⟩ - · rintro ⟨a, ha, hs⟩ - exact ⟨insert a s, hs, subset_insert _ _, card_insert_of_not_mem ha⟩ - · rintro ⟨t, ht, hst, h⟩ - obtain ⟨a, ha⟩ : ∃ a, t \ s = {a} := - card_eq_one.1 (by rw [card_sdiff hst, h, add_tsub_cancel_left]) - exact - ⟨a, fun hat => not_mem_sdiff_of_mem_right hat (ha.superset <| mem_singleton_self a), - by rwa [insert_eq a s, ← ha, sdiff_union_of_subset hst]⟩ -#align finset.mem_shadow_iff_exists_mem_card_add_one Finset.mem_shadow_iff_exists_mem_card_add_one - /-- Being in the shadow of `𝒜` means we have a superset in `𝒜`. -/ -theorem exists_subset_of_mem_shadow (hs : s ∈ ∂ 𝒜) : ∃ t ∈ 𝒜, s ⊆ t := +lemma exists_subset_of_mem_shadow (hs : t ∈ ∂ 𝒜) : ∃ s ∈ 𝒜, t ⊆ s := let ⟨t, ht, hst⟩ := mem_shadow_iff_exists_mem_card_add_one.1 hs ⟨t, ht, hst.1⟩ #align finset.exists_subset_of_mem_shadow Finset.exists_subset_of_mem_shadow -/-- `t ∈ ∂^k 𝒜` iff `t` is exactly `k` elements less than something in `𝒜`. -/ -theorem mem_shadow_iff_exists_mem_card_add : - s ∈ ∂ ^[k] 𝒜 ↔ ∃ t ∈ 𝒜, s ⊆ t ∧ t.card = s.card + k := by - induction' k with k ih generalizing 𝒜 s - · refine' ⟨fun hs => ⟨s, hs, Subset.refl _, rfl⟩, _⟩ - rintro ⟨t, ht, hst, hcard⟩ - rwa [eq_of_subset_of_card_le hst hcard.le] - simp only [exists_prop, Function.comp_apply, Function.iterate_succ] - refine' ih.trans _ - clear ih - constructor - · rintro ⟨t, ht, hst, hcardst⟩ - obtain ⟨u, hu, htu, hcardtu⟩ := mem_shadow_iff_exists_mem_card_add_one.1 ht - refine' ⟨u, hu, hst.trans htu, _⟩ - rw [hcardtu, hcardst] - rfl - · rintro ⟨t, ht, hst, hcard⟩ - obtain ⟨u, hsu, hut, hu⟩ := - Finset.exists_intermediate_set k - (by - rw [add_comm, hcard] - exact le_succ _) - hst - rw [add_comm] at hu - refine' ⟨u, mem_shadow_iff_exists_mem_card_add_one.2 ⟨t, ht, hut, _⟩, hsu, hu⟩ - rw [hcard, hu] - rfl -#align finset.mem_shadow_iff_exists_mem_card_add Finset.mem_shadow_iff_exists_mem_card_add - end Shadow open FinsetFamily @@ -198,48 +208,83 @@ theorem upShadow_monotone : Monotone (upShadow : Finset (Finset α) → Finset ( fun _ _ => sup_mono #align finset.up_shadow_monotone Finset.upShadow_monotone -/-- `s` is in the upper shadow of `𝒜` iff there is a `t ∈ 𝒜` from which we can remove one element -to get `s`. -/ -theorem mem_upShadow_iff : s ∈ ∂⁺ 𝒜 ↔ ∃ t ∈ 𝒜, ∃ (a : _) (_ : a ∉ t), insert a t = s := by - simp_rw [upShadow, mem_sup, mem_image, exists_prop, mem_compl] +/-- `t` is in the upper shadow of `𝒜` iff there is a `s ∈ 𝒜` from which we can remove one element +to get `t`. -/ +lemma mem_upShadow_iff : t ∈ ∂⁺ 𝒜 ↔ ∃ s ∈ 𝒜, ∃ a, a ∉ s ∧ insert a s = t := by + simp_rw [upShadow, mem_sup, mem_image, mem_compl] #align finset.mem_up_shadow_iff Finset.mem_upShadow_iff theorem insert_mem_upShadow (hs : s ∈ 𝒜) (ha : a ∉ s) : insert a s ∈ ∂⁺ 𝒜 := mem_upShadow_iff.2 ⟨s, hs, a, ha, rfl⟩ #align finset.insert_mem_up_shadow Finset.insert_mem_upShadow -/-- The upper shadow of a family of `r`-sets is a family of `r + 1`-sets. -/ -protected theorem Set.Sized.upShadow (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : - (∂⁺ 𝒜 : Set (Finset α)).Sized (r + 1) := by - intro A h - obtain ⟨A, hA, i, hi, rfl⟩ := mem_upShadow_iff.1 h - rw [card_insert_of_not_mem hi, h𝒜 hA] -#align finset.set.sized.up_shadow Finset.Set.Sized.upShadow +/-- `t` is in the upper shadow of `𝒜` iff `t` is exactly one element more than something from `𝒜`. + +See also `Finset.mem_upShadow_iff_exists_mem_card_add_one`. -/ +lemma mem_upShadow_iff_exists_sdiff : t ∈ ∂⁺ 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ (t \ s).card = 1 := by + simp_rw [mem_upShadow_iff, ←covby_iff_card_sdiff_eq_one, covby_iff_exists_insert] /-- `t` is in the upper shadow of `𝒜` iff we can remove an element from it so that the resulting finset is in `𝒜`. -/ -theorem mem_upShadow_iff_erase_mem : s ∈ ∂⁺ 𝒜 ↔ ∃ a ∈ s, s.erase a ∈ 𝒜 := by - refine' mem_upShadow_iff.trans ⟨_, _⟩ - · rintro ⟨s, hs, a, ha, rfl⟩ - refine' ⟨a, mem_insert_self a s, _⟩ - rwa [erase_insert ha] - · rintro ⟨a, ha, hs⟩ - exact ⟨s.erase a, hs, a, not_mem_erase _ _, insert_erase ha⟩ +lemma mem_upShadow_iff_erase_mem : t ∈ ∂⁺ 𝒜 ↔ ∃ a, a ∈ t ∧ erase t a ∈ 𝒜 := by + simp_rw [mem_upShadow_iff_exists_sdiff, ←covby_iff_card_sdiff_eq_one, covby_iff_exists_erase] + aesop #align finset.mem_up_shadow_iff_erase_mem Finset.mem_upShadow_iff_erase_mem -/-- `s ∈ ∂⁺ 𝒜` iff `s` is exactly one element less than something from `𝒜`. -/ -theorem mem_upShadow_iff_exists_mem_card_add_one : - s ∈ ∂⁺ 𝒜 ↔ ∃ t ∈ 𝒜, t ⊆ s ∧ t.card + 1 = s.card := by - refine' mem_upShadow_iff_erase_mem.trans ⟨_, _⟩ - · rintro ⟨a, ha, hs⟩ - exact ⟨s.erase a, hs, erase_subset _ _, card_erase_add_one ha⟩ - · rintro ⟨t, ht, hts, h⟩ - obtain ⟨a, ha⟩ : ∃ a, s \ t = {a} := - card_eq_one.1 (by rw [card_sdiff hts, ← h, add_tsub_cancel_left]) - refine' ⟨a, sdiff_subset _ _ ((ha.ge : _ ⊆ _) <| mem_singleton_self a), _⟩ - rwa [← sdiff_singleton_eq_erase, ← ha, sdiff_sdiff_eq_self hts] +/-- `t` is in the upper shadow of `𝒜` iff `t` is exactly one element less than something from `𝒜`. + +See also `Finset.mem_upShadow_iff_exists_sdiff`. -/ +lemma mem_upShadow_iff_exists_mem_card_add_one : + t ∈ ∂⁺ 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ t.card = s.card + 1 := by + refine mem_upShadow_iff_exists_sdiff.trans $ exists_congr fun t ↦ and_congr_right fun _ ↦ + and_congr_right fun hst ↦ ?_ + rw [card_sdiff hst, tsub_eq_iff_eq_add_of_le, add_comm] + exact card_mono hst #align finset.mem_up_shadow_iff_exists_mem_card_add_one Finset.mem_upShadow_iff_exists_mem_card_add_one +lemma mem_upShadow_iterate_iff_exists_card : + t ∈ ∂⁺^[k] 𝒜 ↔ ∃ u : Finset α, u.card = k ∧ u ⊆ t ∧ t \ u ∈ 𝒜 := by + induction' k with k ih generalizing t + · simp + simp only [mem_upShadow_iff_erase_mem, ih, Function.iterate_succ_apply', card_eq_succ, + subset_erase, erase_sdiff_comm, ←sdiff_insert] + constructor + · rintro ⟨a, hat, u, rfl, ⟨hut, hau⟩, htu⟩ + exact ⟨_, ⟨_, _, hau, rfl, rfl⟩, insert_subset hat hut, htu⟩ + · rintro ⟨_, ⟨a, u, hau, rfl, rfl⟩, hut, htu⟩ + rw [insert_subset_iff] at hut + exact ⟨a, hut.1, _, rfl, ⟨hut.2, hau⟩, htu⟩ + +/-- `t` is in the upper shadow of `𝒜` iff `t` is exactly `k` elements less than something from `𝒜`. + +See also `Finset.mem_upShadow_iff_exists_mem_card_add`. -/ +lemma mem_upShadow_iterate_iff_exists_sdiff : + t ∈ ∂⁺^[k] 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ (t \ s).card = k := by + rw [mem_upShadow_iterate_iff_exists_card] + constructor + · rintro ⟨u, rfl, hut, htu⟩ + exact ⟨_, htu, sdiff_subset _ _, by rw [sdiff_sdiff_eq_self hut]⟩ + · rintro ⟨s, hs, hst, rfl⟩ + exact ⟨_, rfl, sdiff_subset _ _, by rwa [sdiff_sdiff_eq_self hst]⟩ + +/-- `t ∈ ∂⁺^k 𝒜` iff `t` is exactly `k` elements less than something in `𝒜`. + +See also `Finset.mem_upShadow_iterate_iff_exists_sdiff`. -/ +lemma mem_upShadow_iterate_iff_exists_mem_card_add : + t ∈ ∂⁺^[k] 𝒜 ↔ ∃ s ∈ 𝒜, s ⊆ t ∧ t.card = s.card + k := by + refine mem_upShadow_iterate_iff_exists_sdiff.trans $ exists_congr fun t ↦ and_congr_right fun _ ↦ + and_congr_right fun hst ↦ ?_ + rw [card_sdiff hst, tsub_eq_iff_eq_add_of_le, add_comm] + exact card_mono hst + +/-- The upper shadow of a family of `r`-sets is a family of `r + 1`-sets. -/ +protected lemma _root_.Set.Sized.upShadow (h𝒜 : (𝒜 : Set (Finset α)).Sized r) : + (∂⁺ 𝒜 : Set (Finset α)).Sized (r + 1) := by + intro A h + obtain ⟨A, hA, i, hi, rfl⟩ := mem_upShadow_iff.1 h + rw [card_insert_of_not_mem hi, h𝒜 hA] +#align finset.set.sized.up_shadow Set.Sized.upShadow + /-- Being in the upper shadow of `𝒜` means we have a superset in `𝒜`. -/ theorem exists_subset_of_mem_upShadow (hs : s ∈ ∂⁺ 𝒜) : ∃ t ∈ 𝒜, t ⊆ s := let ⟨t, ht, hts, _⟩ := mem_upShadow_iff_exists_mem_card_add_one.1 hs @@ -260,7 +305,7 @@ theorem mem_upShadow_iff_exists_mem_card_add : · rintro ⟨t, ht, hts, hcardst⟩ obtain ⟨u, hu, hut, hcardtu⟩ := mem_upShadow_iff_exists_mem_card_add_one.1 ht refine' ⟨u, hu, hut.trans hts, _⟩ - rw [← hcardst, ← hcardtu, add_right_comm] + rw [← hcardst, hcardtu, add_right_comm] rfl · rintro ⟨t, ht, hts, hcard⟩ obtain ⟨u, htu, hus, hu⟩ := @@ -270,7 +315,7 @@ theorem mem_upShadow_iff_exists_mem_card_add : exact add_le_add_left (succ_le_of_lt (zero_lt_succ _)) _) hts rw [add_comm] at hu - refine' ⟨u, mem_upShadow_iff_exists_mem_card_add_one.2 ⟨t, ht, htu, hu.symm⟩, hus, _⟩ + refine' ⟨u, mem_upShadow_iff_exists_mem_card_add_one.2 ⟨t, ht, htu, hu⟩, hus, _⟩ rw [hu, ← hcard, add_right_comm] rfl #align finset.mem_up_shadow_iff_exists_mem_card_add Finset.mem_upShadow_iff_exists_mem_card_add diff --git a/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean b/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean index 65a506351d768..c060b7a3a5b7d 100644 --- a/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean +++ b/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean @@ -189,7 +189,7 @@ theorem isAdjMatrix_adjMatrix [Zero α] [One α] : (G.adjMatrix α).IsAdjMatrix theorem toGraph_adjMatrix_eq [MulZeroOneClass α] [Nontrivial α] : (G.isAdjMatrix_adjMatrix α).toGraph = G := by ext - simp only [IsAdjMatrix.toGraph_Adj, adjMatrix_apply, ite_eq_left_iff, zero_ne_one] + simp only [IsAdjMatrix.toGraph_adj, adjMatrix_apply, ite_eq_left_iff, zero_ne_one] apply Classical.not_not #align simple_graph.to_graph_adj_matrix_eq SimpleGraph.toGraph_adjMatrix_eq diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index 5ad90ea3aeb79..69f3f95beef11 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -135,6 +135,8 @@ structure SimpleGraph (V : Type u) where #align simple_graph SimpleGraph -- porting note: changed `obviously` to `aesop` in the `structure` +initialize_simps_projections SimpleGraph (Adj → adj) + /-- Constructor for simple graphs using a symmetric irreflexive boolean function. -/ @[simps] def SimpleGraph.mk' {V : Type u} : @@ -2126,7 +2128,7 @@ end Maps @[simps!] def induceUnivIso (G : SimpleGraph V) : G.induce Set.univ ≃g G where toEquiv := Equiv.Set.univ V - map_rel_iff' := by simp only [Equiv.Set.univ, Equiv.coe_fn_mk, comap_Adj, Embedding.coe_subtype, + map_rel_iff' := by simp only [Equiv.Set.univ, Equiv.coe_fn_mk, comap_adj, Embedding.coe_subtype, Subtype.forall, Set.mem_univ, forall_true_left, implies_true] end SimpleGraph diff --git a/Mathlib/Combinatorics/SimpleGraph/Clique.lean b/Mathlib/Combinatorics/SimpleGraph/Clique.lean index 9f3142135f53b..cd054df13128e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Clique.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Clique.lean @@ -56,7 +56,7 @@ theorem isClique_iff_induce_eq : G.IsClique s ↔ G.induce s = ⊤ := by constructor · intro h ext ⟨v, hv⟩ ⟨w, hw⟩ - simp only [comap_Adj, Subtype.coe_mk, top_adj, Ne.def, Subtype.mk_eq_mk] + simp only [comap_adj, Subtype.coe_mk, top_adj, Ne.def, Subtype.mk_eq_mk] exact ⟨Adj.ne, h hv hw⟩ · intro h v hv w hw hne have h2 : (G.induce s).Adj ⟨v, hv⟩ ⟨w, hw⟩ = _ := rfl @@ -173,7 +173,7 @@ theorem not_cliqueFree_of_top_embedding {n : ℕ} (f : (⊤ : SimpleGraph (Fin n simp only [coe_map, Set.mem_image, coe_univ, Set.mem_univ, true_and_iff] at hv hw obtain ⟨v', rfl⟩ := hv obtain ⟨w', rfl⟩ := hw - simp only [coe_sort_coe, RelEmbedding.coe_toEmbedding, comap_Adj, Function.Embedding.coe_subtype, + simp only [coe_sort_coe, RelEmbedding.coe_toEmbedding, comap_adj, Function.Embedding.coe_subtype, f.map_adj_iff, top_adj, ne_eq, Subtype.mk.injEq, RelEmbedding.inj] -- This used to be the end of the proof before leanprover/lean4#2644 erw [Function.Embedding.coe_subtype, f.map_adj_iff] @@ -210,7 +210,7 @@ theorem not_cliqueFree_card_of_top_embedding [Fintype α] (f : (⊤ : SimpleGrap theorem cliqueFree_bot (h : 2 ≤ n) : (⊥ : SimpleGraph α).CliqueFree n := by intro t ht have := le_trans h (isNClique_bot_iff.1 ht).1 - simp only at this + contradiction #align simple_graph.clique_free_bot SimpleGraph.cliqueFree_bot theorem CliqueFree.mono (h : m ≤ n) : G.CliqueFree m → G.CliqueFree n := by @@ -226,7 +226,7 @@ theorem CliqueFree.anti (h : G ≤ H) : H.CliqueFree n → G.CliqueFree n := /-- See `SimpleGraph.cliqueFree_of_chromaticNumber_lt` for a tighter bound. -/ theorem cliqueFree_of_card_lt [Fintype α] (hc : card α < n) : G.CliqueFree n := by by_contra h - refine' Nat.lt_le_antisymm hc _ + refine' Nat.lt_le_asymm hc _ rw [cliqueFree_iff, not_isEmpty_iff] at h simpa only [Fintype.card_fin] using Fintype.card_le_of_embedding h.some.toEmbedding #align simple_graph.clique_free_of_card_lt SimpleGraph.cliqueFree_of_card_lt @@ -237,7 +237,7 @@ theorem cliqueFree_completeMultipartiteGraph {ι : Type*} [Fintype ι] (V : ι rw [cliqueFree_iff, isEmpty_iff] intro f obtain ⟨v, w, hn, he⟩ := exists_ne_map_eq_of_card_lt (Sigma.fst ∘ f) (by simp [hc]) - rw [← top_adj, ← f.map_adj_iff, comap_Adj, top_adj] at hn + rw [← top_adj, ← f.map_adj_iff, comap_adj, top_adj] at hn exact absurd he hn /-- Clique-freeness is preserved by `replaceVertex`. -/ diff --git a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean index d6f2658fbf5c0..5e3813ebb0c5a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean @@ -351,7 +351,7 @@ theorem chromaticNumber_eq_card_of_forall_surj [Fintype α] (C : G.Coloring α) specialize h (G.recolorOfEmbedding f C') have h1 : Function.Surjective f := Function.Surjective.of_comp h have h2 := Fintype.card_le_of_surjective _ h1 - exact Nat.lt_le_antisymm hc h2 + exact Nat.lt_le_asymm hc h2 #align simple_graph.chromatic_number_eq_card_of_forall_surj SimpleGraph.chromaticNumber_eq_card_of_forall_surj theorem chromaticNumber_bot [Nonempty V] : (⊥ : SimpleGraph V).chromaticNumber = 1 := by @@ -446,7 +446,7 @@ protected theorem Colorable.cliqueFree {n m : ℕ} (hc : G.Colorable n) (hm : n by_contra h simp only [CliqueFree, isNClique_iff, not_forall, Classical.not_not] at h obtain ⟨s, h, rfl⟩ := h - exact Nat.lt_le_antisymm hm (h.card_le_of_colorable hc) + exact Nat.lt_le_asymm hm (h.card_le_of_colorable hc) #align simple_graph.colorable.clique_free SimpleGraph.Colorable.cliqueFree -- TODO eliminate `Finite V` constraint once chromatic numbers are refactored. diff --git a/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean b/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean index c39d379d10a77..3136f0fea5b3c 100644 --- a/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean +++ b/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean @@ -67,8 +67,8 @@ theorem incMatrix_apply [Zero R] [One R] {a : α} {e : Sym2 α} : /-- Entries of the incidence matrix can be computed given additional decidable instances. -/ theorem incMatrix_apply' [Zero R] [One R] [DecidableEq α] [DecidableRel G.Adj] {a : α} {e : Sym2 α} : G.incMatrix R a e = if e ∈ G.incidenceSet a then 1 else 0 := by - unfold incMatrix Set.indicator -- Porting note: was `convert rfl` - simp only [Pi.one_apply] + unfold incMatrix Set.indicator + convert rfl #align simple_graph.inc_matrix_apply' SimpleGraph.incMatrix_apply' section MulZeroOneClass diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean index 7ce16573f6aba..29b165534b363 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean @@ -65,16 +65,14 @@ open SzemerediRegularity variable {α : Type*} [DecidableEq α] [Fintype α] {P : Finpartition (univ : Finset α)} {u : Finset α} {ε : ℝ} -local notation3 (prettyPrint := false) - "m" => (card α / stepBound P.parts.card : ℕ) +local notation3 "m" => (card α / stepBound P.parts.card : ℕ) -local notation3 (prettyPrint := false) - "a" => (card α / P.parts.card - m * 4 ^ P.parts.card : ℕ) +local notation3 "a" => (card α / P.parts.card - m * 4 ^ P.parts.card : ℕ) namespace SzemerediRegularity.Positivity private theorem eps_pos {ε : ℝ} {n : ℕ} (h : 100 ≤ (4 : ℝ) ^ n * ε ^ 5) : 0 < ε := - (Odd.pow_pos_iff (by norm_num)).mp + (Odd.pow_pos_iff (by decide)).mp (pos_of_mul_pos_right ((show 0 < (100 : ℝ) by norm_num).trans_le h) (by positivity)) private theorem m_pos [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) : 0 < m := @@ -126,7 +124,7 @@ theorem eps_pow_five_pos (hPε : 100 ≤ (4 : ℝ) ^ P.parts.card * ε ^ 5) : #align szemeredi_regularity.eps_pow_five_pos SzemerediRegularity.eps_pow_five_pos theorem eps_pos (hPε : 100 ≤ (4 : ℝ) ^ P.parts.card * ε ^ 5) : 0 < ε := - (Odd.pow_pos_iff (by norm_num)).mp (eps_pow_five_pos hPε) + (Odd.pow_pos_iff (by decide)).mp (eps_pow_five_pos hPε) #align szemeredi_regularity.eps_pos SzemerediRegularity.eps_pos theorem hundred_div_ε_pow_five_le_m [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) @@ -150,7 +148,7 @@ theorem a_add_one_le_four_pow_parts_card : a + 1 ≤ 4 ^ P.parts.card := by rw [stepBound, ← Nat.div_div_eq_div_mul] conv_rhs => rw [← Nat.sub_add_cancel h] rw [add_le_add_iff_right, tsub_le_iff_left, ← Nat.add_sub_assoc h] - exact Nat.le_pred_of_lt (Nat.lt_div_mul_add h) + exact Nat.le_sub_one_of_lt (Nat.lt_div_mul_add h) #align szemeredi_regularity.a_add_one_le_four_pow_parts_card SzemerediRegularity.a_add_one_le_four_pow_parts_card theorem card_aux₁ (hucard : u.card = m * 4 ^ P.parts.card + a) : diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean index a3ce7371e835d..0ec26e2331238 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean @@ -48,8 +48,7 @@ namespace SzemerediRegularity variable {α : Type*} [Fintype α] {P : Finpartition (univ : Finset α)} (hP : P.IsEquipartition) (G : SimpleGraph α) (ε : ℝ) {U : Finset α} (hU : U ∈ P.parts) (V : Finset α) -local notation3 (prettyPrint := false) - "m" => (card α / stepBound P.parts.card : ℕ) +local notation3 "m" => (card α / stepBound P.parts.card : ℕ) /-! ### Definitions @@ -451,7 +450,10 @@ private theorem edgeDensity_star_not_uniform [Nonempty α] set s : ℝ := ↑(G.edgeDensity (G.nonuniformWitness ε U V) (G.nonuniformWitness ε V U)) set t : ℝ := ↑(G.edgeDensity U V) have hrs : |r - s| ≤ ε / 5 := abs_density_star_sub_density_le_eps hPε hε₁ hUVne hUV - have hst : ε ≤ |s - t| := by exact_mod_cast G.nonuniformWitness_spec hUVne hUV + have hst : ε ≤ |s - t| := by + -- After leanprover/lean4#2734, we need to do the zeta reduction before `norm_cast`. + unfold_let s t + exact_mod_cast G.nonuniformWitness_spec hUVne hUV have hpr : |p - r| ≤ ε ^ 5 / 49 := average_density_near_total_density hPα hPε hε₁ star_subset_chunk star_subset_chunk have hqt : |q - t| ≤ ε ^ 5 / 49 := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean index c75d9f087c903..d16579fa938ef 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean @@ -45,8 +45,7 @@ open scoped BigOperators Classical SzemerediRegularity.Positivity variable {α : Type*} [Fintype α] {P : Finpartition (univ : Finset α)} (hP : P.IsEquipartition) (G : SimpleGraph α) (ε : ℝ) -local notation3 (prettyPrint := false) - "m" => (card α / stepBound P.parts.card : ℕ) +local notation3 "m" => (card α / stepBound P.parts.card : ℕ) namespace SzemerediRegularity diff --git a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean index 8c39f7aa9a2b0..656e03e16d171 100644 --- a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean +++ b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean @@ -225,7 +225,8 @@ theorem IsSRGWith.matrix_eq {α : Type*} [Semiring α] (h : G.IsSRGWith n k ℓ simp [commonNeighbors, ← neighborFinset_def, h.regular v] · simp only [Matrix.one_apply_ne' hn.symm, ne_eq, hn] by_cases ha : G.Adj v w <;> - simp only [ha, ite_true, ite_false, add_zero, zero_add, nsmul_eq_mul, smul_zero, mul_one] + simp only [ha, ite_true, ite_false, add_zero, zero_add, nsmul_eq_mul, smul_zero, mul_one, + not_true_eq_false, not_false_eq_true, and_false, and_self] · rw [h.of_adj v w ha] · rw [h.of_not_adj v w hn ha] diff --git a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean index 9518e8ea41992..11ab0642b56f4 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean @@ -65,6 +65,8 @@ structure Subgraph {V : Type u} (G : SimpleGraph V) where symm : Symmetric Adj := by aesop_graph -- Porting note: Originally `by obviously` #align simple_graph.subgraph SimpleGraph.Subgraph +initialize_simps_projections SimpleGraph.Subgraph (Adj → adj) + variable {ι : Sort*} {V : Type u} {W : Type v} /-- The one-vertex subgraph. -/ @@ -661,7 +663,7 @@ theorem map_sup {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') {H H' : ext1 · simp only [Set.image_union, map_verts, verts_sup] · ext - simp only [Relation.Map, map_Adj, sup_adj] + simp only [Relation.Map, map_adj, sup_adj] constructor · rintro ⟨a, b, h | h, rfl, rfl⟩ · exact Or.inl ⟨_, _, h, rfl, rfl⟩ @@ -688,7 +690,7 @@ theorem comap_monotone {G' : SimpleGraph W} (f : G →g G') : Monotone (Subgraph simp only [comap_verts, Set.mem_preimage] apply h.1 · intro v w - simp (config := { contextual := true }) only [comap_Adj, and_imp, true_and_iff] + simp (config := { contextual := true }) only [comap_adj, and_imp, true_and_iff] intro apply h.2 #align simple_graph.subgraph.comap_monotone SimpleGraph.Subgraph.comap_monotone @@ -698,12 +700,12 @@ theorem map_le_iff_le_comap {G' : SimpleGraph W} (f : G →g G') (H : G.Subgraph refine' ⟨fun h ↦ ⟨fun v hv ↦ _, fun v w hvw ↦ _⟩, fun h ↦ ⟨fun v ↦ _, fun v w ↦ _⟩⟩ · simp only [comap_verts, Set.mem_preimage] exact h.1 ⟨v, hv, rfl⟩ - · simp only [H.adj_sub hvw, comap_Adj, true_and_iff] + · simp only [H.adj_sub hvw, comap_adj, true_and_iff] exact h.2 ⟨v, w, hvw, rfl, rfl⟩ · simp only [map_verts, Set.mem_image, forall_exists_index, and_imp] rintro w hw rfl exact h.1 hw - · simp only [Relation.Map, map_Adj, forall_exists_index, and_imp] + · simp only [Relation.Map, map_adj, forall_exists_index, and_imp] rintro u u' hu rfl rfl exact (h.2 hu).2 #align simple_graph.subgraph.map_le_iff_le_comap SimpleGraph.Subgraph.map_le_iff_le_comap @@ -854,7 +856,7 @@ theorem singletonSubgraph_le_iff (v : V) (H : G.Subgraph) : @[simp] theorem map_singletonSubgraph (f : G →g G') {v : V} : Subgraph.map f (G.singletonSubgraph v) = G'.singletonSubgraph (f v) := by - ext <;> simp only [Relation.Map, Subgraph.map_Adj, singletonSubgraph_Adj, Pi.bot_apply, + ext <;> simp only [Relation.Map, Subgraph.map_adj, singletonSubgraph_adj, Pi.bot_apply, exists_and_left, and_iff_left_iff_imp, IsEmpty.forall_iff, Subgraph.map_verts, singletonSubgraph_verts, Set.image_singleton] exact False.elim @@ -875,7 +877,7 @@ theorem eq_singletonSubgraph_iff_verts_eq (H : G.Subgraph) {v : V} : refine' ⟨fun h ↦ by rw [h, singletonSubgraph_verts], fun h ↦ _⟩ ext · rw [h, singletonSubgraph_verts] - · simp only [Prop.bot_eq_false, singletonSubgraph_Adj, Pi.bot_apply, iff_false_iff] + · simp only [Prop.bot_eq_false, singletonSubgraph_adj, Pi.bot_apply, iff_false_iff] intro ha have ha1 := ha.fst_mem have ha2 := ha.snd_mem @@ -894,7 +896,7 @@ theorem edgeSet_subgraphOfAdj {v w : V} (hvw : G.Adj v w) : (G.subgraphOfAdj hvw).edgeSet = {⟦(v, w)⟧} := by ext e refine' e.ind _ - simp only [eq_comm, Set.mem_singleton_iff, Subgraph.mem_edgeSet, subgraphOfAdj_Adj, iff_self_iff, + simp only [eq_comm, Set.mem_singleton_iff, Subgraph.mem_edgeSet, subgraphOfAdj_adj, iff_self_iff, forall₂_true_iff] #align simple_graph.edge_set_subgraph_of_adj SimpleGraph.edgeSet_subgraphOfAdj @@ -904,7 +906,7 @@ lemma subgraphOfAdj_le_of_adj (H : G.Subgraph) (h : H.Adj v w) : constructor · intro x rintro (rfl | rfl) <;> simp [H.edge_vert h, H.edge_vert h.symm] - · simp only [subgraphOfAdj_Adj, Quotient.eq, Sym2.rel_iff] + · simp only [subgraphOfAdj_adj, Quotient.eq, Sym2.rel_iff] rintro _ _ (⟨rfl, rfl⟩ | ⟨rfl, rfl⟩) <;> simp [h, h.symm] theorem subgraphOfAdj_symm {v w : V} (hvw : G.Adj v w) : @@ -925,7 +927,7 @@ theorem map_subgraphOfAdj (f : G →g G') {v w : V} (hvw : G.Adj v w) : simp · use w simp - · simp only [Relation.Map, Subgraph.map_Adj, subgraphOfAdj_Adj, Quotient.eq, Sym2.rel_iff] + · simp only [Relation.Map, Subgraph.map_adj, subgraphOfAdj_adj, Quotient.eq, Sym2.rel_iff] constructor · rintro ⟨a, b, ⟨rfl, rfl⟩ | ⟨rfl, rfl⟩, rfl, rfl⟩ <;> simp · rintro (⟨rfl, rfl⟩ | ⟨rfl, rfl⟩) @@ -1001,19 +1003,19 @@ protected def restrict {G' : G.Subgraph} : G.Subgraph → G'.coe.Subgraph := Subgraph.comap G'.hom #align simple_graph.subgraph.restrict SimpleGraph.Subgraph.restrict -lemma coeSubgraph_Adj {G' : G.Subgraph} (G'' : G'.coe.Subgraph) (v w : V) : +lemma coeSubgraph_adj {G' : G.Subgraph} (G'' : G'.coe.Subgraph) (v w : V) : (G'.coeSubgraph G'').Adj v w ↔ ∃ (hv : v ∈ G'.verts) (hw : w ∈ G'.verts), G''.Adj ⟨v, hv⟩ ⟨w, hw⟩ := by simp [Relation.Map] -lemma restrict_Adj {G' G'' : G.Subgraph} (v w : G'.verts) : +lemma restrict_adj {G' G'' : G.Subgraph} (v w : G'.verts) : (G'.restrict G'').Adj v w ↔ G'.Adj v w ∧ G''.Adj v w := Iff.rfl theorem restrict_coeSubgraph {G' : G.Subgraph} (G'' : G'.coe.Subgraph) : Subgraph.restrict (Subgraph.coeSubgraph G'') = G'' := by ext · simp - · rw [restrict_Adj, coeSubgraph_Adj] + · rw [restrict_adj, coeSubgraph_adj] simpa using G''.adj_sub #align simple_graph.subgraph.restrict_coe_subgraph SimpleGraph.Subgraph.restrict_coeSubgraph @@ -1033,7 +1035,7 @@ lemma coeSubgraph_restrict_eq {H : G.Subgraph} (H' : G.Subgraph) : Subgraph.coeSubgraph (H.restrict H') = H ⊓ H' := by ext · simp [and_comm] - · simp_rw [coeSubgraph_Adj, restrict_Adj] + · simp_rw [coeSubgraph_adj, restrict_adj] simp only [exists_and_left, exists_prop, ge_iff_le, inf_adj, and_congr_right_iff] intro h simp [H.edge_vert h, H.edge_vert h.symm] @@ -1088,7 +1090,7 @@ theorem deleteEdges_spanningCoe_eq : theorem deleteEdges_coe_eq (s : Set (Sym2 G'.verts)) : G'.coe.deleteEdges s = (G'.deleteEdges (Sym2.map (↑) '' s)).coe := by ext ⟨v, hv⟩ ⟨w, hw⟩ - simp only [SimpleGraph.deleteEdges_adj, coe_Adj, deleteEdges_adj, Set.mem_image, not_exists, + simp only [SimpleGraph.deleteEdges_adj, coe_adj, deleteEdges_adj, Set.mem_image, not_exists, not_and, and_congr_right_iff] intro constructor @@ -1174,7 +1176,7 @@ variable {G' G'' : G.Subgraph} {s s' : Set V} theorem induce_mono (hg : G' ≤ G'') (hs : s ⊆ s') : G'.induce s ≤ G''.induce s' := by constructor · simp [hs] - · simp (config := { contextual := true }) only [induce_Adj, true_and_iff, and_imp] + · simp (config := { contextual := true }) only [induce_adj, true_and_iff, and_imp] intro v w hv hw ha exact ⟨hs hv, hs hw, hg.2 ha⟩ #align simple_graph.subgraph.induce_mono SimpleGraph.Subgraph.induce_mono @@ -1199,7 +1201,7 @@ theorem induce_self_verts : G'.induce G'.verts = G' := by ext · simp · constructor <;> - simp (config := { contextual := true }) only [induce_Adj, imp_true_iff, and_true_iff] + simp (config := { contextual := true }) only [induce_adj, imp_true_iff, and_true_iff] exact fun ha ↦ ⟨G'.edge_vert ha, G'.edge_vert ha.symm⟩ #align simple_graph.subgraph.induce_self_verts SimpleGraph.Subgraph.induce_self_verts @@ -1210,7 +1212,7 @@ lemma le_induce_top_verts : G' ≤ (⊤ : G.Subgraph).induce G'.verts := lemma le_induce_union : G'.induce s ⊔ G'.induce s' ≤ G'.induce (s ∪ s') := by constructor · simp only [verts_sup, induce_verts, Set.Subset.rfl] - · simp only [sup_adj, induce_Adj, Set.mem_union] + · simp only [sup_adj, induce_adj, Set.mem_union] rintro v w (h | h) <;> simp [h] lemma le_induce_union_left : G'.induce s ≤ G'.induce (s ∪ s') := by @@ -1229,10 +1231,10 @@ theorem subgraphOfAdj_eq_induce {v w : V} (hvw : G.Adj v w) : · simp · constructor · intro h - simp only [subgraphOfAdj_Adj, Quotient.eq, Sym2.rel_iff] at h + simp only [subgraphOfAdj_adj, Quotient.eq, Sym2.rel_iff] at h obtain ⟨rfl, rfl⟩ | ⟨rfl, rfl⟩ := h <;> simp [hvw, hvw.symm] · intro h - simp only [induce_Adj, Set.mem_insert_iff, Set.mem_singleton_iff, top_adj] at h + simp only [induce_adj, Set.mem_insert_iff, Set.mem_singleton_iff, top_adj] at h obtain ⟨rfl | rfl, rfl | rfl, ha⟩ := h <;> first |exact (ha.ne rfl).elim|simp #align simple_graph.subgraph.subgraph_of_adj_eq_induce SimpleGraph.Subgraph.subgraphOfAdj_eq_induce diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index 608d67ba074e7..ec26bfaaa6eda 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -499,7 +499,7 @@ theorem rowLens_length_ofRowLens {w : List ℕ} {hw : w.Sorted (· ≥ ·)} (hpo (ofRowLens w hw).rowLens.length = w.length := by simp only [length_rowLens, colLen, Nat.find_eq_iff, mem_cells, mem_ofRowLens, lt_self_iff_false, IsEmpty.exists_iff, Classical.not_not] - refine' ⟨True.intro, fun n hn => ⟨hn, hpos _ (List.get_mem _ _ hn)⟩⟩ + refine' ⟨not_false, fun n hn => ⟨hn, hpos _ (List.get_mem _ _ hn)⟩⟩ #align young_diagram.row_lens_length_of_row_lens YoungDiagram.rowLens_length_ofRowLens -- Porting note: use `List.get` instead of `List.nthLe` because it has been deprecated diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index 6c2fdde4e3475..2a4625465eeec 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -101,7 +101,7 @@ theorem ack_three (n : ℕ) : ack 3 n = 2 ^ (n + 3) - 3 := by Nat.mul_sub_left_distrib, ← Nat.sub_add_comm, two_mul 3, Nat.add_sub_add_right] have H : 2 * 3 ≤ 2 * 2 ^ 3 := by norm_num apply H.trans - simp [pow_le_pow] + simp [pow_le_pow (show 1 ≤ 2 by norm_num)] #align ack_three ack_three theorem ack_pos : ∀ m n, 0 < ack m n diff --git a/Mathlib/Computability/Halting.lean b/Mathlib/Computability/Halting.lean index 2a513d73756e4..9f60ef77c8412 100644 --- a/Mathlib/Computability/Halting.lean +++ b/Mathlib/Computability/Halting.lean @@ -378,7 +378,7 @@ theorem rfindOpt {n} {f : Vector ℕ (n + 1) → ℕ} (hf : @Partrec' (n + 1) f) exists_congr fun a => (and_congr (iff_of_eq _) Iff.rfl).trans (and_congr_right fun h => _) · congr funext n - cases f (n ::ᵥ v) <;> simp [Nat.succ_le_succ]; rfl + cases f (n ::ᵥ v) <;> simp [Nat.succ_le_succ] <;> rfl · have := Nat.rfind_spec h simp only [Part.coe_some, Part.mem_some_iff] at this revert this; cases' f (a ::ᵥ v) with c <;> intro this diff --git a/Mathlib/Computability/Language.lean b/Mathlib/Computability/Language.lean index 94b091c58b300..dacbc12e7e79a 100644 --- a/Mathlib/Computability/Language.lean +++ b/Mathlib/Computability/Language.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Fox Thomson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Fox Thomson -/ -import Mathlib.Algebra.Hom.Ring.Defs import Mathlib.Algebra.Order.Kleene +import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Data.List.Join import Mathlib.Data.Set.Lattice import Mathlib.Tactic.DeriveFintype diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index fd573c762d535..5edab43f57ebe 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -837,7 +837,7 @@ theorem nat_mod : Primrec₂ ((· % ·) : ℕ → ℕ → ℕ) := theorem nat_bodd : Primrec Nat.bodd := (Primrec.beq.comp (nat_mod.comp .id (const 2)) (const 1)).of_eq fun n => by - cases H : n.bodd <;> simp [Nat.mod_two_of_bodd, H] + cases H : n.bodd <;> simp [Nat.mod_two_of_bodd, H]; rfl #align primrec.nat_bodd Primrec.nat_bodd theorem nat_div2 : Primrec Nat.div2 := diff --git a/Mathlib/Computability/TMComputable.lean b/Mathlib/Computability/TMComputable.lean index d414d6d2a6bf7..a79be2891b674 100644 --- a/Mathlib/Computability/TMComputable.lean +++ b/Mathlib/Computability/TMComputable.lean @@ -260,7 +260,7 @@ def idComputableInPolyTime {α : Type} (ea : FinEncoding α) : outputsFun _ := { steps := 1 evals_in_steps := rfl - steps_le_m := by simp only [Polynomial.eval_one] } + steps_le_m := by simp only [Polynomial.eval_one, le_refl] } #align turing.id_computable_in_poly_time Turing.idComputableInPolyTime instance inhabitedTM2ComputableInPolyTime : diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 2c3872c83fc5f..d16613bcbfc9c 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -1536,7 +1536,7 @@ theorem succ_ok {q s n} {c d : List Γ'} : simp only [Option.mem_def, TM2.stepAux, elim_main, decide_False, elim_update_main, ne_eq, Function.update_noteq, elim_rev, elim_update_rev, decide_True, Function.update_same, cond_true, cond_false] - convert unrev_ok using 2 + convert unrev_ok using 1 simp only [elim_update_rev, elim_rev, elim_main, List.reverseAux_nil, elim_update_main] rfl simp only [trNum, Num.succ, Num.succ'] diff --git a/Mathlib/Control/Functor/Multivariate.lean b/Mathlib/Control/Functor/Multivariate.lean index ca5f0aca283d4..6337351fee9b7 100644 --- a/Mathlib/Control/Functor/Multivariate.lean +++ b/Mathlib/Control/Functor/Multivariate.lean @@ -10,12 +10,12 @@ import Mathlib.Data.TypeVec /-! -Functors between the category of tuples of types, and the category Type +# Functors between the category of tuples of types, and the category Type Features: -`MvFunctor n` : the type class of multivariate functors -`f <$$> x` : notation for map +* `MvFunctor n` : the type class of multivariate functors +* `f <$$> x` : notation for map -/ diff --git a/Mathlib/Data/BinaryHeap.lean b/Mathlib/Data/BinaryHeap.lean index 96c9e990747f8..4b86e27cc2484 100644 --- a/Mathlib/Data/BinaryHeap.lean +++ b/Mathlib/Data/BinaryHeap.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Data.Fin.Basic +import Std.Data.Fin.Basic set_option autoImplicit true @@ -49,7 +49,7 @@ def mkHeap (lt : α → α → Bool) (a : Array α) : {a' : Array α // a'.size | i+1, a, h => let h := Nat.lt_of_succ_le h let a' := heapifyDown lt a ⟨i, h⟩ - let ⟨a₂, h₂⟩ := loop i a' ((heapifyDown ..).2.symm ▸ le_of_lt h) + let ⟨a₂, h₂⟩ := loop i a' ((heapifyDown ..).2.symm ▸ Nat.le_of_lt h) ⟨a₂, h₂.trans a'.2⟩ loop (a.size / 2) a (Nat.div_le_self ..) @@ -61,9 +61,9 @@ Given an array which is a max-heap, push item `i` up to restore the max-heap pro def heapifyUp (lt : α → α → Bool) (a : Array α) (i : Fin a.size) : {a' : Array α // a'.size = a.size} := if i0 : i.1 = 0 then ⟨a, rfl⟩ else - have : (i.1 - 1) / 2 < i := lt_of_le_of_lt (Nat.div_le_self ..) $ - Nat.sub_lt (Nat.pos_iff_ne_zero.2 i0) Nat.one_pos - let j := ⟨(i.1 - 1) / 2, lt_trans this i.2⟩ + have : (i.1 - 1) / 2 < i := Nat.lt_of_le_of_lt (Nat.div_le_self ..) $ + Nat.sub_lt (Nat.pos_of_ne_zero i0) Nat.zero_lt_one + let j := ⟨(i.1 - 1) / 2, Nat.lt_trans this i.2⟩ if lt (a.get j) (a.get i) then let a' := a.swap i j let ⟨a₂, h₂⟩ := heapifyUp lt a' ⟨j.1, by rw [a.size_swap i j]; exact j.2⟩ diff --git a/Mathlib/Data/Bitvec/Lemmas.lean b/Mathlib/Data/Bitvec/Lemmas.lean index bc995d403f452..9e961fabc48f9 100644 --- a/Mathlib/Data/Bitvec/Lemmas.lean +++ b/Mathlib/Data/Bitvec/Lemmas.lean @@ -102,7 +102,7 @@ theorem toNat_lt {n : ℕ} (v : Bitvec n) : v.toNat < 2 ^ n := by -- Porting note: removed `ac_mono`, `mono` calls · rw [add_assoc] apply Nat.add_le_add_left - cases head <;> simp only + cases head <;> decide · rw [← left_distrib] rw [mul_comm _ 2] apply Nat.mul_le_mul_left @@ -112,14 +112,14 @@ theorem toNat_lt {n : ℕ} (v : Bitvec n) : v.toNat < 2 ^ n := by theorem addLsb_div_two {x b} : addLsb x b / 2 = x := by cases b <;> simp only [Nat.add_mul_div_left, addLsb, ← two_mul, add_comm, Nat.succ_pos', - Nat.mul_div_right, gt_iff_lt, zero_add, cond] + Nat.mul_div_right, gt_iff_lt, zero_add, zero_lt_two, cond] norm_num #align bitvec.add_lsb_div_two Bitvec.addLsb_div_two theorem decide_addLsb_mod_two {x b} : decide (addLsb x b % 2 = 1) = b := by cases b <;> simp only [Bool.decide_iff, Nat.add_mul_mod_self_left, addLsb, ← two_mul, add_comm, - Bool.decide_False, Nat.mul_mod_right, zero_add, cond, zero_ne_one] + Bool.decide_False, Nat.mul_mod_right, zero_add, cond, zero_ne_one]; rfl #align bitvec.to_bool_add_lsb_mod_two Bitvec.decide_addLsb_mod_two theorem ofNat_toNat {n : ℕ} (v : Bitvec n) : Bitvec.ofNat n v.toNat = v := by diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 1e42361e9c90c..28b4753db3afc 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -20,8 +20,6 @@ bool, boolean, Bool, De Morgan -/ -set_option autoImplicit true - namespace Bool theorem decide_True {h} : @decide True h = true := @@ -83,11 +81,11 @@ theorem eq_iff_eq_true_iff {a b : Bool} : a = b ↔ ((a = true) ↔ (b = true)) cases a <;> cases b <;> simp -- Porting note: new theorem -theorem beq_eq_decide_eq [DecidableEq α] +theorem beq_eq_decide_eq {α} [DecidableEq α] (a b : α) : (a == b) = decide (a = b) := rfl -- Porting note: new theorem -theorem beq_comm [BEq α] [LawfulBEq α] {a b : α} : (a == b) = (b == a) := +theorem beq_comm {α} [BEq α] [LawfulBEq α] {a b : α} : (a == b) = (b == a) := eq_iff_eq_true_iff.2 (by simp [@eq_comm α]) @[simp] @@ -147,12 +145,8 @@ theorem eq_true_of_ne_false : ∀ {a : Bool}, a ≠ false → a = true := by dec theorem eq_false_of_ne_true : ∀ {a : Bool}, a ≠ true → a = false := by decide #align bool.eq_ff_of_ne_tt Bool.eq_false_of_ne_true -theorem or_comm : ∀ a b, (a || b) = (b || a) := by decide #align bool.bor_comm Bool.or_comm - #align bool.bor_assoc Bool.or_assoc - -theorem or_left_comm : ∀ a b c, (a || (b || c)) = (b || (a || c)) := by decide #align bool.bor_left_comm Bool.or_left_comm theorem or_inl {a b : Bool} (H : a) : a || b := by simp [H] @@ -161,12 +155,8 @@ theorem or_inl {a b : Bool} (H : a) : a || b := by simp [H] theorem or_inr {a b : Bool} (H : b) : a || b := by cases a <;> simp [H] #align bool.bor_inr Bool.or_inr -theorem and_comm : ∀ a b, (a && b) = (b && a) := by decide #align bool.band_comm Bool.and_comm - #align bool.band_assoc Bool.and_assoc - -theorem and_left_comm : ∀ a b c, (a && (b && c)) = (b && (a && c)) := by decide #align bool.band_left_comm Bool.and_left_comm theorem and_elim_left : ∀ {a b : Bool}, a && b → a := by decide @@ -178,22 +168,10 @@ theorem and_intro : ∀ {a b : Bool}, a → b → a && b := by decide theorem and_elim_right : ∀ {a b : Bool}, a && b → b := by decide #align bool.band_elim_right Bool.and_elim_right -theorem and_or_distrib_left (a b c : Bool) : (a && (b || c)) = (a && b || a && c) := by - cases a <;> simp #align bool.band_bor_distrib_left Bool.and_or_distrib_left - -theorem and_or_distrib_right (a b c : Bool) : ((a || b) && c) = (a && c || b && c) := by - cases a <;> cases b <;> cases c <;> simp #align bool.band_bor_distrib_right Bool.and_or_distrib_right - -theorem or_and_distrib_left (a b c : Bool) : (a || b && c) = ((a || b) && (a || c)) := by - cases a <;> simp #align bool.bor_band_distrib_left Bool.or_and_distrib_left - -theorem or_and_distrib_right (a b c : Bool) : (a && b || c) = ((a || c) && (b || c)) := by - cases a <;> cases b <;> cases c <;> simp #align bool.bor_band_distrib_right Bool.or_and_distrib_right - #align bool.bnot_ff Bool.not_false #align bool.bnot_tt Bool.not_true @@ -243,60 +221,27 @@ theorem eq_false_of_not_eq_true' {a : Bool} : !a = true → a = false := by cases a <;> decide #align bool.eq_ff_of_bnot_eq_tt Bool.eq_false_of_not_eq_true' -@[simp] -theorem and_not_self : ∀ x, (x && !x) = false := by decide #align bool.band_bnot_self Bool.and_not_self - -@[simp] -theorem not_and_self : ∀ x, (!x && x) = false := by decide #align bool.bnot_band_self Bool.not_and_self - -@[simp] -theorem or_not_self : ∀ x, (x || !x) = true := by decide #align bool.bor_bnot_self Bool.or_not_self - -@[simp] -theorem not_or_self : ∀ x, (!x || x) = true := by decide #align bool.bnot_bor_self Bool.not_or_self theorem bne_eq_xor : bne = xor := by funext a b; revert a b; decide -theorem xor_comm : ∀ a b, xor a b = xor b a := by decide #align bool.bxor_comm Bool.xor_comm -@[simp] -theorem xor_assoc : ∀ a b c, xor (xor a b) c = xor a (xor b c) := by decide +attribute [simp] xor_assoc #align bool.bxor_assoc Bool.xor_assoc -theorem xor_left_comm : ∀ a b c, xor a (xor b c) = xor b (xor a c) := by decide #align bool.bxor_left_comm Bool.xor_left_comm +#align bool.bxor_bnot_left Bool.not_xor +#align bool.bxor_bnot_right Bool.xor_not -@[simp] -theorem xor_not_left : ∀ a, xor (!a) a = true := by decide -#align bool.bxor_bnot_left Bool.xor_not_left - -@[simp] -theorem xor_not_right : ∀ a, xor a (!a) = true := by decide -#align bool.bxor_bnot_right Bool.xor_not_right - -@[simp] -theorem xor_not_not : ∀ a b, xor (!a) (!b) = xor a b := by decide -#align bool.bxor_bnot_bnot Bool.xor_not_not - -@[simp] -theorem xor_false_left : ∀ a, xor false a = a := by decide -#align bool.bxor_ff_left Bool.xor_false_left - -@[simp] -theorem xor_false_right : ∀ a, xor a false = a := by decide -#align bool.bxor_ff_right Bool.xor_false_right +#align bool.bxor_bnot_bnot Bool.not_xor_not -theorem and_xor_distrib_left (a b c : Bool) : (a && xor b c) = xor (a && b) (a && c) := by - cases a <;> simp +#align bool.bxor_ff_left Bool.false_xor +#align bool.bxor_ff_right Bool.xor_false #align bool.band_bxor_distrib_left Bool.and_xor_distrib_left - -theorem and_xor_distrib_right (a b c : Bool) : (xor a b && c) = xor (a && c) (b && c) := by - cases a <;> cases b <;> cases c <;> simp #align bool.band_bxor_distrib_right Bool.and_xor_distrib_right theorem xor_iff_ne : ∀ {x y : Bool}, xor x y = true ↔ x ≠ y := by decide @@ -304,42 +249,28 @@ theorem xor_iff_ne : ∀ {x y : Bool}, xor x y = true ↔ x ≠ y := by decide /-! ### De Morgan's laws for booleans-/ -@[simp] -theorem not_and : ∀ a b : Bool, (!(a && b)) = (!a || !b) := by decide +attribute [simp] not_and #align bool.bnot_band Bool.not_and -@[simp] -theorem not_or : ∀ a b : Bool, (!(a || b)) = (!a && !b) := by decide +attribute [simp] not_or #align bool.bnot_bor Bool.not_or -theorem not_inj : ∀ {a b : Bool}, (!a) = !b → a = b := by decide #align bool.bnot_inj Bool.not_inj --- Porting note: having to unfold here is not pretty. --- There is a discussion on zulip about this at --- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/LinearOrder.20in.20mathlib3.2F4/near/308228493 instance linearOrder : LinearOrder Bool where - le := fun a b ↦ a = false ∨ b = true - le_refl := by unfold LE.le; decide - le_trans := by unfold LE.le; decide - le_antisymm := by unfold LE.le Preorder.toLE; decide - le_total := by unfold LE.le Preorder.toLE PartialOrder.toPreorder; decide - decidableLE := by unfold LE.le Preorder.toLE PartialOrder.toPreorder; exact inferInstance - decidableEq := inferInstance - max := or - max_def := λ a b => by cases a <;> cases b <;> decide - min := and - min_def := λ a b => by cases a <;> cases b <;> decide + le_refl := by decide + le_trans := by decide + le_antisymm := by decide + le_total := by decide + decidableLE := inferInstance + lt_iff_le_not_le := by decide + max_def := by decide + min_def := by decide #align bool.linear_order Bool.linearOrder -@[simp] -theorem false_le {x : Bool} : false ≤ x := - Or.intro_left _ rfl -#align bool.ff_le Bool.false_le +attribute [simp] Bool.max_eq_or Bool.min_eq_and -@[simp] -theorem le_true {x : Bool} : x ≤ true := - Or.intro_right _ rfl +#align bool.ff_le Bool.false_le #align bool.le_tt Bool.le_true theorem lt_iff : ∀ {x y : Bool}, x < y ↔ x = false ∧ y = true := by decide @@ -384,17 +315,15 @@ def ofNat (n : Nat) : Bool := theorem ofNat_le_ofNat {n m : Nat} (h : n ≤ m) : ofNat n ≤ ofNat m := by simp only [ofNat, ne_eq, _root_.decide_not] cases Nat.decEq n 0 with - | isTrue hn => rw [decide_eq_true hn]; exact false_le + | isTrue hn => rw [decide_eq_true hn]; exact Bool.false_le _ | isFalse hn => cases Nat.decEq m 0 with - | isFalse hm => rw [decide_eq_false hm]; exact le_true + | isFalse hm => rw [decide_eq_false hm]; exact Bool.le_true _ | isTrue hm => subst hm; have h := le_antisymm h (Nat.zero_le n); contradiction #align bool.of_nat_le_of_nat Bool.ofNat_le_ofNat theorem toNat_le_toNat {b₀ b₁ : Bool} (h : b₀ ≤ b₁) : toNat b₀ ≤ toNat b₁ := by - cases h with - | inl h => subst h; exact Nat.zero_le _ - | inr h => subst h; cases b₀ <;> simp + cases b₀ <;> cases b₁ <;> simp_all (config := { decide := true }) #align bool.to_nat_le_to_nat Bool.toNat_le_toNat theorem ofNat_toNat (b : Bool) : ofNat (toNat b) = b := by diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 370e0eacde586..a2f3bd1fb0849 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -226,7 +226,7 @@ theorem ofReal_add (r s : ℝ) : ((r + s : ℝ) : ℂ) = r + s := #align complex.of_real_add Complex.ofReal_add @[simp, norm_cast] -theorem ofReal_bit0 (r : ℝ) : ((bit0 r : ℝ) : ℂ) = bit0 (r : ℂ) := +theorem ofReal_bit0 (r : ℝ) : ((bit0 r : ℝ) : ℂ) = bit0 (r : ℂ) := ext_iff.2 <| by simp [bit0] #align complex.of_real_bit0 Complex.ofReal_bit0 @@ -638,6 +638,11 @@ theorem normSq_one : normSq 1 = 1 := normSq.map_one #align complex.norm_sq_one Complex.normSq_one +@[simp] +theorem normSq_ofNat (n : ℕ) [n.AtLeastTwo] : + normSq (no_index (OfNat.ofNat n : ℂ)) = OfNat.ofNat n * OfNat.ofNat n := by + simp [normSq] + @[simp] theorem normSq_I : normSq I = 1 := by simp [normSq] set_option linter.uppercaseLean3 false in @@ -900,6 +905,38 @@ theorem ofReal_rat_cast (n : ℚ) : ((n : ℝ) : ℂ) = (n : ℂ) := map_ratCast ofReal n #align complex.of_real_rat_cast Complex.ofReal_rat_cast +lemma div_ofReal (z : ℂ) (x : ℝ) : z / x = ⟨z.re / x, z.im / x⟩ := by + simp_rw [div_eq_inv_mul, ← ofReal_inv, ofReal_mul'] + +lemma div_nat_cast (z : ℂ) (n : ℕ) : z / n = ⟨z.re / n, z.im / n⟩ := by + exact_mod_cast div_ofReal z n + +lemma div_int_cast (z : ℂ) (n : ℤ) : z / n = ⟨z.re / n, z.im / n⟩ := by + exact_mod_cast div_ofReal z n + +lemma div_rat_cast (z : ℂ) (x : ℚ) : z / x = ⟨z.re / x, z.im / x⟩ := by + exact_mod_cast div_ofReal z x + +lemma div_ofNat (z : ℂ) (n : ℕ) [n.AtLeastTwo] : + z / OfNat.ofNat n = ⟨z.re / OfNat.ofNat n, z.im / OfNat.ofNat n⟩ := + div_nat_cast z n + +@[simp] lemma div_ofReal_re (z : ℂ) (x : ℝ) : (z / x).re = z.re / x := by rw [div_ofReal] +@[simp] lemma div_ofReal_im (z : ℂ) (x : ℝ) : (z / x).im = z.im / x := by rw [div_ofReal] +@[simp] lemma div_nat_cast_re (z : ℂ) (n : ℕ) : (z / n).re = z.re / n := by rw [div_nat_cast] +@[simp] lemma div_nat_cast_im (z : ℂ) (n : ℕ) : (z / n).im = z.im / n := by rw [div_nat_cast] +@[simp] lemma div_int_cast_re (z : ℂ) (n : ℤ) : (z / n).re = z.re / n := by rw [div_int_cast] +@[simp] lemma div_int_cast_im (z : ℂ) (n : ℤ) : (z / n).im = z.im / n := by rw [div_int_cast] +@[simp] lemma div_rat_cast_re (z : ℂ) (x : ℚ) : (z / x).re = z.re / x := by rw [div_rat_cast] +@[simp] lemma div_rat_cast_im (z : ℂ) (x : ℚ) : (z / x).im = z.im / x := by rw [div_rat_cast] + +@[simp] +lemma div_ofNat_re (z : ℂ) (n : ℕ) [n.AtLeastTwo] : + (z / no_index (OfNat.ofNat n)).re = z.re / OfNat.ofNat n := div_nat_cast_re z n + +@[simp] +lemma div_ofNat_im (z : ℂ) (n : ℕ) [n.AtLeastTwo] : + (z / no_index (OfNat.ofNat n)).im = z.im / OfNat.ofNat n := div_nat_cast_im z n /-! ### Characteristic zero -/ @@ -990,11 +1027,13 @@ nonrec theorem abs_of_nonneg {r : ℝ} (h : 0 ≤ r) : Complex.abs r = r := (Complex.abs_ofReal _).trans (abs_of_nonneg h) #align complex.abs_of_nonneg Complex.abs_of_nonneg -theorem abs_of_nat (n : ℕ) : Complex.abs n = n := - calc - Complex.abs n = Complex.abs (n : ℝ) := by rw [ofReal_nat_cast] - _ = _ := Complex.abs_of_nonneg (Nat.cast_nonneg n) -#align complex.abs_of_nat Complex.abs_of_nat +theorem abs_natCast (n : ℕ) : Complex.abs n = n := Complex.abs_of_nonneg (Nat.cast_nonneg n) +#align complex.abs_of_nat Complex.abs_natCast + +@[simp] +theorem abs_ofNat (n : ℕ) [n.AtLeastTwo] : + Complex.abs (no_index (OfNat.ofNat n : ℂ)) = OfNat.ofNat n := + abs_natCast n theorem mul_self_abs (z : ℂ) : Complex.abs z * Complex.abs z = normSq z := Real.mul_self_sqrt (normSq_nonneg _) @@ -1019,11 +1058,7 @@ theorem abs_I : Complex.abs I = 1 := by simp [Complex.abs] set_option linter.uppercaseLean3 false in #align complex.abs_I Complex.abs_I -@[simp] -theorem abs_two : Complex.abs 2 = 2 := - calc - Complex.abs 2 = Complex.abs (2 : ℝ) := rfl - _ = (2 : ℝ) := Complex.abs_of_nonneg (by norm_num) +theorem abs_two : Complex.abs 2 = 2 := abs_ofNat 2 #align complex.abs_two Complex.abs_two @[simp] @@ -1089,6 +1124,14 @@ theorem abs_im_lt_abs {z : ℂ} : |z.im| < Complex.abs z ↔ z.re ≠ 0 := by simpa using @abs_re_lt_abs (z * I) #align complex.abs_im_lt_abs Complex.abs_im_lt_abs +@[simp] +lemma abs_re_eq_abs {z : ℂ} : |z.re| = abs z ↔ z.im = 0 := + not_iff_not.1 <| (abs_re_le_abs z).lt_iff_ne.symm.trans abs_re_lt_abs + +@[simp] +lemma abs_im_eq_abs {z : ℂ} : |z.im| = abs z ↔ z.re = 0 := + not_iff_not.1 <| (abs_im_le_abs z).lt_iff_ne.symm.trans abs_im_lt_abs + @[simp] theorem abs_abs (z : ℂ) : |Complex.abs z| = Complex.abs z := _root_.abs_of_nonneg (AbsoluteValue.nonneg _ z) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 27ef2936bcfb7..08557413e1145 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1303,6 +1303,9 @@ theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := eq_sub_iff_add_eq.2 <| sin_sq_add_cos_sq _ #align real.sin_sq Real.sin_sq +lemma sin_sq_eq_half_sub : sin x ^ 2 = 1 / 2 - cos (2 * x) / 2 := by + rw [sin_sq, cos_sq, ← sub_sub, sub_half] + theorem abs_sin_eq_sqrt_one_sub_cos_sq (x : ℝ) : |sin x| = sqrt (1 - cos x ^ 2) := by rw [← sin_sq, sqrt_sq_eq_abs] #align real.abs_sin_eq_sqrt_one_sub_cos_sq Real.abs_sin_eq_sqrt_one_sub_cos_sq diff --git a/Mathlib/Data/Complex/Module.lean b/Mathlib/Data/Complex/Module.lean index 085514d0a279e..8f833afe4071b 100644 --- a/Mathlib/Data/Complex/Module.lean +++ b/Mathlib/Data/Complex/Module.lean @@ -479,7 +479,7 @@ lemma IsSelfAdjoint.coe_realPart {x : A} (hx : IsSelfAdjoint x) : (ℜ x : A) = x := hx.coe_selfAdjointPart_apply ℝ -lemma IsSelfAdjoint.imaginaryPart {x : A} (hx : IsSelfAdjoint x) : +nonrec lemma IsSelfAdjoint.imaginaryPart {x : A} (hx : IsSelfAdjoint x) : ℑ x = 0 := by rw [imaginaryPart, LinearMap.comp_apply, hx.skewAdjointPart_apply _, map_zero] diff --git a/Mathlib/Data/Complex/Order.lean b/Mathlib/Data/Complex/Order.lean index c0a96ef7a5ec9..f043f6e020bbf 100644 --- a/Mathlib/Data/Complex/Order.lean +++ b/Mathlib/Data/Complex/Order.lean @@ -69,7 +69,6 @@ theorem real_le_real {x y : ℝ} : (x : ℂ) ≤ (y : ℂ) ↔ x ≤ y := by sim theorem real_lt_real {x y : ℝ} : (x : ℂ) < (y : ℂ) ↔ x < y := by simp [lt_def, ofReal'] #align complex.real_lt_real Complex.real_lt_real - @[simp, norm_cast] theorem zero_le_real {x : ℝ} : (0 : ℂ) ≤ (x : ℂ) ↔ 0 ≤ x := real_le_real @@ -96,10 +95,24 @@ theorem not_lt_zero_iff {z : ℂ} : ¬z < 0 ↔ 0 ≤ z.re ∨ z.im ≠ 0 := not_lt_iff #align complex.not_lt_zero_iff Complex.not_lt_zero_iff -theorem eq_re_ofReal_le {r : ℝ} {z : ℂ} (hz : (r : ℂ) ≤ z) : z = z.re := by +theorem eq_re_of_ofReal_le {r : ℝ} {z : ℂ} (hz : (r : ℂ) ≤ z) : z = z.re := by ext rfl simp only [← (Complex.le_def.1 hz).2, Complex.zero_im, Complex.ofReal_im] -#align complex.eq_re_of_real_le Complex.eq_re_ofReal_le +#align complex.eq_re_of_real_le Complex.eq_re_of_ofReal_le + +@[simp] +lemma re_eq_abs {z : ℂ} : z.re = abs z ↔ 0 ≤ z := + have : 0 ≤ abs z := map_nonneg abs z + ⟨fun h ↦ ⟨h.symm ▸ this, (abs_re_eq_abs.1 <| h.symm ▸ _root_.abs_of_nonneg this).symm⟩, + fun ⟨h₁, h₂⟩ ↦ by rw [← abs_re_eq_abs.2 h₂.symm, _root_.abs_of_nonneg h₁]⟩ + +@[simp] +lemma neg_re_eq_abs {z : ℂ} : -z.re = abs z ↔ z ≤ 0 := by + rw [← neg_re, ← abs.map_neg, re_eq_abs] + exact neg_nonneg.and <| eq_comm.trans neg_eq_zero + +@[simp] +lemma re_eq_neg_abs {z : ℂ} : z.re = -abs z ↔ z ≤ 0 := by rw [← neg_eq_iff_eq_neg, neg_re_eq_abs] end Complex diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index fcb273d3dc9bb..bcb95ce0da54e 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -149,13 +149,17 @@ theorem coe_ne_top (a : ℕ) : (a : ℕ∞) ≠ ⊤ := theorem top_sub_coe (a : ℕ) : (⊤ : ℕ∞) - a = ⊤ := WithTop.top_sub_coe +@[simp] +theorem zero_lt_top : (0 : ℕ∞) < ⊤ := + WithTop.zero_lt_top + --Porting note: new theorem copied from `WithTop` theorem sub_top (a : ℕ∞) : a - ⊤ = 0 := WithTop.sub_top @[simp] theorem coe_toNat_eq_self : ENat.toNat (n : ℕ∞) = n ↔ n ≠ ⊤ := - ENat.recTopCoe (by simp) (fun _ => by simp [toNat_coe]) n + ENat.recTopCoe (by decide) (fun _ => by simp [toNat_coe]) n #align enat.coe_to_nat_eq_self ENat.coe_toNat_eq_self alias ⟨_, coe_toNat⟩ := coe_toNat_eq_self diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 570fd2c1156c5..37b98bad627e5 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -506,7 +506,7 @@ theorem val_one'' {n : ℕ} : ((1 : Fin (n + 1)) : ℕ) = 1 % (n + 1) := #align fin.mk_one Fin.mk_one instance nontrivial {n : ℕ} : Nontrivial (Fin (n + 2)) where - exists_pair_ne := ⟨0, 1, (ne_iff_vne 0 1).mpr (by simp only [val_one, val_zero])⟩ + exists_pair_ne := ⟨0, 1, (ne_iff_vne 0 1).mpr (by simp [val_one, val_zero])⟩ theorem nontrivial_iff_two_le : Nontrivial (Fin n) ↔ 2 ≤ n := by rcases n with (_ | _ | n) <;> @@ -1218,6 +1218,9 @@ protected theorem coe_sub (a b : Fin n) : ((a - b : Fin n) : ℕ) = (a + (n - b) theorem coe_fin_one (a : Fin 1) : (a : ℕ) = 0 := by simp [Subsingleton.elim a 0] #align fin.coe_fin_one Fin.coe_fin_one +lemma eq_one_of_neq_zero (i : Fin 2) (hi : i ≠ 0) : i = 1 := + fin_two_eq_of_eq_zero_iff (by simpa only [one_eq_zero_iff, succ.injEq, iff_false] using hi) + @[simp] theorem coe_neg_one : ↑(-1 : Fin (n + 1)) = n := by cases n @@ -1265,7 +1268,7 @@ theorem coe_sub_iff_lt {n : ℕ} {a b : Fin n} : (↑(a - b) : ℕ) = n + a - b theorem lt_sub_one_iff {n : ℕ} {k : Fin (n + 2)} : k < k - 1 ↔ k = 0 := by rcases k with ⟨_ | k, hk⟩ simp only [zero_eq, zero_eta, zero_sub, lt_iff_val_lt_val, val_zero, coe_neg_one, add_pos_iff, - or_true] + _root_.zero_lt_one, or_true] have : (k + 1 + (n + 1)) % (n + 2) = k % (n + 2) := by rw [add_right_comm, add_assoc, add_mod_right] simp [lt_iff_val_lt_val, ext_iff, Fin.coe_sub, succ_eq_add_one, this, @@ -1424,7 +1427,7 @@ theorem succAbove_castLT {x y : Fin (n + 1)} (h : x < y) theorem succAbove_pred {x y : Fin (n + 1)} (h : x < y) (hy : y ≠ 0 := (x.zero_le.trans_lt h).ne') : x.succAbove (y.pred hy) = y := by rw [succAbove_above, succ_pred] - simpa [le_iff_val_le_val] using Nat.le_pred_of_lt h + simpa [le_iff_val_le_val] using Nat.le_sub_one_of_lt h #align fin.succ_above_pred Fin.succAbove_pred theorem castLT_succAbove {x : Fin n} {y : Fin (n + 1)} (h : castSucc x < y) @@ -1664,7 +1667,7 @@ theorem succAbove_predAbove {p : Fin n} {i : Fin (n + 1)} (h : i ≠ castSucc p) rw [if_neg] · simp · simp only [pred, Fin.mk_lt_mk, not_lt] - exact Nat.le_pred_of_lt (h.symm.lt_of_le H) + exact Nat.le_sub_one_of_lt (h.symm.lt_of_le H) · exact lt_of_le_of_ne H h.symm #align fin.succ_above_pred_above Fin.succAbove_predAbove @@ -1679,7 +1682,7 @@ theorem predAbove_succAbove (p : Fin n) (i : Fin n) : dsimp split_ifs with h₁ h₂ h₃ · simp only [← val_fin_lt, not_lt] at h₁ h₂ - exact (lt_le_antisymm h₁ (le_of_lt h₂)).elim + exact (Nat.lt_le_asymm h₁ (le_of_lt h₂)).elim · rfl · rfl · simp only [← val_fin_lt, not_lt] at h₁ h₃ diff --git a/Mathlib/Data/Fin/Tuple/Monotone.lean b/Mathlib/Data/Fin/Tuple/Monotone.lean index 7c48630f7e14b..51d4c6d892045 100644 --- a/Mathlib/Data/Fin/Tuple/Monotone.lean +++ b/Mathlib/Data/Fin/Tuple/Monotone.lean @@ -38,12 +38,12 @@ theorem monotone_vecCons : Monotone (vecCons a f) ↔ a ≤ f 0 ∧ Monotone f : --Porting note: new lemma, in Lean3 would be proven by `Subsingleton.monotone` @[simp] -theorem monotone_vecEmpty : Monotone (vecCons a vecEmpty) +theorem monotone_vecEmpty : Monotone ![a] | ⟨0, _⟩, ⟨0, _⟩, _ => le_refl _ --Porting note: new lemma, in Lean3 would be proven by `Subsingleton.strictMono` @[simp] -theorem strictMono_vecEmpty : StrictMono (vecCons a vecEmpty) +theorem strictMono_vecEmpty : StrictMono ![a] | ⟨0, _⟩, ⟨0, _⟩, h => (irrefl _ h).elim @[simp] @@ -82,4 +82,5 @@ theorem Antitone.vecCons (hf : Antitone f) (ha : f 0 ≤ a) : Antitone (vecCons antitone_vecCons.2 ⟨ha, hf⟩ #align antitone.vec_cons Antitone.vecCons -example : Monotone ![1, 2, 2, 3] := by simp +-- NOTE: was "by simp", but simp lemmas were not being used +example : Monotone ![1, 2, 2, 3] := by decide diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index 6a40727a4ad56..eabbc1fd9527a 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -264,7 +264,7 @@ section EquivProd /-- The disjoint union of antidiagonal tuples `Σ n, antidiagonalTuple k n` is equivalent to the `k`-tuple `Fin k → ℕ`. This is such an equivalence, obtained by mapping `(n, x)` to `x`. -This is the tuple version of `Finset.Nat.sigmaAntidiagonalEquivProd`. -/ +This is the tuple version of `Finset.sigmaAntidiagonalEquivProd`. -/ @[simps] def sigmaAntidiagonalTupleEquivTuple (k : ℕ) : (Σ n, antidiagonalTuple k n) ≃ (Fin k → ℕ) where diff --git a/Mathlib/Data/Finset/Antidiagonal.lean b/Mathlib/Data/Finset/Antidiagonal.lean new file mode 100644 index 0000000000000..e7c0e8f21ad13 --- /dev/null +++ b/Mathlib/Data/Finset/Antidiagonal.lean @@ -0,0 +1,195 @@ +/- +Copyright (c) 2023 Antoine Chambert-Loir and María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández, Bhavik Mehta, Eric Wieser +-/ + +import Mathlib.Data.Finset.Basic +import Mathlib.Data.Finsupp.Defs +import Mathlib.Data.Finsupp.Interval +import Mathlib.Algebra.Order.Sub.Defs + +/-! # Antidiagonal with values in general types + +We define a type class `Finset.HasAntidiagonal A` which contains a function +`antidiagonal : A → Finset (A × A)` such that `antidiagonal n` +is the finset of all pairs adding to `n`, as witnessed by `mem_antidiagonal`. + +When `A` is a canonically ordered add monoid with locally finite order +this typeclass can be instantiated with `Finset.antidiagonalOfLocallyFinite`. +This applies in particular when `A` is `ℕ`, more generally or `σ →₀ ℕ`, +or even `ι →₀ A` under the additional assumption `OrderedSub A` +that make it a canonically ordered add monoid. +(In fact, we would just need an `AddMonoid` with a compatible order, +finite `Iic`, such that if `a + b = n`, then `a, b ≤ n`, +and any finiteness condition would be OK.) + +For computational reasons it is better to manually provide instances for `ℕ` +and `σ →₀ ℕ`, to avoid quadratic runtime performance. +These instances are provided as `Finset.Nat.instHasAntidiagonal` and `Finsupp.instHasAntidiagonal`. +This is why `Finset.antidiagonalOfLocallyFinite` is an `abbrev` and not an `instance`. + +This definition does not exactly match with that of `Multiset.antidiagonal` +defined in `Mathlib.Data.Multiset.Antidiagonal`, because of the multiplicities. +Indeed, by counting multiplicities, `Multiset α` is equivalent to `α →₀ ℕ`, +but `Finset.antidiagonal` and `Multiset.antidiagonal` will return different objects. +For example, for `s : Multiset ℕ := {0,0,0}`, `Multiset.antidiagonal s` has 8 elements +but `Finset.antidiagonal s` has only 4. + +```lean +def s : Multiset ℕ := {0, 0, 0} +#eval (Finset.antidiagonal s).card -- 4 +#eval Multiset.card (Multiset.antidiagonal s) -- 8 +``` + +## TODO + +* Define `HasMulAntidiagonal` (for monoids). + For `PNat`, we will recover the set of divisors of a strictly positive integer. +-/ + +namespace Finset + +open scoped BigOperators + +open Function + +/-- The class of additive monoids with an antidiagonal -/ +class HasAntidiagonal (A : Type*) [AddMonoid A] where + /-- The antidiagonal of an element `n` is the finset of pairs `(i, j)` such that `i + j = n`. -/ + antidiagonal : A → Finset (A × A) + /-- A pair belongs to `antidiagonal n` iff the sum of its components is equal to `n`. -/ + mem_antidiagonal {n} {a} : a ∈ antidiagonal n ↔ a.fst + a.snd = n + +export HasAntidiagonal (antidiagonal mem_antidiagonal) + +attribute [simp] mem_antidiagonal + +variable {A : Type*} + +/-- All `HasAntidiagonal` instances are equal -/ +instance [AddMonoid A] : Subsingleton (HasAntidiagonal A) := + ⟨by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + congr with n xy + rw [ha, hb]⟩ + +-- The goal of this lemma is to allow to rewrite antidiagonal +-- when the decidability instances obsucate Lean +lemma hasAntidiagonal_congr (A : Type*) [AddMonoid A] + [H1 : HasAntidiagonal A] [H2 : HasAntidiagonal A] : + H1.antidiagonal = H2.antidiagonal := by congr! + +theorem swap_mem_antidiagonal [AddCommMonoid A] [HasAntidiagonal A] {n : A} {xy : A × A}: + xy.swap ∈ antidiagonal n ↔ xy ∈ antidiagonal n := by + simp [add_comm] + +@[simp] theorem map_prodComm_antidiagonal [AddCommMonoid A] [HasAntidiagonal A] {n : A} : + (antidiagonal n).map (Equiv.prodComm A A) = antidiagonal n := + Finset.ext fun ⟨a, b⟩ => by simp [add_comm] + +/-- See also `Finset.map_prodComm_antidiagonal`. -/ +@[simp] theorem map_swap_antidiagonal [AddCommMonoid A] [HasAntidiagonal A] {n : A} : + (antidiagonal n).map ⟨Prod.swap, Prod.swap_injective⟩ = antidiagonal n := + map_prodComm_antidiagonal +#align finset.nat.map_swap_antidiagonal Finset.map_swap_antidiagonal + +/-- A point in the antidiagonal is determined by its first co-ordinate. -/ +theorem antidiagonal_congr [AddCancelMonoid A] [HasAntidiagonal A] {n : A} {p q : A × A} + (hp : p ∈ antidiagonal n) (hq : q ∈ antidiagonal n) : p = q ↔ p.fst = q.fst := by + refine' ⟨congr_arg Prod.fst, fun h ↦ Prod.ext h ((add_right_inj q.fst).mp _)⟩ + rw [mem_antidiagonal] at hp hq + rw [hq, ← h, hp] +#align finset.nat.antidiagonal_congr Finset.antidiagonal_congr + +/-- A point in the antidiagonal is determined by its first co-ordinate (subtype version of +`Finset.antidiagonal_congr`). This lemma is used by the `ext` tactic. -/ +@[ext] theorem antidiagonal_subtype_ext [AddCancelMonoid A] [HasAntidiagonal A] + {n : A} {p q : antidiagonal n} (h : p.val.1 = q.val.1) : + p = q := Subtype.ext ((antidiagonal_congr p.prop q.prop).mpr h) + +section CanonicallyOrderedAddCommMonoid +variable [CanonicallyOrderedAddCommMonoid A] [HasAntidiagonal A] + +@[simp] +theorem antidiagonal_zero : antidiagonal (0 : A) = {(0, 0)} := by + ext ⟨x, y⟩ + simp + +theorem antidiagonal.fst_le {n : A} {kl : A × A} (hlk : kl ∈ antidiagonal n) : kl.1 ≤ n := by + rw [le_iff_exists_add] + use kl.2 + rwa [mem_antidiagonal, eq_comm] at hlk +#align finset.nat.antidiagonal.fst_le Finset.antidiagonal.fst_le + +theorem antidiagonal.snd_le {n : A} {kl : A × A} (hlk : kl ∈ antidiagonal n) : kl.2 ≤ n := by + rw [le_iff_exists_add] + use kl.1 + rwa [mem_antidiagonal, eq_comm, add_comm] at hlk +#align finset.nat.antidiagonal.snd_le Finset.antidiagonal.snd_le + +end CanonicallyOrderedAddCommMonoid + +section OrderedSub +variable [CanonicallyOrderedAddCommMonoid A] [Sub A] [OrderedSub A] +variable [ContravariantClass A A (· + ·) (· ≤ ·)] +variable [HasAntidiagonal A] + +theorem filter_fst_eq_antidiagonal (n m : A) [DecidablePred (· = m)] [Decidable (m ≤ n)] : + filter (fun x : A × A ↦ x.fst = m) (antidiagonal n) = if m ≤ n then {(m, n - m)} else ∅ := by + ext ⟨a, b⟩ + suffices a = m → (a + b = n ↔ m ≤ n ∧ b = n - m) by + rw [mem_filter, mem_antidiagonal, apply_ite (fun n ↦ (a, b) ∈ n), mem_singleton, + Prod.mk.inj_iff, ite_prop_iff_or] + simpa [ ← and_assoc, @and_right_comm _ (a = _), and_congr_left_iff] + rintro rfl + constructor + · rintro rfl + exact ⟨le_add_right le_rfl, (add_tsub_cancel_left _ _).symm⟩ + · rintro ⟨h, rfl⟩ + exact add_tsub_cancel_of_le h +#align finset.nat.filter_fst_eq_antidiagonal Finset.filter_fst_eq_antidiagonal + +theorem filter_snd_eq_antidiagonal (n m : A) [DecidablePred (· = m)] [Decidable (m ≤ n)] : + filter (fun x : A × A ↦ x.snd = m) (antidiagonal n) = if m ≤ n then {(n - m, m)} else ∅ := by + have : (fun x : A × A ↦ (x.snd = m)) ∘ Prod.swap = fun x : A × A ↦ x.fst = m := by + ext; simp + rw [← map_swap_antidiagonal, filter_map] + simp [this, filter_fst_eq_antidiagonal, apply_ite (Finset.map _)] +#align finset.nat.filter_snd_eq_antidiagonal Finset.filter_snd_eq_antidiagonal + +end OrderedSub + +/-- The disjoint union of antidiagonals `Σ (n : A), antidiagonal n` is equivalent to the product + `A × A`. This is such an equivalence, obtained by mapping `(n, (k, l))` to `(k, l)`. -/ +@[simps] +def sigmaAntidiagonalEquivProd [AddMonoid A] [HasAntidiagonal A] : + (Σ n : A, antidiagonal n) ≃ A × A where + toFun x := x.2 + invFun x := ⟨x.1 + x.2, x, mem_antidiagonal.mpr rfl⟩ + left_inv := by + rintro ⟨n, ⟨k, l⟩, h⟩ + rw [mem_antidiagonal] at h + exact Sigma.subtype_ext h rfl + right_inv x := rfl +#align finset.nat.sigma_antidiagonal_equiv_prod Finset.sigmaAntidiagonalEquivProd +#align finset.nat.sigma_antidiagonal_equiv_prod_symm_apply_fst Finset.sigmaAntidiagonalEquivProd_symm_apply_fst +#align finset.nat.sigma_antidiagonal_equiv_prod_symm_apply_snd_coe Finset.sigmaAntidiagonalEquivProd_symm_apply_snd_coe +#align finset.nat.sigma_antidiagonal_equiv_prod_apply Finset.sigmaAntidiagonalEquivProd_apply + +variable {A : Type*} + [CanonicallyOrderedAddCommMonoid A] + [LocallyFiniteOrder A] [DecidableEq A] + +/-- In a canonically ordered add monoid, the antidiagonal can be construct by filtering. + +Note that this is not an instance, as for some times a more efficient algorithm is available. -/ +abbrev antidiagonalOfLocallyFinite : HasAntidiagonal A where + antidiagonal n := Finset.filter (fun uv => uv.fst + uv.snd = n) (Finset.product (Iic n) (Iic n)) + mem_antidiagonal {n} {a} := by + simp only [Prod.forall, mem_filter, and_iff_right_iff_imp] + intro h; rw [← h] + erw [mem_product, mem_Iic, mem_Iic] + exact ⟨le_self_add, le_add_self⟩ + +end Finset diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index d6e8fbc565b43..2c3a447124edb 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura, Jeremy Avigad, Minchao Wu, Mario Carneiro -/ import Mathlib.Data.Multiset.FinsetOps import Mathlib.Data.Set.Lattice +import Mathlib.Order.Cover #align_import data.finset.basic from "leanprover-community/mathlib"@"442a83d738cb208d3600056c489be16900ba701d" @@ -927,6 +928,17 @@ theorem ssubset_iff_exists_cons_subset : s ⊂ t ↔ ∃ (a : _) (h : a ∉ s), exact ⟨a, ht, cons_subset.2 ⟨hs, h.subset⟩⟩ #align finset.ssubset_iff_exists_cons_subset Finset.ssubset_iff_exists_cons_subset +lemma covby_cons (ha : a ∉ s) : s ⋖ cons a s ha := + Covby.of_image ⟨⟨_, coe_injective⟩, coe_subset⟩ <| by + simpa using Set.covby_insert (show a ∉ (s : Set α) from ha) + +lemma covby_iff_exists_cons : s ⋖ t ↔ ∃ a, ∃ ha : a ∉ s, cons a s ha = t := by + refine ⟨fun hst ↦ ?_, ?_⟩ + · obtain ⟨a, ha, hast⟩ := ssubset_iff_exists_cons_subset.1 hst.1 + exact ⟨a, ha, eq_of_le_of_not_lt hast <| hst.2 <| ssubset_cons _⟩ + · rintro ⟨a, ha, rfl⟩ + exact covby_cons ha + end Cons /-! ### disjoint -/ @@ -1191,7 +1203,7 @@ theorem insert_subset_iff : insert a s ⊆ t ↔ a ∈ t ∧ s ⊆ t := by theorem insert_subset (ha : a ∈ t) (hs : s ⊆ t) : insert a s ⊆ t := insert_subset_iff.mpr ⟨ha,hs⟩ -theorem subset_insert (a : α) (s : Finset α) : s ⊆ insert a s := fun _b => mem_insert_of_mem +@[simp] theorem subset_insert (a : α) (s : Finset α) : s ⊆ insert a s := fun _b => mem_insert_of_mem #align finset.subset_insert Finset.subset_insert theorem insert_subset_insert (a : α) {s t : Finset α} (h : s ⊆ t) : insert a s ⊆ insert a t := @@ -1311,6 +1323,11 @@ theorem disjoint_insert_right : Disjoint s (insert a t) ↔ a ∉ s ∧ Disjoint disjoint_comm.trans <| by rw [disjoint_insert_left, _root_.disjoint_comm] #align finset.disjoint_insert_right Finset.disjoint_insert_right +lemma covby_insert (ha : a ∉ s) : s ⋖ insert a s := by simpa using covby_cons ha + +lemma covby_iff_exists_insert : t ⋖ s ↔ ∃ a, a ∉ t ∧ insert a t = s := by + simp [covby_iff_exists_cons] + end Insert /-! ### Lattice structure -/ @@ -1658,7 +1675,7 @@ theorem inter_right_comm (s₁ s₂ s₃ : Finset α) : s₁ ∩ s₂ ∩ s₃ = @[simp] theorem inter_self (s : Finset α) : s ∩ s = s := - ext fun _ => mem_inter.trans <| and_self_iff _ + ext fun _ => mem_inter.trans <| and_self_iff #align finset.inter_self Finset.inter_self @[simp] @@ -1954,7 +1971,7 @@ theorem erase_cons_of_ne {a b : α} {s : Finset α} (ha : a ∉ s) (hb : a ≠ b simp only [cons_eq_insert, erase_insert_of_ne hb] #align finset.erase_cons_of_ne Finset.erase_cons_of_ne -theorem insert_erase {a : α} {s : Finset α} (h : a ∈ s) : insert a (erase s a) = s := +@[simp] theorem insert_erase (h : a ∈ s) : insert a (erase s a) = s := ext fun x => by simp only [mem_insert, mem_erase, or_and_left, dec_em, true_and_iff] apply or_iff_right_of_imp @@ -1963,7 +1980,7 @@ theorem insert_erase {a : α} {s : Finset α} (h : a ∈ s) : insert a (erase s #align finset.insert_erase Finset.insert_erase lemma erase_eq_iff_eq_insert (hs : a ∈ s) (ht : a ∉ t) : erase s a = t ↔ s = insert a t := by - have := insert_erase hs; aesop + aesop lemma insert_erase_invOn : Set.InvOn (insert a) (λ s ↦ erase s a) {s : Finset α | a ∈ s} {s : Finset α | a ∉ s} := @@ -2055,6 +2072,9 @@ theorem erase_injOn' (a : α) : { s : Finset α | a ∈ s }.InjOn fun s => erase fun s hs t ht (h : s.erase a = _) => by rw [← insert_erase hs, ← insert_erase ht, h] #align finset.erase_inj_on' Finset.erase_injOn' +lemma covby_iff_exists_erase : t ⋖ s ↔ ∃ a, a ∈ s ∧ t = s.erase a := + covby_iff_exists_insert.trans $ exists_congr fun a ↦ by aesop + end Erase /-! ### sdiff -/ @@ -2218,7 +2238,7 @@ theorem insert_sdiff_of_mem (s : Finset α) {x : α} (h : x ∈ t) : insert x s exact Set.insert_diff_of_mem _ h #align finset.insert_sdiff_of_mem Finset.insert_sdiff_of_mem -lemma insert_sdiff_cancel (ha : a ∉ s) : insert a s \ s = {a} := by +@[simp] lemma insert_sdiff_cancel (ha : a ∉ s) : insert a s \ s = {a} := by rw [insert_sdiff_of_not_mem _ ha, Finset.sdiff_self, insert_emptyc_eq] @[simp] @@ -2672,6 +2692,9 @@ instance decidableExistsAndFinset {p : α → Prop} [_hp : ∀ (a), Decidable (p Decidable (∃ a ∈ s, p a) := decidable_of_iff (∃ (a : _) (_ : a ∈ s), p a) (by simp) +instance decidableExistsAndFinsetCoe {p : α → Prop} [DecidablePred p] : + Decidable (∃ a ∈ (s : Set α), p a) := decidableExistsAndFinset + /-- decidable equality for functions whose domain is bounded by finsets -/ instance decidableEqPiFinset {β : α → Type*} [_h : ∀ a, DecidableEq (β a)] : DecidableEq (∀ a ∈ s, β a) := diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 60dabbd2eafe2..523f8fb3c0565 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -624,8 +624,10 @@ lemma exists_of_one_lt_card_pi {ι : Type*} {α : ι → Type*} [∀ i, Decidabl obtain rfl | hne := eq_or_ne (a2 i) ai exacts [⟨a1, h1, hne⟩, ⟨a2, h2, hne⟩] -theorem card_eq_succ [DecidableEq α] : - s.card = n + 1 ↔ ∃ a t, a ∉ t ∧ insert a t = s ∧ t.card = n := +section DecidableEq +variable [DecidableEq α] + +theorem card_eq_succ : s.card = n + 1 ↔ ∃ a t, a ∉ t ∧ insert a t = s ∧ t.card = n := ⟨fun h => let ⟨a, has⟩ := card_pos.mp (h.symm ▸ Nat.zero_lt_succ _ : 0 < s.card) ⟨a, s.erase a, s.not_mem_erase a, insert_erase has, by @@ -633,7 +635,7 @@ theorem card_eq_succ [DecidableEq α] : fun ⟨a, t, hat, s_eq, n_eq⟩ => s_eq ▸ n_eq ▸ card_insert_of_not_mem hat⟩ #align finset.card_eq_succ Finset.card_eq_succ -theorem card_eq_two [DecidableEq α] : s.card = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := by +theorem card_eq_two : s.card = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := by constructor · rw [card_eq_succ] simp_rw [card_eq_one] @@ -643,8 +645,7 @@ theorem card_eq_two [DecidableEq α] : s.card = 2 ↔ ∃ x y, x ≠ y ∧ s = { exact card_doubleton h #align finset.card_eq_two Finset.card_eq_two -theorem card_eq_three [DecidableEq α] : - s.card = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ z ∧ s = {x, y, z} := by +theorem card_eq_three : s.card = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ z ∧ s = {x, y, z} := by constructor · rw [card_eq_succ] simp_rw [card_eq_two] @@ -656,6 +657,18 @@ theorem card_eq_three [DecidableEq α] : or_self_iff, card_singleton] #align finset.card_eq_three Finset.card_eq_three +lemma covby_iff_card_sdiff_eq_one : t ⋖ s ↔ t ⊆ s ∧ (s \ t).card = 1 := by + rw [covby_iff_exists_insert] + constructor + · rintro ⟨a, ha, rfl⟩ + simp [*] + · simp_rw [card_eq_one] + rintro ⟨hts, a, ha⟩ + refine ⟨a, (mem_sdiff.1 $ superset_of_eq ha $ mem_singleton_self _).2, ?_⟩ + rw [insert_eq, ←ha, sdiff_union_of_subset hts] + +end DecidableEq + /-! ### Inductions -/ diff --git a/Mathlib/Data/Finset/Image.lean b/Mathlib/Data/Finset/Image.lean index 54cf5a15c993e..9644e027cb155 100644 --- a/Mathlib/Data/Finset/Image.lean +++ b/Mathlib/Data/Finset/Image.lean @@ -3,7 +3,7 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Minchao Wu, Mario Carneiro -/ -import Mathlib.Algebra.Hom.Embedding +import Mathlib.Algebra.Group.Embedding import Mathlib.Data.Fin.Basic import Mathlib.Data.Finset.Basic import Mathlib.Data.Int.Order.Basic diff --git a/Mathlib/Data/Finset/Interval.lean b/Mathlib/Data/Finset/Interval.lean index 14ec96913ee86..1f9cc8a7f8399 100644 --- a/Mathlib/Data/Finset/Interval.lean +++ b/Mathlib/Data/Finset/Interval.lean @@ -24,7 +24,7 @@ out of `Finset α` in terms of `Finset.insert` -/ -variable {α : Type*} +variable {α β : Type*} namespace Finset @@ -132,50 +132,29 @@ theorem card_Iio_finset : (Iio s).card = 2 ^ s.card - 1 := by end Decidable -variable {s t : Finset α} +variable [Preorder β] {s t : Finset α} {f : Finset α → β} section Cons -lemma covby_cons {i : α} (hi : i ∉ s) : s ⋖ cons i s hi := - Covby.of_image ⟨⟨((↑) : Finset α → Set α), coe_injective⟩, coe_subset⟩ <| by - simp only [RelEmbedding.coe_mk, Function.Embedding.coeFn_mk, coe_cons, mem_coe] - exact_mod_cast Set.covby_insert (show i ∉ (s : Set α) from hi) +/-- A function `f` from `Finset α` is monotone if and only if `f s ≤ f (cons a s ha)` for all `s` +and `a ∉ s`. -/ +lemma monotone_iff_forall_le_cons : Monotone f ↔ ∀ s, ∀ ⦃a⦄ (ha), f s ≤ f (cons a s ha) := by + classical simp [monotone_iff_forall_covby, covby_iff_exists_cons] -lemma covby_iff : s ⋖ t ↔ ∃ i : α, ∃ hi : i ∉ s, t = cons i s hi := by - constructor - · intro hst - obtain ⟨i, hi, his⟩ := ssubset_iff_exists_cons_subset.mp hst.1 - exact ⟨i, hi, .symm <| eq_of_le_of_not_lt his <| hst.2 <| ssubset_cons hi⟩ - · rintro ⟨i, hi, rfl⟩ - exact covby_cons hi - -/-- A function `f` from `Finset α` is monotone if and only if `f s ≤ f (cons i s hi)` for all -`s` and `i ∉ s`. -/ -theorem monotone_iff {β : Type*} [Preorder β] (f : Finset α → β) : - Monotone f ↔ ∀ s : Finset α, ∀ {i} (hi : i ∉ s), f s ≤ f (cons i s hi) := by - classical - simp only [monotone_iff_forall_covby, covby_iff, forall_exists_index, and_imp] - aesop - -/-- A function `f` from `Finset α` is strictly monotone if and only if `f s < f (insert i s)` for -all `s` and `i ∉ s`. -/ -theorem strictMono_iff {β : Type*} [Preorder β] (f : Finset α → β) : - StrictMono f ↔ ∀ s : Finset α, ∀ {i} (hi : i ∉ s), f s < f (cons i s hi) := by - classical - simp only [strictMono_iff_forall_covby, covby_iff, forall_exists_index, and_imp] - aesop - -/-- A function `f` from `Finset α` is antitone if and only if `f (cons i s hi) ≤ f s` for all -`s` and `i ∉ s`. -/ -theorem antitone_iff {β : Type*} [Preorder β] (f : Finset α → β) : - Antitone f ↔ ∀ s : Finset α, ∀ {i} (hi : i ∉ s), f (cons i s hi) ≤ f s := - monotone_iff (β := βᵒᵈ) f - -/-- A function `f` from `Finset α` is strictly antitone if and only if `f (cons i s hi) < f s` for -all `s` and `i ∉ s`. -/ -theorem strictAnti_iff {β : Type*} [Preorder β] (f : Finset α → β) : - StrictAnti f ↔ ∀ s : Finset α, ∀ {i} (hi : i ∉ s), f (cons i s hi) < f s := - strictMono_iff (β := βᵒᵈ) f +/-- A function `f` from `Finset α` is antitone if and only if `f (cons a s ha) ≤ f s` for all +`s` and `a ∉ s`. -/ +lemma antitone_iff_forall_cons_le : Antitone f ↔ ∀ s ⦃a⦄ ha, f (cons a s ha) ≤ f s := + monotone_iff_forall_le_cons (β := βᵒᵈ) + +/-- A function `f` from `Finset α` is strictly monotone if and only if `f s < f (cons a s ha)` for +all `s` and `a ∉ s`. -/ +lemma strictMono_iff_forall_lt_cons : StrictMono f ↔ ∀ s ⦃a⦄ ha, f s < f (cons a s ha) := by + classical simp [strictMono_iff_forall_covby, covby_iff_exists_cons] + +/-- A function `f` from `Finset α` is strictly antitone if and only if `f (cons a s ha) < f s` for +all `s` and `a ∉ s`. -/ +lemma strictAnti_iff_forall_cons_lt : StrictAnti f ↔ ∀ s ⦃a⦄ ha, f (cons a s ha) < f s := + strictMono_iff_forall_lt_cons (β := βᵒᵈ) end Cons @@ -183,35 +162,25 @@ section Insert variable [DecidableEq α] -lemma covby_insert {i : α} (hi : i ∉ s) : s ⋖ insert i s := by - simpa using covby_cons hi - -lemma covby_iff' : s ⋖ t ↔ ∃ i : α, i ∉ s ∧ t = insert i s := by - simp [covby_iff] - -/-- A function `f` from `Finset α` is monotone if and only if `f s ≤ f (insert i s)` for all -`s` and `i ∉ s`. -/ -theorem monotone_iff' {β : Type*} [Preorder β] (f : Finset α → β) : - Monotone f ↔ ∀ s : Finset α, ∀ {i} (_hi : i ∉ s), f s ≤ f (insert i s) := by - simp [monotone_iff] - -/-- A function `f` from `Finset α` is strictly monotone if and only if `f s < f (insert i s)` for -all `s` and `i ∉ s`. -/ -theorem strictMono_iff' {β : Type*} [Preorder β] (f : Finset α → β) : - StrictMono f ↔ ∀ s : Finset α, ∀ {i} (_hi : i ∉ s), f s < f (insert i s) := by - simp [strictMono_iff] - -/-- A function `f` from `Finset α` is antitone if and only if `f (insert i s) ≤ f s` for all -`s` and `i ∉ s`. -/ -theorem antitone_iff' {β : Type*} [Preorder β] (f : Finset α → β) : - Antitone f ↔ ∀ s : Finset α, ∀ {i} (_hi : i ∉ s), f (insert i s) ≤ f s := - monotone_iff' (β := βᵒᵈ) f - -/-- A function `f` from `Finset α` is strictly antitone if and only if `f (insert i s) < f s` for -all `s` and `i ∉ s`. -/ -theorem strictAnti_iff' {β : Type*} [Preorder β] (f : Finset α → β) : - StrictAnti f ↔ ∀ s : Finset α, ∀ {i} (_hi : i ∉ s), f (insert i s) < f s := - strictMono_iff' (β := βᵒᵈ) f +/-- A function `f` from `Finset α` is monotone if and only if `f s ≤ f (insert a s)` for all `s` and +`a ∉ s`. -/ +lemma monotone_iff_forall_le_insert : Monotone f ↔ ∀ s ⦃a⦄, a ∉ s → f s ≤ f (insert a s) := by + simp [monotone_iff_forall_le_cons] + +/-- A function `f` from `Finset α` is antitone if and only if `f (insert a s) ≤ f s` for all +`s` and `a ∉ s`. -/ +lemma antitone_iff_forall_insert_le : Antitone f ↔ ∀ s ⦃a⦄, a ∉ s → f (insert a s) ≤ f s := + monotone_iff_forall_le_insert (β := βᵒᵈ) + +/-- A function `f` from `Finset α` is strictly monotone if and only if `f s < f (insert a s)` for +all `s` and `a ∉ s`. -/ +lemma strictMono_iff_forall_lt_insert : StrictMono f ↔ ∀ s ⦃a⦄, a ∉ s → f s < f (insert a s) := by + simp [strictMono_iff_forall_lt_cons] + +/-- A function `f` from `Finset α` is strictly antitone if and only if `f (insert a s) < f s` for +all `s` and `a ∉ s`. -/ +lemma strictAnti_iff_forall_lt_insert : StrictAnti f ↔ ∀ s ⦃a⦄, a ∉ s → f (insert a s) < f s := + strictMono_iff_forall_lt_insert (β := βᵒᵈ) end Insert diff --git a/Mathlib/Data/Finset/LocallyFinite.lean b/Mathlib/Data/Finset/LocallyFinite.lean index fc4aa6ceeb80c..057d9c41bbfe0 100644 --- a/Mathlib/Data/Finset/LocallyFinite.lean +++ b/Mathlib/Data/Finset/LocallyFinite.lean @@ -357,12 +357,12 @@ theorem Ico_filter_le_of_left_le {a b c : α} [DecidablePred ((· ≤ ·) c)] (h theorem Icc_filter_lt_of_lt_right {a b c : α} [DecidablePred (· < c)] (h : b < c) : (Icc a b).filter (· < c) = Icc a b := - filter_true_of_mem fun _ hx => lt_of_le_of_lt (mem_Icc.1 hx).2 h + filter_true_of_mem fun _ hx => lt_of_le_of_lt (mem_Icc.1 hx).2 h #align finset.Icc_filter_lt_of_lt_right Finset.Icc_filter_lt_of_lt_right theorem Ioc_filter_lt_of_lt_right {a b c : α} [DecidablePred (· < c)] (h : b < c) : (Ioc a b).filter (· < c) = Ioc a b := - filter_true_of_mem fun _ hx => lt_of_le_of_lt (mem_Ioc.1 hx).2 h + filter_true_of_mem fun _ hx => lt_of_le_of_lt (mem_Ioc.1 hx).2 h #align finset.Ioc_filter_lt_of_lt_right Finset.Ioc_filter_lt_of_lt_right theorem Iic_filter_lt_of_lt_right {α} [Preorder α] [LocallyFiniteOrderBot α] {a c : α} diff --git a/Mathlib/Data/Finset/NatAntidiagonal.lean b/Mathlib/Data/Finset/NatAntidiagonal.lean index 62586272350ad..ae6a9c67e5890 100644 --- a/Mathlib/Data/Finset/NatAntidiagonal.lean +++ b/Mathlib/Data/Finset/NatAntidiagonal.lean @@ -5,6 +5,7 @@ Authors: Johan Commelin -/ import Mathlib.Data.Finset.Card import Mathlib.Data.Multiset.NatAntidiagonal +import Mathlib.Data.Finset.Antidiagonal #align_import data.finset.nat_antidiagonal from "leanprover-community/mathlib"@"9003f28797c0664a49e4179487267c494477d853" @@ -17,7 +18,8 @@ generally for sums going from `0` to `n`. ## Notes -This refines files `Data.List.NatAntidiagonal` and `Data.Multiset.NatAntidiagonal`. +This refines files `Data.List.NatAntidiagonal` and `Data.Multiset.NatAntidiagonal`, providing an +instance enabling `Finset.antidiagonal` on `Nat`. -/ open Function @@ -28,15 +30,10 @@ namespace Nat /-- The antidiagonal of a natural number `n` is the finset of pairs `(i, j)` such that `i + j = n`. -/ -def antidiagonal (n : ℕ) : Finset (ℕ × ℕ) := - ⟨Multiset.Nat.antidiagonal n, Multiset.Nat.nodup_antidiagonal n⟩ -#align finset.nat.antidiagonal Finset.Nat.antidiagonal - -/-- A pair (i, j) is contained in the antidiagonal of `n` if and only if `i + j = n`. -/ -@[simp] -theorem mem_antidiagonal {n : ℕ} {x : ℕ × ℕ} : x ∈ antidiagonal n ↔ x.1 + x.2 = n := by - rw [antidiagonal, mem_def, Multiset.Nat.mem_antidiagonal] -#align finset.nat.mem_antidiagonal Finset.Nat.mem_antidiagonal +instance instHasAntidiagonal : HasAntidiagonal ℕ where + antidiagonal n := ⟨Multiset.Nat.antidiagonal n, Multiset.Nat.nodup_antidiagonal n⟩ + mem_antidiagonal {n} {xy} := by + rw [mem_def, Multiset.Nat.mem_antidiagonal] /-- The cardinality of the antidiagonal of `n` is `n + 1`. -/ @[simp] @@ -44,7 +41,8 @@ theorem card_antidiagonal (n : ℕ) : (antidiagonal n).card = n + 1 := by simp [ #align finset.nat.card_antidiagonal Finset.Nat.card_antidiagonal /-- The antidiagonal of `0` is the list `[(0, 0)]` -/ -@[simp] +-- nolint as this is for dsimp +@[simp, nolint simpNF] theorem antidiagonal_zero : antidiagonal 0 = {(0, 0)} := rfl #align finset.nat.antidiagonal_zero Finset.Nat.antidiagonal_zero @@ -83,66 +81,12 @@ theorem antidiagonal_succ_succ' {n : ℕ} : rfl #align finset.nat.antidiagonal_succ_succ' Finset.Nat.antidiagonal_succ_succ' -/-- See also `Finset.map.map_prodComm_antidiagonal`. -/ -@[simp] theorem map_swap_antidiagonal {n : ℕ} : - (antidiagonal n).map ⟨Prod.swap, Prod.swap_injective⟩ = antidiagonal n := - eq_of_veq <| by simp [antidiagonal, Multiset.Nat.map_swap_antidiagonal] -#align finset.nat.map_swap_antidiagonal Finset.Nat.map_swap_antidiagonal - -@[simp] theorem map_prodComm_antidiagonal {n : ℕ} : - (antidiagonal n).map (Equiv.prodComm ℕ ℕ) = antidiagonal n := - map_swap_antidiagonal - -/-- A point in the antidiagonal is determined by its first co-ordinate. -/ -theorem antidiagonal_congr {n : ℕ} {p q : ℕ × ℕ} (hp : p ∈ antidiagonal n) - (hq : q ∈ antidiagonal n) : p = q ↔ p.fst = q.fst := by - refine' ⟨congr_arg Prod.fst, fun h ↦ Prod.ext h ((add_right_inj q.fst).mp _)⟩ - rw [mem_antidiagonal] at hp hq - rw [hq, ← h, hp] -#align finset.nat.antidiagonal_congr Finset.Nat.antidiagonal_congr - -/-- A point in the antidiagonal is determined by its first co-ordinate (subtype version of -`antidiagonal_congr`). This lemma is used by the `ext` tactic. -/ -@[ext] theorem antidiagonal_subtype_ext {n : ℕ} {p q : antidiagonal n} (h : p.val.1 = q.val.1) : - p = q := Subtype.ext ((antidiagonal_congr p.prop q.prop).mpr h) - -theorem antidiagonal.fst_le {n : ℕ} {kl : ℕ × ℕ} (hlk : kl ∈ antidiagonal n) : kl.1 ≤ n := by - rw [le_iff_exists_add] - use kl.2 - rwa [mem_antidiagonal, eq_comm] at hlk -#align finset.nat.antidiagonal.fst_le Finset.Nat.antidiagonal.fst_le - theorem antidiagonal.fst_lt {n : ℕ} {kl : ℕ × ℕ} (hlk : kl ∈ antidiagonal n) : kl.1 < n + 1 := Nat.lt_succ_of_le $ antidiagonal.fst_le hlk -theorem antidiagonal.snd_le {n : ℕ} {kl : ℕ × ℕ} (hlk : kl ∈ antidiagonal n) : kl.2 ≤ n := by - rw [le_iff_exists_add] - use kl.1 - rwa [mem_antidiagonal, eq_comm, add_comm] at hlk -#align finset.nat.antidiagonal.snd_le Finset.Nat.antidiagonal.snd_le - theorem antidiagonal.snd_lt {n : ℕ} {kl : ℕ × ℕ} (hlk : kl ∈ antidiagonal n) : kl.2 < n + 1 := Nat.lt_succ_of_le $ antidiagonal.snd_le hlk -theorem filter_fst_eq_antidiagonal (n m : ℕ) : - filter (fun x : ℕ × ℕ ↦ x.fst = m) (antidiagonal n) = if m ≤ n then {(m, n - m)} else ∅ := by - ext ⟨x, y⟩ - simp only [mem_filter, Nat.mem_antidiagonal] - split_ifs with h - · simp (config := { contextual := true }) [and_comm, eq_tsub_iff_add_eq_of_le h, add_comm] - · rw [not_le] at h - simp only [not_mem_empty, iff_false_iff, not_and, decide_eq_true_eq] - exact fun hn => ne_of_lt (lt_of_le_of_lt (le_self_add.trans hn.le) h) -#align finset.nat.filter_fst_eq_antidiagonal Finset.Nat.filter_fst_eq_antidiagonal - -theorem filter_snd_eq_antidiagonal (n m : ℕ) : - filter (fun x : ℕ × ℕ ↦ x.snd = m) (antidiagonal n) = if m ≤ n then {(n - m, m)} else ∅ := by - have : (fun x : ℕ × ℕ ↦ (x.snd = m)) ∘ Prod.swap = fun x : ℕ × ℕ ↦ x.fst = m := by - ext; simp - rw [← map_swap_antidiagonal, filter_map] - simp [this, filter_fst_eq_antidiagonal, apply_ite (Finset.map _)] -#align finset.nat.filter_snd_eq_antidiagonal Finset.Nat.filter_snd_eq_antidiagonal - @[simp] lemma antidiagonal_filter_snd_le_of_le {n k : ℕ} (h : k ≤ n) : (antidiagonal n).filter (fun a ↦ a.snd ≤ k) = (antidiagonal k).map (Embedding.prodMap ⟨_, add_left_injective (n - k)⟩ (Embedding.refl ℕ)) := by @@ -192,26 +136,6 @@ theorem filter_snd_eq_antidiagonal (n m : ℕ) : ext ⟨i, j⟩ simpa using aux₂ i j -section EquivProd - -/-- The disjoint union of antidiagonals `Σ (n : ℕ), antidiagonal n` is equivalent to the product - `ℕ × ℕ`. This is such an equivalence, obtained by mapping `(n, (k, l))` to `(k, l)`. -/ -@[simps] -def sigmaAntidiagonalEquivProd : (Σn : ℕ, antidiagonal n) ≃ ℕ × ℕ where - toFun x := x.2 - invFun x := ⟨x.1 + x.2, x, mem_antidiagonal.mpr rfl⟩ - left_inv := by - rintro ⟨n, ⟨k, l⟩, h⟩ - rw [mem_antidiagonal] at h - exact Sigma.subtype_ext h rfl - right_inv x := rfl -#align finset.nat.sigma_antidiagonal_equiv_prod Finset.Nat.sigmaAntidiagonalEquivProd -#align finset.nat.sigma_antidiagonal_equiv_prod_symm_apply_fst Finset.Nat.sigmaAntidiagonalEquivProd_symm_apply_fst -#align finset.nat.sigma_antidiagonal_equiv_prod_symm_apply_snd_coe Finset.Nat.sigmaAntidiagonalEquivProd_symm_apply_snd_coe -#align finset.nat.sigma_antidiagonal_equiv_prod_apply Finset.Nat.sigmaAntidiagonalEquivProd_apply - -end EquivProd - /-- The set `antidiagonal n` is equivalent to `Fin (n+1)`, via the first projection. --/ @[simps] def antidiagonalEquivFin (n : ℕ) : antidiagonal n ≃ Fin (n + 1) where diff --git a/Mathlib/Data/Finset/NoncommProd.lean b/Mathlib/Data/Finset/NoncommProd.lean index 116b9e8901b6a..4d6d7a48fe920 100644 --- a/Mathlib/Data/Finset/NoncommProd.lean +++ b/Mathlib/Data/Finset/NoncommProd.lean @@ -3,9 +3,9 @@ Copyright (c) 2021 Yakov Pechersky. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yakov Pechersky -/ -import Mathlib.Data.Fintype.Card -import Mathlib.Algebra.Hom.Commute import Mathlib.Algebra.BigOperators.Basic +import Mathlib.Algebra.Group.Commute.Hom +import Mathlib.Data.Fintype.Card #align_import data.finset.noncomm_prod from "leanprover-community/mathlib"@"509de852e1de55e1efa8eacfa11df0823f26f226" @@ -175,6 +175,14 @@ theorem noncommProd_add (s t : Multiset α) (comm) : #align multiset.noncomm_prod_add Multiset.noncommProd_add #align multiset.noncomm_sum_add Multiset.noncommSum_add +@[to_additive] +lemma noncommProd_induction (s : Multiset α) (comm) + (p : α → Prop) (hom : ∀ a b, p a → p b → p (a * b)) (unit : p 1) (base : ∀ x ∈ s, p x) : + p (s.noncommProd comm) := by + induction' s using Quotient.inductionOn with l + simp only [quot_mk_to_coe, noncommProd_coe, mem_coe] at base ⊢ + exact l.prod_induction p hom unit base + @[to_additive] protected theorem noncommProd_map_aux [MonoidHomClass F α β] (s : Multiset α) (comm : { x | x ∈ s }.Pairwise Commute) (f : F) : { x | x ∈ s.map f }.Pairwise Commute := by @@ -218,6 +226,25 @@ theorem noncommProd_commute (s : Multiset α) (comm) (y : α) (h : ∀ x ∈ s, #align multiset.noncomm_prod_commute Multiset.noncommProd_commute #align multiset.noncomm_sum_add_commute Multiset.noncommSum_addCommute +theorem mul_noncommProd_erase [DecidableEq α] (s : Multiset α) {a : α} (h : a ∈ s) (comm) + (comm' := fun x hx y hy hxy ↦ comm (s.mem_of_mem_erase hx) (s.mem_of_mem_erase hy) hxy) : + a * (s.erase a).noncommProd comm' = s.noncommProd comm := by + induction' s using Quotient.inductionOn with l + simp only [quot_mk_to_coe, mem_coe, coe_erase, noncommProd_coe] at comm h ⊢ + suffices ∀ x ∈ l, ∀ y ∈ l, x * y = y * x by rw [List.prod_erase_of_comm h this] + intro x hx y hy + rcases eq_or_ne x y with rfl | hxy; rfl + exact comm hx hy hxy + +theorem noncommProd_erase_mul [DecidableEq α] (s : Multiset α) {a : α} (h : a ∈ s) (comm) + (comm' := fun x hx y hy hxy ↦ comm (s.mem_of_mem_erase hx) (s.mem_of_mem_erase hy) hxy) : + (s.erase a).noncommProd comm' * a = s.noncommProd comm := by + suffices ∀ b ∈ erase s a, Commute a b by + rw [← (noncommProd_commute (s.erase a) comm' a this).eq, mul_noncommProd_erase s h comm comm'] + intro b hb + rcases eq_or_ne a b with rfl | hab; rfl + exact comm h (mem_of_mem_erase hb) hab + end Multiset namespace Finset @@ -245,6 +272,14 @@ def noncommProd (s : Finset α) (f : α → β) #align finset.noncomm_prod Finset.noncommProd #align finset.noncomm_sum Finset.noncommSum +@[to_additive] +lemma noncommProd_induction (s : Finset α) (f : α → β) (comm) + (p : β → Prop) (hom : ∀ a b, p a → p b → p (a * b)) (unit : p 1) (base : ∀ x ∈ s, p (f x)) : + p (s.noncommProd f comm) := by + refine Multiset.noncommProd_induction _ _ _ hom unit fun b hb ↦ ?_ + obtain (⟨a, ha : a ∈ s, rfl : f a = b⟩) := by simpa using hb + exact base a ha + @[to_additive (attr := congr)] theorem noncommProd_congr {s₁ s₂ : Finset α} {f g : α → β} (h₁ : s₁ = s₂) (h₂ : ∀ x ∈ s₂, f x = g x) (comm) : @@ -342,6 +377,20 @@ theorem noncommProd_commute (s : Finset α) (f : α → β) (comm) (y : β) #align finset.noncomm_prod_commute Finset.noncommProd_commute #align finset.noncomm_sum_add_commute Finset.noncommSum_addCommute +theorem mul_noncommProd_erase [DecidableEq α] (s : Finset α) {a : α} (h : a ∈ s) (f : α → β) (comm) + (comm' := fun x hx y hy hxy ↦ comm (s.mem_of_mem_erase hx) (s.mem_of_mem_erase hy) hxy) : + f a * (s.erase a).noncommProd f comm' = s.noncommProd f comm := by + classical + simpa only [← Multiset.map_erase_of_mem _ _ h] using + Multiset.mul_noncommProd_erase (s.1.map f) (Multiset.mem_map_of_mem f h) _ + +theorem noncommProd_erase_mul [DecidableEq α] (s : Finset α) {a : α} (h : a ∈ s) (f : α → β) (comm) + (comm' := fun x hx y hy hxy ↦ comm (s.mem_of_mem_erase hx) (s.mem_of_mem_erase hy) hxy) : + (s.erase a).noncommProd f comm' * f a = s.noncommProd f comm := by + classical + simpa only [← Multiset.map_erase_of_mem _ _ h] using + Multiset.noncommProd_erase_mul (s.1.map f) (Multiset.mem_map_of_mem f h) _ + @[to_additive] theorem noncommProd_eq_prod {β : Type*} [CommMonoid β] (s : Finset α) (f : α → β) : (noncommProd s f fun _ _ _ _ _ => Commute.all _ _) = s.prod f := by diff --git a/Mathlib/Data/Finset/Option.lean b/Mathlib/Data/Finset/Option.lean index 69f58fa332fdd..453a2483a5e66 100644 --- a/Mathlib/Data/Finset/Option.lean +++ b/Mathlib/Data/Finset/Option.lean @@ -75,6 +75,10 @@ theorem mem_insertNone {s : Finset α} : ∀ {o : Option α}, o ∈ insertNone s theorem some_mem_insertNone {s : Finset α} {a : α} : some a ∈ insertNone s ↔ a ∈ s := by simp #align finset.some_mem_insert_none Finset.some_mem_insertNone +lemma none_mem_insertNone {s : Finset α} : none ∈ insertNone s := by simp + +lemma insertNone_nonempty {s : Finset α} : insertNone s |>.Nonempty := ⟨none, none_mem_insertNone⟩ + @[simp] theorem card_insertNone (s : Finset α) : s.insertNone.card = s.card + 1 := by simp [insertNone] #align finset.card_insert_none Finset.card_insertNone diff --git a/Mathlib/Data/Finset/Powerset.lean b/Mathlib/Data/Finset/Powerset.lean index a48c38a32b6a2..37696b7f88740 100644 --- a/Mathlib/Data/Finset/Powerset.lean +++ b/Mathlib/Data/Finset/Powerset.lean @@ -312,7 +312,7 @@ theorem powersetCard_sup [DecidableEq α] (u : Finset α) (n : ℕ) (hn : n < u. · intro x hx simp only [mem_biUnion, exists_prop, id.def] obtain ⟨t, ht⟩ : ∃ t, t ∈ powersetCard n (u.erase x) := powersetCard_nonempty - (le_trans (Nat.le_pred_of_lt hn) pred_card_le_card_erase) + (le_trans (Nat.le_sub_one_of_lt hn) pred_card_le_card_erase) · refine' ⟨insert x t, _, mem_insert_self _ _⟩ rw [← insert_erase hx, powersetCard_succ_insert (not_mem_erase _ _)] exact mem_union_right _ (mem_image_of_mem _ ht) diff --git a/Mathlib/Data/Finset/Sort.lean b/Mathlib/Data/Finset/Sort.lean index 9540952f0c14f..b83b95db7ae78 100644 --- a/Mathlib/Data/Finset/Sort.lean +++ b/Mathlib/Data/Finset/Sort.lean @@ -125,7 +125,7 @@ theorem sorted_last_eq_max'_aux (s : Finset α) · have : s.max' H ∈ l := (Finset.mem_sort (α := α) (· ≤ ·)).mpr (s.max'_mem H) obtain ⟨i, hi⟩ : ∃ i, l.get i = s.max' H := List.mem_iff_get.1 this rw [← hi] - exact (s.sort_sorted (· ≤ ·)).rel_nthLe_of_le _ _ (Nat.le_pred_of_lt i.prop) + exact (s.sort_sorted (· ≤ ·)).rel_nthLe_of_le _ _ (Nat.le_sub_one_of_lt i.prop) #align finset.sorted_last_eq_max'_aux Finset.sorted_last_eq_max'_aux theorem sorted_last_eq_max' {s : Finset α} diff --git a/Mathlib/Data/Finset/Sum.lean b/Mathlib/Data/Finset/Sum.lean index d852001b5a447..efe24a464cff7 100644 --- a/Mathlib/Data/Finset/Sum.lean +++ b/Mathlib/Data/Finset/Sum.lean @@ -6,7 +6,7 @@ Authors: Yaël Dillies import Mathlib.Data.Multiset.Sum import Mathlib.Data.Finset.Card -#align_import data.finset.sum from "leanprover-community/mathlib"@"9003f28797c0664a49e4179487267c494477d853" +#align_import data.finset.sum from "leanprover-community/mathlib"@"48a058d7e39a80ed56858505719a0b2197900999" /-! # Disjoint sum of finsets @@ -79,6 +79,10 @@ theorem inr_mem_disjSum : inr b ∈ s.disjSum t ↔ b ∈ t := Multiset.inr_mem_disjSum #align finset.inr_mem_disj_sum Finset.inr_mem_disjSum +@[simp] +theorem disjSum_eq_empty : s.disjSum t = ∅ ↔ s = ∅ ∧ t = ∅ := by simp [ext_iff] +#align finset.disj_sum_eq_empty Finset.disjSum_eq_empty + theorem disjSum_mono (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : s₁.disjSum t₁ ⊆ s₂.disjSum t₂ := val_le_iff.1 <| Multiset.disjSum_mono (val_le_iff.2 hs) (val_le_iff.2 ht) #align finset.disj_sum_mono Finset.disjSum_mono diff --git a/Mathlib/Data/Finset/Sups.lean b/Mathlib/Data/Finset/Sups.lean index c0418dfa04c4a..c4704148c0923 100644 --- a/Mathlib/Data/Finset/Sups.lean +++ b/Mathlib/Data/Finset/Sups.lean @@ -7,7 +7,7 @@ import Mathlib.Data.Finset.NAry import Mathlib.Data.Finset.Slice import Mathlib.Data.Set.Sups -#align_import data.finset.sups from "leanprover-community/mathlib"@"20715f4ac6819ef2453d9e5106ecd086a5dc2a5e" +#align_import data.finset.sups from "leanprover-community/mathlib"@"8818fdefc78642a7e6afcd20be5c184f3c7d9699" /-! # Set family operations @@ -37,6 +37,9 @@ We define the following notation in locale `FinsetFamily`: [B. Bollobás, *Combinatorics*][bollobas1986] -/ +#align finset.decidable_pred_mem_upper_closure instDecidablePredMemUpperClosure +#align finset.decidable_pred_mem_lower_closure instDecidablePredMemLowerClosure + open Function open SetFamily @@ -186,6 +189,8 @@ lemma subset_sups_self : s ⊆ s ⊻ s := fun _a ha ↦ mem_sups.2 ⟨_, ha, _, lemma sups_subset_self : s ⊻ s ⊆ s ↔ SupClosed (s : Set α) := sups_subset_iff @[simp] lemma sups_eq_self : s ⊻ s = s ↔ SupClosed (s : Set α) := by simp [←coe_inj] +@[simp] lemma univ_sups_univ [Fintype α] : (univ : Finset α) ⊻ univ = univ := by simp + lemma filter_sups_le [@DecidableRel α (· ≤ ·)] (s t : Finset α) (a : α) : (s ⊻ t).filter (· ≤ a) = s.filter (· ≤ a) ⊻ t.filter (· ≤ a) := by simp only [←coe_inj, coe_filter, coe_sups, ←mem_coe, Set.sep_sups_le] @@ -223,6 +228,8 @@ theorem sups_sups_sups_comm : s ⊻ t ⊻ (u ⊻ v) = s ⊻ u ⊻ (t ⊻ v) := image₂_image₂_image₂_comm sup_sup_sup_comm #align finset.sups_sups_sups_comm Finset.sups_sups_sups_comm +#align finset.filter_sups_le Finset.filter_sups_le + end Sups section Infs @@ -366,6 +373,8 @@ lemma subset_infs_self : s ⊆ s ⊼ s := fun _a ha ↦ mem_infs.2 ⟨_, ha, _, lemma infs_self_subset : s ⊼ s ⊆ s ↔ InfClosed (s : Set α) := infs_subset_iff @[simp] lemma infs_self : s ⊼ s = s ↔ InfClosed (s : Set α) := by simp [←coe_inj] +@[simp] lemma univ_infs_univ [Fintype α] : (univ : Finset α) ⊼ univ = univ := by simp + lemma filter_infs_le [@DecidableRel α (· ≤ ·)] (s t : Finset α) (a : α) : (s ⊼ t).filter (a ≤ ·) = s.filter (a ≤ ·) ⊼ t.filter (a ≤ ·) := by simp only [←coe_inj, coe_filter, coe_infs, ←mem_coe, Set.sep_infs_le] @@ -403,26 +412,12 @@ theorem infs_infs_infs_comm : s ⊼ t ⊼ (u ⊼ v) = s ⊼ u ⊼ (t ⊼ v) := image₂_image₂_image₂_comm inf_inf_inf_comm #align finset.infs_infs_infs_comm Finset.infs_infs_infs_comm +#align finset.filter_infs_ge Finset.filter_infs_le + end Infs open FinsetFamily -@[simp] lemma powerset_union (s t : Finset α) : (s ∪ t).powerset = s.powerset ⊻ t.powerset := by - ext u - simp only [mem_sups, mem_powerset, le_eq_subset, sup_eq_union] - refine ⟨fun h ↦ ⟨_, inter_subset_left _ u, _, inter_subset_left _ u, ?_⟩, ?_⟩ - · rwa [←inter_distrib_right, inter_eq_right] - · rintro ⟨v, hv, w, hw, rfl⟩ - exact union_subset_union hv hw - -@[simp] lemma powerset_inter (s t : Finset α) : (s ∩ t).powerset = s.powerset ⊼ t.powerset := by - ext u - simp only [mem_infs, mem_powerset, le_eq_subset, inf_eq_inter] - refine ⟨fun h ↦ ⟨_, inter_subset_left _ u, _, inter_subset_left _ u, ?_⟩, ?_⟩ - · rwa [←inter_inter_distrib_right, inter_eq_right] - · rintro ⟨v, hv, w, hw, rfl⟩ - exact inter_subset_inter hv hw - section DistribLattice variable [DistribLattice α] (s t u : Finset α) @@ -445,6 +440,36 @@ theorem infs_sups_subset_right : (t ⊻ u) ⊼ s ⊆ t ⊼ s ⊻ u ⊼ s := end DistribLattice +section Finset +variable {𝒜 ℬ : Finset (Finset α)} {s t : Finset α} {a : α} + +@[simp] lemma powerset_union (s t : Finset α) : (s ∪ t).powerset = s.powerset ⊻ t.powerset := by + ext u + simp only [mem_sups, mem_powerset, le_eq_subset, sup_eq_union] + refine ⟨fun h ↦ ⟨_, inter_subset_left _ u, _, inter_subset_left _ u, ?_⟩, ?_⟩ + · rwa [←inter_distrib_right, inter_eq_right] + · rintro ⟨v, hv, w, hw, rfl⟩ + exact union_subset_union hv hw + +@[simp] lemma powerset_inter (s t : Finset α) : (s ∩ t).powerset = s.powerset ⊼ t.powerset := by + ext u + simp only [mem_infs, mem_powerset, le_eq_subset, inf_eq_inter] + refine ⟨fun h ↦ ⟨_, inter_subset_left _ u, _, inter_subset_left _ u, ?_⟩, ?_⟩ + · rwa [←inter_inter_distrib_right, inter_eq_right] + · rintro ⟨v, hv, w, hw, rfl⟩ + exact inter_subset_inter hv hw + +@[simp] lemma powerset_sups_powerset_self (s : Finset α) : + s.powerset ⊻ s.powerset = s.powerset := by simp [←powerset_union] + +@[simp] lemma powerset_infs_powerset_self (s : Finset α) : + s.powerset ⊼ s.powerset = s.powerset := by simp [←powerset_inter] + +lemma union_mem_sups : s ∈ 𝒜 → t ∈ ℬ → s ∪ t ∈ 𝒜 ⊻ ℬ := sup_mem_sups +lemma inter_mem_infs : s ∈ 𝒜 → t ∈ ℬ → s ∩ t ∈ 𝒜 ⊼ ℬ := inf_mem_infs + +end Finset + section DisjSups variable [SemilatticeSup α] [OrderBot α] [@DecidableRel α Disjoint] (s s₁ s₂ t t₁ t₂ u : Finset α) diff --git a/Mathlib/Data/Finset/Sym.lean b/Mathlib/Data/Finset/Sym.lean index ebc3314b27923..52e7636af293d 100644 --- a/Mathlib/Data/Finset/Sym.lean +++ b/Mathlib/Data/Finset/Sym.lean @@ -100,7 +100,7 @@ theorem sym2_singleton (a : α) : ({a} : Finset α).sym2 = {Sym2.diag a} := by @[simp] theorem diag_mem_sym2_mem_iff : (∀ b, b ∈ Sym2.diag a → b ∈ s) ↔ a ∈ s := by rw [← mem_sym2_iff] - exact mk'_mem_sym2_iff.trans <| and_self_iff _ + exact mk'_mem_sym2_iff.trans <| and_self_iff theorem diag_mem_sym2_iff : Sym2.diag a ∈ s.sym2 ↔ a ∈ s := by simp [diag_mem_sym2_mem_iff] #align finset.diag_mem_sym2_iff Finset.diag_mem_sym2_iff diff --git a/Mathlib/Data/Finsupp/AList.lean b/Mathlib/Data/Finsupp/AList.lean index e82a8b6a4aa77..98ed6de508117 100644 --- a/Mathlib/Data/Finsupp/AList.lean +++ b/Mathlib/Data/Finsupp/AList.lean @@ -76,17 +76,15 @@ noncomputable def lookupFinsupp (l : AList fun _x : α => M) : α →₀ M @[simp] theorem lookupFinsupp_apply [DecidableEq α] (l : AList fun _x : α => M) (a : α) : l.lookupFinsupp a = (l.lookup a).getD 0 := by - -- porting note: was `convert rfl` - simp only [lookupFinsupp, ne_eq, Finsupp.coe_mk]; congr + convert rfl; congr #align alist.lookup_finsupp_apply AList.lookupFinsupp_apply @[simp] theorem lookupFinsupp_support [DecidableEq α] [DecidableEq M] (l : AList fun _x : α => M) : l.lookupFinsupp.support = (l.1.filter fun x => Sigma.snd x ≠ 0).keys.toFinset := by - -- porting note: was `convert rfl` - simp only [lookupFinsupp, ne_eq, Finsupp.coe_mk]; congr - · apply Subsingleton.elim - · funext; congr + convert rfl; congr + · apply Subsingleton.elim + · funext; congr #align alist.lookup_finsupp_support AList.lookupFinsupp_support theorem lookupFinsupp_eq_iff_of_ne_zero [DecidableEq α] {l : AList fun _x : α => M} {a : α} {x : M} diff --git a/Mathlib/Data/Finsupp/Antidiagonal.lean b/Mathlib/Data/Finsupp/Antidiagonal.lean index af8de7d4977d6..c3a9c042631c3 100644 --- a/Mathlib/Data/Finsupp/Antidiagonal.lean +++ b/Mathlib/Data/Finsupp/Antidiagonal.lean @@ -17,8 +17,6 @@ The antidiagonal of `s : α →₀ ℕ` consists of all pairs `(t₁, t₂) : (α →₀ ℕ) × (α →₀ ℕ)` such that `t₁ + t₂ = s`. -/ - - open BigOperators namespace Finsupp @@ -39,53 +37,18 @@ def antidiagonal' (f : α →₀ ℕ) : (α →₀ ℕ) × (α →₀ ℕ) → /-- The antidiagonal of `s : α →₀ ℕ` is the finset of all pairs `(t₁, t₂) : (α →₀ ℕ) × (α →₀ ℕ)` such that `t₁ + t₂ = s`. -/ -def antidiagonal (f : α →₀ ℕ) : Finset ((α →₀ ℕ) × (α →₀ ℕ)) := f.antidiagonal'.support -#align finsupp.antidiagonal Finsupp.antidiagonal - -@[simp] -theorem mem_antidiagonal {f : α →₀ ℕ} {p : (α →₀ ℕ) × (α →₀ ℕ)} : - p ∈ antidiagonal f ↔ p.1 + p.2 = f := by - rcases p with ⟨p₁, p₂⟩ - simp [antidiagonal, antidiagonal', ← and_assoc, Multiset.toFinsupp_eq_iff, +instance instHasAntidiagonal : HasAntidiagonal (α →₀ ℕ) where + antidiagonal f := f.antidiagonal'.support + mem_antidiagonal {f} {p} := by + rcases p with ⟨p₁, p₂⟩ + simp [antidiagonal', ← and_assoc, Multiset.toFinsupp_eq_iff, ← Multiset.toFinsupp_eq_iff (f := f)] -#align finsupp.mem_antidiagonal Finsupp.mem_antidiagonal -theorem swap_mem_antidiagonal {n : α →₀ ℕ} {f : (α →₀ ℕ) × (α →₀ ℕ)} : - f.swap ∈ antidiagonal n ↔ f ∈ antidiagonal n := by - simp only [mem_antidiagonal, add_comm, Prod.swap] -#align finsupp.swap_mem_antidiagonal Finsupp.swap_mem_antidiagonal +#align finsupp.antidiagonal_filter_fst_eq Finset.filter_fst_eq_antidiagonal +#align finsupp.antidiagonal_filter_snd_eq Finset.filter_snd_eq_antidiagonal -theorem antidiagonal_filter_fst_eq (f g : α →₀ ℕ) - [D : ∀ p : (α →₀ ℕ) × (α →₀ ℕ), Decidable (p.1 = g)] : - ((antidiagonal f).filter fun p ↦ p.1 = g) = if g ≤ f then {(g, f - g)} else ∅ := by - ext ⟨a, b⟩ - suffices a = g → (a + b = f ↔ g ≤ f ∧ b = f - g) by - simpa [apply_ite (fun f ↦ (a, b) ∈ f), mem_filter, mem_antidiagonal, mem_singleton, - Prod.mk.inj_iff, ← and_assoc, @and_right_comm _ (a = _), and_congr_left_iff] - rintro rfl - constructor - · rintro rfl - exact ⟨le_add_right le_rfl, (add_tsub_cancel_left _ _).symm⟩ - · rintro ⟨h, rfl⟩ - exact add_tsub_cancel_of_le h -#align finsupp.antidiagonal_filter_fst_eq Finsupp.antidiagonal_filter_fst_eq - -theorem antidiagonal_filter_snd_eq (f g : α →₀ ℕ) - [D : ∀ p : (α →₀ ℕ) × (α →₀ ℕ), Decidable (p.2 = g)] : - ((antidiagonal f).filter fun p ↦ p.2 = g) = if g ≤ f then {(f - g, g)} else ∅ := by - ext ⟨a, b⟩ - suffices b = g → (a + b = f ↔ g ≤ f ∧ a = f - g) by - simpa [apply_ite (fun f ↦ (a, b) ∈ f), mem_filter, mem_antidiagonal, mem_singleton, - Prod.mk.inj_iff, ← and_assoc, and_congr_left_iff] - rintro rfl - constructor - · rintro rfl - exact ⟨le_add_left le_rfl, (add_tsub_cancel_right _ _).symm⟩ - · rintro ⟨h, rfl⟩ - exact tsub_add_cancel_of_le h -#align finsupp.antidiagonal_filter_snd_eq Finsupp.antidiagonal_filter_snd_eq - -@[simp] +-- nolint as this is for dsimp +@[simp, nolint simpNF] theorem antidiagonal_zero : antidiagonal (0 : α →₀ ℕ) = singleton (0, 0) := rfl #align finsupp.antidiagonal_zero Finsupp.antidiagonal_zero @@ -101,10 +64,10 @@ theorem prod_antidiagonal_swap {M : Type*} [CommMonoid M] (n : α →₀ ℕ) @[simp] theorem antidiagonal_single (a : α) (n : ℕ) : - antidiagonal (single a n) = (Finset.Nat.antidiagonal n).map + antidiagonal (single a n) = (antidiagonal n).map (Function.Embedding.prodMap ⟨_, single_injective a⟩ ⟨_, single_injective a⟩) := by ext ⟨x, y⟩ - simp only [mem_antidiagonal, mem_map, Nat.mem_antidiagonal, Function.Embedding.coe_prodMap, + simp only [mem_antidiagonal, mem_map, mem_antidiagonal, Function.Embedding.coe_prodMap, Function.Embedding.coeFn_mk, Prod_map, Prod.mk.injEq, Prod.exists] constructor · intro h diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index d2136cd1b196e..f26d195b77fa5 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Scott Morrison -/ import Mathlib.Algebra.BigOperators.Finsupp -import Mathlib.Algebra.Hom.GroupAction import Mathlib.Algebra.Regular.SMul import Mathlib.Data.Finset.Preimage import Mathlib.Data.Finsupp.Notation import Mathlib.Data.Rat.BigOperators import Mathlib.Data.Set.Countable +import Mathlib.GroupTheory.GroupAction.Hom #align_import data.finsupp.basic from "leanprover-community/mathlib"@"f69db8cecc668e2d5894d7e9bfc491da60db3b9f" diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 1e0c9ef129aa3..7c5010e102a41 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -835,7 +835,7 @@ def embDomain (f : α ↪ β) (v : α →₀ M) : β →₀ M where · simp only [h, true_iff_iff, Ne.def] rw [← not_mem_support_iff, not_not] classical apply Finset.choose_mem - · simp only [h, Ne.def, ne_self_iff_false] + · simp only [h, Ne.def, ne_self_iff_false, not_true_eq_false] #align finsupp.emb_domain Finsupp.embDomain @[simp] diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 8c9a53aecdb52..bee30a7edd088 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -1217,10 +1217,10 @@ end Trunc namespace Multiset -variable [Fintype α] [DecidableEq α] [Fintype β] +variable [Fintype α] [Fintype β] @[simp] -theorem count_univ (a : α) : count a Finset.univ.val = 1 := +theorem count_univ [DecidableEq α] (a : α) : count a Finset.univ.val = 1 := count_eq_one_of_mem Finset.univ.nodup (Finset.mem_univ _) #align multiset.count_univ Multiset.count_univ @@ -1229,6 +1229,15 @@ theorem map_univ_val_equiv (e : α ≃ β) : map e univ.val = univ.val := by rw [←congr_arg Finset.val (Finset.map_univ_equiv e), Finset.map_val, Equiv.coe_toEmbedding] +/-- For functions on finite sets, they are bijections iff they map universes into universes. -/ +@[simp] +theorem bijective_iff_map_univ_eq_univ (f : α → β) : + f.Bijective ↔ map f (Finset.univ : Finset α).val = univ.val := + ⟨fun bij ↦ congr_arg (·.val) (map_univ_equiv <| Equiv.ofBijective f bij), + fun eq ↦ ⟨ + fun a₁ a₂ ↦ inj_on_of_nodup_map (eq.symm ▸ univ.nodup) _ (mem_univ a₁) _ (mem_univ a₂), + fun b ↦ have ⟨a, _, h⟩ := mem_map.mp (eq.symm ▸ mem_univ_val b); ⟨a, h⟩⟩⟩ + end Multiset /-- Auxiliary definition to show `exists_seq_of_forall_finset_exists`. -/ diff --git a/Mathlib/Data/Fintype/BigOperators.lean b/Mathlib/Data/Fintype/BigOperators.lean index 8564dc3be013d..7466691f6c9a2 100644 --- a/Mathlib/Data/Fintype/BigOperators.lean +++ b/Mathlib/Data/Fintype/BigOperators.lean @@ -7,6 +7,7 @@ import Mathlib.Data.Fintype.Option import Mathlib.Data.Fintype.Powerset import Mathlib.Data.Fintype.Sigma import Mathlib.Data.Fintype.Sum +import Mathlib.Data.Fintype.Prod import Mathlib.Data.Fintype.Vector import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.BigOperators.Option @@ -295,4 +296,26 @@ theorem Fintype.prod_sum_type (f : Sum α₁ α₂ → M) : #align fintype.prod_sum_type Fintype.prod_sum_type #align fintype.sum_sum_type Fintype.sum_sum_type +@[to_additive (attr := simp) Fintype.sum_prod_type] +theorem Fintype.prod_prod_type [CommMonoid γ] {f : α₁ × α₂ → γ} : + ∏ x, f x = ∏ x, ∏ y, f (x, y) := + Finset.prod_product + +/-- An uncurried version of `Finset.prod_prod_type`. -/ +@[to_additive Fintype.sum_prod_type' "An uncurried version of `Finset.sum_prod_type`"] +theorem Fintype.prod_prod_type' [CommMonoid γ] {f : α₁ → α₂ → γ} : + ∏ x : α₁ × α₂, f x.1 x.2 = ∏ x, ∏ y, f x y := + Finset.prod_product' + +@[to_additive Fintype.sum_prod_type_right] +theorem Fintype.prod_prod_type_right [CommMonoid γ] {f : α₁ × α₂ → γ} : + ∏ x, f x = ∏ y, ∏ x, f (x, y) := + Finset.prod_product_right + +/-- An uncurried version of `Finset.prod_prod_type_right`. -/ +@[to_additive Fintype.sum_prod_type_right' "An uncurried version of `Finset.sum_prod_type_right`"] +theorem Fintype.prod_prod_type_right' [CommMonoid γ] {f : α₁ → α₂ → γ} : + ∏ x : α₁ × α₂, f x.1 x.2 = ∏ y, ∏ x, f x y := + Finset.prod_product_right' + end diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index a118a7098ddce..0a6336f1463cf 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -1203,6 +1203,58 @@ theorem not_surjective_finite_infinite {α β} [Finite α] [Infinite β] (f : α fun hf => (Infinite.of_surjective f hf).not_finite ‹_› #align not_surjective_finite_infinite not_surjective_finite_infinite +section Ranges + +/-- For any `c : List ℕ` whose sum is at most `Fintype.card α`, + we can find `o : List (List α)` whose members have no duplicate, + whose lengths given by `c`, and which are pairwise disjoint -/ +theorem List.exists_pw_disjoint_with_card {α : Type*} [Fintype α] + {c : List ℕ} (hc : c.sum ≤ Fintype.card α) : + ∃ o : List (List α), + o.map length = c ∧ (∀ s ∈ o, s.Nodup) ∧ Pairwise List.Disjoint o := by + let klift (n : ℕ) (hn : n < Fintype.card α) : Fin (Fintype.card α) := + (⟨n, hn⟩ : Fin (Fintype.card α)) + let klift' (l : List ℕ) (hl : ∀ a ∈ l, a < Fintype.card α) : + List (Fin (Fintype.card α)) := List.pmap klift l hl + have hc'_lt : ∀ l ∈ c.ranges, ∀ n ∈ l, n < Fintype.card α := by + intro l hl n hn + apply lt_of_lt_of_le _ hc + rw [← mem_mem_ranges_iff_lt_sum] + exact ⟨l, hl, hn⟩ + let l := (ranges c).pmap klift' hc'_lt + have hl : ∀ (a : List ℕ) (ha : a ∈ c.ranges), + (klift' a (hc'_lt a ha)).map Fin.valEmbedding = a := by + intro a ha + conv_rhs => rw [← List.map_id a] + rw [List.map_pmap] + simp only [Fin.valEmbedding_apply, Fin.val_mk, List.pmap_eq_map, List.map_id'', List.map_id] + use l.map (List.map (Fintype.equivFin α).symm) + constructor + · -- length + rw [← ranges_length c] + simp only [map_map, map_pmap, Function.comp_apply, length_map, length_pmap, pmap_eq_map] + constructor + · -- nodup + intro s + rw [mem_map] + rintro ⟨t, ht, rfl⟩ + apply Nodup.map (Equiv.injective _) + obtain ⟨u, hu, rfl⟩ := mem_pmap.mp ht + apply Nodup.of_map + rw [hl u hu] + exact ranges_nodup hu + · -- pairwise disjoint + refine Pairwise.map _ (fun s t ↦ disjoint_map (Equiv.injective _)) ?_ + · -- List.Pairwise List.disjoint l + apply Pairwise.pmap (List.ranges_disjoint c) + intro u hu v hv huv + apply disjoint_pmap + · intro a a' ha ha' h + simpa only [Fin.mk_eq_mk] using h + exact huv + +end Ranges + section Trunc /-- A `Fintype` with positive cardinality constructively contains an element. diff --git a/Mathlib/Data/Int/Bitwise.lean b/Mathlib/Data/Int/Bitwise.lean index 5df6f41d88edc..b39481fe6a610 100644 --- a/Mathlib/Data/Int/Bitwise.lean +++ b/Mathlib/Data/Int/Bitwise.lean @@ -53,7 +53,7 @@ theorem bodd_subNatNat (m n : ℕ) : bodd (subNatNat m n) = xor m.bodd n.bodd := @[simp] theorem bodd_negOfNat (n : ℕ) : bodd (negOfNat n) = n.bodd := by - cases n <;> simp + cases n <;> simp (config := {decide := true}) rfl #align int.bodd_neg_of_nat Int.bodd_negOfNat @@ -164,7 +164,7 @@ theorem bit_negSucc (b) (n : ℕ) : bit b -[n+1] = -[Nat.bit (not b) n+1] := by @[simp] theorem bodd_bit (b n) : bodd (bit b n) = b := by rw [bit_val] - cases b <;> cases bodd n <;> simp + cases b <;> cases bodd n <;> simp [(show bodd 2 = false by rfl)] #align int.bodd_bit Int.bodd_bit @[simp, deprecated] @@ -293,12 +293,12 @@ theorem bitwise_xor : bitwise xor = Int.xor := by theorem bitwise_bit (f : Bool → Bool → Bool) (a m b n) : bitwise f (bit a m) (bit b n) = bit (f a b) (bitwise f m n) := by cases' m with m m <;> cases' n with n n <;> - simp only [bitwise, ofNat_eq_coe, bit_coe_nat, natBitwise, Bool.not_false, Bool.not_eq_false', + simp [bitwise, ofNat_eq_coe, bit_coe_nat, natBitwise, Bool.not_false, Bool.not_eq_false', bit_negSucc] - · by_cases h : f false false <;> simp [h] - · by_cases h : f false true <;> simp [h] - · by_cases h : f true false <;> simp [h] - · by_cases h : f true true <;> simp [h] + · by_cases h : f false false <;> simp (config := {decide := true}) [h] + · by_cases h : f false true <;> simp (config := {decide := true}) [h] + · by_cases h : f true false <;> simp (config := {decide := true}) [h] + · by_cases h : f true true <;> simp (config := {decide := true}) [h] #align int.bitwise_bit Int.bitwise_bit @[simp] diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index 19587dc8509e2..e5f87b4c0d47d 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -3,10 +3,10 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Algebra.Ring.Hom.Basic import Mathlib.Data.Int.Order.Basic import Mathlib.Data.Nat.Cast.Commute import Mathlib.Data.Nat.Cast.Order -import Mathlib.Algebra.Hom.Ring.Basic #align_import data.int.cast.lemmas from "leanprover-community/mathlib"@"acebd8d49928f6ed8920e502a6c90674e75bd441" diff --git a/Mathlib/Data/Int/Lemmas.lean b/Mathlib/Data/Int/Lemmas.lean index 55d5ca346a1af..5fa5f5943d0b9 100644 --- a/Mathlib/Data/Int/Lemmas.lean +++ b/Mathlib/Data/Int/Lemmas.lean @@ -122,7 +122,7 @@ attribute [local simp] Int.zero_div theorem div2_bit (b n) : div2 (bit b n) = n := by rw [bit_val, div2_val, add_comm, Int.add_mul_ediv_left, (_ : (_ / 2 : ℤ) = 0), zero_add] cases b - · simp + · decide · show ofNat _ = _ rw [Nat.div_eq_of_lt] <;> simp · decide diff --git a/Mathlib/Data/Int/ModEq.lean b/Mathlib/Data/Int/ModEq.lean index a446a520cd44c..b0b3171818481 100644 --- a/Mathlib/Data/Int/ModEq.lean +++ b/Mathlib/Data/Int/ModEq.lean @@ -129,7 +129,7 @@ protected theorem mul_left' (h : a ≡ b [ZMOD n]) : c * a ≡ c * b [ZMOD c * n obtain hc | rfl | hc := lt_trichotomy c 0 · rw [← neg_modEq_neg, ← modEq_neg, ← neg_mul, ← neg_mul, ← neg_mul] simp only [ModEq, mul_emod_mul_of_pos _ _ (neg_pos.2 hc), h.eq] - · simp + · simp only [zero_mul]; rfl · simp only [ModEq, mul_emod_mul_of_pos _ _ hc, h.eq] #align int.modeq.mul_left' Int.ModEq.mul_left' diff --git a/Mathlib/Data/Int/Parity.lean b/Mathlib/Data/Int/Parity.lean index 6965c6a9c0470..d1961d01dd35f 100644 --- a/Mathlib/Data/Int/Parity.lean +++ b/Mathlib/Data/Int/Parity.lean @@ -27,6 +27,9 @@ theorem emod_two_ne_one : ¬n % 2 = 1 ↔ n % 2 = 0 := by cases' emod_two_eq_zero_or_one n with h h <;> simp [h] #align int.mod_two_ne_one Int.emod_two_ne_one +@[simp] +theorem one_emod_two : (1 : Int) % 2 = 1 := rfl + -- `EuclideanDomain.mod_eq_zero` uses (2 ∣ n) as normal form @[local simp] theorem emod_two_ne_zero : ¬n % 2 = 0 ↔ n % 2 = 1 := by @@ -102,6 +105,7 @@ theorem even_add : Even (m + n) ↔ (Even m ↔ Even n) := by cases' emod_two_eq_zero_or_one m with h₁ h₁ <;> cases' emod_two_eq_zero_or_one n with h₂ h₂ <;> simp [even_iff, h₁, h₂, Int.add_emod] + rfl #align int.even_add Int.even_add theorem even_add' : Even (m + n) ↔ (Odd m ↔ Odd n) := by @@ -279,8 +283,8 @@ theorem two_mul_ediv_two_of_odd (h : Odd n) : 2 * (n / 2) = n - 1 := -- Here are examples of how `parity_simps` can be used with `Int`. example (m n : ℤ) (h : Even m) : ¬Even (n + 3) ↔ Even (m ^ 2 + m + n) := by - simp [*, (by decide : ¬2 = 0), parity_simps] + simp (config := {decide := true}) [*, (by decide : ¬2 = 0), parity_simps] -example : ¬Even (25394535 : ℤ) := by simp +example : ¬Even (25394535 : ℤ) := by decide end Int diff --git a/Mathlib/Data/Int/Range.lean b/Mathlib/Data/Int/Range.lean index 91905365d8e9e..cc7e2b9ac47e1 100644 --- a/Mathlib/Data/Int/Range.lean +++ b/Mathlib/Data/Int/Range.lean @@ -49,7 +49,7 @@ instance decidableLELE (P : Int → Prop) [DecidablePred P] (m n : ℤ) : apply decidable_of_iff (∀ r ∈ range m (n + 1), P r) apply Iff.intro <;> intros h _ _ · intro _; apply h - simp_all only [mem_range_iff, and_imp, lt_add_one_iff] + simp_all only [mem_range_iff, and_imp, and_self, lt_add_one_iff] · simp_all only [mem_range_iff, and_imp, lt_add_one_iff] #align int.decidable_le_le Int.decidableLELE diff --git a/Mathlib/Data/Int/Sqrt.lean b/Mathlib/Data/Int/Sqrt.lean index 84cd4db1bf7d0..0178909058a11 100644 --- a/Mathlib/Data/Int/Sqrt.lean +++ b/Mathlib/Data/Int/Sqrt.lean @@ -37,4 +37,9 @@ theorem sqrt_nonneg (n : ℤ) : 0 ≤ sqrt n := coe_nat_nonneg _ #align int.sqrt_nonneg Int.sqrt_nonneg +/-- `IsSquare` can be decided on `ℤ` by checking against the square root. -/ +instance : DecidablePred (IsSquare : ℤ → Prop) := + fun m => decidable_of_iff' (sqrt m * sqrt m = m) <| by + simp_rw [←exists_mul_self m, IsSquare, eq_comm] + end Int diff --git a/Mathlib/Data/Int/Units.lean b/Mathlib/Data/Int/Units.lean index 8667113994db0..5a69a99b1033f 100644 --- a/Mathlib/Data/Int/Units.lean +++ b/Mathlib/Data/Int/Units.lean @@ -122,7 +122,7 @@ theorem isUnit_add_isUnit_eq_isUnit_add_isUnit {a b c d : ℤ} (ha : IsUnit a) ( rw [isUnit_iff] at ha hb hc hd cases ha <;> cases hb <;> cases hc <;> cases hd <;> subst a <;> subst b <;> subst c <;> subst d <;> - simp + simp (config := {decide := true}) #align int.is_unit_add_is_unit_eq_is_unit_add_is_unit Int.isUnit_add_isUnit_eq_isUnit_add_isUnit theorem eq_one_or_neg_one_of_mul_eq_neg_one {z w : ℤ} (h : z * w = -1) : z = 1 ∨ z = -1 := diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 5277fce521e84..35988ef5c14d4 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -237,7 +237,7 @@ instance instSingletonList : Singleton α (List α) := ⟨fun x => [x]⟩ instance [DecidableEq α] : Insert α (List α) := ⟨List.insert⟩ -- ADHOC Porting note: instance from Lean3 core -instance [DecidableEq α]: IsLawfulSingleton α (List α) := +instance [DecidableEq α] : IsLawfulSingleton α (List α) := { insert_emptyc_eq := fun x => show (if x ∈ ([] : List α) then [] else [x]) = [x] from if_neg (not_mem_nil _) } @@ -1888,13 +1888,13 @@ theorem take_take : ∀ (n m) (l : List α), take n (take m l) = take (min n m) | 0, m, l => by rw [zero_min, take_zero, take_zero] | succ n, succ m, nil => by simp only [take_nil] | succ n, succ m, a :: l => by - simp only [take, min_succ_succ, take_take n m l] + simp only [take, succ_min_succ, take_take n m l] #align list.take_take List.take_take theorem take_replicate (a : α) : ∀ n m : ℕ, take n (replicate m a) = replicate (min n m) a | n, 0 => by simp | 0, m => by simp - | succ n, succ m => by simp [min_succ_succ, take_replicate] + | succ n, succ m => by simp [succ_min_succ, take_replicate] #align list.take_replicate List.take_replicate theorem map_take {α β : Type*} (f : α → β) : @@ -1986,7 +1986,7 @@ theorem take_eq_take : | _ :: xs, 0, 0 => by simp | x :: xs, m + 1, 0 => by simp | x :: xs, 0, n + 1 => by simp [@eq_comm ℕ 0] - | x :: xs, m + 1, n + 1 => by simp [Nat.min_succ_succ, take_eq_take] + | x :: xs, m + 1, n + 1 => by simp [Nat.succ_min_succ, take_eq_take] #align list.take_eq_take List.take_eq_take theorem take_add (l : List α) (m n : ℕ) : l.take (m + n) = l.take m ++ (l.drop m).take n := by @@ -2576,7 +2576,7 @@ theorem nthLe_succ_scanl {i : ℕ} {h : i + 1 < (scanl f b l).length} : induction i generalizing b l with | zero => cases l - · simp only [length, zero_add, scanl_nil] at h + · simp only [length, zero_eq, lt_self_iff_false] at h · simp [scanl_cons, singleton_append, nthLe_zero_scanl, nthLe_cons] | succ i hi => cases l @@ -3331,7 +3331,7 @@ theorem reduceOption_map {l : List (Option α)} {f : α → β} : reduceOption (map (Option.map f) l) = map f (reduceOption l) := by induction' l with hd tl hl · simp only [reduceOption_nil, map_nil] - ·cases hd <;> + · cases hd <;> simpa [true_and_iff, Option.map_some', map, eq_self_iff_true, reduceOption_cons_of_some] using hl #align list.reduce_option_map List.reduceOption_map @@ -3343,7 +3343,7 @@ theorem reduceOption_append (l l' : List (Option α)) : theorem reduceOption_length_le (l : List (Option α)) : l.reduceOption.length ≤ l.length := by induction' l with hd tl hl - · simp only [reduceOption_nil, length] + · simp [reduceOption_nil, length] · cases hd · exact Nat.le_succ_of_le hl · simpa only [length, add_le_add_iff_right, reduceOption_cons_of_some] using hl @@ -4422,4 +4422,36 @@ theorem getI_zero_eq_headI : l.getI 0 = l.headI := by cases l <;> rfl end getI +section Disjoint + +variable {α β : Type*} + +/-- The images of disjoint maps under a map are disjoint -/ +theorem disjoint_map {f : α → β} {s t : List α} (hf : Function.Injective f) + (h : Disjoint s t) : Disjoint (s.map f) (t.map f) := by + simp only [Disjoint] + intro b hbs hbt + rw [mem_map] at hbs hbt + obtain ⟨a, ha, rfl⟩ := hbs + apply h ha + obtain ⟨a', ha', ha''⟩ := hbt + rw [hf ha''.symm]; exact ha' + +/-- The images of disjoint lists under a partially defined map are disjoint -/ +theorem disjoint_pmap {p : α → Prop} {f : ∀ a : α, p a → β} {s t : List α} + (hs : ∀ a ∈ s, p a) (ht : ∀ a ∈ t, p a) + (hf : ∀ (a a' : α) (ha : p a) (ha' : p a'), f a ha = f a' ha' → a = a') + (h : Disjoint s t) : + Disjoint (s.pmap f hs) (t.pmap f ht) := by + simp only [Disjoint] + intro b hbs hbt + rw [mem_pmap] at hbs hbt + obtain ⟨a, ha, rfl⟩ := hbs + apply h ha + obtain ⟨a', ha', ha''⟩ := hbt + rw [hf a a' (hs a ha) (ht a' ha') ha''.symm] + exact ha' + +end Disjoint + end List diff --git a/Mathlib/Data/List/BigOperators/Basic.lean b/Mathlib/Data/List/BigOperators/Basic.lean index 65bc366cd25c6..0e16729b630e4 100644 --- a/Mathlib/Data/List/BigOperators/Basic.lean +++ b/Mathlib/Data/List/BigOperators/Basic.lean @@ -46,6 +46,15 @@ theorem prod_cons : (a :: l).prod = a * l.prod := #align list.prod_cons List.prod_cons #align list.sum_cons List.sum_cons +@[to_additive] +lemma prod_induction + (p : M → Prop) (hom : ∀ a b, p a → p b → p (a * b)) (unit : p 1) (base : ∀ x ∈ l, p x) : + p l.prod := by + induction' l with a l ih; simpa + rw [List.prod_cons] + simp only [Bool.not_eq_true, List.mem_cons, forall_eq_or_imp] at base + exact hom _ _ (base.1) (ih base.2) + @[to_additive (attr := simp)] theorem prod_append : (l₁ ++ l₂).prod = l₁.prod * l₂.prod := calc @@ -170,14 +179,12 @@ theorem prod_take_mul_prod_drop : ∀ (L : List M) (i : ℕ), (L.take i).prod * @[to_additive (attr := simp)] theorem prod_take_succ : - ∀ (L : List M) (i : ℕ) (p), (L.take (i + 1)).prod = (L.take i).prod * L.nthLe i p + ∀ (L : List M) (i : ℕ) (p), (L.take (i + 1)).prod = (L.take i).prod * L.get ⟨i, p⟩ | [], i, p => by cases p | h :: t, 0, _ => rfl | h :: t, n + 1, p => by dsimp - rw [prod_cons, prod_cons, prod_take_succ t n (Nat.lt_of_succ_lt_succ p), mul_assoc, - nthLe_cons, dif_neg (Nat.add_one_ne_zero _)] - simp + rw [prod_cons, prod_cons, prod_take_succ t n (Nat.lt_of_succ_lt_succ p), mul_assoc] #align list.prod_take_succ List.prod_take_succ #align list.sum_take_succ List.sum_take_succ @@ -422,14 +429,13 @@ theorem prod_reverse_noncomm : ∀ L : List G, L.reverse.prod = (L.map fun x => #align list.prod_reverse_noncomm List.prod_reverse_noncomm #align list.sum_reverse_noncomm List.sum_reverse_noncomm -set_option linter.deprecated false in /-- Counterpart to `List.prod_take_succ` when we have an inverse operation -/ @[to_additive (attr := simp) "Counterpart to `List.sum_take_succ` when we have a negation operation"] theorem prod_drop_succ : - ∀ (L : List G) (i : ℕ) (p), (L.drop (i + 1)).prod = (L.nthLe i p)⁻¹ * (L.drop i).prod + ∀ (L : List G) (i : ℕ) (p), (L.drop (i + 1)).prod = (L.get ⟨i, p⟩)⁻¹ * (L.drop i).prod | [], i, p => False.elim (Nat.not_lt_zero _ p) - | x :: xs, 0, _ => by simp [nthLe] + | x :: xs, 0, _ => by simp | x :: xs, i + 1, p => prod_drop_succ xs i _ #align list.prod_drop_succ List.prod_drop_succ #align list.sum_drop_succ List.sum_drop_succ @@ -451,7 +457,7 @@ theorem prod_inv : ∀ L : List G, L.prod⁻¹ = (L.map fun x => x⁻¹).prod /-- Alternative version of `List.prod_set` when the list is over a group -/ @[to_additive "Alternative version of `List.sum_set` when the list is over a group"] theorem prod_set' (L : List G) (n : ℕ) (a : G) : - (L.set n a).prod = L.prod * if hn : n < L.length then (L.nthLe n hn)⁻¹ * a else 1 := by + (L.set n a).prod = L.prod * if hn : n < L.length then (L.get ⟨n, hn⟩)⁻¹ * a else 1 := by refine (prod_set L n a).trans ?_ split_ifs with hn · rw [mul_comm _ a, mul_assoc a, prod_drop_succ L n hn, mul_comm _ (drop n L).prod, ← @@ -545,13 +551,20 @@ theorem sum_le_foldr_max [AddMonoid M] [AddMonoid N] [LinearOrder N] (f : M → exact (hadd _ _).trans (max_le_max le_rfl IH) #align list.sum_le_foldr_max List.sum_le_foldr_max +@[to_additive] +theorem prod_erase_of_comm [DecidableEq M] [Monoid M] {a} {l : List M} (ha : a ∈ l) + (comm : ∀ x ∈ l, ∀ y ∈ l, x * y = y * x) : + a * (l.erase a).prod = l.prod := by + induction' l with b l ih; simp at ha + obtain rfl | ⟨ne, h⟩ := Decidable.List.eq_or_ne_mem_of_mem ha; simp + rw [List.erase, beq_false_of_ne ne.symm, List.prod_cons, List.prod_cons, ← mul_assoc, + comm a ha b (l.mem_cons_self b), mul_assoc, + ih h fun x hx y hy ↦ comm _ (List.mem_cons_of_mem b hx) _ (List.mem_cons_of_mem b hy)] + @[to_additive (attr := simp)] -theorem prod_erase [DecidableEq M] [CommMonoid M] {a} : - ∀ {l : List M}, a ∈ l → a * (l.erase a).prod = l.prod - | b :: l, h => by - obtain rfl | ⟨ne, h⟩ := Decidable.List.eq_or_ne_mem_of_mem h - · simp only [List.erase, if_pos, prod_cons, beq_self_eq_true] - · simp only [List.erase, beq_false_of_ne ne.symm, prod_cons, prod_erase h, mul_left_comm a b] +theorem prod_erase [DecidableEq M] [CommMonoid M] {a} {l : List M} (ha : a ∈ l) : + a * (l.erase a).prod = l.prod := + prod_erase_of_comm ha fun x _ y _ ↦ mul_comm x y #align list.prod_erase List.prod_erase #align list.sum_erase List.sum_erase @@ -597,7 +610,7 @@ If desired, we could add a class stating that `default = 0`. /-- This relies on `default ℕ = 0`. -/ theorem headI_add_tail_sum (L : List ℕ) : L.headI + L.tail.sum = L.sum := by - cases L <;> simp + cases L <;> simp; rfl #align list.head_add_tail_sum List.headI_add_tail_sum /-- This relies on `default ℕ = 0`. -/ diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index a6378404a3ba4..368b6cdf504f9 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -354,7 +354,7 @@ theorem chain'_iff_get {R} : ∀ {l : List α}, Chain' R l ↔ rw [← and_forall_succ, chain'_cons, chain'_iff_get] simp only [length_cons, get_cons_succ, Fin.zero_eta, get_cons_zero, zero_add, Fin.mk_one, get_cons_cons_one, succ_sub_succ_eq_sub, nonpos_iff_eq_zero, add_eq_zero_iff, and_false, - tsub_zero, add_pos_iff, or_true, forall_true_left, and_congr_right_iff] + tsub_zero, add_pos_iff, zero_lt_one, or_true, forall_true_left, and_congr_right_iff] dsimp [succ_sub_one] exact fun _ => ⟨fun h i hi => h i (Nat.lt_of_succ_lt_succ hi), fun h i hi => h i (Nat.succ_lt_succ hi)⟩ @@ -441,6 +441,31 @@ theorem relationReflTransGen_of_exists_chain (l : List α) (hl₁ : Chain r a l) Relation.ReflTransGen.refl #align list.relation_refl_trans_gen_of_exists_chain List.relationReflTransGen_of_exists_chain +theorem Chain'.cons_of_le [LinearOrder α] {a : α} {as m : List α} + (ha : List.Chain' (· > ·) (a :: as)) (hm : List.Chain' (· > ·) m) (hmas : m ≤ as) : + List.Chain' (· > ·) (a :: m) := by + cases m with + | nil => simp only [List.chain'_singleton] + | cons b bs => + apply hm.cons + cases as with + | nil => + simp only [le_iff_lt_or_eq, or_false] at hmas + exact (List.Lex.not_nil_right (·<·) _ hmas).elim + | cons a' as => + rw [List.chain'_cons] at ha + refine gt_of_gt_of_ge ha.1 ?_ + rw [le_iff_lt_or_eq] at hmas + cases' hmas with hmas hmas + · by_contra' hh + rw [← not_le] at hmas + apply hmas + apply le_of_lt + exact (List.lt_iff_lex_lt _ _).mp (List.lt.head _ _ hh) + · simp only [List.cons.injEq] at hmas + rw [ge_iff_le, le_iff_lt_or_eq] + exact Or.inr hmas.1 + end List diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 88fa6f2664aa3..d2fa37f01ded9 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -410,7 +410,7 @@ theorem prev_reverse_eq_next (l : List α) (h : Nodup l) (x : α) (hx : x ∈ l) length_reverse, Nat.mod_eq_of_lt (tsub_lt_self lpos Nat.succ_pos'), tsub_tsub_cancel_of_le (Nat.succ_le_of_lt lpos)] rw [← nthLe_reverse] - · simp [tsub_tsub_cancel_of_le (Nat.le_pred_of_lt hk)] + · simp [tsub_tsub_cancel_of_le (Nat.le_sub_one_of_lt hk)] · simpa using (Nat.sub_le _ _).trans_lt (tsub_lt_self lpos Nat.succ_pos') · simpa #align list.prev_reverse_eq_next List.prev_reverse_eq_next diff --git a/Mathlib/Data/List/EditDistance/Defs.lean b/Mathlib/Data/List/EditDistance/Defs.lean index a11ebf91b4417..02186dc68b252 100644 --- a/Mathlib/Data/List/EditDistance/Defs.lean +++ b/Mathlib/Data/List/EditDistance/Defs.lean @@ -59,6 +59,28 @@ def defaultCost [DecidableEq α] : Cost α α ℕ where instance [DecidableEq α] : Inhabited (Cost α α ℕ) := ⟨defaultCost⟩ +/-- +Cost structure given by a function. +Delete and insert cost the same, and substitution costs the greater value. +-/ +@[simps] +def weightCost (f : α → ℕ) : Cost α α ℕ where + delete a := f a + insert b := f b + substitute a b := max (f a) (f b) + +/-- +Cost structure for strings, where cost is the length of the token. +-/ +@[simps!] +def stringLengthCost : Cost String String ℕ := weightCost String.length + +/-- +Cost structure for strings, where cost is the log base 2 length of the token. +-/ +@[simps!] +def stringLogLengthCost : Cost String String ℕ := weightCost fun s => Nat.log2 (s.length + 1) + variable (C : Cost α β δ) /-- @@ -111,7 +133,7 @@ theorem impl_length (d : {r : List δ // 0 < r.length}) (w : d.1.length = xs.len | ⟨d₁ :: d₂ :: ds, _⟩, w => dsimp congr 1 - refine ih ⟨d₂ :: ds, (by simp)⟩ (by simpa using w) + exact ih ⟨d₂ :: ds, (by simp)⟩ (by simpa using w) end Levenshtein diff --git a/Mathlib/Data/List/Func.lean b/Mathlib/Data/List/Func.lean index 92bdf575de901..0c0ff33a59db9 100644 --- a/Mathlib/Data/List/Func.lean +++ b/Mathlib/Data/List/Func.lean @@ -109,7 +109,7 @@ theorem length_set : ∀ {m : ℕ} {as : List α}, as {m ↦ a}.length = max as. have := @length_set m [] simp [set, length, @length_set m, Nat.zero_max] | m + 1, _ :: as => by - simp [set, length, @length_set m, Nat.max_succ_succ] + simp [set, length, @length_set m, Nat.succ_max_succ] #align list.func.length_set List.Func.length_set -- porting note : @[simp] has been removed since `#lint` says this is @@ -306,7 +306,7 @@ theorem length_pointwise {f : α → β → γ} : | _ :: as, [] => by simp only [pointwise, length, length_map, max_eq_left (Nat.zero_le (length as + 1))] | _ :: as, _ :: bs => by - simp only [pointwise, length, Nat.max_succ_succ, @length_pointwise _ as bs] + simp only [pointwise, length, Nat.succ_max_succ, @length_pointwise _ as bs] #align list.func.length_pointwise List.Func.length_pointwise end Func diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index 6eb0edaea1148..d6e882e1aa7df 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -22,9 +22,9 @@ All those (except `insert`) are defined in `Mathlib.Data.List.Defs`. ## Notation -`l₁ <+: l₂`: `l₁` is a prefix of `l₂`. -`l₁ <:+ l₂`: `l₁` is a suffix of `l₂`. -`l₁ <:+: l₂`: `l₁` is an infix of `l₂`. +* `l₁ <+: l₂`: `l₁` is a prefix of `l₂`. +* `l₁ <:+ l₂`: `l₁` is a suffix of `l₂`. +* `l₁ <:+: l₂`: `l₁` is an infix of `l₂`. -/ open Nat @@ -259,7 +259,7 @@ theorem prefix_take_le_iff {L : List (List (Option α))} (hm : m < L.length) : | zero => refine' iff_of_false _ (zero_lt_succ _).not_le rw [take_zero, take_nil] - simp only [take] + simp only [take, not_false_eq_true] | succ n => simp only [length] at hm have specializedIH := @IH n ls (Nat.lt_of_succ_lt_succ hm) diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index 25136a1b3a160..a8fb8984e4e69 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -45,17 +45,17 @@ theorem zero_bot (n : ℕ) : Ico 0 n = range n := by rw [Ico, tsub_zero, range_e @[simp] theorem length (n m : ℕ) : length (Ico n m) = m - n := by dsimp [Ico] - simp only [length_range'] + simp [length_range', autoParam] #align list.Ico.length List.Ico.length theorem pairwise_lt (n m : ℕ) : Pairwise (· < ·) (Ico n m) := by dsimp [Ico] - simp only [pairwise_lt_range'] + simp [pairwise_lt_range', autoParam] #align list.Ico.pairwise_lt List.Ico.pairwise_lt theorem nodup (n m : ℕ) : Nodup (Ico n m) := by dsimp [Ico] - simp only [nodup_range'] + simp [nodup_range', autoParam] #align list.Ico.nodup List.Ico.nodup @[simp] diff --git a/Mathlib/Data/List/Join.lean b/Mathlib/Data/List/Join.lean index df9eebf7a47a8..49652b41e51db 100644 --- a/Mathlib/Data/List/Join.lean +++ b/Mathlib/Data/List/Join.lean @@ -153,11 +153,10 @@ theorem drop_take_succ_join_eq_nthLe (L : List (List α)) {i : ℕ} (hi : i < L. simp [take_sum_join, this, drop_sum_join, drop_take_succ_eq_cons_nthLe _ hi] #align list.drop_take_succ_join_eq_nth_le List.drop_take_succ_join_eq_nthLe -set_option linter.deprecated false in /-- Auxiliary lemma to control elements in a join. -/ @[deprecated] theorem sum_take_map_length_lt1 (L : List (List α)) {i j : ℕ} (hi : i < L.length) - (hj : j < (nthLe L i hi).length) : + (hj : j < (L.get ⟨i, hi⟩).length) : ((L.map length).take i).sum + j < ((L.map length).take (i + 1)).sum := by simp [hi, sum_take_succ, hj] #align list.sum_take_map_length_lt1 List.sum_take_map_length_lt1 @@ -228,4 +227,20 @@ theorem join_reverse (L : List (List α)) : simpa [reverse_reverse, map_reverse] using congr_arg List.reverse (reverse_join L.reverse) #align list.join_reverse List.join_reverse +/-- Any member of `l : List (List α))` is a sublist of `l.join` -/ +lemma sublist_join (l : List (List α)) {s : List α} (hs : s ∈ l) : + List.Sublist s (l.join) := by + induction l with + | nil => + exfalso + exact not_mem_nil s hs + | cons t m ht => + cases mem_cons.mp hs with + | inl h => + rw [h] + simp only [join_cons, sublist_append_left] + | inr h => + simp only [join_cons] + exact sublist_append_of_sublist_right (ht h) + end List diff --git a/Mathlib/Data/List/Lattice.lean b/Mathlib/Data/List/Lattice.lean index 5902bcaca37ad..b442e56bd2580 100644 --- a/Mathlib/Data/List/Lattice.lean +++ b/Mathlib/Data/List/Lattice.lean @@ -134,12 +134,12 @@ theorem inter_nil (l : List α) : [] ∩ l = [] := @[simp] theorem inter_cons_of_mem (l₁ : List α) (h : a ∈ l₂) : (a :: l₁) ∩ l₂ = a :: l₁ ∩ l₂ := by - simp only [Inter.inter, List.inter, filter_cons_of_pos, h] + simp [Inter.inter, List.inter, h] #align list.inter_cons_of_mem List.inter_cons_of_mem @[simp] theorem inter_cons_of_not_mem (l₁ : List α) (h : a ∉ l₂) : (a :: l₁) ∩ l₂ = l₁ ∩ l₂ := by - simp only [Inter.inter, List.inter, filter_cons_of_neg, h] + simp [Inter.inter, List.inter, h] #align list.inter_cons_of_not_mem List.inter_cons_of_not_mem theorem mem_of_mem_inter_left : a ∈ l₁ ∩ l₂ → a ∈ l₁ := diff --git a/Mathlib/Data/List/Lex.lean b/Mathlib/Data/List/Lex.lean index 3380921e8bfd5..157769c78c640 100644 --- a/Mathlib/Data/List/Lex.lean +++ b/Mathlib/Data/List/Lex.lean @@ -212,14 +212,27 @@ theorem lt_iff_lex_lt [LinearOrder α] (l l' : List α) : lt l l' ↔ Lex (· < | @cons a as bs _ ih => apply lt.tail <;> simp [ih] | @rel a as b bs h => apply lt.head; assumption -theorem head!_le_of_lt [LinearOrder α] [Inhabited α] (l l' : List α) (h : lt l' l) (hl' : l' ≠ []) : +theorem head_le_of_lt [LinearOrder α] {a a' : α} {l l' : List α} (h : (a' :: l') < (a :: l)) : + a' ≤ a := by + by_contra hh + simp only [not_le] at hh + exact List.Lex.isAsymm.aux _ _ _ (List.Lex.rel hh) h + +theorem head!_le_of_lt [LinearOrder α] [Inhabited α] (l l' : List α) (h : l' < l) (hl' : l' ≠ []) : l'.head! ≤ l.head! := by - rw [lt_iff_lex_lt l' l] at h + replace h : List.Lex (· < ·) l' l := h by_cases hl : l = [] - · simp only [hl, List.Lex.not_nil_right] at h - · by_contra hh - have := List.Lex.rel (r := (·<·)) (l₁ := l.tail) (l₂ := l'.tail) (not_le.mp hh) - rw [List.cons_head!_tail hl', List.cons_head!_tail hl] at this - exact asymm h this + · simp [hl] at h + · rw [← List.cons_head!_tail hl', ← List.cons_head!_tail hl] at h + exact head_le_of_lt h + +theorem cons_le_cons [LinearOrder α] (a : α) {l l' : List α} (h : l' ≤ l) : + a :: l' ≤ a :: l := by + rw [le_iff_lt_or_eq] at h ⊢ + refine h.imp ?_ (congr_arg _) + intro h + have haa : ¬(a < a) := gt_irrefl a + exact (List.lt_iff_lex_lt _ _).mp + (List.lt.tail haa haa ((List.lt_iff_lex_lt _ _).mpr h)) end List diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index 7a279dc09e66f..cb34a55947e94 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -449,6 +449,15 @@ theorem minimum_of_length_pos_le_iff (h : 0 < l.length) : minimum_of_length_pos h ≤ b ↔ l.minimum ≤ b := le_maximum_of_length_pos_iff (α := αᵒᵈ) h +theorem maximum_of_length_pos_mem (h : 0 < l.length) : + maximum_of_length_pos h ∈ l := by + apply maximum_mem + simp only [coe_maximum_of_length_pos] + +theorem minimum_of_length_pos_mem (h : 0 < l.length) : + minimum_of_length_pos h ∈ l := + maximum_of_length_pos_mem (α := αᵒᵈ) h + theorem le_maximum_of_length_pos_of_mem (h : a ∈ l) (w : 0 < l.length) : a ≤ l.maximum_of_length_pos w := by simp only [le_maximum_of_length_pos_iff] diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index dd1dc5d18d635..90542a4208bd7 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -275,9 +275,11 @@ theorem filter_append_perm (p : α → Bool) (l : List α) : induction' l with x l ih · rfl · by_cases h : p x - · simp only [h, filter_cons_of_pos, filter_cons_of_neg, not_true, not_false_iff, cons_append] + · simp only [h, filter_cons_of_pos, filter_cons_of_neg, not_true, not_false_iff, cons_append, + decide_False] exact ih.cons x - · simp only [h, filter_cons_of_neg, not_false_iff, filter_cons_of_pos] + · simp only [h, filter_cons_of_neg, not_false_iff, filter_cons_of_pos, cons_append, + not_false_eq_true, decide_True] refine' Perm.trans _ (ih.cons x) exact perm_append_comm.trans (perm_append_comm.cons _) #align list.filter_append_perm List.filter_append_perm @@ -1197,10 +1199,9 @@ theorem perm_of_mem_permutationsAux : refine' permutationsAux.rec (by simp) _ introv IH1 IH2 m rw [permutationsAux_cons, permutations, mem_foldr_permutationsAux2] at m - rcases m with (m | ⟨l₁, l₂, m, _, e⟩) + rcases m with (m | ⟨l₁, l₂, m, _, rfl⟩) · exact (IH1 _ m).trans perm_middle - · subst e - have p : l₁ ++ l₂ ~ is := by + · have p : l₁ ++ l₂ ~ is := by simp [permutations] at m cases' m with e m · simp [e] diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index b19a20e013d3f..16942639d2398 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kenny Lau, Scott Morrison -/ import Mathlib.Data.List.Chain +import Mathlib.Data.List.Join import Mathlib.Data.List.Nodup import Mathlib.Data.List.Zip @@ -75,14 +76,15 @@ theorem nthLe_range'_1 {n m} (i) (H : i < (range' n m).length) : #align list.range_eq_nil List.range_eq_nil theorem pairwise_lt_range (n : ℕ) : Pairwise (· < ·) (range n) := by - simp only [range_eq_range', pairwise_lt_range'] + simp (config := {decide := true}) only [range_eq_range', pairwise_lt_range'] #align list.pairwise_lt_range List.pairwise_lt_range theorem pairwise_le_range (n : ℕ) : Pairwise (· ≤ ·) (range n) := Pairwise.imp (@le_of_lt ℕ _) (pairwise_lt_range _) #align list.pairwise_le_range List.pairwise_le_range -theorem nodup_range (n : ℕ) : Nodup (range n) := by simp only [range_eq_range', nodup_range'] +theorem nodup_range (n : ℕ) : Nodup (range n) := by + simp (config := {decide := true}) only [range_eq_range', nodup_range'] #align list.nodup_range List.nodup_range #align list.range_sublist List.range_sublist #align list.range_subset List.range_subset @@ -234,4 +236,71 @@ theorem nthLe_finRange {n : ℕ} {i : ℕ} (h) : have h₂ : (finRange k).get ⟨i, by simp⟩ = i := get_finRange _ simpa using (Nodup.get_inj_iff (nodup_finRange k)).mp (Eq.trans h₁ h₂.symm) +section Ranges + +/-- From `l : List ℕ`, construct `l.ranges : List (List ℕ)` such that + `l.ranges.map List.length = l` and `l.ranges.join = range l.sum` +* Example: `[1,2,3].ranges = [[0],[1,2],[3,4,5]]` -/ +def ranges : List ℕ → List (List ℕ) + | [] => nil + | a::l => range a::(ranges l).map (map (Nat.add a)) + +/-- The members of `l.ranges` are pairwise disjoint -/ +theorem ranges_disjoint (l : List ℕ) : + Pairwise Disjoint (ranges l) := by + induction l with + | nil => exact Pairwise.nil + | cons a l hl => + simp only [ranges, pairwise_cons] + constructor + · intro s hs + obtain ⟨s', _, rfl⟩ := mem_map.mp hs + intro u hu + rw [mem_map] + rintro ⟨v, _, rfl⟩ + rw [mem_range] at hu + exact lt_iff_not_le.mp hu le_self_add + · rw [pairwise_map] + apply Pairwise.imp _ hl + intro u v + apply disjoint_map + exact fun u v => Nat.add_left_cancel +#align list.ranges_disjoint List.ranges_disjoint + +/-- The lengths of the members of `l.ranges` are those given by `l` -/ +theorem ranges_length (l : List ℕ) : + l.ranges.map length = l := by + induction l with + | nil => simp only [ranges, map_nil] + | cons a l hl => -- (a :: l) + simp only [map, length_range, map_map, cons.injEq, true_and] + conv_rhs => rw [← hl] + apply map_congr + intro s _ + simp only [Function.comp_apply, length_map] + +theorem ranges_join (l : List ℕ) : + l.ranges.join = range l.sum := by + induction l with + | nil => exact rfl + | cons a l hl => + simp only [sum_cons, join] + rw [← map_join, hl] + rw [range_add] + rfl + +/-- Any entry of any member of `l.ranges` is strictly smaller than `l.sum` -/ +theorem mem_mem_ranges_iff_lt_sum (l : List ℕ) {n : ℕ} : + (∃ s ∈ List.ranges l, n ∈ s) ↔ n < l.sum := by + rw [← mem_range, ← ranges_join, mem_join] + + /-- The members of `l.ranges` have no duplicate -/ +theorem ranges_nodup {l s : List ℕ} (hs : s ∈ ranges l) : + s.Nodup := by + refine (List.pairwise_join.mp ?_).1 s hs + rw [ranges_join] + exact nodup_range (sum l) + +end Ranges + end List diff --git a/Mathlib/Data/List/Rdrop.lean b/Mathlib/Data/List/Rdrop.lean index f91c332d0f803..ab4856e998b08 100644 --- a/Mathlib/Data/List/Rdrop.lean +++ b/Mathlib/Data/List/Rdrop.lean @@ -229,10 +229,10 @@ theorem rtakeWhile_eq_nil_iff : rtakeWhile p l = [] ↔ ∀ hl : l ≠ [], ¬p ( · simp only [rtakeWhile, takeWhile, reverse_nil, true_iff] intro f; contradiction · simp only [rtakeWhile, reverse_append, takeWhile, reverse_eq_nil, getLast_append, ne_eq, - append_eq_nil, and_false, forall_true_left] + append_eq_nil, and_false, not_false_eq_true, forall_true_left] refine' ⟨fun h => _ , fun h => _⟩ - · intro pa; simp only [pa] at h - · simp only [h] + · intro pa; simp [pa] at h + · simp [h] #align list.rtake_while_eq_nil_iff List.rtakeWhile_eq_nil_iff theorem mem_rtakeWhile_imp {x : α} (hx : x ∈ rtakeWhile p l) : p x := by diff --git a/Mathlib/Data/MLList/BestFirst.lean b/Mathlib/Data/MLList/BestFirst.lean index 0600fcebaacf9..d36620a4a961a 100644 --- a/Mathlib/Data/MLList/BestFirst.lean +++ b/Mathlib/Data/MLList/BestFirst.lean @@ -4,7 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Std.Data.MLList.Basic -import Std.Data.RBMap.Basic +import Mathlib.Data.Finset.Preimage +import Mathlib.Data.Prod.Lex +import Mathlib.Order.Estimator /-! # Best first search @@ -23,43 +25,242 @@ Options: * `maxDepth` allows bounding the search depth * `maxQueued` implements "beam" search, by discarding elements from the priority queue when it grows too large -* `removeDuplicates` maintains an `RBSet` of previously visited nodes; +* `removeDuplicatesBy?` maintains an `RBSet` of previously visited nodes; otherwise if the graph is not a tree nodes may be visited multiple times. -/ set_option autoImplicit true +open Std EstimatorData Estimator Set -variable {α : Type u} [Monad m] [Alternative m] [Ord α] +/-! +We begin by defining a best-first queue of `MLList`s. +This is a somewhat baroque data structure designed for the application in this file +(and in particularly for the implementation of `rewrite_search`). +If someone would like to generalize appropriately that would be great. + +We want to maintain a priority queue of `MLList m β`, each indexed by some `a : α` with a priority. +(One could simplify matters here by simply flattening this out to a priority queue of pairs `α × β`, +with the priority determined by the `α` factor. +However the lazyness of `MLList` is essential to performance here: +we will extract elements from these lists one at a time, +and only when they at the head of the queue. +If another item arrives at the head of the queue, +we may not need to continue calculate the previous head's elements.) + +To complicate matters, the priorities might be expensive to calculate, +so we instead keep track of a lower bound (where less is better) for each such `a : α`. +The priority queue maintains the `MLList m β` in order of the current best lower bound for the +corresponding `a : α`. +When we insert a new `α × MLList m β` into the queue, we have to provide a lower bound, +and we just insert it at a position depending on the estimate. +When it is time to pop a `β` off the queue, we iteratively improve the lower bound for the +front element of the queue, until we decide that either it must be the least element, +or we can exchange it with the second element of the queue and continue. + +A `BestFirstQueue prio ε m β maxSize` consists of an `RBMap`, +where the keys are in `BestFirstNode prio ε` +and the values are `MLList m β`. + +A `BestFirstNode prio ε` consists of a `key : α` and an estimator `ε : key`. +Here `ε` provides the current best lower bound for `prio key : Thunk ω`. +(The actual priority is hidden behind a `Thunk` to avoid evaluating it, in case it is expensive.) + +We ask for the type classes `LinearOrder ω` and `∀ a : α, Estimator (prio a) (ε a)`. +This later typeclass ensures that we can always produce progressively better estimates +for a priority. We also need a `WellFounded` instance to ensure that improving estimates terminates. + +This whole structure is designed around the problem of searching rewrite graphs, +prioritising according to edit distances (either between sides of an equation, +or from a goal to a target). Edit distance computations are particularly suited to this design +because the usual algorithm for computing them produces improving lower bounds at each step. + +With this design, it is okay if we visit nodes with very large edit distances: +while these would be expensive to compute, we never actually finish the computation +except in cases where the node arrives at the front of the queue. +-/ +section + +/-- A node in a `BestFirstQueue`. -/ +structure BestFirstNode (prio : α → Thunk ω) (ε : α → Type) where + /-- The data to store at a node, from which we can calculate a priority using `prio`. -/ + key : α + /-- An estimator for the priority of the key. + (We will assume we have `[∀ a : α, Estimator (prio a) (ε a)]`.) -/ + estimator : ε key + +variable {α : Type} {prio : α → Thunk ω} {ε : α → Type} [LinearOrder ω] + [∀ a, Estimator (prio a) (ε a)] + [I : ∀ a : α, WellFoundedGT (range (bound (prio a) : ε a → ω))] + {m : Type → Type} [Monad m] {β : Type} + +/-- Calculate the current best lower bound for the the priority of a node. -/ +def BestFirstNode.estimate (n : BestFirstNode prio ε) : ω := bound (prio n.key) n.estimator + +instance [Ord ω] [Ord α] : Ord (BestFirstNode prio ε) where + compare := + compareLex + (compareOn BestFirstNode.estimate) + (compareOn BestFirstNode.key) + +set_option linter.unusedVariables false in +variable (prio ε m β) [Ord ω] [Ord α] in +/-- A queue of `MLList m β`s, lazily prioritized by lower bounds. -/ +@[nolint unusedArguments] +def BestFirstQueue (maxSize : Option Nat) := RBMap (BestFirstNode prio ε) (MLList m β) compare + +variable [Ord ω] [Ord α] {maxSize : Option Nat} + +namespace BestFirstQueue + +/-- +Add a new `MLList m β` to the `BestFirstQueue`, and if this takes the size above `maxSize`, +eject a `MLList` from the tail of the queue. +-/ +-- Note this ejects the element with the greatest estimated priority, +-- not necessarily the greatest priority! +def insertAndEject + (q : BestFirstQueue prio ε m β maxSize) (n : BestFirstNode prio ε) (l : MLList m β) : + BestFirstQueue prio ε m β maxSize := + match maxSize with + | none => q.insert n l + | some max => + if q.size < max then + q.insert n l + else + match q.max with + | none => RBMap.empty + | some m => q.insert n l |>.erase m.1 -open Std MLList +/-- +By improving priority estimates as needed, and permuting elements, +ensure that the first element of the queue has the greatest priority. +-/ +partial def ensureFirstIsBest (q : BestFirstQueue prio ε m β maxSize) : + m (BestFirstQueue prio ε m β maxSize) := do + let s := @toStream (RBMap _ _ _) _ _ q + match s.next? with + | none => + -- The queue is empty, nothing to do. + return q + | some ((n, l), s') => match s'.next? with + | none => do + -- There's only one element in the queue, no reordering necessary. + return q + | some ((m, _), _) => + -- `n` is the first index, `m` is the second index. + -- We need to improve our estimate of the priority for `n` to make sure + -- it really should come before `m`. + match improveUntil (prio n.key) (m.estimate < ·) n.estimator with + | .error none => + -- If we couldn't improve the estimate at all, it is exact, and hence the best element. + return q + | .error (some e') => + -- If we improve the estimate, but it is still at most the estimate for `m`, + -- this is the best element, so all we need to do is store the updated estimate. + return q.erase n |>.insert ⟨n.key, e'⟩ l + | .ok e' => + -- If we improved the estimate and it becomes greater than the estimate for `m`, + -- we re-insert `n` with its new estimate, and then try again. + ensureFirstIsBest (q.erase n |>.insert ⟨n.key, e'⟩ l) /-- -Auxiliary function for `bestFirstSearch`, that updates the internal state, -consisting of a priority queue of triples `α × Nat × MLList m α`. -We remove the next element from the list contained in the best triple -(discarding the triple if there is no next element), -enqueue it and return it. +Pop a `β` off the `MLList m β` with lowest priority, +also returning the index in `α` and the current best lower bound for its priority. +This may require improving estimates of priorities and shuffling the queue. -/ --- The return type has `× List α` rather than just `× Option α` so `bestFirstSearch` can use `fixl`. -def bestFirstSearchAux - (f : Nat → α → MLList m α) (maxQueued : Option Nat := none) : - RBMap α (Nat × MLList m α) compare → m (RBMap α (Nat × MLList m α) compare × List α) := -fun s => do - match s.min with - | none => failure - | some (a, (n, L)) => - match ← uncons L with - | none => pure (s.erase a, []) - | some (b, L') => do - let s' := s.insert a (n, L') |>.insert b (n + 1, f (n+1) b) - let s' := match maxQueued with - | some q => if s'.size > q then - match s'.max with | some x => s'.erase x.1 | none => unreachable! - else - s' - | none => s' - pure (s', [b]) +partial def popWithBound (q : BestFirstQueue prio ε m β maxSize) : + m (Option (((a : α) × (ε a) × β) × BestFirstQueue prio ε m β maxSize)) := do + let q' ← ensureFirstIsBest q + match q'.min with + | none => + -- The queue is empty, nothing to return. + return none + | some (n, l) => + match ← l.uncons with + | none => + -- The `MLList` associated to `n` was actually empty, so we remove `n` and try again. + popWithBound (q'.erase n) + | some (b, l') => + -- Return the initial element `b` along with the current estimator, + -- and replace the `MLList` associated with `n` with its tail. + return some (⟨n.key, n.estimator, b⟩, q'.modify n fun _ => l') + +/-- +Pop a `β` off the `MLList m β` with lowest priority, +also returning the index in `α` and the value of the current best lower bound for its priority. +-/ +def popWithPriority (q : BestFirstQueue prio ε m β maxSize) : + m (Option (((α × ω) × β) × BestFirstQueue prio ε m β maxSize)) := do + match ← q.popWithBound with + | none => pure none + | some (⟨a, e, b⟩, q') => pure (some (((a, bound (prio a) e), b), q')) + +/-- +Pop a `β` off the `MLList m β` with lowest priority. +-/ +def pop (q : BestFirstQueue prio ε m β maxSize) : + m (Option ((α × β) × BestFirstQueue prio ε m β maxSize)) := do + match ← q.popWithBound with + | none => pure none + | some (⟨a, _, b⟩, q') => pure (some ((a, b), q')) + +/-- +Convert a `BestFirstQueue` to a `MLList ((α × ω) × β)`, by popping off all elements, +recording also the values in `ω` of the best current lower bounds. +-/ +-- This is not used in the algorithms below, but can be useful for debugging. +partial def toMLListWithPriority (q : BestFirstQueue prio ε m β maxSize) : MLList m ((α × ω) × β) := + .squash fun _ => do + match ← q.popWithPriority with + | none => pure .nil + | some (p, q') => pure <| MLList.cons p q'.toMLListWithPriority + +/-- +Convert a `BestFirstQueue` to a `MLList (α × β)`, by popping off all elements. +-/ +def toMLList (q : BestFirstQueue prio ε m β maxSize) : MLList m (α × β) := + q.toMLListWithPriority.map fun t => (t.1.1, t.2) + +end BestFirstQueue + +open MLList + +variable {m : Type → Type} [Monad m] [Alternative m] [∀ a, Bot (ε a)] (prio ε) + +/-- +Core implementation of `bestFirstSearch`, that works by iteratively updating an internal state, +consisting of a priority queue of `MLList m α`. + +At each step we pop an element off the queue, +compute its children (lazily) and put these back on the queue. +-/ +def impl (maxSize : Option Nat) (f : α → MLList m α) (a : α) : MLList m α := + let init : BestFirstQueue prio ε m α maxSize := RBMap.single ⟨a, ⊥⟩ (f a) + cons a (iterate go |>.runState' init) +where + /-- A single step of the best first search. + Pop an element, and insert its children back into the queue, + with a trivial estimator for their priority. -/ + go : StateT (BestFirstQueue prio ε m α maxSize) m α := fun s => do + match ← s.pop with + | none => failure + | some ((_, b), s') => pure (b, s'.insertAndEject ⟨b, ⊥⟩ (f b)) + +/-- +Wrapper for `impl` implementing the `maxDepth` option. +-/ +def implMaxDepth (maxSize : Option Nat) (maxDepth : Option Nat) (f : α → MLList m α) (a : α) : + MLList m α := + match maxDepth with + | none => impl prio ε maxSize f a + | some max => + let f' : α ×ₗ Nat → MLList m (α × Nat) := fun ⟨a, n⟩ => + if max < n then + nil + else + (f a).map fun a' => (a', n + 1) + impl (fun p : α ×ₗ Nat => prio p.1) (fun p : α ×ₗ Nat => ε p.1) maxSize f' (a, 0) |>.map (·.1) /-- A lazy list recording the best first search of a graph generated by a function @@ -76,21 +277,64 @@ This implements a "beam" search, which may be incomplete but uses bounded memory The option `removeDuplicates` keeps an `RBSet` of previously visited nodes. Otherwise, if the graph is not a tree then nodes will be visited multiple times. + +This version allows specifying a custom priority function `prio : α → Thunk ω` +along with estimators `ε : α → Type` equipped with `[∀ a, Estimator (prio a) (ε a)]` +that control the behaviour of the priority queue. +This function returns values `a : α` that have +the lowest possible `prio a` amongst unvisited neighbours of visited nodes, +but lazily estimates these priorities to avoid unnecessary computations. +-/ +def bestFirstSearchCore (f : α → MLList m α) (a : α) + (β : Type _) [Ord β] (removeDuplicatesBy? : Option (α → β) := none) + (maxQueued : Option Nat := none) (maxDepth : Option Nat := none) : + MLList m α := + match removeDuplicatesBy? with + | some g => + let f' : α → MLList (StateT (RBSet β compare) m) α := fun a => + (f a).liftM >>= fun a' => do + let b := g a' + guard !(← get).contains b + modify fun s => s.insert b + pure a' + implMaxDepth prio ε maxQueued maxDepth f' a |>.runState' (RBSet.empty.insert (g a)) + | none => + implMaxDepth prio ε maxQueued maxDepth f a + +end + +variable [Monad m] [Alternative m] [LinearOrder α] + +/-- A local instance that enables using "the actual value" as a priority estimator, +for simple use cases. -/ +local instance instOrderBotEq : OrderBot { x : α // x = a } where + bot := ⟨a, rfl⟩ + bot_le := by aesop + +/-- +A lazy list recording the best first search of a graph generated by a function +`f : α → MLList m α`. + +We maintain a priority queue of visited-but-not-exhausted nodes, +and at each step take the next child of the highest priority node in the queue. + +The option `maxDepth` limits the search depth. + +The option `maxQueued` bounds the size of the priority queue, +discarding the lowest priority nodes as needed. +This implements a "beam" search, which may be incomplete but uses bounded memory. + +The option `removeDuplicates` keeps an `RBSet` of previously visited nodes. +Otherwise, if the graph is not a tree then nodes will be visited multiple times. + +This function returns values `a : α` that are least in the `[LinearOrder α]` +amongst unvisited neighbours of visited nodes. -/ +-- Although the core implementation lazily computes estimates of priorities, +-- this version does not take advantage of those features. def bestFirstSearch (f : α → MLList m α) (a : α) - (maxDepth : Option Nat := none) (maxQueued : Option Nat := none) (removeDuplicates := true) : + (maxQueued : Option Nat := none) (maxDepth : Option Nat := none) (removeDuplicates := true) : MLList m α := -let f' : Nat → α → MLList m α := match maxDepth with -| none => fun _ a => f a -| some d => fun n a => if d < n then nil else f a -if removeDuplicates then - let f'' : Nat → α → MLList (StateT.{u} (RBSet α compare) m) α := fun n a => - (f' n a).liftM >>= fun b => do - let s ← get - if s.contains b then failure - set <| s.insert b - pure b - cons a (fixl (bestFirstSearchAux f'' maxQueued) (RBMap.single a (0, f'' 0 a))) - |>.runState' (RBSet.empty.insert a) -else - cons a (fixl (bestFirstSearchAux f' maxQueued) (RBMap.single a (0, f' 0 a))) + bestFirstSearchCore Thunk.pure (fun a : α => { x // x = a }) f a + (β := α) (removeDuplicatesBy? := if removeDuplicates then some id else none) + maxQueued maxDepth diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index ac2a6026d11f3..5da9e0349183a 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -187,28 +187,6 @@ def conjTranspose [Star α] (M : Matrix m n α) : Matrix n m α := @[inherit_doc] scoped postfix:1024 "ᴴ" => Matrix.conjTranspose -/-- `Matrix.col u` is the column matrix whose entries are given by `u`. -/ -def col (w : m → α) : Matrix m Unit α := - of fun x _ => w x -#align matrix.col Matrix.col - --- TODO: set as an equation lemma for `col`, see mathlib4#3024 -@[simp] -theorem col_apply (w : m → α) (i j) : col w i j = w i := - rfl -#align matrix.col_apply Matrix.col_apply - -/-- `Matrix.row u` is the row matrix whose entries are given by `u`. -/ -def row (v : n → α) : Matrix Unit n α := - of fun _ y => v y -#align matrix.row Matrix.row - --- TODO: set as an equation lemma for `row`, see mathlib4#3024 -@[simp] -theorem row_apply (v : n → α) (i j) : row v i j = v j := - rfl -#align matrix.row_apply Matrix.row_apply - instance inhabited [Inhabited α] : Inhabited (Matrix m n α) := -- Porting note: this instance was called `Pi.inhabited` in lean3-core, which is much -- nicer than the name `instInhabitedForAll_1` it got in lean4-core... @@ -481,6 +459,47 @@ theorem diagonal_smul [Monoid R] [AddMonoid α] [DistribMulAction R α] (r : R) simp [h] #align matrix.diagonal_smul Matrix.diagonal_smul +@[simp] +theorem diagonal_neg [DecidableEq n] [NegZeroClass α] (d : n → α) : + -diagonal d = diagonal fun i => -d i := by + ext i j + by_cases h : i = j <;> + simp [h] +#align matrix.diagonal_neg Matrix.diagonal_neg + +@[simp] +theorem diagonal_sub [SubNegZeroMonoid α] (d₁ d₂ : n → α) : + diagonal d₁ - diagonal d₂ = diagonal fun i => d₁ i - d₂ i := by + ext i j + by_cases h : i = j <;> + simp [h] + +instance [Zero α] [NatCast α] : NatCast (Matrix n n α) where + natCast m := diagonal fun _ => m + +@[norm_cast] +theorem diagonal_natCast [Zero α] [NatCast α] (m : ℕ) : diagonal (fun _ : n => (m : α)) = m := rfl + +@[norm_cast] +theorem diagonal_natCast' [Zero α] [NatCast α] (m : ℕ) : diagonal ((m : n → α)) = m := rfl + +-- See note [no_index around OfNat.ofNat] +theorem diagonal_ofNat [Zero α] [NatCast α] (m : ℕ) [m.AtLeastTwo] : + diagonal (fun _ : n => no_index (OfNat.ofNat m : α)) = OfNat.ofNat m := rfl + +-- See note [no_index around OfNat.ofNat] +theorem diagonal_ofNat' [Zero α] [NatCast α] (m : ℕ) [m.AtLeastTwo] : + diagonal (no_index (OfNat.ofNat m : n → α)) = OfNat.ofNat m := rfl + +instance [Zero α] [IntCast α] : IntCast (Matrix n n α) where + intCast m := diagonal fun _ => m + +@[norm_cast] +theorem diagonal_intCast [Zero α] [IntCast α] (m : ℤ) : diagonal (fun _ : n => (m : α)) = m := rfl + +@[norm_cast] +theorem diagonal_intCast' [Zero α] [IntCast α] (m : ℤ) : diagonal ((m : n → α)) = m := rfl + variable (n α) /-- `Matrix.diagonal` as an `AddMonoidHom`. -/ @@ -561,6 +580,30 @@ theorem one_eq_pi_single {i j} : (1 : Matrix n n α) i j = Pi.single (f := fun _ end One +instance instAddMonoidWithOne [AddMonoidWithOne α] : AddMonoidWithOne (Matrix n n α) where + natCast_zero := show diagonal _ = _ by + rw [Nat.cast_zero, diagonal_zero] + natCast_succ n := show diagonal _ = diagonal _ + _ by + rw [Nat.cast_succ, ←diagonal_add, diagonal_one] + +instance instAddGroupWithOne [AddGroupWithOne α] : AddGroupWithOne (Matrix n n α) where + intCast_ofNat n := show diagonal _ = diagonal _ by + rw [Int.cast_ofNat] + intCast_negSucc n := show diagonal _ = -(diagonal _) by + rw [Int.cast_negSucc, diagonal_neg] + __ := addGroup + __ := instAddMonoidWithOne + +instance instAddCommMonoidWithOne [AddCommMonoidWithOne α] : + AddCommMonoidWithOne (Matrix n n α) where + __ := addCommMonoid + __ := instAddMonoidWithOne + +instance instAddCommGroupWithOne [AddCommGroupWithOne α] : + AddCommGroupWithOne (Matrix n n α) where + __ := addCommGroup + __ := instAddGroupWithOne + section Numeral set_option linter.deprecated false @@ -932,12 +975,6 @@ theorem mul_apply' [Fintype m] [Mul α] [AddCommMonoid α] {M : Matrix l m α} { rfl #align matrix.mul_apply' Matrix.mul_apply' -@[simp] -theorem diagonal_neg [DecidableEq n] [AddGroup α] (d : n → α) : - -diagonal d = diagonal fun i => -d i := - ((diagonalAddMonoidHom n α).map_neg d).symm -#align matrix.diagonal_neg Matrix.diagonal_neg - theorem sum_apply [AddCommMonoid α] (i : m) (j : n) (s : Finset β) (g : β → Matrix m n α) : (∑ c in s, g c) i j = ∑ c in s, g c i j := (congr_fun (s.sum_apply i g) j).trans (s.sum_apply j _) @@ -1040,12 +1077,6 @@ theorem smul_eq_diagonal_mul [Fintype m] [DecidableEq m] (M : Matrix m n α) (a simp #align matrix.smul_eq_diagonal_mul Matrix.smul_eq_diagonal_mul -@[simp] -theorem diag_col_mul_row (a b : n → α) : diag (col a * row b) = a * b := by - ext - simp [Matrix.mul_apply, col, row] -#align matrix.diag_col_mul_row Matrix.diag_col_mul_row - /-- Left multiplication by a matrix, as an `AddMonoidHom` from matrices to matrices. -/ @[simps] def addMonoidHomMulLeft [Fintype m] (M : Matrix l m α) : Matrix m n α →+ Matrix l n α where @@ -1105,18 +1136,10 @@ protected theorem mul_one [Fintype n] [DecidableEq n] (M : Matrix m n α) : #align matrix.mul_one Matrix.mul_one instance nonAssocSemiring [Fintype n] [DecidableEq n] : NonAssocSemiring (Matrix n n α) := - { Matrix.nonUnitalNonAssocSemiring with + { Matrix.nonUnitalNonAssocSemiring, Matrix.instAddCommMonoidWithOne with one := 1 one_mul := Matrix.one_mul - mul_one := Matrix.mul_one - natCast := fun n => diagonal fun _ => n - natCast_zero := by - ext - simp [Nat.cast] - natCast_succ := fun n => by - ext i j - by_cases i = j <;> - simp [Nat.cast, *]} + mul_one := Matrix.mul_one } @[simp] theorem map_mul [Fintype n] {L : Matrix m n α} {M : Matrix n o α} [NonAssocSemiring β] @@ -1199,11 +1222,11 @@ instance instNonUnitalRing [Fintype n] [NonUnitalRing α] : NonUnitalRing (Matri instance instNonAssocRing [Fintype n] [DecidableEq n] [NonAssocRing α] : NonAssocRing (Matrix n n α) := - { Matrix.nonAssocSemiring, Matrix.addCommGroup with } + { Matrix.nonAssocSemiring, Matrix.instAddCommGroupWithOne with } #align matrix.non_assoc_ring Matrix.instNonAssocRing instance instRing [Fintype n] [DecidableEq n] [Ring α] : Ring (Matrix n n α) := - { Matrix.semiring, Matrix.addCommGroup with } + { Matrix.semiring, Matrix.instAddCommGroupWithOne with } #align matrix.ring Matrix.instRing section Semiring @@ -1655,13 +1678,6 @@ theorem vecMulVec_apply [Mul α] (w : m → α) (v : n → α) (i j) : vecMulVec rfl #align matrix.vec_mul_vec_apply Matrix.vecMulVec_apply -theorem vecMulVec_eq [Mul α] [AddCommMonoid α] (w : m → α) (v : n → α) : - vecMulVec w v = col w * row v := by - ext - simp only [vecMulVec, mul_apply, Fintype.univ_punit, Finset.sum_singleton] - rfl -#align matrix.vec_mul_vec_eq Matrix.vecMulVec_eq - section NonUnitalNonAssocSemiring variable [NonUnitalNonAssocSemiring α] @@ -2670,309 +2686,6 @@ def subDownLeft {d u l r : Nat} (A : Matrix (Fin (u + d)) (Fin (l + r)) α) : subDown (subLeft A) #align matrix.sub_down_left Matrix.subDownLeft -section RowCol - -/-! -### `row_col` section - -Simplification lemmas for `Matrix.row` and `Matrix.col`. --/ - - -open Matrix - -theorem col_injective : Function.Injective (col : (m → α) → _) := - fun _x _y h => funext fun i => congr_fun₂ h i () - -@[simp] theorem col_inj {v w : m → α} : col v = col w ↔ v = w := col_injective.eq_iff - -@[simp] theorem col_zero [Zero α] : col (0 : m → α) = 0 := rfl - -@[simp] theorem col_eq_zero [Zero α] (v : m → α) : col v = 0 ↔ v = 0 := col_inj - -@[simp] -theorem col_add [Add α] (v w : m → α) : col (v + w) = col v + col w := by - ext - rfl -#align matrix.col_add Matrix.col_add - -@[simp] -theorem col_smul [SMul R α] (x : R) (v : m → α) : col (x • v) = x • col v := by - ext - rfl -#align matrix.col_smul Matrix.col_smul - -theorem row_injective : Function.Injective (row : (n → α) → _) := - fun _x _y h => funext fun j => congr_fun₂ h () j - -@[simp] theorem row_inj {v w : n → α} : row v = row w ↔ v = w := row_injective.eq_iff - -@[simp] theorem row_zero [Zero α] : row (0 : n → α) = 0 := rfl - -@[simp] theorem row_eq_zero [Zero α] (v : n → α) : row v = 0 ↔ v = 0 := row_inj - -@[simp] -theorem row_add [Add α] (v w : m → α) : row (v + w) = row v + row w := by - ext - rfl -#align matrix.row_add Matrix.row_add - -@[simp] -theorem row_smul [SMul R α] (x : R) (v : m → α) : row (x • v) = x • row v := by - ext - rfl -#align matrix.row_smul Matrix.row_smul - -@[simp] -theorem transpose_col (v : m → α) : (Matrix.col v)ᵀ = Matrix.row v := by - ext - rfl -#align matrix.transpose_col Matrix.transpose_col - -@[simp] -theorem transpose_row (v : m → α) : (Matrix.row v)ᵀ = Matrix.col v := by - ext - rfl -#align matrix.transpose_row Matrix.transpose_row - -@[simp] -theorem conjTranspose_col [Star α] (v : m → α) : (col v)ᴴ = row (star v) := by - ext - rfl -#align matrix.conj_transpose_col Matrix.conjTranspose_col - -@[simp] -theorem conjTranspose_row [Star α] (v : m → α) : (row v)ᴴ = col (star v) := by - ext - rfl -#align matrix.conj_transpose_row Matrix.conjTranspose_row - -theorem row_vecMul [Fintype m] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : m → α) : - Matrix.row (Matrix.vecMul v M) = Matrix.row v * M := by - ext - rfl -#align matrix.row_vec_mul Matrix.row_vecMul - -theorem col_vecMul [Fintype m] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : m → α) : - Matrix.col (Matrix.vecMul v M) = (Matrix.row v * M)ᵀ := by - ext - rfl -#align matrix.col_vec_mul Matrix.col_vecMul - -theorem col_mulVec [Fintype n] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : n → α) : - Matrix.col (Matrix.mulVec M v) = M * Matrix.col v := by - ext - rfl -#align matrix.col_mul_vec Matrix.col_mulVec - -theorem row_mulVec [Fintype n] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : n → α) : - Matrix.row (Matrix.mulVec M v) = (M * Matrix.col v)ᵀ := by - ext - rfl -#align matrix.row_mul_vec Matrix.row_mulVec - -@[simp] -theorem row_mul_col_apply [Fintype m] [Mul α] [AddCommMonoid α] (v w : m → α) (i j) : - (row v * col w) i j = v ⬝ᵥ w := - rfl -#align matrix.row_mul_col_apply Matrix.row_mul_col_apply - -end RowCol - -section Update - -/-- Update, i.e. replace the `i`th row of matrix `A` with the values in `b`. -/ -def updateRow [DecidableEq m] (M : Matrix m n α) (i : m) (b : n → α) : Matrix m n α := - of <| Function.update M i b -#align matrix.update_row Matrix.updateRow - -/-- Update, i.e. replace the `j`th column of matrix `A` with the values in `b`. -/ -def updateColumn [DecidableEq n] (M : Matrix m n α) (j : n) (b : m → α) : Matrix m n α := - of fun i => Function.update (M i) j (b i) -#align matrix.update_column Matrix.updateColumn - -variable {M : Matrix m n α} {i : m} {j : n} {b : n → α} {c : m → α} - -@[simp] -theorem updateRow_self [DecidableEq m] : updateRow M i b i = b := - -- Porting note: (implicit arg) added `(β := _)` - Function.update_same (β := fun _ => (n → α)) i b M -#align matrix.update_row_self Matrix.updateRow_self - -@[simp] -theorem updateColumn_self [DecidableEq n] : updateColumn M j c i j = c i := - -- Porting note: (implicit arg) added `(β := _)` - Function.update_same (β := fun _ => α) j (c i) (M i) -#align matrix.update_column_self Matrix.updateColumn_self - -@[simp] -theorem updateRow_ne [DecidableEq m] {i' : m} (i_ne : i' ≠ i) : updateRow M i b i' = M i' := - -- Porting note: (implicit arg) added `(β := _)` - Function.update_noteq (β := fun _ => (n → α)) i_ne b M -#align matrix.update_row_ne Matrix.updateRow_ne - -@[simp] -theorem updateColumn_ne [DecidableEq n] {j' : n} (j_ne : j' ≠ j) : - updateColumn M j c i j' = M i j' := - -- Porting note: (implicit arg) added `(β := _)` - Function.update_noteq (β := fun _ => α) j_ne (c i) (M i) -#align matrix.update_column_ne Matrix.updateColumn_ne - -theorem updateRow_apply [DecidableEq m] {i' : m} : - updateRow M i b i' j = if i' = i then b j else M i' j := by - by_cases i' = i - · rw [h, updateRow_self, if_pos rfl] - · rw [updateRow_ne h, if_neg h] -#align matrix.update_row_apply Matrix.updateRow_apply - -theorem updateColumn_apply [DecidableEq n] {j' : n} : - updateColumn M j c i j' = if j' = j then c i else M i j' := by - by_cases j' = j - · rw [h, updateColumn_self, if_pos rfl] - · rw [updateColumn_ne h, if_neg h] -#align matrix.update_column_apply Matrix.updateColumn_apply - -@[simp] -theorem updateColumn_subsingleton [Subsingleton n] (A : Matrix m n R) (i : n) (b : m → R) : - A.updateColumn i b = (col b).submatrix id (Function.const n ()) := by - ext x y - simp [updateColumn_apply, Subsingleton.elim i y] -#align matrix.update_column_subsingleton Matrix.updateColumn_subsingleton - -@[simp] -theorem updateRow_subsingleton [Subsingleton m] (A : Matrix m n R) (i : m) (b : n → R) : - A.updateRow i b = (row b).submatrix (Function.const m ()) id := by - ext x y - simp [updateColumn_apply, Subsingleton.elim i x] -#align matrix.update_row_subsingleton Matrix.updateRow_subsingleton - -theorem map_updateRow [DecidableEq m] (f : α → β) : - map (updateRow M i b) f = updateRow (M.map f) i (f ∘ b) := by - ext - rw [updateRow_apply, map_apply, map_apply, updateRow_apply] - exact apply_ite f _ _ _ -#align matrix.map_update_row Matrix.map_updateRow - -theorem map_updateColumn [DecidableEq n] (f : α → β) : - map (updateColumn M j c) f = updateColumn (M.map f) j (f ∘ c) := by - ext - rw [updateColumn_apply, map_apply, map_apply, updateColumn_apply] - exact apply_ite f _ _ _ -#align matrix.map_update_column Matrix.map_updateColumn - -theorem updateRow_transpose [DecidableEq n] : updateRow Mᵀ j c = (updateColumn M j c)ᵀ := by - ext - rw [transpose_apply, updateRow_apply, updateColumn_apply] - rfl -#align matrix.update_row_transpose Matrix.updateRow_transpose - -theorem updateColumn_transpose [DecidableEq m] : updateColumn Mᵀ i b = (updateRow M i b)ᵀ := by - ext - rw [transpose_apply, updateRow_apply, updateColumn_apply] - rfl -#align matrix.update_column_transpose Matrix.updateColumn_transpose - -theorem updateRow_conjTranspose [DecidableEq n] [Star α] : - updateRow Mᴴ j (star c) = (updateColumn M j c)ᴴ := by - rw [conjTranspose, conjTranspose, transpose_map, transpose_map, updateRow_transpose, - map_updateColumn] - rfl -#align matrix.update_row_conj_transpose Matrix.updateRow_conjTranspose - -theorem updateColumn_conjTranspose [DecidableEq m] [Star α] : - updateColumn Mᴴ i (star b) = (updateRow M i b)ᴴ := by - rw [conjTranspose, conjTranspose, transpose_map, transpose_map, updateColumn_transpose, - map_updateRow] - rfl -#align matrix.update_column_conj_transpose Matrix.updateColumn_conjTranspose - -@[simp] -theorem updateRow_eq_self [DecidableEq m] (A : Matrix m n α) (i : m) : A.updateRow i (A i) = A := - Function.update_eq_self i A -#align matrix.update_row_eq_self Matrix.updateRow_eq_self - -@[simp] -theorem updateColumn_eq_self [DecidableEq n] (A : Matrix m n α) (i : n) : - (A.updateColumn i fun j => A j i) = A := - funext fun j => Function.update_eq_self i (A j) -#align matrix.update_column_eq_self Matrix.updateColumn_eq_self - -theorem diagonal_updateColumn_single [DecidableEq n] [Zero α] (v : n → α) (i : n) (x : α) : - (diagonal v).updateColumn i (Pi.single i x) = diagonal (Function.update v i x) := by - ext j k - obtain rfl | hjk := eq_or_ne j k - · rw [diagonal_apply_eq] - obtain rfl | hji := eq_or_ne j i - · rw [updateColumn_self, Pi.single_eq_same, Function.update_same] - · rw [updateColumn_ne hji, diagonal_apply_eq, Function.update_noteq hji] - · rw [diagonal_apply_ne _ hjk] - obtain rfl | hki := eq_or_ne k i - · rw [updateColumn_self, Pi.single_eq_of_ne hjk] - · rw [updateColumn_ne hki, diagonal_apply_ne _ hjk] -#align matrix.diagonal_update_column_single Matrix.diagonal_updateColumn_single - -theorem diagonal_updateRow_single [DecidableEq n] [Zero α] (v : n → α) (i : n) (x : α) : - (diagonal v).updateRow i (Pi.single i x) = diagonal (Function.update v i x) := by - rw [← diagonal_transpose, updateRow_transpose, diagonal_updateColumn_single, diagonal_transpose] -#align matrix.diagonal_update_row_single Matrix.diagonal_updateRow_single - -/-! Updating rows and columns commutes in the obvious way with reindexing the matrix. -/ - - -theorem updateRow_submatrix_equiv [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : l) - (r : o → α) (e : l ≃ m) (f : o ≃ n) : - updateRow (A.submatrix e f) i r = (A.updateRow (e i) fun j => r (f.symm j)).submatrix e f := by - ext i' j - simp only [submatrix_apply, updateRow_apply, Equiv.apply_eq_iff_eq, Equiv.symm_apply_apply] -#align matrix.update_row_submatrix_equiv Matrix.updateRow_submatrix_equiv - -theorem submatrix_updateRow_equiv [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : m) - (r : n → α) (e : l ≃ m) (f : o ≃ n) : - (A.updateRow i r).submatrix e f = updateRow (A.submatrix e f) (e.symm i) fun i => r (f i) := - Eq.trans (by simp_rw [Equiv.apply_symm_apply]) (updateRow_submatrix_equiv A _ _ e f).symm -#align matrix.submatrix_update_row_equiv Matrix.submatrix_updateRow_equiv - -theorem updateColumn_submatrix_equiv [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : o) - (c : l → α) (e : l ≃ m) (f : o ≃ n) : updateColumn (A.submatrix e f) j c = - (A.updateColumn (f j) fun i => c (e.symm i)).submatrix e f := by - simpa only [← transpose_submatrix, updateRow_transpose] using - congr_arg transpose (updateRow_submatrix_equiv Aᵀ j c f e) -#align matrix.update_column_submatrix_equiv Matrix.updateColumn_submatrix_equiv - -theorem submatrix_updateColumn_equiv [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : n) - (c : m → α) (e : l ≃ m) (f : o ≃ n) : (A.updateColumn j c).submatrix e f = - updateColumn (A.submatrix e f) (f.symm j) fun i => c (e i) := - Eq.trans (by simp_rw [Equiv.apply_symm_apply]) (updateColumn_submatrix_equiv A _ _ e f).symm -#align matrix.submatrix_update_column_equiv Matrix.submatrix_updateColumn_equiv - -/-! `reindex` versions of the above `submatrix` lemmas for convenience. -/ - - -theorem updateRow_reindex [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : l) (r : o → α) - (e : m ≃ l) (f : n ≃ o) : - updateRow (reindex e f A) i r = reindex e f (A.updateRow (e.symm i) fun j => r (f j)) := - updateRow_submatrix_equiv _ _ _ _ _ -#align matrix.update_row_reindex Matrix.updateRow_reindex - -theorem reindex_updateRow [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : m) (r : n → α) - (e : m ≃ l) (f : n ≃ o) : - reindex e f (A.updateRow i r) = updateRow (reindex e f A) (e i) fun i => r (f.symm i) := - submatrix_updateRow_equiv _ _ _ _ _ -#align matrix.reindex_update_row Matrix.reindex_updateRow - -theorem updateColumn_reindex [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : o) (c : l → α) - (e : m ≃ l) (f : n ≃ o) : - updateColumn (reindex e f A) j c = reindex e f (A.updateColumn (f.symm j) fun i => c (e i)) := - updateColumn_submatrix_equiv _ _ _ _ _ -#align matrix.update_column_reindex Matrix.updateColumn_reindex - -theorem reindex_updateColumn [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : n) (c : m → α) - (e : m ≃ l) (f : n ≃ o) : - reindex e f (A.updateColumn j c) = updateColumn (reindex e f A) (f j) fun i => c (e.symm i) := - submatrix_updateColumn_equiv _ _ _ _ _ -#align matrix.reindex_update_column Matrix.reindex_updateColumn - -end Update end Matrix diff --git a/Mathlib/Data/Matrix/Notation.lean b/Mathlib/Data/Matrix/Notation.lean index 36714cc1e0ebe..06fe9a1792df6 100644 --- a/Mathlib/Data/Matrix/Notation.lean +++ b/Mathlib/Data/Matrix/Notation.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Eric Wieser -/ import Mathlib.Data.Matrix.Basic +import Mathlib.Data.Matrix.RowCol import Mathlib.Data.Fin.VecNotation import Mathlib.Tactic.FinCases -import Mathlib.Algebra.BigOperators.Fin #align_import data.matrix.notation from "leanprover-community/mathlib"@"a99f85220eaf38f14f94e04699943e185a5e1d1a" diff --git a/Mathlib/Data/Matrix/RowCol.lean b/Mathlib/Data/Matrix/RowCol.lean new file mode 100644 index 0000000000000..556b79eccc3cd --- /dev/null +++ b/Mathlib/Data/Matrix/RowCol.lean @@ -0,0 +1,352 @@ +/- +Copyright (c) 2019 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen, Eric Wieser +-/ +import Mathlib.Data.Matrix.Basic + +/-! +# Row and column matrices + +This file provides results about row and column matrices + +## Main definitions + +* `Matrix.row r : Matrix Unit n α`: a matrix with a single row +* `Matrix.col c : Matrix m Unit α`: a matrix with a single column +* `Matrix.updateRow M i r`: update the `i`th row of `M` to `r` +* `Matrix.updateCol M j c`: update the `j`th column of `M` to `c` + +-/ + +variable {l m n o : Type*} + +universe u v w +variable {R : Type*} {α : Type v} {β : Type w} + +namespace Matrix + +/-- `Matrix.col u` is the column matrix whose entries are given by `u`. -/ +def col (w : m → α) : Matrix m Unit α := + of fun x _ => w x +#align matrix.col Matrix.col + +-- TODO: set as an equation lemma for `col`, see mathlib4#3024 +@[simp] +theorem col_apply (w : m → α) (i j) : col w i j = w i := + rfl +#align matrix.col_apply Matrix.col_apply + +/-- `Matrix.row u` is the row matrix whose entries are given by `u`. -/ +def row (v : n → α) : Matrix Unit n α := + of fun _ y => v y +#align matrix.row Matrix.row + +-- TODO: set as an equation lemma for `row`, see mathlib4#3024 +@[simp] +theorem row_apply (v : n → α) (i j) : row v i j = v j := + rfl +#align matrix.row_apply Matrix.row_apply + +theorem col_injective : Function.Injective (col : (m → α) → _) := + fun _x _y h => funext fun i => congr_fun₂ h i () + +@[simp] theorem col_inj {v w : m → α} : col v = col w ↔ v = w := col_injective.eq_iff + +@[simp] theorem col_zero [Zero α] : col (0 : m → α) = 0 := rfl + +@[simp] theorem col_eq_zero [Zero α] (v : m → α) : col v = 0 ↔ v = 0 := col_inj + +@[simp] +theorem col_add [Add α] (v w : m → α) : col (v + w) = col v + col w := by + ext + rfl +#align matrix.col_add Matrix.col_add + +@[simp] +theorem col_smul [SMul R α] (x : R) (v : m → α) : col (x • v) = x • col v := by + ext + rfl +#align matrix.col_smul Matrix.col_smul + +theorem row_injective : Function.Injective (row : (n → α) → _) := + fun _x _y h => funext fun j => congr_fun₂ h () j + +@[simp] theorem row_inj {v w : n → α} : row v = row w ↔ v = w := row_injective.eq_iff + +@[simp] theorem row_zero [Zero α] : row (0 : n → α) = 0 := rfl + +@[simp] theorem row_eq_zero [Zero α] (v : n → α) : row v = 0 ↔ v = 0 := row_inj + +@[simp] +theorem row_add [Add α] (v w : m → α) : row (v + w) = row v + row w := by + ext + rfl +#align matrix.row_add Matrix.row_add + +@[simp] +theorem row_smul [SMul R α] (x : R) (v : m → α) : row (x • v) = x • row v := by + ext + rfl +#align matrix.row_smul Matrix.row_smul + +@[simp] +theorem transpose_col (v : m → α) : (Matrix.col v)ᵀ = Matrix.row v := by + ext + rfl +#align matrix.transpose_col Matrix.transpose_col + +@[simp] +theorem transpose_row (v : m → α) : (Matrix.row v)ᵀ = Matrix.col v := by + ext + rfl +#align matrix.transpose_row Matrix.transpose_row + +@[simp] +theorem conjTranspose_col [Star α] (v : m → α) : (col v)ᴴ = row (star v) := by + ext + rfl +#align matrix.conj_transpose_col Matrix.conjTranspose_col + +@[simp] +theorem conjTranspose_row [Star α] (v : m → α) : (row v)ᴴ = col (star v) := by + ext + rfl +#align matrix.conj_transpose_row Matrix.conjTranspose_row + +theorem row_vecMul [Fintype m] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : m → α) : + Matrix.row (Matrix.vecMul v M) = Matrix.row v * M := by + ext + rfl +#align matrix.row_vec_mul Matrix.row_vecMul + +theorem col_vecMul [Fintype m] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : m → α) : + Matrix.col (Matrix.vecMul v M) = (Matrix.row v * M)ᵀ := by + ext + rfl +#align matrix.col_vec_mul Matrix.col_vecMul + +theorem col_mulVec [Fintype n] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : n → α) : + Matrix.col (Matrix.mulVec M v) = M * Matrix.col v := by + ext + rfl +#align matrix.col_mul_vec Matrix.col_mulVec + +theorem row_mulVec [Fintype n] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : n → α) : + Matrix.row (Matrix.mulVec M v) = (M * Matrix.col v)ᵀ := by + ext + rfl +#align matrix.row_mul_vec Matrix.row_mulVec + +@[simp] +theorem row_mul_col_apply [Fintype m] [Mul α] [AddCommMonoid α] (v w : m → α) (i j) : + (row v * col w) i j = v ⬝ᵥ w := + rfl +#align matrix.row_mul_col_apply Matrix.row_mul_col_apply + +@[simp] +theorem diag_col_mul_row [Mul α] [AddCommMonoid α] (a b : n → α) : + diag (col a * row b) = a * b := by + ext + simp [Matrix.mul_apply, col, row] +#align matrix.diag_col_mul_row Matrix.diag_col_mul_row + +theorem vecMulVec_eq [Mul α] [AddCommMonoid α] (w : m → α) (v : n → α) : + vecMulVec w v = col w * row v := by + ext + simp only [vecMulVec, mul_apply, Fintype.univ_punit, Finset.sum_singleton] + rfl +#align matrix.vec_mul_vec_eq Matrix.vecMulVec_eq + +/-! ### Updating rows and columns -/ + +/-- Update, i.e. replace the `i`th row of matrix `A` with the values in `b`. -/ +def updateRow [DecidableEq m] (M : Matrix m n α) (i : m) (b : n → α) : Matrix m n α := + of <| Function.update M i b +#align matrix.update_row Matrix.updateRow + +/-- Update, i.e. replace the `j`th column of matrix `A` with the values in `b`. -/ +def updateColumn [DecidableEq n] (M : Matrix m n α) (j : n) (b : m → α) : Matrix m n α := + of fun i => Function.update (M i) j (b i) +#align matrix.update_column Matrix.updateColumn + +variable {M : Matrix m n α} {i : m} {j : n} {b : n → α} {c : m → α} + +@[simp] +theorem updateRow_self [DecidableEq m] : updateRow M i b i = b := + -- Porting note: (implicit arg) added `(β := _)` + Function.update_same (β := fun _ => (n → α)) i b M +#align matrix.update_row_self Matrix.updateRow_self + +@[simp] +theorem updateColumn_self [DecidableEq n] : updateColumn M j c i j = c i := + -- Porting note: (implicit arg) added `(β := _)` + Function.update_same (β := fun _ => α) j (c i) (M i) +#align matrix.update_column_self Matrix.updateColumn_self + +@[simp] +theorem updateRow_ne [DecidableEq m] {i' : m} (i_ne : i' ≠ i) : updateRow M i b i' = M i' := + -- Porting note: (implicit arg) added `(β := _)` + Function.update_noteq (β := fun _ => (n → α)) i_ne b M +#align matrix.update_row_ne Matrix.updateRow_ne + +@[simp] +theorem updateColumn_ne [DecidableEq n] {j' : n} (j_ne : j' ≠ j) : + updateColumn M j c i j' = M i j' := + -- Porting note: (implicit arg) added `(β := _)` + Function.update_noteq (β := fun _ => α) j_ne (c i) (M i) +#align matrix.update_column_ne Matrix.updateColumn_ne + +theorem updateRow_apply [DecidableEq m] {i' : m} : + updateRow M i b i' j = if i' = i then b j else M i' j := by + by_cases i' = i + · rw [h, updateRow_self, if_pos rfl] + · rw [updateRow_ne h, if_neg h] +#align matrix.update_row_apply Matrix.updateRow_apply + +theorem updateColumn_apply [DecidableEq n] {j' : n} : + updateColumn M j c i j' = if j' = j then c i else M i j' := by + by_cases j' = j + · rw [h, updateColumn_self, if_pos rfl] + · rw [updateColumn_ne h, if_neg h] +#align matrix.update_column_apply Matrix.updateColumn_apply + +@[simp] +theorem updateColumn_subsingleton [Subsingleton n] (A : Matrix m n R) (i : n) (b : m → R) : + A.updateColumn i b = (col b).submatrix id (Function.const n ()) := by + ext x y + simp [updateColumn_apply, Subsingleton.elim i y] +#align matrix.update_column_subsingleton Matrix.updateColumn_subsingleton + +@[simp] +theorem updateRow_subsingleton [Subsingleton m] (A : Matrix m n R) (i : m) (b : n → R) : + A.updateRow i b = (row b).submatrix (Function.const m ()) id := by + ext x y + simp [updateColumn_apply, Subsingleton.elim i x] +#align matrix.update_row_subsingleton Matrix.updateRow_subsingleton + +theorem map_updateRow [DecidableEq m] (f : α → β) : + map (updateRow M i b) f = updateRow (M.map f) i (f ∘ b) := by + ext + rw [updateRow_apply, map_apply, map_apply, updateRow_apply] + exact apply_ite f _ _ _ +#align matrix.map_update_row Matrix.map_updateRow + +theorem map_updateColumn [DecidableEq n] (f : α → β) : + map (updateColumn M j c) f = updateColumn (M.map f) j (f ∘ c) := by + ext + rw [updateColumn_apply, map_apply, map_apply, updateColumn_apply] + exact apply_ite f _ _ _ +#align matrix.map_update_column Matrix.map_updateColumn + +theorem updateRow_transpose [DecidableEq n] : updateRow Mᵀ j c = (updateColumn M j c)ᵀ := by + ext + rw [transpose_apply, updateRow_apply, updateColumn_apply] + rfl +#align matrix.update_row_transpose Matrix.updateRow_transpose + +theorem updateColumn_transpose [DecidableEq m] : updateColumn Mᵀ i b = (updateRow M i b)ᵀ := by + ext + rw [transpose_apply, updateRow_apply, updateColumn_apply] + rfl +#align matrix.update_column_transpose Matrix.updateColumn_transpose + +theorem updateRow_conjTranspose [DecidableEq n] [Star α] : + updateRow Mᴴ j (star c) = (updateColumn M j c)ᴴ := by + rw [conjTranspose, conjTranspose, transpose_map, transpose_map, updateRow_transpose, + map_updateColumn] + rfl +#align matrix.update_row_conj_transpose Matrix.updateRow_conjTranspose + +theorem updateColumn_conjTranspose [DecidableEq m] [Star α] : + updateColumn Mᴴ i (star b) = (updateRow M i b)ᴴ := by + rw [conjTranspose, conjTranspose, transpose_map, transpose_map, updateColumn_transpose, + map_updateRow] + rfl +#align matrix.update_column_conj_transpose Matrix.updateColumn_conjTranspose + +@[simp] +theorem updateRow_eq_self [DecidableEq m] (A : Matrix m n α) (i : m) : A.updateRow i (A i) = A := + Function.update_eq_self i A +#align matrix.update_row_eq_self Matrix.updateRow_eq_self + +@[simp] +theorem updateColumn_eq_self [DecidableEq n] (A : Matrix m n α) (i : n) : + (A.updateColumn i fun j => A j i) = A := + funext fun j => Function.update_eq_self i (A j) +#align matrix.update_column_eq_self Matrix.updateColumn_eq_self + +theorem diagonal_updateColumn_single [DecidableEq n] [Zero α] (v : n → α) (i : n) (x : α) : + (diagonal v).updateColumn i (Pi.single i x) = diagonal (Function.update v i x) := by + ext j k + obtain rfl | hjk := eq_or_ne j k + · rw [diagonal_apply_eq] + obtain rfl | hji := eq_or_ne j i + · rw [updateColumn_self, Pi.single_eq_same, Function.update_same] + · rw [updateColumn_ne hji, diagonal_apply_eq, Function.update_noteq hji] + · rw [diagonal_apply_ne _ hjk] + obtain rfl | hki := eq_or_ne k i + · rw [updateColumn_self, Pi.single_eq_of_ne hjk] + · rw [updateColumn_ne hki, diagonal_apply_ne _ hjk] +#align matrix.diagonal_update_column_single Matrix.diagonal_updateColumn_single + +theorem diagonal_updateRow_single [DecidableEq n] [Zero α] (v : n → α) (i : n) (x : α) : + (diagonal v).updateRow i (Pi.single i x) = diagonal (Function.update v i x) := by + rw [← diagonal_transpose, updateRow_transpose, diagonal_updateColumn_single, diagonal_transpose] +#align matrix.diagonal_update_row_single Matrix.diagonal_updateRow_single + +/-! Updating rows and columns commutes in the obvious way with reindexing the matrix. -/ + + +theorem updateRow_submatrix_equiv [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : l) + (r : o → α) (e : l ≃ m) (f : o ≃ n) : + updateRow (A.submatrix e f) i r = (A.updateRow (e i) fun j => r (f.symm j)).submatrix e f := by + ext i' j + simp only [submatrix_apply, updateRow_apply, Equiv.apply_eq_iff_eq, Equiv.symm_apply_apply] +#align matrix.update_row_submatrix_equiv Matrix.updateRow_submatrix_equiv + +theorem submatrix_updateRow_equiv [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : m) + (r : n → α) (e : l ≃ m) (f : o ≃ n) : + (A.updateRow i r).submatrix e f = updateRow (A.submatrix e f) (e.symm i) fun i => r (f i) := + Eq.trans (by simp_rw [Equiv.apply_symm_apply]) (updateRow_submatrix_equiv A _ _ e f).symm +#align matrix.submatrix_update_row_equiv Matrix.submatrix_updateRow_equiv + +theorem updateColumn_submatrix_equiv [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : o) + (c : l → α) (e : l ≃ m) (f : o ≃ n) : updateColumn (A.submatrix e f) j c = + (A.updateColumn (f j) fun i => c (e.symm i)).submatrix e f := by + simpa only [← transpose_submatrix, updateRow_transpose] using + congr_arg transpose (updateRow_submatrix_equiv Aᵀ j c f e) +#align matrix.update_column_submatrix_equiv Matrix.updateColumn_submatrix_equiv + +theorem submatrix_updateColumn_equiv [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : n) + (c : m → α) (e : l ≃ m) (f : o ≃ n) : (A.updateColumn j c).submatrix e f = + updateColumn (A.submatrix e f) (f.symm j) fun i => c (e i) := + Eq.trans (by simp_rw [Equiv.apply_symm_apply]) (updateColumn_submatrix_equiv A _ _ e f).symm +#align matrix.submatrix_update_column_equiv Matrix.submatrix_updateColumn_equiv + +/-! `reindex` versions of the above `submatrix` lemmas for convenience. -/ + + +theorem updateRow_reindex [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : l) (r : o → α) + (e : m ≃ l) (f : n ≃ o) : + updateRow (reindex e f A) i r = reindex e f (A.updateRow (e.symm i) fun j => r (f j)) := + updateRow_submatrix_equiv _ _ _ _ _ +#align matrix.update_row_reindex Matrix.updateRow_reindex + +theorem reindex_updateRow [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : m) (r : n → α) + (e : m ≃ l) (f : n ≃ o) : + reindex e f (A.updateRow i r) = updateRow (reindex e f A) (e i) fun i => r (f.symm i) := + submatrix_updateRow_equiv _ _ _ _ _ +#align matrix.reindex_update_row Matrix.reindex_updateRow + +theorem updateColumn_reindex [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : o) (c : l → α) + (e : m ≃ l) (f : n ≃ o) : + updateColumn (reindex e f A) j c = reindex e f (A.updateColumn (f.symm j) fun i => c (e i)) := + updateColumn_submatrix_equiv _ _ _ _ _ +#align matrix.update_column_reindex Matrix.updateColumn_reindex + +theorem reindex_updateColumn [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : n) (c : m → α) + (e : m ≃ l) (f : n ≃ o) : + reindex e f (A.updateColumn j c) = updateColumn (reindex e f A) (f j) fun i => c (e.symm i) := + submatrix_updateColumn_equiv _ _ _ _ _ +#align matrix.reindex_update_column Matrix.reindex_updateColumn diff --git a/Mathlib/Data/Matroid/Basic.lean b/Mathlib/Data/Matroid/Basic.lean new file mode 100644 index 0000000000000..b7636ef907075 --- /dev/null +++ b/Mathlib/Data/Matroid/Basic.lean @@ -0,0 +1,922 @@ +/- +Copyright (c) 2023 Peter Nelson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Peter Nelson +-/ +import Mathlib.Data.Set.Card +import Mathlib.Order.Minimal +import Mathlib.Data.Matroid.Init + +/-! +# Matroids + +A `Matroid` is a structure that combinatorially abstracts +the notion of linear independence and dependence; +matroids have connections with graph theory, discrete optimization, +additive combinatorics and algebraic geometry. +Mathematically, a matroid `M` is a structure on a set `E` comprising a +collection of subsets of `E` called the bases of `M`, +where the bases are required to obey certain axioms. + +This file gives a definition of a matroid `M` in terms of its bases, +and some API relating independent sets (subsets of bases) and the notion of a +basis of a set `X` (a maximal independent subset of `X`). + +## Main definitions + +* a `Matroid α` on a type `α` is a structure comprising a 'ground set' + and a suitably behaved 'base' predicate. + +Given `M : Matroid α` ... + +* `M.E` denotes the ground set of `M`, which has type `Set α` +* For `B : Set α`, `M.Base B` means that `B` is a base of `M`. +* For `I : Set α`, `M.Indep I` means that `I` is independent in `M` + (that is, `I` is contained in a base of `M`). +* For `D : Set α`, `M.Dep D` means that `D` is contained in the ground set of `M` + but isn't independent. +* For `I : Set α` and `X : Set α`, `M.Basis I X` means that `I` is a maximal independent + subset of `X`. +* `M.Finite` means that `M` has finite ground set. +* `M.Nonempty` means that the ground set of `M` is nonempty. +* `FiniteRk M` means that the bases of `M` are finite. +* `InfiniteRk M` means that the bases of `M` are infinite. +* `RkPos M` means that the bases of `M` are nonempty. +* `Finitary M` means that a set is independent if and only if all its finite subsets are + independent. + +* `aesop_mat` : a tactic designed to prove `X ⊆ M.E` for some set `X` and matroid `M`. + +## Implementation details + +There are a few design decisions worth discussing. + +### Finiteness + The first is that our matroids are allowed to be infinite. + Unlike with many mathematical structures, this isn't such an obvious choice. + Finite matroids have been studied since the 1930's, + and there was never controversy as to what is and isn't an example of a finite matroid - + in fact, surprisingly many apparently different definitions of a matroid + give rise to the same class of objects. + + However, generalizing different definitions of a finite matroid + to the infinite in the obvious way (i.e. by simply allowing the ground set to be infinite) + gives a number of different notions of 'infinite matroid' that disagree with each other, + and that all lack nice properties. + Many different competing notions of infinite matroid were studied through the years; + in fact, the problem of which definition is the best was only really solved in 2013, + when Bruhn et al. [2] showed that there is a unique 'reasonable' notion of an infinite matroid; + these objects had been previously called 'B-matroids'. + These are defined by adding one carefully chosen axiom to the standard set, + and adapting existing axioms to not mention set cardinalities; + they enjoy nearly all the nice properties of standard finite matroids. + + Even though at least 90% of the literature is on finite matroids, + B-matroids are the definition we use, because they allow for additional generality, + nearly all theorems are still true and just as easy to state, + and (hopefully) the more general definition will prevent the need for a costly future refactor. + The disadvantage is that developing API for the finite case is harder work + (for instance, it is harder to prove that something is a matroid in the first place, + and one must deal with `ℕ∞` rather than `ℕ`). + For serious work on finite matroids, we provide the typeclasses + `[M.Finite]` and `[FiniteRk M]` and associated API. + +### Cardinality + Just as with bases of a vector space, + all bases of a finite matroid `M` are finite and have the same cardinality; + this cardinality is an important invariant known as the 'rank' of `M`. + For infinite matroids, bases are not in general equicardinal; + in fact the equicardinality of bases of infinite matroids is independent of ZFC [3]. + What is still true is that either all bases are finite and equicardinal, + or all bases are infinite. This means that the natural notion of 'size' + for a set in matroid theory is given by the function `Set.encard`, which + is the cardinality as a term in `ℕ∞`. We use this function extensively + in building the API; it is preferable to both `Set.ncard` and `Finset.card` + because it allows infinite sets to be handled without splitting into cases. + +### The ground `Set` + A last place where we make a consequential choice is making the ground set of a matroid + a structure field of type `Set α` (where `α` is the type of 'possible matroid elements') + rather than just having a type `α` of all the matroid elements. + This is because of how common it is to simultaneously consider + a number of matroids on different but related ground sets. + For example, a matroid `M` on ground set `E` can have its structure + 'restricted' to some subset `R ⊆ E` to give a smaller matroid `M ↾ R` with ground set `R`. + A statement like `(M ↾ R₁) ↾ R₂ = M ↾ R₂` is mathematically obvious. + But if the ground set of a matroid is a type, this doesn't typecheck, + and is only true up to canonical isomorphism. + Restriction is just the tip of the iceberg here; + one can also 'contract' and 'delete' elements and sets of elements + in a matroid to give a smaller matroid, + and in practice it is common to make statements like `M₁.E = M₂.E ∩ M₃.E` and + `((M ⟋ e) ↾ R) ⟋ C = M ⟋ (C ∪ {e}) ↾ R`. + Such things are a nightmare to work with unless `=` is actually propositional equality + (especially because the relevant coercions are usually between sets and not just elements). + + So the solution is that the ground set `M.E` has type `Set α`, + and there are elements of type `α` that aren't in the matroid. + The tradeoff is that for many statements, one now has to add + hypotheses of the form `X ⊆ M.E` to make sure than `X` is actually 'in the matroid', + rather than letting a 'type of matroid elements' take care of this invisibly. + It still seems that this is worth it. + The tactic `aesop_mat` exists specifically to discharge such goals + with minimal fuss (using default values). + The tactic works fairly well, but has room for improvement. + Even though the carrier set is written `M.E`, + we mirror common informal practice by referring explicitly to the `ground` set + rather than the notation `E` in lemma names. + + A related decision is to not have matroids themselves be a typeclass. + This would make things be notationally simpler + (having `Base` in the presence of `[Matroid α]` rather than `M.Base` for a term `M : Matroid α`) + but is again just too awkward when one has multiple matroids on the same type. + In fact, in regular written mathematics, + it is normal to explicitly indicate which matroid something is happening in, + so our notation mirrors common practice. + +## References + +[1] The standard text on matroid theory +[J. G. Oxley, Matroid Theory, Oxford University Press, New York, 2011.] + +[2] The robust axiomatic definition of infinite matroids +[H. Bruhn, R. Diestel, M. Kriesell, R. Pendavingh, P. Wollan, Axioms for infinite matroids, + Adv. Math 239 (2013), 18-46] + +[3] Equicardinality of matroid bases is independent of ZFC. +[N. Bowler, S. Geschke, Self-dual uniform matroids on infinite sets, + Proc. Amer. Math. Soc. 144 (2016), 459-471] +-/ + +set_option autoImplicit true + +open Set + +/-- A predicate `P` on sets satisfies the exchange property if, for all `X` and `Y` satisfying `P` + and all `a ∈ X \ Y`, there exists `b ∈ Y \ X` so that swapping `a` for `b` in `X` maintains `P`.-/ +def Matroid.ExchangeProperty {α : Type _} (P : Set α → Prop) : Prop := + ∀ X Y, P X → P Y → ∀ a ∈ X \ Y, ∃ b ∈ Y \ X, P (insert b (X \ {a})) + +/-- A set `X` has the maximal subset property for a predicate `P` if every subset of `X` satisfying + `P` is contained in a maximal subset of `X` satisfying `P`. -/ +def Matroid.ExistsMaximalSubsetProperty {α : Type _} (P : Set α → Prop) (X : Set α) : Prop := + ∀ I, P I → I ⊆ X → (maximals (· ⊆ ·) {Y | P Y ∧ I ⊆ Y ∧ Y ⊆ X}).Nonempty + +/-- A `Matroid α` is a ground set `E` of type `Set α`, and a nonempty collection of its subsets + satisfying the exchange property and the maximal subset property. Each such set is called a + `Base` of `M`. -/ +@[ext] structure Matroid (α : Type _) where + /-- `M` has a ground set `E`. -/ + (E : Set α) + /-- `M` has a predicate `Base` definining its bases -/ + (Base : Set α → Prop) + /-- There is at least one `Base` -/ + (exists_base' : ∃ B, Base B) + /-- For any bases `B`, `B'` and `e ∈ B \ B'`, there is some `f ∈ B' \ B` for which `B-e+f` + is a base -/ + (base_exchange' : Matroid.ExchangeProperty Base) + /-- Every subset `I` of a set `X` for which `I` is contained in a base is contained in a maximal + subset of `X` that is contained in a base. -/ + (maximality' : ∀ X, X ⊆ E → Matroid.ExistsMaximalSubsetProperty (∃ B, Base B ∧ · ⊆ B) X) + /-- every base is contained in the ground set -/ + (subset_ground' : ∀ B, Base B → B ⊆ E) + +namespace Matroid + +variable {α : Type _} {M : Matroid α} + +attribute [pp_dot] Base E + +/-- Typeclass for a matroid having finite ground set. Just a wrapper for `M.E.Finite`-/ +protected class Finite (M : Matroid α) : Prop where + /-- The ground set is finite -/ + (ground_finite : M.E.Finite) + +/-- Typeclass for a matroid having nonempty ground set. Just a wrapper for `M.E.Nonempty`-/ +protected class Nonempty (M : Matroid α) : Prop where + /-- The ground set is nonempty -/ + (ground_nonempty : M.E.Nonempty) + +theorem ground_nonempty (M : Matroid α) [M.Nonempty] : M.E.Nonempty := + Nonempty.ground_nonempty + +theorem ground_finite (M : Matroid α) [M.Finite] : M.E.Finite := + Finite.ground_finite + +theorem set_finite (M : Matroid α) [M.Finite] (X : Set α) (hX : X ⊆ M.E := by aesop) : X.Finite := + M.ground_finite.subset hX + +instance finite_of_finite [Finite α] {M : Matroid α} : M.Finite := + ⟨Set.toFinite _⟩ + +/-- A `FiniteRk` matroid is one whose bases are finite -/ +class FiniteRk (M : Matroid α) : Prop where + /-- There is a finite base -/ + exists_finite_base : ∃ B, M.Base B ∧ B.Finite + +instance finiteRk_of_finite (M : Matroid α) [M.Finite] : FiniteRk M := + ⟨ M.exists_base'.imp (fun B hB ↦ ⟨hB, M.set_finite B (M.subset_ground' _ hB)⟩) ⟩ + +/-- An `InfiniteRk` matroid is one whose bases are infinite. -/ +class InfiniteRk (M : Matroid α) : Prop where + /-- There is an infinite base -/ + exists_infinite_base : ∃ B, M.Base B ∧ B.Infinite + +/-- A `RkPos` matroid is one whose bases are nonempty. -/ +class RkPos (M : Matroid α) : Prop where + /-- The empty set isn't a base -/ + empty_not_base : ¬M.Base ∅ + +theorem rkPos_iff_empty_not_base : M.RkPos ↔ ¬M.Base ∅ := + ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ + +section exchange + +namespace ExchangeProperty + +variable {Base : Set α → Prop} (exch : ExchangeProperty Base) + +/-- A family of sets with the exchange property is an antichain. -/ +theorem antichain (hB : Base B) (hB' : Base B') (h : B ⊆ B') : B = B' := + h.antisymm (fun x hx ↦ by_contra + (fun hxB ↦ let ⟨_, hy, _⟩ := exch B' B hB' hB x ⟨hx, hxB⟩; hy.2 <| h hy.1)) + +theorem encard_diff_le_aux (exch : ExchangeProperty Base) (hB₁ : Base B₁) (hB₂ : Base B₂) : + (B₁ \ B₂).encard ≤ (B₂ \ B₁).encard := by + obtain (he | hinf | ⟨e, he, hcard⟩) := + (B₂ \ B₁).eq_empty_or_encard_eq_top_or_encard_diff_singleton_lt + · rw [exch.antichain hB₂ hB₁ (diff_eq_empty.mp he)] + · exact le_top.trans_eq hinf.symm + + obtain ⟨f, hf, hB'⟩ := exch B₂ B₁ hB₂ hB₁ e he + + have : encard (insert f (B₂ \ {e}) \ B₁) < encard (B₂ \ B₁) := by + rw [insert_diff_of_mem _ hf.1, diff_diff_comm]; exact hcard + + have hencard := encard_diff_le_aux exch hB₁ hB' + rw [insert_diff_of_mem _ hf.1, diff_diff_comm, ←union_singleton, ←diff_diff, diff_diff_right, + inter_singleton_eq_empty.mpr he.2, union_empty] at hencard + + rw [←encard_diff_singleton_add_one he, ←encard_diff_singleton_add_one hf] + exact add_le_add_right hencard 1 +termination_by _ => (B₂ \ B₁).encard + +/-- For any two sets `B₁`, `B₂` in a family with the exchange property, the differences `B₁ \ B₂` +and `B₂ \ B₁` have the same `ℕ∞`-cardinality. -/ +theorem encard_diff_eq (hB₁ : Base B₁) (hB₂ : Base B₂) : (B₁ \ B₂).encard = (B₂ \ B₁).encard := + (encard_diff_le_aux exch hB₁ hB₂).antisymm (encard_diff_le_aux exch hB₂ hB₁) + +/-- Any two sets `B₁`, `B₂` in a family with the exchange property have the same +`ℕ∞`-cardinality. -/ +theorem encard_base_eq (hB₁ : Base B₁) (hB₂ : Base B₂) : B₁.encard = B₂.encard := by +rw [←encard_diff_add_encard_inter B₁ B₂, exch.encard_diff_eq hB₁ hB₂, inter_comm, + encard_diff_add_encard_inter] + +end ExchangeProperty + +end exchange + +section aesop + +/-- The `aesop_mat` tactic attempts to prove a set is contained in the ground set of a matroid. + It uses a `[Matroid]` ruleset, and is allowed to fail. -/ +macro (name := aesop_mat) "aesop_mat" c:Aesop.tactic_clause* : tactic => +`(tactic| + aesop $c* (options := { terminal := true }) + (rule_sets [$(Lean.mkIdent `Matroid):ident])) + +/- We add a number of trivial lemmas (deliberately specialized to statements in terms of the + ground set of a matroid) to the ruleset `Matroid` for `aesop`. -/ + +@[aesop unsafe 5% (rule_sets [Matroid])] +private theorem inter_right_subset_ground (hX : X ⊆ M.E) : + X ∩ Y ⊆ M.E := (inter_subset_left _ _).trans hX + +@[aesop unsafe 5% (rule_sets [Matroid])] +private theorem inter_left_subset_ground (hX : X ⊆ M.E) : + Y ∩ X ⊆ M.E := (inter_subset_right _ _).trans hX + +@[aesop unsafe 5% (rule_sets [Matroid])] +private theorem diff_subset_ground (hX : X ⊆ M.E) : X \ Y ⊆ M.E := + (diff_subset _ _).trans hX + +@[aesop unsafe 10% (rule_sets [Matroid])] +private theorem ground_diff_subset_ground : M.E \ X ⊆ M.E := + diff_subset_ground rfl.subset + +@[aesop unsafe 10% (rule_sets [Matroid])] +private theorem singleton_subset_ground (he : e ∈ M.E) : {e} ⊆ M.E := + singleton_subset_iff.mpr he + +@[aesop unsafe 5% (rule_sets [Matroid])] +private theorem subset_ground_of_subset (hXY : X ⊆ Y) (hY : Y ⊆ M.E) : X ⊆ M.E := + hXY.trans hY + +@[aesop unsafe 5% (rule_sets [Matroid])] +private theorem mem_ground_of_mem_of_subset (hX : X ⊆ M.E) (heX : e ∈ X) : e ∈ M.E := + hX heX + +@[aesop safe (rule_sets [Matroid])] +private theorem insert_subset_ground {e : α} {X : Set α} {M : Matroid α} + (he : e ∈ M.E) (hX : X ⊆ M.E) : insert e X ⊆ M.E := + insert_subset he hX + +@[aesop safe (rule_sets [Matroid])] +private theorem ground_subset_ground {M : Matroid α} : M.E ⊆ M.E := + rfl.subset + +attribute [aesop safe (rule_sets [Matroid])] empty_subset union_subset iUnion_subset + +end aesop +section Base + +@[aesop unsafe 10% (rule_sets [Matroid])] +theorem Base.subset_ground (hB : M.Base B) : B ⊆ M.E := + M.subset_ground' B hB + +theorem exists_base (M : Matroid α) : ∃ B, M.Base B := M.exists_base' + +theorem Base.exchange (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) (hx : e ∈ B₁ \ B₂) : + ∃ y ∈ B₂ \ B₁, M.Base (insert y (B₁ \ {e})) := + M.base_exchange' B₁ B₂ hB₁ hB₂ _ hx + +theorem Base.exchange_mem (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) (hxB₁ : e ∈ B₁) (hxB₂ : e ∉ B₂) : + ∃ y, (y ∈ B₂ ∧ y ∉ B₁) ∧ M.Base (insert y (B₁ \ {e})) := + by simpa using hB₁.exchange hB₂ ⟨hxB₁, hxB₂⟩ + +theorem Base.eq_of_subset_base (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) (hB₁B₂ : B₁ ⊆ B₂) : + B₁ = B₂ := + M.base_exchange'.antichain hB₁ hB₂ hB₁B₂ + +theorem Base.not_base_of_ssubset (hB : M.Base B) (hX : X ⊂ B) : ¬ M.Base X := + fun h ↦ hX.ne (h.eq_of_subset_base hB hX.subset) + +theorem Base.insert_not_base (hB : M.Base B) (heB : e ∉ B) : ¬ M.Base (insert e B) := + fun h ↦ h.not_base_of_ssubset (ssubset_insert heB) hB + +theorem Base.encard_diff_comm (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) : + (B₁ \ B₂).encard = (B₂ \ B₁).encard := + M.base_exchange'.encard_diff_eq hB₁ hB₂ + +theorem Base.ncard_diff_comm (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) : + (B₁ \ B₂).ncard = (B₂ \ B₁).ncard := by + rw [ncard_def, hB₁.encard_diff_comm hB₂, ←ncard_def] + +theorem Base.card_eq_card_of_base (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) : + B₁.encard = B₂.encard := by + rw [M.base_exchange'.encard_base_eq hB₁ hB₂] + +theorem Base.ncard_eq_ncard_of_base (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) : B₁.ncard = B₂.ncard := by + rw [ncard_def B₁, hB₁.card_eq_card_of_base hB₂, ←ncard_def] + +theorem Base.finite_of_finite (hB : M.Base B) (h : B.Finite) (hB' : M.Base B') : B'.Finite := + (finite_iff_finite_of_encard_eq_encard (hB.card_eq_card_of_base hB')).mp h + +theorem Base.infinite_of_infinite (hB : M.Base B) (h : B.Infinite) (hB₁ : M.Base B₁) : + B₁.Infinite := + by_contra (fun hB_inf ↦ (hB₁.finite_of_finite (not_infinite.mp hB_inf) hB).not_infinite h) + +theorem Base.finite [FiniteRk M] (hB : M.Base B) : B.Finite := + let ⟨B₀,hB₀⟩ := ‹FiniteRk M›.exists_finite_base + hB₀.1.finite_of_finite hB₀.2 hB + +theorem Base.infinite [InfiniteRk M] (hB : M.Base B) : B.Infinite := + let ⟨B₀,hB₀⟩ := ‹InfiniteRk M›.exists_infinite_base + hB₀.1.infinite_of_infinite hB₀.2 hB + +theorem empty_not_base [h : RkPos M] : ¬M.Base ∅ := + h.empty_not_base + +theorem Base.nonempty [RkPos M] (hB : M.Base B) : B.Nonempty := by + rw [nonempty_iff_ne_empty]; rintro rfl; exact M.empty_not_base hB + +theorem Base.rkPos_of_nonempty (hB : M.Base B) (h : B.Nonempty) : M.RkPos := by + rw [rkPos_iff_empty_not_base] + intro he + obtain rfl := he.eq_of_subset_base hB (empty_subset B) + simp at h + +theorem Base.finiteRk_of_finite (hB : M.Base B) (hfin : B.Finite) : FiniteRk M := + ⟨⟨B, hB, hfin⟩⟩ + +theorem Base.infiniteRk_of_infinite (hB : M.Base B) (h : B.Infinite) : InfiniteRk M := + ⟨⟨B, hB, h⟩⟩ + +theorem not_finiteRk (M : Matroid α) [InfiniteRk M] : ¬ FiniteRk M := by + intro h; obtain ⟨B,hB⟩ := M.exists_base; exact hB.infinite hB.finite + +theorem not_infiniteRk (M : Matroid α) [FiniteRk M] : ¬ InfiniteRk M := by + intro h; obtain ⟨B,hB⟩ := M.exists_base; exact hB.infinite hB.finite + +theorem finite_or_infiniteRk (M : Matroid α) : FiniteRk M ∨ InfiniteRk M := + let ⟨B, hB⟩ := M.exists_base + B.finite_or_infinite.elim + (Or.inl ∘ hB.finiteRk_of_finite) (Or.inr ∘ hB.infiniteRk_of_infinite) + +theorem Base.diff_finite_comm (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) : + (B₁ \ B₂).Finite ↔ (B₂ \ B₁).Finite := + finite_iff_finite_of_encard_eq_encard (hB₁.encard_diff_comm hB₂) + +theorem Base.diff_infinite_comm (hB₁ : M.Base B₁) (hB₂ : M.Base B₂) : + (B₁ \ B₂).Infinite ↔ (B₂ \ B₁).Infinite := + infinite_iff_infinite_of_encard_eq_encard (hB₁.encard_diff_comm hB₂) + +theorem eq_of_base_iff_base_forall {M₁ M₂ : Matroid α} (hE : M₁.E = M₂.E) + (h : ∀ B, B ⊆ M₁.E → (M₁.Base B ↔ M₂.Base B)) : M₁ = M₂ := by + apply Matroid.ext _ _ hE + ext B + exact ⟨fun h' ↦ (h _ h'.subset_ground).mp h', + fun h' ↦ (h _ (h'.subset_ground.trans_eq hE.symm)).mpr h'⟩ + +theorem base_compl_iff_mem_maximals_disjoint_base (hB : B ⊆ M.E := by aesop_mat) : + M.Base (M.E \ B) ↔ B ∈ maximals (· ⊆ ·) {I | I ⊆ M.E ∧ ∃ B, M.Base B ∧ Disjoint I B} := by + simp_rw [mem_maximals_setOf_iff, and_iff_right hB, and_imp, forall_exists_index] + refine' ⟨fun h ↦ ⟨⟨_, h, disjoint_sdiff_right⟩, + fun I hI B' ⟨hB', hIB'⟩ hBI ↦ hBI.antisymm _⟩, fun ⟨⟨ B', hB', hBB'⟩,h⟩ ↦ _⟩ + · rw [hB'.eq_of_subset_base h, ←subset_compl_iff_disjoint_right, diff_eq, compl_inter, + compl_compl] at hIB' + · exact fun e he ↦ (hIB' he).elim (fun h' ↦ (h' (hI he)).elim) id + rw [subset_diff, and_iff_right hB'.subset_ground, disjoint_comm] + exact disjoint_of_subset_left hBI hIB' + rw [h (diff_subset M.E B') B' ⟨hB', disjoint_sdiff_left⟩] + · simpa [hB'.subset_ground] + simp [subset_diff, hB, hBB'] + +end Base +section dep_indep + +/-- A set is `Indep`endent if it is contained in a `Base`. -/ +@[pp_dot] def Indep (M : Matroid α) (I : Set α) : Prop := ∃ B, M.Base B ∧ I ⊆ B + +/-- A subset of `M.E` is `Dep`endent if it is not `Indep`endent . -/ +@[pp_dot] def Dep (M : Matroid α) (D : Set α) : Prop := ¬M.Indep D ∧ D ⊆ M.E + +theorem indep_iff_subset_base : M.Indep I ↔ ∃ B, M.Base B ∧ I ⊆ B := Iff.rfl + +theorem setOf_indep_eq (M : Matroid α) : {I | M.Indep I} = lowerClosure ({B | M.Base B}) := rfl + +theorem dep_iff : M.Dep D ↔ ¬M.Indep D ∧ D ⊆ M.E := Iff.rfl + +theorem setOf_dep_eq (M : Matroid α) : {D | M.Dep D} = {I | M.Indep I}ᶜ ∩ Iic M.E := rfl + +@[aesop unsafe 30% (rule_sets [Matroid])] +theorem Indep.subset_ground (hI : M.Indep I) : I ⊆ M.E := by + obtain ⟨B, hB, hIB⟩ := hI + exact hIB.trans hB.subset_ground + +@[aesop unsafe 20% (rule_sets [Matroid])] +theorem Dep.subset_ground (hD : M.Dep D) : D ⊆ M.E := + hD.2 + +theorem indep_or_dep (hX : X ⊆ M.E := by aesop_mat) : M.Indep X ∨ M.Dep X := by + rw [Dep, and_iff_left hX] + apply em + +theorem Indep.not_dep (hI : M.Indep I) : ¬ M.Dep I := + fun h ↦ h.1 hI + +theorem Dep.not_indep (hD : M.Dep D) : ¬ M.Indep D := + hD.1 + +theorem dep_of_not_indep (hD : ¬ M.Indep D) (hDE : D ⊆ M.E := by aesop_mat) : M.Dep D := + ⟨hD, hDE⟩ + +theorem indep_of_not_dep (hI : ¬ M.Dep I) (hIE : I ⊆ M.E := by aesop_mat) : M.Indep I := + by_contra (fun h ↦ hI ⟨h, hIE⟩) + +@[simp] theorem not_dep_iff (hX : X ⊆ M.E := by aesop_mat) : ¬ M.Dep X ↔ M.Indep X := by + rw [Dep, and_iff_left hX, not_not] + +@[simp] theorem not_indep_iff (hX : X ⊆ M.E := by aesop_mat) : ¬ M.Indep X ↔ M.Dep X := by + rw [Dep, and_iff_left hX] + +theorem indep_iff_not_dep : M.Indep I ↔ ¬M.Dep I ∧ I ⊆ M.E := by + rw [dep_iff, not_and, not_imp_not] + exact ⟨fun h ↦ ⟨fun _ ↦ h, h.subset_ground⟩, fun h ↦ h.1 h.2⟩ + +theorem Indep.exists_base_superset (hI : M.Indep I) : ∃ B, M.Base B ∧ I ⊆ B := + hI + +theorem Indep.subset (hJ : M.Indep J) (hIJ : I ⊆ J) : M.Indep I := + let ⟨B, hB, hJB⟩ := hJ + ⟨B, hB, hIJ.trans hJB⟩ + +theorem Dep.superset (hD : M.Dep D) (hDX : D ⊆ X) (hXE : X ⊆ M.E := by aesop_mat) : M.Dep X := + dep_of_not_indep (fun hI ↦ (hI.subset hDX).not_dep hD) + +@[simp] theorem empty_indep (M : Matroid α) : M.Indep ∅ := + Exists.elim M.exists_base (fun B hB ↦ ⟨_, hB, B.empty_subset⟩) + +theorem Dep.nonempty (hD : M.Dep D) : D.Nonempty := by + rw [nonempty_iff_ne_empty]; rintro rfl; exact hD.not_indep M.empty_indep + +theorem Indep.finite [FiniteRk M] (hI : M.Indep I) : I.Finite := + let ⟨_, hB, hIB⟩ := hI + hB.finite.subset hIB + +theorem Indep.rkPos_of_nonempty (hI : M.Indep I) (hne : I.Nonempty) : M.RkPos := by + obtain ⟨B, hB, hIB⟩ := hI.exists_base_superset + exact hB.rkPos_of_nonempty (hne.mono hIB) + +theorem Indep.inter_right (hI : M.Indep I) (X : Set α) : M.Indep (I ∩ X) := + hI.subset (inter_subset_left _ _) + +theorem Indep.inter_left (hI : M.Indep I) (X : Set α) : M.Indep (X ∩ I) := + hI.subset (inter_subset_right _ _) + +theorem Indep.diff (hI : M.Indep I) (X : Set α) : M.Indep (I \ X) := + hI.subset (diff_subset _ _) + +theorem Base.indep (hB : M.Base B) : M.Indep B := + ⟨B, hB, subset_rfl⟩ + +theorem Base.eq_of_subset_indep (hB : M.Base B) (hI : M.Indep I) (hBI : B ⊆ I) : B = I := + let ⟨B', hB', hB'I⟩ := hI + hBI.antisymm (by rwa [hB.eq_of_subset_base hB' (hBI.trans hB'I)]) + +theorem base_iff_maximal_indep : M.Base B ↔ M.Indep B ∧ ∀ I, M.Indep I → B ⊆ I → B = I := by + refine' ⟨fun h ↦ ⟨h.indep, fun _ ↦ h.eq_of_subset_indep ⟩, fun h ↦ _⟩ + obtain ⟨⟨B', hB', hBB'⟩, h⟩ := h + rwa [h _ hB'.indep hBB'] + +theorem setOf_base_eq_maximals_setOf_indep : {B | M.Base B} = maximals (· ⊆ ·) {I | M.Indep I} := by + ext B; rw [mem_maximals_setOf_iff, mem_setOf, base_iff_maximal_indep] + +theorem Indep.base_of_maximal (hI : M.Indep I) (h : ∀ J, M.Indep J → I ⊆ J → I = J) : M.Base I := + base_iff_maximal_indep.mpr ⟨hI,h⟩ + +theorem Base.dep_of_ssubset (hB : M.Base B) (h : B ⊂ X) (hX : X ⊆ M.E := by aesop_mat) : M.Dep X := + ⟨λ hX ↦ h.ne (hB.eq_of_subset_indep hX h.subset), hX⟩ + +theorem Base.dep_of_insert (hB : M.Base B) (heB : e ∉ B) (he : e ∈ M.E := by aesop_mat) : + M.Dep (insert e B) := hB.dep_of_ssubset (ssubset_insert heB) (insert_subset he hB.subset_ground) + +/-- If the difference of two Bases is a singleton, then they differ by an insertion/removal -/ +theorem Base.eq_exchange_of_diff_eq_singleton (hB : M.Base B) (hB' : M.Base B') (h : B \ B' = {e}) : + ∃ f ∈ B' \ B, B' = (insert f B) \ {e} := by + obtain ⟨f, hf, hb⟩ := hB.exchange hB' (h.symm.subset (mem_singleton e)) + have hne : f ≠ e := by rintro rfl; exact hf.2 (h.symm.subset (mem_singleton f)).1 + rw [insert_diff_singleton_comm hne] at hb + refine ⟨f, hf, (hb.eq_of_subset_base hB' ?_).symm⟩ + rw [diff_subset_iff, insert_subset_iff, union_comm, ←diff_subset_iff, h, and_iff_left rfl.subset] + exact Or.inl hf.1 + +theorem Base.exchange_base_of_indep (hB : M.Base B) (hf : f ∉ B) + (hI : M.Indep (insert f (B \ {e}))) : M.Base (insert f (B \ {e})) := by + obtain ⟨B', hB', hIB'⟩ := hI.exists_base_superset + have hcard := hB'.encard_diff_comm hB + rw [insert_subset_iff, ←diff_eq_empty, diff_diff_comm, diff_eq_empty, subset_singleton_iff_eq] + at hIB' + obtain ⟨hfB, (h | h)⟩ := hIB' + · rw [h, encard_empty, encard_eq_zero, eq_empty_iff_forall_not_mem] at hcard + exact (hcard f ⟨hfB, hf⟩).elim + rw [h, encard_singleton, encard_eq_one] at hcard + obtain ⟨x, hx⟩ := hcard + obtain (rfl : f = x) := hx.subset ⟨hfB, hf⟩ + simp_rw [←h, ←singleton_union, ←hx, sdiff_sdiff_right_self, inf_eq_inter, inter_comm B, + diff_union_inter] + exact hB' + +theorem Base.exchange_base_of_indep' (hB : M.Base B) (he : e ∈ B) (hf : f ∉ B) + (hI : M.Indep (insert f B \ {e})) : M.Base (insert f B \ {e}) := by + have hfe : f ≠ e := by rintro rfl; exact hf he + rw [←insert_diff_singleton_comm hfe] at * + exact hB.exchange_base_of_indep hf hI + +theorem Base.insert_dep (hB : M.Base B) (h : e ∈ M.E \ B) : M.Dep (insert e B) := by + rw [←not_indep_iff (insert_subset h.1 hB.subset_ground)] + exact h.2 ∘ (fun hi ↦ insert_eq_self.mp (hB.eq_of_subset_indep hi (subset_insert e B)).symm) + +theorem Indep.exists_insert_of_not_base (hI : M.Indep I) (hI' : ¬M.Base I) (hB : M.Base B) : + ∃ e ∈ B \ I, M.Indep (insert e I) := by + obtain ⟨B', hB', hIB'⟩ := hI.exists_base_superset + obtain ⟨x, hxB', hx⟩ := exists_of_ssubset (hIB'.ssubset_of_ne (by (rintro rfl; exact hI' hB'))) + obtain (hxB | hxB) := em (x ∈ B) + · exact ⟨x, ⟨hxB, hx⟩ , ⟨B', hB', insert_subset hxB' hIB'⟩⟩ + obtain ⟨e,he, hBase⟩ := hB'.exchange hB ⟨hxB',hxB⟩ + exact ⟨e, ⟨he.1, not_mem_subset hIB' he.2⟩, + ⟨_, hBase, insert_subset_insert (subset_diff_singleton hIB' hx)⟩⟩ + +theorem ground_indep_iff_base : M.Indep M.E ↔ M.Base M.E := + ⟨fun h ↦ h.base_of_maximal (fun _ hJ hEJ ↦ hEJ.antisymm hJ.subset_ground), Base.indep⟩ + +theorem Base.exists_insert_of_ssubset (hB : M.Base B) (hIB : I ⊂ B) (hB' : M.Base B') : + ∃ e ∈ B' \ I, M.Indep (insert e I) := +(hB.indep.subset hIB.subset).exists_insert_of_not_base + (fun hI ↦ hIB.ne (hI.eq_of_subset_base hB hIB.subset)) hB' + +theorem eq_of_indep_iff_indep_forall {M₁ M₂ : Matroid α} (hE : M₁.E = M₂.E) + (h : ∀ I, I ⊆ M₁.E → (M₁.Indep I ↔ M₂.Indep I)) : M₁ = M₂ := + let h' : ∀ I, M₁.Indep I ↔ M₂.Indep I := fun I ↦ + (em (I ⊆ M₁.E)).elim (h I) (fun h' ↦ iff_of_false (fun hi ↦ h' (hi.subset_ground)) + (fun hi ↦ h' (hi.subset_ground.trans_eq hE.symm))) + eq_of_base_iff_base_forall hE (fun B _ ↦ by simp_rw [base_iff_maximal_indep, h']) + +theorem eq_iff_indep_iff_indep_forall {M₁ M₂ : Matroid α} : + M₁ = M₂ ↔ (M₁.E = M₂.E) ∧ ∀ I, I ⊆ M₁.E → (M₁.Indep I ↔ M₂.Indep I) := +⟨fun h ↦ by (subst h; simp), fun h ↦ eq_of_indep_iff_indep_forall h.1 h.2⟩ + +/-- A `Finitary` matroid is one where a set is independent if and only if it all + its finite subsets are independent, or equivalently a matroid whose circuits are finite. -/ +class Finitary (M : Matroid α) : Prop where + /-- `I` is independent if all its finite subsets are independent. -/ + indep_of_forall_finite : ∀ I, (∀ J, J ⊆ I → J.Finite → M.Indep J) → M.Indep I + +theorem indep_of_forall_finite_subset_indep {M : Matroid α} [Finitary M] (I : Set α) + (h : ∀ J, J ⊆ I → J.Finite → M.Indep J) : M.Indep I := + Finitary.indep_of_forall_finite I h + +theorem indep_iff_forall_finite_subset_indep {M : Matroid α} [Finitary M] : + M.Indep I ↔ ∀ J, J ⊆ I → J.Finite → M.Indep J := + ⟨fun h _ hJI _ ↦ h.subset hJI, Finitary.indep_of_forall_finite I⟩ + +instance finitary_of_finiteRk {M : Matroid α} [FiniteRk M] : Finitary M := +⟨ by + refine fun I hI ↦ I.finite_or_infinite.elim (hI _ Subset.rfl) (fun h ↦ False.elim ?_) + obtain ⟨B, hB⟩ := M.exists_base + obtain ⟨I₀, hI₀I, hI₀fin, hI₀card⟩ := h.exists_subset_ncard_eq (B.ncard + 1) + obtain ⟨B', hB', hI₀B'⟩ := hI _ hI₀I hI₀fin + have hle := ncard_le_of_subset hI₀B' hB'.finite + rw [hI₀card, hB'.ncard_eq_ncard_of_base hB, Nat.add_one_le_iff] at hle + exact hle.ne rfl ⟩ + +/-- Matroids obey the maximality axiom -/ +theorem existsMaximalSubsetProperty_indep (M : Matroid α) : + ∀ X, X ⊆ M.E → ExistsMaximalSubsetProperty M.Indep X := + M.maximality' + +end dep_indep + +section Basis + +/-- A Basis for a set `X ⊆ M.E` is a maximal independent subset of `X` + (Often in the literature, the word 'Basis' is used to refer to what we call a 'Base'). -/ +@[pp_dot] def Basis (M : Matroid α) (I X : Set α) : Prop := + I ∈ maximals (· ⊆ ·) {A | M.Indep A ∧ A ⊆ X} ∧ X ⊆ M.E + +/-- A `Basis'` is a basis without the requirement that `X ⊆ M.E`. This is convenient for some + API building, especially when working with rank and closure. -/ +@[pp_dot] def Basis' (M : Matroid α) (I X : Set α) : Prop := + I ∈ maximals (· ⊆ ·) {A | M.Indep A ∧ A ⊆ X} + +theorem Basis'.indep (hI : M.Basis' I X) : M.Indep I := + hI.1.1 + +theorem Basis.indep (hI : M.Basis I X) : M.Indep I := + hI.1.1.1 + +theorem Basis.subset (hI : M.Basis I X) : I ⊆ X := + hI.1.1.2 + +theorem Basis.basis' (hI : M.Basis I X) : M.Basis' I X := + hI.1 + +theorem Basis'.basis (hI : M.Basis' I X) (hX : X ⊆ M.E := by aesop_mat) : M.Basis I X := + ⟨hI, hX⟩ + +theorem Basis'.subset (hI : M.Basis' I X) : I ⊆ X := + hI.1.2 + +theorem setOf_basis_eq (M : Matroid α) (hX : X ⊆ M.E := by aesop_mat) : + {I | M.Basis I X} = maximals (· ⊆ ·) ({I | M.Indep I} ∩ Iic X) := by + ext I; simp [Matroid.Basis, maximals, iff_true_intro hX] + +@[aesop unsafe 15% (rule_sets [Matroid])] +theorem Basis.subset_ground (hI : M.Basis I X) : X ⊆ M.E := + hI.2 + +theorem Basis.basis_inter_ground (hI : M.Basis I X) : M.Basis I (X ∩ M.E) := by + convert hI + rw [inter_eq_self_of_subset_left hI.subset_ground] + +@[aesop unsafe 15% (rule_sets [Matroid])] +theorem Basis.left_subset_ground (hI : M.Basis I X) : I ⊆ M.E := + hI.indep.subset_ground + +theorem Basis.eq_of_subset_indep (hI : M.Basis I X) (hJ : M.Indep J) (hIJ : I ⊆ J) (hJX : J ⊆ X) : + I = J := + hIJ.antisymm (hI.1.2 ⟨hJ, hJX⟩ hIJ) + +theorem Basis.Finite (hI : M.Basis I X) [FiniteRk M] : I.Finite := hI.indep.finite + +theorem basis_iff' : + M.Basis I X ↔ (M.Indep I ∧ I ⊆ X ∧ ∀ J, M.Indep J → I ⊆ J → J ⊆ X → I = J) ∧ X ⊆ M.E := by + simp [Basis, mem_maximals_setOf_iff, and_assoc, and_congr_left_iff, and_imp, + and_congr_left_iff, and_congr_right_iff, @Imp.swap (_ ⊆ X)] + +theorem basis_iff (hX : X ⊆ M.E := by aesop_mat) : + M.Basis I X ↔ (M.Indep I ∧ I ⊆ X ∧ ∀ J, M.Indep J → I ⊆ J → J ⊆ X → I = J) := +by rw [basis_iff', and_iff_left hX] + +theorem basis'_iff_basis_inter_ground : M.Basis' I X ↔ M.Basis I (X ∩ M.E) := by + rw [Basis', Basis, and_iff_left (inter_subset_right _ _)] + convert Iff.rfl using 3 + ext I + simp only [subset_inter_iff, mem_setOf_eq, and_congr_right_iff, and_iff_left_iff_imp] + exact fun h _ ↦ h.subset_ground + +theorem basis'_iff_basis (hX : X ⊆ M.E := by aesop_mat) : M.Basis' I X ↔ M.Basis I X := by + rw [basis'_iff_basis_inter_ground, inter_eq_self_of_subset_left hX] + +theorem basis_iff_basis'_subset_ground : M.Basis I X ↔ M.Basis' I X ∧ X ⊆ M.E := + ⟨fun h ↦ ⟨h.basis', h.subset_ground⟩, fun h ↦ (basis'_iff_basis h.2).mp h.1⟩ + +theorem Basis'.basis_inter_ground (hIX : M.Basis' I X) : M.Basis I (X ∩ M.E) := + basis'_iff_basis_inter_ground.mp hIX + +theorem Basis'.eq_of_subset_indep (hI : M.Basis' I X) (hJ : M.Indep J) (hIJ : I ⊆ J) + (hJX : J ⊆ X) : I = J := + hIJ.antisymm (hI.2 ⟨hJ, hJX⟩ hIJ) + +theorem basis_iff_mem_maximals (hX : X ⊆ M.E := by aesop_mat): + M.Basis I X ↔ I ∈ maximals (· ⊆ ·) {I | M.Indep I ∧ I ⊆ X} := by + rw [Basis, and_iff_left hX] + +theorem basis_iff_mem_maximals_Prop (hX : X ⊆ M.E := by aesop_mat): + M.Basis I X ↔ I ∈ maximals (· ⊆ ·) (fun I ↦ M.Indep I ∧ I ⊆ X) := + basis_iff_mem_maximals + +theorem Indep.basis_of_maximal_subset (hI : M.Indep I) (hIX : I ⊆ X) + (hmax : ∀ ⦃J⦄, M.Indep J → I ⊆ J → J ⊆ X → J ⊆ I) (hX : X ⊆ M.E := by aesop_mat) : + M.Basis I X := by + rw [basis_iff (by aesop_mat : X ⊆ M.E), and_iff_right hI, and_iff_right hIX] + exact fun J hJ hIJ hJX ↦ hIJ.antisymm (hmax hJ hIJ hJX) + +theorem Basis.basis_subset (hI : M.Basis I X) (hIY : I ⊆ Y) (hYX : Y ⊆ X) : M.Basis I Y := by + rw [basis_iff (hYX.trans hI.subset_ground), and_iff_right hI.indep, and_iff_right hIY] + exact fun J hJ hIJ hJY ↦ hI.eq_of_subset_indep hJ hIJ (hJY.trans hYX) + +@[simp] theorem basis_self_iff_indep : M.Basis I I ↔ M.Indep I := by + rw [basis_iff', and_iff_right rfl.subset, and_assoc, and_iff_left_iff_imp] + exact fun hi ↦ ⟨fun _ _ ↦ subset_antisymm, hi.subset_ground⟩ + +theorem Indep.basis_self (h : M.Indep I) : M.Basis I I := + basis_self_iff_indep.mpr h + +@[simp] theorem basis_empty_iff (M : Matroid α) : M.Basis I ∅ ↔ I = ∅ := + ⟨fun h ↦ subset_empty_iff.mp h.subset, fun h ↦ by (rw [h]; exact M.empty_indep.basis_self)⟩ + +theorem Basis.dep_of_ssubset (hI : M.Basis I X) (hIY : I ⊂ Y) (hYX : Y ⊆ X) : M.Dep Y := by + have : X ⊆ M.E := hI.subset_ground + rw [←not_indep_iff] + exact fun hY ↦ hIY.ne (hI.eq_of_subset_indep hY hIY.subset hYX) + +theorem Basis.insert_dep (hI : M.Basis I X) (he : e ∈ X \ I) : M.Dep (insert e I) := + hI.dep_of_ssubset (ssubset_insert he.2) (insert_subset he.1 hI.subset) + +theorem Basis.mem_of_insert_indep (hI : M.Basis I X) (he : e ∈ X) (hIe : M.Indep (insert e I)) : + e ∈ I := + by_contra (fun heI ↦ (hI.insert_dep ⟨he, heI⟩).not_indep hIe) + +theorem Basis.not_basis_of_ssubset (hI : M.Basis I X) (hJI : J ⊂ I) : ¬ M.Basis J X := + fun h ↦ hJI.ne (h.eq_of_subset_indep hI.indep hJI.subset hI.subset) + +theorem Indep.subset_basis_of_subset (hI : M.Indep I) (hIX : I ⊆ X) (hX : X ⊆ M.E := by aesop_mat) : + ∃ J, M.Basis J X ∧ I ⊆ J := by + obtain ⟨J, ⟨(hJ : M.Indep J),hIJ,hJX⟩, hJmax⟩ := M.maximality' X hX I hI hIX + use J + rw [and_iff_left hIJ, basis_iff, and_iff_right hJ, and_iff_right hJX] + exact fun K hK hJK hKX ↦ hJK.antisymm (hJmax ⟨hK, hIJ.trans hJK, hKX⟩ hJK) + +theorem Indep.subset_basis'_of_subset (hI : M.Indep I) (hIX : I ⊆ X) : + ∃ J, M.Basis' J X ∧ I ⊆ J := by + simp_rw [basis'_iff_basis_inter_ground] + exact hI.subset_basis_of_subset (subset_inter hIX hI.subset_ground) + +theorem exists_basis (M : Matroid α) (X : Set α) (hX : X ⊆ M.E := by aesop_mat) : + ∃ I, M.Basis I X := + let ⟨_, hI, _⟩ := M.empty_indep.subset_basis_of_subset (empty_subset X) + ⟨_,hI⟩ + +theorem exists_basis' (M : Matroid α) (X : Set α) : ∃ I, M.Basis' I X := + let ⟨_, hI, _⟩ := M.empty_indep.subset_basis'_of_subset (empty_subset X) + ⟨_,hI⟩ + +theorem exists_basis_subset_basis (M : Matroid α) (hXY : X ⊆ Y) (hY : Y ⊆ M.E := by aesop_mat) : + ∃ I J, M.Basis I X ∧ M.Basis J Y ∧ I ⊆ J := by + obtain ⟨I, hI⟩ := M.exists_basis X (hXY.trans hY) + obtain ⟨J, hJ, hIJ⟩ := hI.indep.subset_basis_of_subset (hI.subset.trans hXY) + exact ⟨_, _, hI, hJ, hIJ⟩ + +theorem Basis.exists_basis_inter_eq_of_superset (hI : M.Basis I X) (hXY : X ⊆ Y) + (hY : Y ⊆ M.E := by aesop_mat) : ∃ J, M.Basis J Y ∧ J ∩ X = I := by + obtain ⟨J, hJ, hIJ⟩ := hI.indep.subset_basis_of_subset (hI.subset.trans hXY) + refine ⟨J, hJ, subset_antisymm ?_ (subset_inter hIJ hI.subset)⟩ + exact fun e he ↦ hI.mem_of_insert_indep he.2 (hJ.indep.subset (insert_subset he.1 hIJ)) + +theorem exists_basis_union_inter_basis (M : Matroid α) (X Y : Set α) (hX : X ⊆ M.E := by aesop_mat) + (hY : Y ⊆ M.E := by aesop_mat) : ∃ I, M.Basis I (X ∪ Y) ∧ M.Basis (I ∩ Y) Y := + let ⟨J, hJ⟩ := M.exists_basis Y + (hJ.exists_basis_inter_eq_of_superset (subset_union_right X Y)).imp + (fun I hI ↦ ⟨hI.1, by rwa [hI.2]⟩) + +theorem Indep.eq_of_basis (hI : M.Indep I) (hJ : M.Basis J I) : J = I := + hJ.eq_of_subset_indep hI hJ.subset rfl.subset + +theorem Basis.exists_base (hI : M.Basis I X) : ∃ B, M.Base B ∧ I = B ∩ X := + let ⟨B,hB, hIB⟩ := hI.indep + ⟨B, hB, subset_antisymm (subset_inter hIB hI.subset) + (by rw [hI.eq_of_subset_indep (hB.indep.inter_right X) (subset_inter hIB hI.subset) + (inter_subset_right _ _)])⟩ + +@[simp] theorem basis_ground_iff : M.Basis B M.E ↔ M.Base B := by + rw [base_iff_maximal_indep, basis_iff', and_assoc, and_congr_right] + rw [and_iff_left (rfl.subset : M.E ⊆ M.E)] + exact fun h ↦ ⟨fun h' I hI hBI ↦ h'.2 _ hI hBI hI.subset_ground, + fun h' ↦ ⟨h.subset_ground,fun J hJ hBJ _ ↦ h' J hJ hBJ⟩⟩ + +theorem Base.basis_ground (hB : M.Base B) : M.Basis B M.E := + basis_ground_iff.mpr hB + +theorem Indep.basis_iff_forall_insert_dep (hI : M.Indep I) (hIX : I ⊆ X) : + M.Basis I X ↔ ∀ e ∈ X \ I, M.Dep (insert e I) := by + rw [basis_iff', and_iff_right hIX, and_iff_right hI] + refine' ⟨fun h e he ↦ ⟨fun hi ↦ he.2 _, insert_subset (h.2 he.1) hI.subset_ground⟩, + fun h ↦ ⟨fun J hJ hIJ hJX ↦ hIJ.antisymm (fun e heJ ↦ by_contra (fun heI ↦ _)),_⟩⟩ + · exact (h.1 _ hi (subset_insert _ _) (insert_subset he.1 hIX)).symm.subset (mem_insert e I) + · exact (h e ⟨hJX heJ, heI⟩).not_indep (hJ.subset (insert_subset heJ hIJ)) + rw [←diff_union_of_subset hIX, union_subset_iff, and_iff_left hI.subset_ground] + exact fun e he ↦ (h e he).subset_ground (mem_insert _ _) + +theorem Indep.basis_of_forall_insert (hI : M.Indep I) (hIX : I ⊆ X) + (he : ∀ e ∈ X \ I, M.Dep (insert e I)) : M.Basis I X := + (hI.basis_iff_forall_insert_dep hIX).mpr he + +theorem Indep.basis_insert_iff (hI : M.Indep I) : + M.Basis I (insert e I) ↔ M.Dep (insert e I) ∨ e ∈ I := by + simp_rw [hI.basis_iff_forall_insert_dep (subset_insert _ _), dep_iff, insert_subset_iff, + and_iff_left hI.subset_ground, mem_diff, mem_insert_iff, or_and_right, and_not_self, + or_false, and_imp, forall_eq] + tauto + +theorem Basis.iUnion_basis_iUnion {ι : Type _} (X I : ι → Set α) (hI : ∀ i, M.Basis (I i) (X i)) + (h_ind : M.Indep (⋃ i, I i)) : M.Basis (⋃ i, I i) (⋃ i, X i) := by + refine' h_ind.basis_of_forall_insert + (iUnion_subset (fun i ↦ (hI i).subset.trans (subset_iUnion _ _))) _ + rintro e ⟨⟨_, ⟨⟨i, hi, rfl⟩, (hes : e ∈ X i)⟩⟩, he'⟩ + rw [mem_iUnion, not_exists] at he' + refine' ((hI i).insert_dep ⟨hes, he' _⟩).superset (insert_subset_insert (subset_iUnion _ _)) _ + rw [insert_subset_iff, iUnion_subset_iff, and_iff_left (fun i ↦ (hI i).indep.subset_ground)] + exact (hI i).subset_ground hes + +theorem Basis.basis_iUnion {ι : Type _} [Nonempty ι] (X : ι → Set α) + (hI : ∀ i, M.Basis I (X i)) : M.Basis I (⋃ i, X i) := by + convert Basis.iUnion_basis_iUnion X (fun _ ↦ I) (fun i ↦ hI i) _ <;> rw [iUnion_const] + exact (hI (Classical.arbitrary ι)).indep + +theorem Basis.basis_sUnion {Xs : Set (Set α)} (hne : Xs.Nonempty) (h : ∀ X ∈ Xs, M.Basis I X) : + M.Basis I (⋃₀ Xs) := by + rw [sUnion_eq_iUnion] + have := Iff.mpr nonempty_coe_sort hne + exact Basis.basis_iUnion _ fun X ↦ (h X X.prop) + +theorem Indep.basis_setOf_insert_basis (hI : M.Indep I) : + M.Basis I {x | M.Basis I (insert x I)} := by + refine' hI.basis_of_forall_insert (fun e he ↦ (_ : M.Basis _ _)) + (fun e he ↦ ⟨fun hu ↦ he.2 _, he.1.subset_ground⟩) + · rw [insert_eq_of_mem he]; exact hI.basis_self + simpa using (hu.eq_of_basis he.1).symm + +theorem Basis.union_basis_union (hIX : M.Basis I X) (hJY : M.Basis J Y) (h : M.Indep (I ∪ J)) : + M.Basis (I ∪ J) (X ∪ Y) := by + rw [union_eq_iUnion, union_eq_iUnion] + refine' Basis.iUnion_basis_iUnion _ _ _ _ + · simp only [Bool.forall_bool, cond_false, cond_true]; exact ⟨hJY, hIX⟩ + rwa [←union_eq_iUnion] + +theorem Basis.basis_union (hIX : M.Basis I X) (hIY : M.Basis I Y) : M.Basis I (X ∪ Y) := by + convert hIX.union_basis_union hIY _ <;> rw [union_self]; exact hIX.indep + +theorem Basis.basis_union_of_subset (hI : M.Basis I X) (hJ : M.Indep J) (hIJ : I ⊆ J) : + M.Basis J (J ∪ X) := by + convert hJ.basis_self.union_basis_union hI _ <;> + rw [union_eq_self_of_subset_right hIJ] + assumption + +theorem Basis.insert_basis_insert (hI : M.Basis I X) (h : M.Indep (insert e I)) : + M.Basis (insert e I) (insert e X) := by + simp_rw [←union_singleton] at * + exact hI.union_basis_union (h.subset (subset_union_right _ _)).basis_self h + +theorem Base.base_of_basis_superset (hB : M.Base B) (hBX : B ⊆ X) (hIX : M.Basis I X) : + M.Base I := by + by_contra h + obtain ⟨e,heBI,he⟩ := hIX.indep.exists_insert_of_not_base h hB + exact heBI.2 (hIX.mem_of_insert_indep (hBX heBI.1) he) + +theorem Indep.exists_base_subset_union_base (hI : M.Indep I) (hB : M.Base B) : + ∃ B', M.Base B' ∧ I ⊆ B' ∧ B' ⊆ I ∪ B := by + obtain ⟨B', hB', hIB'⟩ := hI.subset_basis_of_subset (subset_union_left I B) + exact ⟨B', hB.base_of_basis_superset (subset_union_right _ _) hB', hIB', hB'.subset⟩ + +theorem Basis.inter_eq_of_subset_indep (hIX : M.Basis I X) (hIJ : I ⊆ J) (hJ : M.Indep J) : + J ∩ X = I := +(subset_inter hIJ hIX.subset).antisymm' + (fun _ he ↦ hIX.mem_of_insert_indep he.2 (hJ.subset (insert_subset he.1 hIJ))) + +theorem Base.basis_of_subset (hX : X ⊆ M.E := by aesop_mat) (hB : M.Base B) (hBX : B ⊆ X) : + M.Basis B X := by + rw [basis_iff, and_iff_right hB.indep, and_iff_right hBX] + exact fun J hJ hBJ _ ↦ hB.eq_of_subset_indep hJ hBJ + +end Basis diff --git a/Mathlib/Data/Matroid/Init.lean b/Mathlib/Data/Matroid/Init.lean new file mode 100644 index 0000000000000..b4f88dbe2e5ae --- /dev/null +++ b/Mathlib/Data/Matroid/Init.lean @@ -0,0 +1,17 @@ +/- +Copyright (c) 2023 Peter Nelson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Peter Nelson +-/ +import Aesop + + +/-! +# Matroid Rule Set + +This module defines the `Matroid` Aesop rule set which is used by the +`aesop_mat` tactic. Aesop rule sets only become visible once the file in which +they're declared is imported, so we must put this declaration into its own file. +-/ + +declare_aesop_rule_sets [Matroid] diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 28ccb1e12d65e..a7fb99e13a8e7 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -869,14 +869,12 @@ theorem strongDownwardInductionOn_eq {p : Multiset α → Sort*} (s : Multiset rw [strongDownwardInduction] #align multiset.strong_downward_induction_on_eq Multiset.strongDownwardInductionOn_eq -/-- Another way of expressing `strongInductionOn`: the `(<)` relation is well-founded. -/ -theorem wellFounded_lt : WellFounded ((· < ·) : Multiset α → Multiset α → Prop) := - Subrelation.wf Multiset.card_lt_of_lt (measure Multiset.card).2 -#align multiset.well_founded_lt Multiset.wellFounded_lt +#align multiset.well_founded_lt wellFounded_lt -instance is_wellFounded_lt : WellFoundedLT (Multiset α) := - ⟨wellFounded_lt⟩ -#align multiset.is_well_founded_lt Multiset.is_wellFounded_lt +/-- Another way of expressing `strongInductionOn`: the `(<)` relation is well-founded. -/ +instance instWellFoundedLT : WellFoundedLT (Multiset α) := + ⟨Subrelation.wf Multiset.card_lt_of_lt (measure Multiset.card).2⟩ +#align multiset.is_well_founded_lt Multiset.instWellFoundedLT /-! ### `Multiset.replicate` -/ @@ -1035,6 +1033,12 @@ theorem cons_erase {s : Multiset α} {a : α} : a ∈ s → a ::ₘ s.erase a = Quot.inductionOn s fun _l h => Quot.sound (perm_cons_erase h).symm #align multiset.cons_erase Multiset.cons_erase +theorem erase_cons_tail_of_mem (h : a ∈ s) : + (b ::ₘ s).erase a = b ::ₘ s.erase a := by + rcases eq_or_ne a b with rfl | hab + · simp [cons_erase h] + · exact s.erase_cons_tail hab.symm + theorem le_cons_erase (s : Multiset α) (a : α) : s ≤ a ::ₘ s.erase a := if h : a ∈ s then le_of_eq (cons_erase h).symm else by rw [erase_of_not_mem h]; apply le_cons_self @@ -1342,6 +1346,13 @@ theorem map_erase [DecidableEq α] [DecidableEq β] (f : α → β) (hf : Functi · rw [s.erase_cons_tail hxy, map_cons, map_cons, (s.map f).erase_cons_tail (hf.ne hxy), ih] #align multiset.map_erase Multiset.map_erase +theorem map_erase_of_mem [DecidableEq α] [DecidableEq β] (f : α → β) + (s : Multiset α) {x : α} (h : x ∈ s) : (s.erase x).map f = (s.map f).erase (f x) := by + induction' s using Multiset.induction_on with y s ih; simp + rcases eq_or_ne y x with rfl | hxy; simp + replace h : x ∈ s := by simpa [hxy.symm] using h + rw [s.erase_cons_tail hxy, map_cons, map_cons, ih h, erase_cons_tail_of_mem (mem_map_of_mem f h)] + theorem map_surjective_of_surjective {f : α → β} (hf : Function.Surjective f) : Function.Surjective (map f) := by intro s diff --git a/Mathlib/Data/MvPolynomial/Basic.lean b/Mathlib/Data/MvPolynomial/Basic.lean index 1065e0d29f368..34dfa8268d988 100644 --- a/Mathlib/Data/MvPolynomial/Basic.lean +++ b/Mathlib/Data/MvPolynomial/Basic.lean @@ -707,8 +707,8 @@ theorem coeff_C_mul (m) (a : R) (p : MvPolynomial σ R) : coeff m (C a * p) = a #align mv_polynomial.coeff_C_mul MvPolynomial.coeff_C_mul theorem coeff_mul [DecidableEq σ] (p q : MvPolynomial σ R) (n : σ →₀ ℕ) : - coeff n (p * q) = ∑ x in antidiagonal n, coeff x.1 p * coeff x.2 q := - AddMonoidAlgebra.mul_apply_antidiagonal p q _ _ mem_antidiagonal + coeff n (p * q) = ∑ x in Finset.antidiagonal n, coeff x.1 p * coeff x.2 q := + AddMonoidAlgebra.mul_apply_antidiagonal p q _ _ Finset.mem_antidiagonal #align mv_polynomial.coeff_mul MvPolynomial.coeff_mul @[simp] diff --git a/Mathlib/Data/MvPolynomial/Division.lean b/Mathlib/Data/MvPolynomial/Division.lean index b7ddacdc2f8bd..634ae6990b14a 100644 --- a/Mathlib/Data/MvPolynomial/Division.lean +++ b/Mathlib/Data/MvPolynomial/Division.lean @@ -232,8 +232,8 @@ theorem monomial_dvd_monomial {r s : R} {i j : σ →₀ ℕ} : · exact ⟨Or.inr hi, _, hj⟩ · exact ⟨Or.inl hj, hj.symm ▸ dvd_zero _⟩ -- Porting note: two goals remain at this point in Lean 4 - · simp_all only [or_true, dvd_mul_right] - · simp_all only [ite_self, le_refl, ite_true, dvd_mul_right] + · simp_all only [or_true, dvd_mul_right, and_self] + · simp_all only [ite_self, le_refl, ite_true, dvd_mul_right, or_false, and_self] · rintro ⟨h | hij, d, rfl⟩ · simp_rw [h, monomial_zero, dvd_zero] · refine' ⟨monomial (j - i) d, _⟩ @@ -252,7 +252,7 @@ theorem X_dvd_X [Nontrivial R] {i j : σ} : (X i : MvPolynomial σ R) ∣ (X j : MvPolynomial σ R) ↔ i = j := by refine' monomial_one_dvd_monomial_one.trans _ simp_rw [Finsupp.single_le_iff, Nat.one_le_iff_ne_zero, Finsupp.single_apply_ne_zero, - and_true_iff] + ne_eq, not_false_eq_true, and_true] set_option linter.uppercaseLean3 false in #align mv_polynomial.X_dvd_X MvPolynomial.X_dvd_X diff --git a/Mathlib/Data/MvPolynomial/Variables.lean b/Mathlib/Data/MvPolynomial/Variables.lean index f8d2bf97934ce..d6c217d4fe3c1 100644 --- a/Mathlib/Data/MvPolynomial/Variables.lean +++ b/Mathlib/Data/MvPolynomial/Variables.lean @@ -348,7 +348,7 @@ theorem vars_mul [DecidableEq σ] (φ ψ : MvPolynomial σ R) : (φ * ψ).vars cases hd rw [Finset.sum_eq_zero] rintro ⟨d₁, d₂⟩ H - rw [Finsupp.mem_antidiagonal] at H + rw [Finset.mem_antidiagonal] at H subst H obtain H | H : i ∈ d₁.support ∨ i ∈ d₂.support := by simpa only [Finset.mem_union] using Finsupp.support_add hi @@ -686,7 +686,7 @@ theorem totalDegree_smul_le [CommSemiring S] [DistribMulAction R S] (a : R) (f : theorem totalDegree_pow (a : MvPolynomial σ R) (n : ℕ) : (a ^ n).totalDegree ≤ n * a.totalDegree := by induction' n with n ih - · simp only [Nat.zero_eq, zero_mul, pow_zero, totalDegree_one] + · simp only [Nat.zero_eq, pow_zero, totalDegree_one, zero_mul, le_refl] rw [pow_succ] calc totalDegree (a * a ^ n) ≤ a.totalDegree + (a ^ n).totalDegree := totalDegree_mul _ _ diff --git a/Mathlib/Data/Nat/Basic.lean b/Mathlib/Data/Nat/Basic.lean index ee0dbac57f61f..f85079ea9cb3c 100644 --- a/Mathlib/Data/Nat/Basic.lean +++ b/Mathlib/Data/Nat/Basic.lean @@ -165,7 +165,6 @@ theorem eq_of_le_of_lt_succ {n m : ℕ} (h₁ : n ≤ m) (h₂ : m < n + 1) : m -- Moved to Std #align nat.one_add Nat.one_add -@[simp] theorem succ_pos' {n : ℕ} : 0 < succ n := succ_pos n #align nat.succ_pos' Nat.succ_pos' @@ -192,13 +191,7 @@ theorem one_lt_succ_succ (n : ℕ) : 1 < n.succ.succ := -- Porting note: Nat.succ_le_succ_iff is in Std -theorem max_succ_succ {m n : ℕ} : max (succ m) (succ n) = succ (max m n) := by - by_cases h1 : m ≤ n - rw [max_eq_right h1, max_eq_right (succ_le_succ h1)] - · rw [not_le] at h1 - have h2 := le_of_lt h1 - rw [max_eq_left h2, max_eq_left (succ_le_succ h2)] -#align nat.max_succ_succ Nat.max_succ_succ +#align nat.max_succ_succ Nat.succ_max_succ theorem not_succ_lt_self {n : ℕ} : ¬succ n < n := not_lt_of_ge (Nat.le_succ _) diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 51031a5b26d68..673f89a0cec19 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -86,8 +86,9 @@ lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by have h1 x : (x + x) % 2 = 0 := by rw [← two_mul, mul_comm]; apply mul_mod_left have h2 x : (x + x + 1) % 2 = 1 := by rw [← two_mul, add_comm]; apply add_mul_mod_self_left have h3 x : (x + x) / 2 = x := by rw [← two_mul, mul_comm]; apply mul_div_left _ zero_lt_two - have h4 x : (x + x + 1) / 2 = x := by rw [← two_mul, add_comm]; simp [add_mul_div_left] - cases a <;> cases b <;> simp [h1, h2, h3, h4] <;> split_ifs <;> simp_all + have h4 x : (x + x + 1) / 2 = x := by rw [← two_mul, add_comm]; simp [add_mul_div_left]; rfl + cases a <;> cases b <;> simp [h1, h2, h3, h4] <;> split_ifs + <;> simp_all (config := {decide := true}) #align nat.bitwise_bit Nat.bitwise_bit lemma bit_mod_two (a : Bool) (x : ℕ) : @@ -304,6 +305,7 @@ theorem testBit_two_pow_of_ne {n m : ℕ} (hm : n ≠ m) : testBit (2 ^ n) m = f rw [(rfl : succ 0 = 1)] simp only [ge_iff_le, tsub_le_iff_right, pow_succ, bodd_mul, Bool.and_eq_false_eq_eq_false_or_eq_false, or_true] + exact Or.inr rfl #align nat.test_bit_two_pow_of_ne Nat.testBit_two_pow_of_ne theorem testBit_two_pow (n m : ℕ) : testBit (2 ^ n) m = (n = m) := by @@ -349,8 +351,7 @@ theorem xor_comm (n m : ℕ) : n ^^^ m = m ^^^ n := #align nat.lxor_comm Nat.xor_comm @[simp] -theorem zero_xor (n : ℕ) : 0 ^^^ n = n := by - simp only [HXor.hXor, Xor.xor, xor, bitwise_zero_left, ite_true] +theorem zero_xor (n : ℕ) : 0 ^^^ n = n := by simp [HXor.hXor, Xor.xor, xor] #align nat.zero_lxor Nat.zero_xor @[simp] @@ -359,22 +360,22 @@ theorem xor_zero (n : ℕ) : n ^^^ 0 = n := by simp [HXor.hXor, Xor.xor, xor] @[simp] theorem zero_land (n : ℕ) : 0 &&& n = 0 := by - simp only [HAnd.hAnd, AndOp.and, land, bitwise_zero_left, ite_false]; + simp only [HAnd.hAnd, AndOp.and, land, bitwise_zero_left, ite_false, Bool.and_true]; #align nat.zero_land Nat.zero_land @[simp] theorem land_zero (n : ℕ) : n &&& 0 = 0 := by - simp only [HAnd.hAnd, AndOp.and, land, bitwise_zero_right, ite_false] + simp only [HAnd.hAnd, AndOp.and, land, bitwise_zero_right, ite_false, Bool.and_false] #align nat.land_zero Nat.land_zero @[simp] theorem zero_lor (n : ℕ) : 0 ||| n = n := by - simp only [HOr.hOr, OrOp.or, lor, bitwise_zero_left, ite_true] + simp only [HOr.hOr, OrOp.or, lor, bitwise_zero_left, ite_true, Bool.or_true] #align nat.zero_lor Nat.zero_lor @[simp] theorem lor_zero (n : ℕ) : n ||| 0 = n := by - simp only [HOr.hOr, OrOp.or, lor, bitwise_zero_right, ite_true] + simp only [HOr.hOr, OrOp.or, lor, bitwise_zero_right, ite_true, Bool.or_false] #align nat.lor_zero Nat.lor_zero @@ -477,7 +478,7 @@ theorem xor_trichotomy {a b c : ℕ} (h : a ≠ b ^^^ c) : rfl ) h fun j hj => by -- Porting note: this was originally `simp [hi' _ hj]` - rw [Nat.testBit_xor, hi' _ hj, Bool.bne_eq_xor, Bool.xor_false_right, eq_self_iff_true] + rw [Nat.testBit_xor, hi' _ hj, Bool.bne_eq_xor, Bool.xor_false, eq_self_iff_true] trivial #align nat.lxor_trichotomy Nat.xor_trichotomy diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index 972eaf056f2ba..aa08422357cc2 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -3,9 +3,9 @@ Copyright (c) 2014 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Hom.Ring.Defs -import Mathlib.Algebra.Hom.Group.Basic import Mathlib.Algebra.Divisibility.Basic +import Mathlib.Algebra.Group.Hom.Basic +import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Data.Nat.Basic #align_import data.nat.cast.basic from "leanprover-community/mathlib"@"acebd8d49928f6ed8920e502a6c90674e75bd441" diff --git a/Mathlib/Data/Nat/Cast/Defs.lean b/Mathlib/Data/Nat/Cast/Defs.lean index 890e6086097f3..b408bd87e8e32 100644 --- a/Mathlib/Data/Nat/Cast/Defs.lean +++ b/Mathlib/Data/Nat/Cast/Defs.lean @@ -47,6 +47,12 @@ class Nat.AtLeastTwo (n : ℕ) : Prop where instance instNatAtLeastTwo : Nat.AtLeastTwo (n + 2) where prop := Nat.succ_le_succ $ Nat.succ_le_succ $ Nat.zero_le _ +lemma Nat.AtLeastTwo.ne_zero (n : ℕ) [h : n.AtLeastTwo] : n ≠ 0 := by + rintro rfl; exact absurd h.1 (by decide) + +lemma Nat.AtLeastTwo.ne_one (n : ℕ) [h : n.AtLeastTwo] : n ≠ 1 := by + rintro rfl; exact absurd h.1 (by decide) + /-- Recognize numeric literals which are at least `2` as terms of `R` via `Nat.cast`. This instance is what makes things like `37 : R` type check. Note that `0` and `1` are not needed because they are recognized as terms of `R` (at least when `R` is an `AddMonoidWithOne`) through diff --git a/Mathlib/Data/Nat/Cast/Order.lean b/Mathlib/Data/Nat/Cast/Order.lean index 86f97e66dd24e..5362d69c5b1b9 100644 --- a/Mathlib/Data/Nat/Cast/Order.lean +++ b/Mathlib/Data/Nat/Cast/Order.lean @@ -34,12 +34,12 @@ theorem mono_cast : Monotone (Nat.cast : ℕ → α) := rw [Nat.cast_succ]; exact le_add_of_nonneg_right zero_le_one #align nat.mono_cast Nat.mono_cast --- See also `Nat.cast_nonneg`, specialised for an `OrderedSemiring`. +/-- See also `Nat.cast_nonneg`, specialised for an `OrderedSemiring`. -/ @[simp low] theorem cast_nonneg' (n : ℕ) : 0 ≤ (n : α) := @Nat.cast_zero α _ ▸ mono_cast (Nat.zero_le n) --- Specialisation of `Nat.cast_nonneg'`, which seems to be easier for Lean to use. +/-- Specialisation of `Nat.cast_nonneg'`, which seems to be easier for Lean to use. -/ @[simp] theorem cast_nonneg {α} [OrderedSemiring α] (n : ℕ) : 0 ≤ (n : α) := cast_nonneg' n @@ -63,11 +63,11 @@ theorem cast_add_one_pos (n : ℕ) : 0 < (n : α) + 1 := zero_lt_one.trans_le <| le_add_of_nonneg_left n.cast_nonneg' #align nat.cast_add_one_pos Nat.cast_add_one_pos --- See also `Nat.cast_pos`, a specialisation of this to an `OrderedSemiring`. +/-- See also `Nat.cast_pos`, a specialisation of this to an `OrderedSemiring`. -/ @[simp low] theorem cast_pos' {n : ℕ} : (0 : α) < n ↔ 0 < n := by cases n <;> simp [cast_add_one_pos] --- Specialisation of `Nat.cast_pos'`, which seems to be easier for Lean to use. +/-- Specialisation of `Nat.cast_pos'`, which seems to be easier for Lean to use. -/ @[simp] theorem cast_pos {α} [OrderedSemiring α] [Nontrivial α] {n : ℕ} : (0 : α) < n ↔ 0 < n := cast_pos' #align nat.cast_pos Nat.cast_pos diff --git a/Mathlib/Data/Nat/Choose/Multinomial.lean b/Mathlib/Data/Nat/Choose/Multinomial.lean index e36fa5b43e1c2..9f2ec24bf3956 100644 --- a/Mathlib/Data/Nat/Choose/Multinomial.lean +++ b/Mathlib/Data/Nat/Choose/Multinomial.lean @@ -126,7 +126,7 @@ theorem binomial_succ_succ [DecidableEq α] (h : a ≠ b) : multinomial {a, b} (Function.update f a (f a).succ) + multinomial {a, b} (Function.update f b (f b).succ) := by simp only [binomial_eq_choose, Function.update_apply, - h, Ne.def, ite_true, ite_false] + h, Ne.def, ite_true, ite_false, not_false_eq_true] rw [if_neg h.symm] rw [add_succ, choose_succ_succ, succ_add_eq_succ_add] ring diff --git a/Mathlib/Data/Nat/Choose/Sum.lean b/Mathlib/Data/Nat/Choose/Sum.lean index 0249d526a3a3b..59bbde2807745 100644 --- a/Mathlib/Data/Nat/Choose/Sum.lean +++ b/Mathlib/Data/Nat/Choose/Sum.lean @@ -72,7 +72,7 @@ theorem add_pow (h : Commute x y) (n : ℕ) : /-- A version of `Commute.add_pow` that avoids ℕ-subtraction by summing over the antidiagonal and also with the binomial coefficient applied via scalar action of ℕ. -/ theorem add_pow' (h : Commute x y) (n : ℕ) : - (x + y) ^ n = ∑ m in Nat.antidiagonal n, choose n m.fst • (x ^ m.fst * y ^ m.snd) := by + (x + y) ^ n = ∑ m in antidiagonal n, choose n m.fst • (x ^ m.fst * y ^ m.snd) := by simp_rw [Finset.Nat.sum_antidiagonal_eq_sum_range_succ fun m p ↦ choose n m • (x ^ m * y ^ p), _root_.nsmul_eq_mul, cast_comm, h.add_pow] #align commute.add_pow' Commute.add_pow' @@ -196,9 +196,9 @@ theorem prod_pow_choose_succ {M : Type*} [CommMonoid M] (f : ℕ → ℕ → M) -- porting note: new lemma @[to_additive sum_antidiagonal_choose_succ_nsmul] theorem prod_antidiagonal_pow_choose_succ {M : Type*} [CommMonoid M] (f : ℕ → ℕ → M) (n : ℕ) : - (∏ ij in Nat.antidiagonal (n + 1), f ij.1 ij.2 ^ (n + 1).choose ij.1) = - (∏ ij in Nat.antidiagonal n, f ij.1 (ij.2 + 1) ^ n.choose ij.1) * - ∏ ij in Nat.antidiagonal n, f (ij.1 + 1) ij.2 ^ n.choose ij.2 := by + (∏ ij in antidiagonal (n + 1), f ij.1 ij.2 ^ (n + 1).choose ij.1) = + (∏ ij in antidiagonal n, f ij.1 (ij.2 + 1) ^ n.choose ij.1) * + ∏ ij in antidiagonal n, f (ij.1 + 1) ij.2 ^ n.choose ij.2 := by simp only [Nat.prod_antidiagonal_eq_prod_range_succ_mk, prod_pow_choose_succ] have : ∀ i ∈ range (n + 1), i ≤ n := fun i hi ↦ by simpa [Nat.lt_succ_iff] using hi congr 1 @@ -220,9 +220,9 @@ theorem sum_choose_succ_mul (f : ℕ → ℕ → R) (n : ℕ) : /-- The sum along the antidiagonal of `(n+1).choose i * f i j` can be split into two sums along the antidiagonal at rank `n`, respectively of `n.choose i * f i (j+1)` and `n.choose j * f (i+1) j`. -/ theorem sum_antidiagonal_choose_succ_mul (f : ℕ → ℕ → R) (n : ℕ) : - (∑ ij in Nat.antidiagonal (n + 1), ((n + 1).choose ij.1 : R) * f ij.1 ij.2) = - (∑ ij in Nat.antidiagonal n, (n.choose ij.1 : R) * f ij.1 (ij.2 + 1)) + - ∑ ij in Nat.antidiagonal n, (n.choose ij.2 : R) * f (ij.1 + 1) ij.2 := by + (∑ ij in antidiagonal (n + 1), ((n + 1).choose ij.1 : R) * f ij.1 ij.2) = + (∑ ij in antidiagonal n, (n.choose ij.1 : R) * f ij.1 (ij.2 + 1)) + + ∑ ij in antidiagonal n, (n.choose ij.2 : R) * f (ij.1 + 1) ij.2 := by simpa only [nsmul_eq_mul] using sum_antidiagonal_choose_succ_nsmul f n #align finset.sum_antidiagonal_choose_succ_mul Finset.sum_antidiagonal_choose_succ_mul diff --git a/Mathlib/Data/Nat/Choose/Vandermonde.lean b/Mathlib/Data/Nat/Choose/Vandermonde.lean index 17284ebf60a7f..da89b2785398f 100644 --- a/Mathlib/Data/Nat/Choose/Vandermonde.lean +++ b/Mathlib/Data/Nat/Choose/Vandermonde.lean @@ -23,7 +23,7 @@ https://en.wikipedia.org/wiki/Vandermonde%27s_identity#Algebraic_proof . open BigOperators -open Polynomial Finset.Nat +open Polynomial Finset Finset.Nat /-- Vandermonde's identity -/ theorem Nat.add_choose_eq (m n k : ℕ) : diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 6bd1fa22bab23..5d6d5e0c10ced 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -357,6 +357,7 @@ theorem getLast_digit_ne_zero (b : ℕ) {m : ℕ} (hm : m ≠ 0) : · cases hm rfl rename ℕ => m simp only [digits_one, List.getLast_replicate_succ m 1] + exact Nat.one_ne_zero revert hm apply Nat.strongInductionOn m intro n IH hn @@ -555,7 +556,7 @@ theorem sub_one_mul_sum_div_pow_eq_sub_sum_digits have w₂' := fun (h : tl ≠ []) ↦ (List.getLast_cons h) ▸ h_ne_zero have ih := ih (w₂' h') w₁' simp only [self_div_pow_eq_ofDigits_drop _ _ h, digits_ofDigits p h tl w₁' w₂', - succ_eq_one_add] at ih + ←Nat.one_add] at ih have := sum_singleton (fun x ↦ ofDigits p <| tl.drop x) tl.length rw [← Ico_succ_singleton, List.drop_length, ofDigits] at this have h₁ : 1 ≤ tl.length := List.length_pos.mpr h' @@ -569,7 +570,7 @@ theorem sub_one_mul_sum_div_pow_eq_sub_sum_digits · simp [ofDigits_one] · simp [lt_one_iff.mp h] cases L - · simp + · rfl · simp [ofDigits] theorem sub_one_mul_sum_log_div_pow_eq_sub_sum_digits (n : ℕ) : @@ -592,7 +593,7 @@ theorem sub_one_mul_sum_log_div_pow_eq_sub_sum_digits (n : ℕ) : theorem digits_two_eq_bits (n : ℕ) : digits 2 n = n.bits.map fun b => cond b 1 0 := by induction' n using Nat.binaryRecFromOne with b n h ih · simp - · simp + · rfl rw [bits_append_bit _ _ fun hn => absurd hn h] cases b · rw [digits_def' one_lt_two] @@ -689,7 +690,7 @@ theorem ofDigits_neg_one : theorem modEq_eleven_digits_sum (n : ℕ) : n ≡ ((digits 10 n).map fun n : ℕ => (n : ℤ)).alternatingSum [ZMOD 11] := by - have t := zmodeq_ofDigits_digits 11 10 (-1 : ℤ) (by unfold Int.ModEq; norm_num) n + have t := zmodeq_ofDigits_digits 11 10 (-1 : ℤ) (by unfold Int.ModEq; rfl) n rwa [ofDigits_neg_one] at t #align nat.modeq_eleven_digits_sum Nat.modEq_eleven_digits_sum diff --git a/Mathlib/Data/Nat/Factorial/BigOperators.lean b/Mathlib/Data/Nat/Factorial/BigOperators.lean index 64f65882c7a4f..642ba3f562abb 100644 --- a/Mathlib/Data/Nat/Factorial/BigOperators.lean +++ b/Mathlib/Data/Nat/Factorial/BigOperators.lean @@ -32,7 +32,7 @@ theorem prod_factorial_pos : 0 < ∏ i in s, (f i)! := theorem prod_factorial_dvd_factorial_sum : (∏ i in s, (f i)!) ∣ (∑ i in s, f i)! := by classical induction' s using Finset.induction with a' s' has ih - · simp only [Finset.sum_empty, Finset.prod_empty, factorial] + · simp only [prod_empty, factorial, dvd_refl] · simp only [Finset.prod_insert has, Finset.sum_insert has] refine' dvd_trans (mul_dvd_mul_left (f a')! ih) _ apply Nat.factorial_mul_factorial_dvd_factorial_add diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 07135e187cf6b..f2c25b4c51df4 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -115,11 +115,11 @@ theorem factorization_inj : Set.InjOn factorization { x : ℕ | x ≠ 0 } := fun #align nat.factorization_inj Nat.factorization_inj @[simp] -theorem factorization_zero : factorization 0 = 0 := by simp [factorization] +theorem factorization_zero : factorization 0 = 0 := by decide #align nat.factorization_zero Nat.factorization_zero @[simp] -theorem factorization_one : factorization 1 = 0 := by simp [factorization] +theorem factorization_one : factorization 1 = 0 := by decide #align nat.factorization_one Nat.factorization_one /-- The support of `n.factorization` is exactly `n.factors.toFinset` -/ diff --git a/Mathlib/Data/Nat/Fib.lean b/Mathlib/Data/Nat/Fib.lean index 097c4ddeae6b6..078af88ff9be0 100644 --- a/Mathlib/Data/Nat/Fib.lean +++ b/Mathlib/Data/Nat/Fib.lean @@ -308,7 +308,7 @@ theorem fib_dvd (m n : ℕ) (h : m ∣ n) : fib m ∣ fib n := by #align nat.fib_dvd Nat.fib_dvd theorem fib_succ_eq_sum_choose : - ∀ n : ℕ, fib (n + 1) = ∑ p in Finset.Nat.antidiagonal n, choose p.1 p.2 := + ∀ n : ℕ, fib (n + 1) = ∑ p in Finset.antidiagonal n, choose p.1 p.2 := twoStepInduction rfl rfl fun n h1 h2 => by rw [fib_add_two, h1, h2, Finset.Nat.antidiagonal_succ_succ', Finset.Nat.antidiagonal_succ'] simp [choose_succ_succ, Finset.sum_add_distrib, add_left_comm] diff --git a/Mathlib/Data/Nat/GCD/Basic.lean b/Mathlib/Data/Nat/GCD/Basic.lean index a5f67e6a0bb25..844f7e7a69e49 100644 --- a/Mathlib/Data/Nat/GCD/Basic.lean +++ b/Mathlib/Data/Nat/GCD/Basic.lean @@ -30,7 +30,8 @@ theorem gcd_greatest {a b d : ℕ} (hda : d ∣ a) (hdb : d ∣ b) (hd : ∀ e : (dvd_antisymm (hd _ (gcd_dvd_left a b) (gcd_dvd_right a b)) (dvd_gcd hda hdb)).symm #align nat.gcd_greatest Nat.gcd_greatest --- Lemmas where one argument consists of addition of a multiple of the other +/-! Lemmas where one argument consists of addition of a multiple of the other -/ + @[simp] theorem gcd_add_mul_right_right (m n k : ℕ) : gcd m (n + k * m) = gcd m n := by simp [gcd_rec m (n + k * m), gcd_rec m n] @@ -69,7 +70,8 @@ theorem gcd_mul_left_add_left (m n k : ℕ) : gcd (n * k + m) n = gcd m n := by rw [gcd_comm, gcd_mul_left_add_right, gcd_comm] #align nat.gcd_mul_left_add_left Nat.gcd_mul_left_add_left --- Lemmas where one argument consists of an addition of the other +/-! Lemmas where one argument consists of an addition of the other -/ + @[simp] theorem gcd_add_self_right (m n : ℕ) : gcd m (n + m) = gcd m n := Eq.trans (by rw [one_mul]) (gcd_add_mul_right_right m n 1) @@ -89,7 +91,8 @@ theorem gcd_self_add_right (m n : ℕ) : gcd m (m + n) = gcd m n := by rw [add_comm, gcd_add_self_right] #align nat.gcd_self_add_right Nat.gcd_self_add_right --- Lemmas where one argument consists of a subtraction of the other +/-! Lemmas where one argument consists of a subtraction of the other -/ + @[simp] theorem gcd_sub_self_left {m n : ℕ} (h : m ≤ n) : gcd (n - m) m = gcd n m := by calc diff --git a/Mathlib/Data/Nat/Hyperoperation.lean b/Mathlib/Data/Nat/Hyperoperation.lean index ad7d9bfcbf240..ad4daea7483e7 100644 --- a/Mathlib/Data/Nat/Hyperoperation.lean +++ b/Mathlib/Data/Nat/Hyperoperation.lean @@ -119,7 +119,7 @@ theorem hyperoperation_ge_four_zero (n k : ℕ) : hyperoperation (n + 4) 0 k = if Even k then 1 else 0 := by induction' k with kk kih · rw [hyperoperation_ge_three_eq_one] - simp only [even_zero, if_true] + simp only [Nat.zero_eq, even_zero, if_true] · rw [hyperoperation_recursion] rw [kih] simp_rw [Nat.even_add_one] diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index 66bd58ec8ce52..003252fd98157 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -142,7 +142,7 @@ theorem log_eq_iff {b m n : ℕ} (h : m ≠ 0 ∨ 1 < b ∧ n ≠ 0) : · simpa only [log_of_left_le_one hb, hm.symm, false_iff_iff, not_and, not_lt] using le_trans (pow_le_pow_of_le_one' hb m.le_succ) · simpa only [log_zero_right, hm.symm, nonpos_iff_eq_zero, false_iff, not_and, not_lt, - add_pos_iff, or_true, pow_eq_zero_iff] using pow_eq_zero + add_pos_iff, zero_lt_one, or_true, pow_eq_zero_iff] using pow_eq_zero #align nat.log_eq_iff Nat.log_eq_iff theorem log_eq_of_pow_le_of_lt_pow {b m n : ℕ} (h₁ : b ^ m ≤ n) (h₂ : n < b ^ (m + 1)) : diff --git a/Mathlib/Data/Nat/ModEq.lean b/Mathlib/Data/Nat/ModEq.lean index da8a9bbea7bd7..4f0d0831b943d 100644 --- a/Mathlib/Data/Nat/ModEq.lean +++ b/Mathlib/Data/Nat/ModEq.lean @@ -526,7 +526,7 @@ theorem odd_of_mod_four_eq_three {n : ℕ} : n % 4 = 3 → n % 2 = 1 := by theorem odd_mod_four_iff {n : ℕ} : n % 2 = 1 ↔ n % 4 = 1 ∨ n % 4 = 3 := have help : ∀ m : ℕ, m < 4 → m % 2 = 1 → m = 1 ∨ m = 3 := by decide ⟨fun hn => - help (n % 4) (mod_lt n (by norm_num)) <| (mod_mod_of_dvd n (by norm_num : 2 ∣ 4)).trans hn, + help (n % 4) (mod_lt n (by norm_num)) <| (mod_mod_of_dvd n (by decide : 2 ∣ 4)).trans hn, fun h => Or.elim h odd_of_mod_four_eq_one odd_of_mod_four_eq_three⟩ #align nat.odd_mod_four_iff Nat.odd_mod_four_iff diff --git a/Mathlib/Data/Nat/Order/Basic.lean b/Mathlib/Data/Nat/Order/Basic.lean index 9746f7fab7c00..ea58274bb664b 100644 --- a/Mathlib/Data/Nat/Order/Basic.lean +++ b/Mathlib/Data/Nat/Order/Basic.lean @@ -112,8 +112,6 @@ theorem eq_zero_of_mul_le (hb : 2 ≤ n) (h : n * m ≤ m) : m = 0 := eq_zero_of_double_le <| le_trans (Nat.mul_le_mul_right _ hb) h #align nat.eq_zero_of_mul_le Nat.eq_zero_of_mul_le -theorem zero_max : max 0 n = n := - max_eq_right (zero_le _) #align nat.zero_max Nat.zero_max @[simp] diff --git a/Mathlib/Data/Nat/Parity.lean b/Mathlib/Data/Nat/Parity.lean index fdfc1cf9f53c5..939f00cb1674f 100644 --- a/Mathlib/Data/Nat/Parity.lean +++ b/Mathlib/Data/Nat/Parity.lean @@ -278,9 +278,8 @@ end example (m n : ℕ) (h : Even m) : ¬Even (n + 3) ↔ Even (m ^ 2 + m + n) := by simp [*, two_ne_zero, parity_simps] -/- Porting note: the `simp` lemmas about `bit*` no longer apply, but `simp` in Lean 4 currently -simplifies decidable propositions. This may change in the future. -/ -example : ¬Even 25394535 := by simp only +/- Porting note: the `simp` lemmas about `bit*` no longer apply. -/ +example : ¬Even 25394535 := by decide end Nat diff --git a/Mathlib/Data/Nat/PartENat.lean b/Mathlib/Data/Nat/PartENat.lean index 674b4b61b2b79..07b354e0451b3 100644 --- a/Mathlib/Data/Nat/PartENat.lean +++ b/Mathlib/Data/Nat/PartENat.lean @@ -3,9 +3,9 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Hom.Equiv.Basic -import Mathlib.Data.Part +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Data.ENat.Lattice +import Mathlib.Data.Part import Mathlib.Tactic.NormNum #align_import data.nat.part_enat from "leanprover-community/mathlib"@"3ff3f2d6a3118b8711063de7111a0d77a53219a8" @@ -524,7 +524,7 @@ theorem add_eq_top_iff {a b : PartENat} : a + b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := refine PartENat.casesOn a ?_ ?_ <;> refine PartENat.casesOn b ?_ ?_ <;> simp [top_add, add_top] - simp only [←Nat.cast_add, PartENat.natCast_ne_top, forall_const] + simp only [←Nat.cast_add, PartENat.natCast_ne_top, forall_const, not_false_eq_true] #align part_enat.add_eq_top_iff PartENat.add_eq_top_iff protected theorem add_right_cancel_iff {a b c : PartENat} (hc : c ≠ ⊤) : a + c = b + c ↔ a = b := by @@ -647,7 +647,7 @@ lemma ofENat_some (n : ℕ) : ofENat (Option.some n) = ↑n := rfl @[simp, norm_cast] theorem toWithTop_ofENat (n : ℕ∞) {_ : Decidable (n : PartENat).Dom} : toWithTop (↑n) = n := by induction n with - | none => simp + | none => simp; rfl | some n => simp only [toWithTop_natCast', ofENat_some] rfl @@ -755,7 +755,7 @@ theorem lt_wf : @WellFounded PartENat (· < ·) := by classical change WellFounded fun a b : PartENat => a < b simp_rw [← withTopEquiv_lt] - exact InvImage.wf _ (WithTop.wellFounded_lt Nat.lt_wfRel.wf) + exact InvImage.wf _ wellFounded_lt #align part_enat.lt_wf PartENat.lt_wf instance : WellFoundedLT PartENat := diff --git a/Mathlib/Data/Nat/Prime.lean b/Mathlib/Data/Nat/Prime.lean index d75e9d556bf39..ba79502f9cb12 100644 --- a/Mathlib/Data/Nat/Prime.lean +++ b/Mathlib/Data/Nat/Prime.lean @@ -671,7 +671,7 @@ theorem coprime_or_dvd_of_prime {p} (pp : Prime p) (i : ℕ) : Coprime p i ∨ p #align nat.coprime_or_dvd_of_prime Nat.coprime_or_dvd_of_prime theorem coprime_of_lt_prime {n p} (n_pos : 0 < n) (hlt : n < p) (pp : Prime p) : Coprime p n := - (coprime_or_dvd_of_prime pp n).resolve_right fun h => lt_le_antisymm hlt (le_of_dvd n_pos h) + (coprime_or_dvd_of_prime pp n).resolve_right fun h => Nat.lt_le_asymm hlt (le_of_dvd n_pos h) #align nat.coprime_of_lt_prime Nat.coprime_of_lt_prime theorem eq_or_coprime_of_le_prime {n p} (n_pos : 0 < n) (hle : n ≤ p) (pp : Prime p) : diff --git a/Mathlib/Data/Nat/Sqrt.lean b/Mathlib/Data/Nat/Sqrt.lean index 419cd9fee93f5..cbd3e6973d366 100644 --- a/Mathlib/Data/Nat/Sqrt.lean +++ b/Mathlib/Data/Nat/Sqrt.lean @@ -3,6 +3,7 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Johannes Hölzl, Mario Carneiro -/ +import Mathlib.Algebra.Parity import Mathlib.Data.Int.Order.Basic import Mathlib.Data.Nat.Size import Mathlib.Data.Nat.ForSqrt @@ -46,8 +47,8 @@ To turn this into a lean proof we need to manipulate, use properties of natural -/ private theorem sqrt_isSqrt (n : ℕ) : IsSqrt n (sqrt n) := by match n with - | 0 => simp [IsSqrt] - | 1 => simp [IsSqrt] + | 0 => simp [IsSqrt, sqrt] + | 1 => simp [IsSqrt, sqrt] | n + 2 => have h : ¬ (n + 2) ≤ 1 := by simp simp only [IsSqrt, sqrt, h, ite_false] @@ -176,6 +177,11 @@ theorem exists_mul_self' (x : ℕ) : (∃ n, n ^ 2 = x) ↔ sqrt x ^ 2 = x := by simpa only [pow_two] using exists_mul_self x #align nat.exists_mul_self' Nat.exists_mul_self' +/-- `IsSquare` can be decided on `ℕ` by checking against the square root. -/ +instance : DecidablePred (IsSquare : ℕ → Prop) := + fun m => decidable_of_iff' (Nat.sqrt m * Nat.sqrt m = m) <| by + simp_rw [←Nat.exists_mul_self m, IsSquare, eq_comm] + theorem sqrt_mul_sqrt_lt_succ (n : ℕ) : sqrt n * sqrt n < n + 1 := lt_succ_iff.mpr (sqrt_le _) #align nat.sqrt_mul_sqrt_lt_succ Nat.sqrt_mul_sqrt_lt_succ diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index 2efd9b9c22eab..29628a804b196 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -254,7 +254,7 @@ instance : DecidablePred (Squarefree : ℕ → Prop) := fun _ => decidable_of_iff' _ squarefree_iff_minSqFac theorem squarefree_two : Squarefree 2 := by - rw [squarefree_iff_nodup_factors] <;> norm_num + rw [squarefree_iff_nodup_factors] <;> decide #align nat.squarefree_two Nat.squarefree_two theorem divisors_filter_squarefree_of_squarefree {n : ℕ} (hn : Squarefree n) : @@ -345,7 +345,7 @@ theorem sq_mul_squarefree_of_pos {n : ℕ} (hn : 0 < n) : rintro x ⟨y, hy⟩ rw [Nat.isUnit_iff] by_contra hx - refine' lt_le_antisymm _ (Finset.le_max' S ((b * x) ^ 2) _) + refine' Nat.lt_le_asymm _ (Finset.le_max' S ((b * x) ^ 2) _) -- Porting note: these two goals were in the opposite order in Lean 3 · convert lt_mul_of_one_lt_right hlts (one_lt_pow 2 x zero_lt_two (one_lt_iff_ne_zero_and_ne_one.mpr ⟨fun h => by simp_all, hx⟩)) diff --git a/Mathlib/Data/Nat/Totient.lean b/Mathlib/Data/Nat/Totient.lean index e04564b3a3b11..d9bd8fc38db26 100644 --- a/Mathlib/Data/Nat/Totient.lean +++ b/Mathlib/Data/Nat/Totient.lean @@ -245,7 +245,7 @@ theorem card_units_zmod_lt_sub_one {p : ℕ} (hp : 1 < p) [Fintype (ZMod p)ˣ] : Fintype.card (ZMod p)ˣ ≤ p - 1 := by haveI : NeZero p := ⟨(pos_of_gt hp).ne'⟩ rw [ZMod.card_units_eq_totient p] - exact Nat.le_pred_of_lt (Nat.totient_lt p hp) + exact Nat.le_sub_one_of_lt (Nat.totient_lt p hp) #align nat.card_units_zmod_lt_sub_one Nat.card_units_zmod_lt_sub_one theorem prime_iff_card_units (p : ℕ) [Fintype (ZMod p)ˣ] : diff --git a/Mathlib/Data/Nat/WithBot.lean b/Mathlib/Data/Nat/WithBot.lean index f523f068f8d40..7de807e457a8a 100644 --- a/Mathlib/Data/Nat/WithBot.lean +++ b/Mathlib/Data/Nat/WithBot.lean @@ -34,7 +34,8 @@ theorem add_eq_zero_iff {n m : WithBot ℕ} : n + m = 0 ↔ n = 0 ∧ m = 0 := b theorem add_eq_one_iff {n m : WithBot ℕ} : n + m = 1 ↔ n = 0 ∧ m = 1 ∨ n = 1 ∧ m = 0 := by rcases n, m with ⟨_ | _, _ | _⟩ - any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; aesop + any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; + aesop (simp_options := { decide := true }) repeat' erw [WithBot.coe_eq_coe] exact Nat.add_eq_one_iff #align nat.with_bot.add_eq_one_iff Nat.WithBot.add_eq_one_iff @@ -42,7 +43,8 @@ theorem add_eq_one_iff {n m : WithBot ℕ} : n + m = 1 ↔ n = 0 ∧ m = 1 ∨ n theorem add_eq_two_iff {n m : WithBot ℕ} : n + m = 2 ↔ n = 0 ∧ m = 2 ∨ n = 1 ∧ m = 1 ∨ n = 2 ∧ m = 0 := by rcases n, m with ⟨_ | _, _ | _⟩ - any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; aesop + any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; + aesop (simp_options := { decide := true }) repeat' erw [WithBot.coe_eq_coe] exact Nat.add_eq_two_iff #align nat.with_bot.add_eq_two_iff Nat.WithBot.add_eq_two_iff @@ -50,7 +52,8 @@ theorem add_eq_two_iff {n m : WithBot ℕ} : theorem add_eq_three_iff {n m : WithBot ℕ} : n + m = 3 ↔ n = 0 ∧ m = 3 ∨ n = 1 ∧ m = 2 ∨ n = 2 ∧ m = 1 ∨ n = 3 ∧ m = 0 := by rcases n, m with ⟨_ | _, _ | _⟩ - any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; aesop + any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; + aesop (simp_options := { decide := true }) repeat' erw [WithBot.coe_eq_coe] exact Nat.add_eq_three_iff #align nat.with_bot.add_eq_three_iff Nat.WithBot.add_eq_three_iff diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index a888f840872a6..8a57888b52aa2 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -250,7 +250,7 @@ theorem bit1_succ : ∀ n : Num, n.bit1.succ = n.succ.bit0 #align num.bit1_succ Num.bit1_succ theorem ofNat'_succ : ∀ {n}, ofNat' (n + 1) = ofNat' n + 1 := - @(Nat.binaryRec (by simp) fun b n ih => by + @(Nat.binaryRec (by simp; rfl) fun b n ih => by cases b · erw [ofNat'_bit true n, ofNat'_bit] simp only [← bit1_of_bit1, ← bit0_of_bit0, cond, _root_.bit1] diff --git a/Mathlib/Data/Option/Basic.lean b/Mathlib/Data/Option/Basic.lean index dc631e06da0ee..c9609459f7871 100644 --- a/Mathlib/Data/Option/Basic.lean +++ b/Mathlib/Data/Option/Basic.lean @@ -328,7 +328,7 @@ theorem getD_default_eq_iget [Inhabited α] (o : Option α) : @[simp] theorem guard_eq_some' {p : Prop} [Decidable p] (u) : _root_.guard p = some u ↔ p := by cases u - by_cases h : p <;> simp [_root_.guard, h] + by_cases h : p <;> simp [_root_.guard, h]; rfl #align option.guard_eq_some' Option.guard_eq_some' theorem liftOrGet_choice {f : α → α → α} (h : ∀ a b, f a b = a ∨ f a b = b) : diff --git a/Mathlib/Data/Ordmap/Ordset.lean b/Mathlib/Data/Ordmap/Ordset.lean index c9224969e1b44..1b88800c6b9c8 100644 --- a/Mathlib/Data/Ordmap/Ordset.lean +++ b/Mathlib/Data/Ordmap/Ordset.lean @@ -408,7 +408,7 @@ theorem Sized.dual_iff {t : Ordnode α} : Sized (.dual t) ↔ Sized t := theorem Sized.rotateL {l x r} (hl : @Sized α l) (hr : Sized r) : Sized (rotateL l x r) := by cases r; · exact hl.node' hr - rw [rotateL]; split_ifs + rw [Ordnode.rotateL]; split_ifs · exact hl.node3L hr.2.1 hr.2.2 · exact hl.node4L hr.2.1 hr.2.2 #align ordnode.sized.rotate_l Ordnode.Sized.rotateL @@ -1247,7 +1247,7 @@ theorem Valid'.rotateL {l} {x : α} {r o₁ o₂} (hl : Valid' o₁ l x) (hr : V have ablem : ∀ {a b : ℕ}, 1 ≤ a → a + b ≤ 2 → b ≤ 1 := by intros; linarith have hlp : size l > 0 → ¬size rl + size rr ≤ 1 := fun l0 hb => absurd (le_trans (le_trans (Nat.mul_le_mul_left _ l0) H2) hb) (by decide) - rw [rotateL]; split_ifs with h + rw [Ordnode.rotateL]; split_ifs with h · have rr0 : size rr > 0 := (mul_lt_mul_left (by decide)).1 (lt_of_le_of_lt (Nat.zero_le _) h : ratio * 0 < _) suffices BalancedSz (size l) (size rl) ∧ BalancedSz (size l + size rl + 1) (size rr) by diff --git a/Mathlib/Data/PFunctor/Univariate/Basic.lean b/Mathlib/Data/PFunctor/Univariate/Basic.lean index e97540bd76146..b5782de027a0c 100644 --- a/Mathlib/Data/PFunctor/Univariate/Basic.lean +++ b/Mathlib/Data/PFunctor/Univariate/Basic.lean @@ -18,7 +18,7 @@ pfunctor/M.lean.) -- "W", "Idx" set_option linter.uppercaseLean3 false -universe u v v₁ v₂ +universe u v v₁ v₂ v₃ /-- A polynomial functor `P` is given by a type `A` and a family `B` of types over `A`. `P` maps any type `α` to a new type `P α`, which is defined as the sigma type `Σ x, P.B x → α`. @@ -40,7 +40,7 @@ namespace PFunctor instance : Inhabited PFunctor := ⟨⟨default, default⟩⟩ -variable (P : PFunctor.{u}) {α : Type v₁} {β : Type v₂} +variable (P : PFunctor.{u}) {α : Type v₁} {β : Type v₂} {γ : Type v₃} /-- Applying `P` to an object of `Type` -/ @[coe] @@ -62,22 +62,30 @@ instance Obj.inhabited [Inhabited P.A] [Inhabited α] : Inhabited (P α) := instance : Functor.{v, max u v} P.Obj where map := @map P -protected theorem map_eq {α β : Type u} (f : α → β) (a : P.A) (g : P.B a → α) : - @Functor.map P.Obj _ _ _ f ⟨a, g⟩ = ⟨a, f ∘ g⟩ := +/-- We prefer `PFunctor.map` to `Functor.map` because it is universe-polymorphic. -/ +@[simp] +theorem map_eq_map {α β : Type v} (f : α → β) (x : P α) : f <$> x = P.map f x := + rfl + +@[simp] +protected theorem map_eq (f : α → β) (a : P.A) (g : P.B a → α) : + P.map f ⟨a, g⟩ = ⟨a, f ∘ g⟩ := rfl #align pfunctor.map_eq PFunctor.map_eq -protected theorem id_map {α : Type v} : ∀ x : P α, id <$> x = id x := fun ⟨_a, _b⟩ => rfl +@[simp] +protected theorem id_map : ∀ x : P α, P.map id x = x := fun ⟨_, _⟩ => rfl #align pfunctor.id_map PFunctor.id_map -protected theorem comp_map {α β γ : Type v} (f : α → β) (g : β → γ) : - ∀ x : P α, (g ∘ f) <$> x = g <$> f <$> x := fun ⟨_a, _b⟩ => rfl -#align pfunctor.comp_map PFunctor.comp_map +@[simp] +protected theorem map_map (f : α → β) (g : β → γ) : + ∀ x : P α, P.map g (P.map f x) = P.map (g ∘ f) x := fun ⟨_, _⟩ => rfl +#align pfunctor.comp_map PFunctor.map_map instance : LawfulFunctor.{v, max u v} P.Obj where map_const := rfl - id_map := @PFunctor.id_map P - comp_map := @PFunctor.comp_map P + id_map x := P.id_map x + comp_map f g x := P.map_map f g x |>.symm /-- re-export existing definition of W-types and adapt it to a packaged definition of polynomial functor -/ @@ -142,12 +150,12 @@ def Obj.iget [DecidableEq P.A] {α} [Inhabited α] (x : P α) (i : P.Idx) : α : #align pfunctor.obj.iget PFunctor.Obj.iget @[simp] -theorem fst_map {α β : Type u} (x : P α) (f : α → β) : (f <$> x).1 = x.1 := by cases x; rfl +theorem fst_map (x : P α) (f : α → β) : (P.map f x).1 = x.1 := by cases x; rfl #align pfunctor.fst_map PFunctor.fst_map @[simp] -theorem iget_map [DecidableEq P.A] {α β : Type u} [Inhabited α] [Inhabited β] (x : P α) - (f : α → β) (i : P.Idx) (h : i.1 = x.1) : (f <$> x).iget i = f (x.iget i) := by +theorem iget_map [DecidableEq P.A] [Inhabited α] [Inhabited β] (x : P α) + (f : α → β) (i : P.Idx) (h : i.1 = x.1) : (P.map f x).iget i = f (x.iget i) := by simp only [Obj.iget, fst_map, *, dif_pos, eq_self_iff_true] cases x rfl @@ -192,7 +200,7 @@ theorem liftp_iff {α : Type u} (p : α → Prop) (x : P α) : · rintro ⟨y, hy⟩ cases' h : y with a f refine' ⟨a, fun i => (f i).val, _, fun i => (f i).property⟩ - rw [← hy, h, PFunctor.map_eq] + rw [← hy, h, map_eq_map, PFunctor.map_eq] congr rintro ⟨a, f, xeq, pf⟩ use ⟨a, fun i => ⟨f i, pf i⟩⟩ diff --git a/Mathlib/Data/PFunctor/Univariate/M.lean b/Mathlib/Data/PFunctor/Univariate/M.lean index 3b5a9a8ccb47d..31072132ad0bc 100644 --- a/Mathlib/Data/PFunctor/Univariate/M.lean +++ b/Mathlib/Data/PFunctor/Univariate/M.lean @@ -572,14 +572,14 @@ theorem iselect_cons [DecidableEq F.A] [Inhabited (M F)] (ps : Path F) {a} (f : set_option linter.uppercaseLean3 false in #align pfunctor.M.iselect_cons PFunctor.M.iselect_cons -theorem corec_def {X} (f : X → F X) (x₀ : X) : M.corec f x₀ = M.mk (M.corec f <$> f x₀) := by +theorem corec_def {X} (f : X → F X) (x₀ : X) : M.corec f x₀ = M.mk (F.map (M.corec f) (f x₀)) := by dsimp only [M.corec, M.mk] congr with n cases' n with n · dsimp only [sCorec, Approx.sMk] · dsimp only [sCorec, Approx.sMk] cases h : f x₀ - dsimp only [(· <$> ·), PFunctor.map] + dsimp only [PFunctor.map] congr set_option linter.uppercaseLean3 false in #align pfunctor.M.corec_def PFunctor.M.corec_def @@ -712,9 +712,9 @@ def corecOn {X : Type*} (x₀ : X) (f : X → F X) : M F := set_option linter.uppercaseLean3 false in #align pfunctor.M.corec_on PFunctor.M.corecOn -variable {P : PFunctor.{u}} {α : Type u} +variable {P : PFunctor.{u}} {α : Type*} -theorem dest_corec (g : α → P α) (x : α) : M.dest (M.corec g x) = M.corec g <$> g x := by +theorem dest_corec (g : α → P α) (x : α) : M.dest (M.corec g x) = P.map (M.corec g) (g x) := by rw [corec_def, dest_mk] set_option linter.uppercaseLean3 false in #align pfunctor.M.dest_corec PFunctor.M.dest_corec @@ -766,7 +766,7 @@ theorem bisim_equiv (R : M P → M P → Prop) set_option linter.uppercaseLean3 false in #align pfunctor.M.bisim_equiv PFunctor.M.bisim_equiv -theorem corec_unique (g : α → P α) (f : α → M P) (hyp : ∀ x, M.dest (f x) = f <$> g x) : +theorem corec_unique (g : α → P α) (f : α → M P) (hyp : ∀ x, M.dest (f x) = P.map f (g x)) : f = M.corec g := by ext x apply bisim' (fun _ => True) _ _ _ _ trivial @@ -796,7 +796,7 @@ def corec' {α : Type u} (F : ∀ {X : Type u}, (α → X) → α → Sum (M P) let y := a >>= F (rec ∘ Sum.inr) match y with | Sum.inr y => y - | Sum.inl y => (rec ∘ Sum.inl) <$> M.dest y) + | Sum.inl y => P.map (rec ∘ Sum.inl) (M.dest y)) (@Sum.inr (M P) _ x) set_option linter.uppercaseLean3 false in #align pfunctor.M.corec' PFunctor.M.corec' diff --git a/Mathlib/Data/PSigma/Order.lean b/Mathlib/Data/PSigma/Order.lean index c56dcb9735ea0..881ca4a5ebe33 100644 --- a/Mathlib/Data/PSigma/Order.lean +++ b/Mathlib/Data/PSigma/Order.lean @@ -37,7 +37,8 @@ namespace PSigma /-- The notation `Σₗ' i, α i` refers to a sigma type which is locally equipped with the lexicographic order.-/ -notation3 "Σₗ' "(...)", "r:(scoped p => _root_.Lex (PSigma p)) => r +-- TODO: make `Lex` be `Sort u -> Sort u` so we can remove `.{_+1, _+1}` +notation3 "Σₗ' "(...)", "r:(scoped p => _root_.Lex (PSigma.{_+1, _+1} p)) => r namespace Lex diff --git a/Mathlib/Data/Polynomial/Coeff.lean b/Mathlib/Data/Polynomial/Coeff.lean index d4ecf54629c5d..9ac343613bdfc 100644 --- a/Mathlib/Data/Polynomial/Coeff.lean +++ b/Mathlib/Data/Polynomial/Coeff.lean @@ -111,13 +111,13 @@ theorem coeff_sum [Semiring S] (n : ℕ) (f : ℕ → R → S[X]) : #align polynomial.coeff_sum Polynomial.coeff_sum /-- Decomposes the coefficient of the product `p * q` as a sum -over `Nat.antidiagonal`. A version which sums over `range (n + 1)` can be obtained +over `antidiagonal`. A version which sums over `range (n + 1)` can be obtained by using `Finset.Nat.sum_antidiagonal_eq_sum_range_succ`. -/ theorem coeff_mul (p q : R[X]) (n : ℕ) : - coeff (p * q) n = ∑ x in Nat.antidiagonal n, coeff p x.1 * coeff q x.2 := by + coeff (p * q) n = ∑ x in antidiagonal n, coeff p x.1 * coeff q x.2 := by rcases p with ⟨p⟩; rcases q with ⟨q⟩ simp_rw [← ofFinsupp_mul, coeff] - exact AddMonoidAlgebra.mul_apply_antidiagonal p q n _ Nat.mem_antidiagonal + exact AddMonoidAlgebra.mul_apply_antidiagonal p q n _ Finset.mem_antidiagonal #align polynomial.coeff_mul Polynomial.coeff_mul @[simp] @@ -211,7 +211,7 @@ theorem support_binomial {k m : ℕ} (hkm : k ≠ m) {x y : R} (hx : x ≠ 0) (h apply subset_antisymm (support_binomial' k m x y) simp_rw [insert_subset_iff, singleton_subset_iff, mem_support_iff, coeff_add, coeff_C_mul, coeff_X_pow_self, mul_one, coeff_X_pow, if_neg hkm, if_neg hkm.symm, mul_zero, zero_add, - add_zero, Ne.def, hx, hy] + add_zero, Ne.def, hx, hy, not_false_eq_true, and_true] #align polynomial.support_binomial Polynomial.support_binomial theorem support_trinomial {k m n : ℕ} (hkm : k < m) (hmn : m < n) {x y z : R} (hx : x ≠ 0) @@ -221,7 +221,7 @@ theorem support_trinomial {k m n : ℕ} (hkm : k < m) (hmn : m < n) {x y z : R} simp_rw [insert_subset_iff, singleton_subset_iff, mem_support_iff, coeff_add, coeff_C_mul, coeff_X_pow_self, mul_one, coeff_X_pow, if_neg hkm.ne, if_neg hkm.ne', if_neg hmn.ne, if_neg hmn.ne', if_neg (hkm.trans hmn).ne, if_neg (hkm.trans hmn).ne', mul_zero, add_zero, - zero_add, Ne.def, hx, hy, hz] + zero_add, Ne.def, hx, hy, hz, not_false_eq_true, and_true] #align polynomial.support_trinomial Polynomial.support_trinomial theorem card_support_binomial {k m : ℕ} (h : k ≠ m) {x y : R} (hx : x ≠ 0) (hy : y ≠ 0) : @@ -247,10 +247,10 @@ theorem coeff_mul_X_pow (p : R[X]) (n d : ℕ) : rw [coeff_X_pow, if_neg, mul_zero] rintro rfl apply h2 - rw [Nat.mem_antidiagonal, add_right_cancel_iff] at h1 + rw [mem_antidiagonal, add_right_cancel_iff] at h1 subst h1 rfl - · exact fun h1 => (h1 (Nat.mem_antidiagonal.2 rfl)).elim + · exact fun h1 => (h1 (mem_antidiagonal.2 rfl)).elim #align polynomial.coeff_mul_X_pow Polynomial.coeff_mul_X_pow @[simp] @@ -264,7 +264,7 @@ theorem coeff_mul_X_pow' (p : R[X]) (n d : ℕ) : · rw [← tsub_add_cancel_of_le h, coeff_mul_X_pow, add_tsub_cancel_right] · refine' (coeff_mul _ _ _).trans (Finset.sum_eq_zero fun x hx => _) rw [coeff_X_pow, if_neg, mul_zero] - exact ((le_of_add_le_right (Finset.Nat.mem_antidiagonal.mp hx).le).trans_lt <| not_le.mp h).ne + exact ((le_of_add_le_right (mem_antidiagonal.mp hx).le).trans_lt <| not_le.mp h).ne #align polynomial.coeff_mul_X_pow' Polynomial.coeff_mul_X_pow' theorem coeff_X_pow_mul' (p : R[X]) (n d : ℕ) : diff --git a/Mathlib/Data/Polynomial/Degree/Definitions.lean b/Mathlib/Data/Polynomial/Degree/Definitions.lean index d602e0aba009a..d4c47e7be2ce9 100644 --- a/Mathlib/Data/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Data/Polynomial/Degree/Definitions.lean @@ -55,7 +55,7 @@ def degree (p : R[X]) : WithBot ℕ := #align polynomial.degree Polynomial.degree theorem degree_lt_wf : WellFounded fun p q : R[X] => degree p < degree q := - InvImage.wf degree (WithBot.wellFounded_lt Nat.lt_wfRel.wf) + InvImage.wf degree wellFounded_lt #align polynomial.degree_lt_wf Polynomial.degree_lt_wf instance : WellFoundedRelation R[X] := @@ -916,12 +916,12 @@ theorem coeff_mul_degree_add_degree (p q : R[X]) : coeff (p * q) (natDegree p + natDegree q) = leadingCoeff p * leadingCoeff q := calc coeff (p * q) (natDegree p + natDegree q) = - ∑ x in Nat.antidiagonal (natDegree p + natDegree q), coeff p x.1 * coeff q x.2 := + ∑ x in antidiagonal (natDegree p + natDegree q), coeff p x.1 * coeff q x.2 := coeff_mul _ _ _ _ = coeff p (natDegree p) * coeff q (natDegree q) := by refine' Finset.sum_eq_single (natDegree p, natDegree q) _ _ · rintro ⟨i, j⟩ h₁ h₂ - rw [Nat.mem_antidiagonal] at h₁ + rw [mem_antidiagonal] at h₁ by_cases H : natDegree p < i · rw [coeff_eq_zero_of_degree_lt (lt_of_le_of_lt degree_le_natDegree (WithBot.coe_lt_coe.2 H)), @@ -944,7 +944,7 @@ theorem coeff_mul_degree_add_degree (p q : R[X]) : · intro H exfalso apply H - rw [Nat.mem_antidiagonal] + rw [mem_antidiagonal] #align polynomial.coeff_mul_degree_add_degree Polynomial.coeff_mul_degree_add_degree theorem degree_mul' (h : leadingCoeff p * leadingCoeff q ≠ 0) : @@ -1098,14 +1098,14 @@ theorem coeff_mul_add_eq_of_natDegree_le {df dg : ℕ} {g : R[X]} (hdf : natDegree f ≤ df) (hdg : natDegree g ≤ dg) : (f * g).coeff (df + dg) = f.coeff df * g.coeff dg := by rw [coeff_mul, Finset.sum_eq_single_of_mem (df, dg)] - · rw [Finset.Nat.mem_antidiagonal] + · rw [mem_antidiagonal] rintro ⟨df', dg'⟩ hmem hne obtain h | hdf' := lt_or_le df df' · rw [coeff_eq_zero_of_natDegree_lt (hdf.trans_lt h), zero_mul] obtain h | hdg' := lt_or_le dg dg' · rw [coeff_eq_zero_of_natDegree_lt (hdg.trans_lt h), mul_zero] obtain ⟨rfl, rfl⟩ := - (add_eq_add_iff_eq_and_eq hdf' hdg').mp (Finset.Nat.mem_antidiagonal.1 hmem) + (add_eq_add_iff_eq_and_eq hdf' hdg').mp (mem_antidiagonal.1 hmem) exact (hne rfl).elim theorem zero_le_degree_iff : 0 ≤ degree p ↔ p ≠ 0 := by diff --git a/Mathlib/Data/Polynomial/Degree/TrailingDegree.lean b/Mathlib/Data/Polynomial/Degree/TrailingDegree.lean index 503da1c656e2c..507f2a4792a0c 100644 --- a/Mathlib/Data/Polynomial/Degree/TrailingDegree.lean +++ b/Mathlib/Data/Polynomial/Degree/TrailingDegree.lean @@ -48,7 +48,7 @@ def trailingDegree (p : R[X]) : ℕ∞ := #align polynomial.trailing_degree Polynomial.trailingDegree theorem trailingDegree_lt_wf : WellFounded fun p q : R[X] => trailingDegree p < trailingDegree q := - InvImage.wf trailingDegree (WithTop.wellFounded_lt Nat.lt_wfRel.2) + InvImage.wf trailingDegree wellFounded_lt #align polynomial.trailing_degree_lt_wf Polynomial.trailingDegree_lt_wf /-- `natTrailingDegree p` forces `trailingDegree p` to `ℕ`, by defining @@ -363,7 +363,7 @@ theorem le_trailingDegree_mul : p.trailingDegree + q.trailingDegree ≤ (p * q). (add_le_add (min_le (mem_support_iff.mpr (left_ne_zero_of_mul hpq))) (min_le (mem_support_iff.mpr (right_ne_zero_of_mul hpq)))).trans (le_of_eq _) - rwa [← WithTop.coe_add, WithTop.coe_eq_coe, ← Nat.mem_antidiagonal] + rwa [← WithTop.coe_add, WithTop.coe_eq_coe, ← mem_antidiagonal] #align polynomial.le_trailing_degree_mul Polynomial.le_trailingDegree_mul theorem le_natTrailingDegree_mul (h : p * q ≠ 0) : @@ -384,9 +384,9 @@ theorem coeff_mul_natTrailingDegree_add_natTrailingDegree : (p * q).coeff rw [coeff_mul] refine' Finset.sum_eq_single (p.natTrailingDegree, q.natTrailingDegree) _ fun h => - (h (Nat.mem_antidiagonal.mpr rfl)).elim + (h (mem_antidiagonal.mpr rfl)).elim rintro ⟨i, j⟩ h₁ h₂ - rw [Nat.mem_antidiagonal] at h₁ + rw [mem_antidiagonal] at h₁ by_cases hi : i < p.natTrailingDegree · rw [coeff_eq_zero_of_lt_natTrailingDegree hi, zero_mul] by_cases hj : j < q.natTrailingDegree diff --git a/Mathlib/Data/Polynomial/Derivation.lean b/Mathlib/Data/Polynomial/Derivation.lean index ad8bceb8591b6..100ef936804c5 100644 --- a/Mathlib/Data/Polynomial/Derivation.lean +++ b/Mathlib/Data/Polynomial/Derivation.lean @@ -6,7 +6,7 @@ Authors: Kevin Buzzard, Richard Hill import Mathlib.Data.Polynomial.Derivative import Mathlib.RingTheory.Derivation.Basic import Mathlib.Data.Polynomial.AlgebraMap - +import Mathlib.Data.Polynomial.Module /-! # Derivations of univariate polynomials @@ -20,7 +20,7 @@ noncomputable section namespace Polynomial -open scoped BigOperators +section CommSemiring variable {R A : Type*} [CommSemiring R] @@ -92,4 +92,59 @@ def mkDerivationEquiv : A ≃ₗ[R] Derivation R R[X] A := @[simp] lemma mkDerivationEquiv_symm_apply (D : Derivation R R[X] A) : (mkDerivationEquiv R).symm D = D X := rfl +end CommSemiring end Polynomial + +namespace Derivation + +variable {R A M : Type*} [CommSemiring R] [CommSemiring A] [Algebra R A] [AddCommMonoid M] + [Module A M] [Module R M] [IsScalarTower R A M] (d : Derivation R A M) (a : A) + +open Polynomial Module + +/-- +For a derivation `d : A → M` and an element `a : A`, `d.compAEval a` is the +derivation of `R[X]` which takes a polynomial `f` to `d(aeval a f)`. + +This derivation takes values in `Module.AEval R M a`, which is `M`, regarded as an +`R[X]`-module, with the action of a polynomial `f` defined by `f • m = (aeval a f) • m`. +-/ +/- +Note: `compAEval` is not defined using `Derivation.compAlgebraMap`. +This because `A` is not an `R[X]` algebra and it would be messy to create an algebra instance +within the definition. +-/ +@[simps] +def compAEval : Derivation R R[X] <| AEval R M a where + toFun f := AEval.of R M a (d (aeval a f)) + map_add' := by simp + map_smul' := by simp + leibniz' := by simp [AEval.of_aeval_smul] + map_one_eq_zero' := by simp + +/-- + A form of the chain rule: if `f` is a polynomial over `R` + and `d : A → M` is an `R`-derivation then for all `a : A` we have + $$ d(f(a)) = f' (a) d a. $$ + The equation is in the `R[X]`-module `Module.AEval R M a`. + For the same equation in `M`, see `Derivation.compAEval_eq`. +-/ +theorem compAEval_eq (d : Derivation R A M) (f : R[X]) : + d.compAEval a f = derivative f • (AEval.of R M a (d a)) := by + rw [←mkDerivation_apply] + congr + apply derivation_ext + simp + +/-- + A form of the chain rule: if `f` is a polynomial over `R` + and `d : A → M` is an `R`-derivation then for all `a : A` we have + $$ d(f(a)) = f' (a) d a. $$ + The equation is in `M`. For the same equation in `Module.AEval R M a`, + see `Derivation.compAEval_eq`. +-/ +theorem comp_aeval_eq (d : Derivation R A M) (f : R[X]) : + d (aeval a f) = aeval a (derivative f) • d a := + calc + _ = (AEval.of R M a).symm (d.compAEval a f) := rfl + _ = _ := by simp [-compAEval_apply, compAEval_eq] diff --git a/Mathlib/Data/Polynomial/Derivative.lean b/Mathlib/Data/Polynomial/Derivative.lean index 8609f48f059b3..ba62878b67843 100644 --- a/Mathlib/Data/Polynomial/Derivative.lean +++ b/Mathlib/Data/Polynomial/Derivative.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Scott Morrison, Jens Wagemaker -/ -import Mathlib.Algebra.Hom.Iterate +import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Data.Polynomial.Eval #align_import data.polynomial.derivative from "leanprover-community/mathlib"@"bbeb185db4ccee8ed07dc48449414ebfa39cb821" @@ -216,7 +216,7 @@ theorem natDegree_derivative_lt {p : R[X]} (hp : p.natDegree ≠ 0) : theorem natDegree_derivative_le (p : R[X]) : p.derivative.natDegree ≤ p.natDegree - 1 := by by_cases p0 : p.natDegree = 0 · simp [p0, derivative_of_natDegree_zero] - · exact Nat.le_pred_of_lt (natDegree_derivative_lt p0) + · exact Nat.le_sub_one_of_lt (natDegree_derivative_lt p0) #align polynomial.nat_degree_derivative_le Polynomial.natDegree_derivative_le theorem natDegree_iterate_derivative (p : R[X]) (k : ℕ) : diff --git a/Mathlib/Data/Polynomial/Div.lean b/Mathlib/Data/Polynomial/Div.lean index a72a8076d7d33..0c7df6380fe60 100644 --- a/Mathlib/Data/Polynomial/Div.lean +++ b/Mathlib/Data/Polynomial/Div.lean @@ -74,7 +74,7 @@ theorem multiplicity_finite_of_degree_pos_of_monic (hp : (0 : WithBot ℕ) < deg have hpn0' : leadingCoeff p ^ (natDegree q + 1) ≠ 0 := hpn1.symm ▸ zn0.symm have hpnr0 : leadingCoeff (p ^ (natDegree q + 1)) * leadingCoeff r ≠ 0 := by simp only [leadingCoeff_pow' hpn0', leadingCoeff_eq_zero, hpn1, one_pow, one_mul, Ne.def, - hr0] + hr0, not_false_eq_true] have hnp : 0 < natDegree p := Nat.cast_lt.1 <| by rw [← degree_eq_natDegree hp0]; exact hp have := congr_arg natDegree hr diff --git a/Mathlib/Data/Polynomial/EraseLead.lean b/Mathlib/Data/Polynomial/EraseLead.lean index 2a916fb970dfc..d2bbe05b6cba3 100644 --- a/Mathlib/Data/Polynomial/EraseLead.lean +++ b/Mathlib/Data/Polynomial/EraseLead.lean @@ -204,7 +204,7 @@ theorem eraseLead_natDegree_lt_or_eraseLead_eq_zero (f : R[X]) : theorem eraseLead_natDegree_le (f : R[X]) : (eraseLead f).natDegree ≤ f.natDegree - 1 := by rcases f.eraseLead_natDegree_lt_or_eraseLead_eq_zero with (h | h) - · exact Nat.le_pred_of_lt h + · exact Nat.le_sub_one_of_lt h · simp only [h, natDegree_zero, zero_le] #align polynomial.erase_lead_nat_degree_le Polynomial.eraseLead_natDegree_le diff --git a/Mathlib/Data/Polynomial/FieldDivision.lean b/Mathlib/Data/Polynomial/FieldDivision.lean index f0a7fc6b6eda8..5aff519ea1db8 100644 --- a/Mathlib/Data/Polynomial/FieldDivision.lean +++ b/Mathlib/Data/Polynomial/FieldDivision.lean @@ -176,7 +176,7 @@ theorem isUnit_iff_degree_eq_zero : IsUnit p ↔ degree p = 0 := ⟨degree_eq_zero_of_isUnit, fun h => have : degree p ≤ 0 := by simp [*, le_refl] have hc : coeff p 0 ≠ 0 := fun hc => by - rw [eq_C_of_degree_le_zero this, hc] at h; simp at h + rw [eq_C_of_degree_le_zero this, hc] at h; simp only [map_zero] at h; contradiction isUnit_iff_dvd_one.2 ⟨C (coeff p 0)⁻¹, by conv in p => rw [eq_C_of_degree_le_zero this] @@ -403,7 +403,7 @@ theorem exists_root_of_degree_eq_one (h : degree p = 1) : ∃ x, IsRoot p x := have : p.coeff 1 ≠ 0 := by have h' := natDegree_eq_of_degree_eq_some h change natDegree p = 1 at h'; rw [←h'] - exact mt leadingCoeff_eq_zero.1 fun h0 => by simp [h0] at h + exact mt leadingCoeff_eq_zero.1 fun h0 => by simp [h0] at h; contradiction conv in p => rw [eq_X_add_C_of_degree_le_one (show degree p ≤ 1 by rw [h])] simp [IsRoot, mul_div_cancel' _ this]⟩ #align polynomial.exists_root_of_degree_eq_one Polynomial.exists_root_of_degree_eq_one @@ -486,7 +486,7 @@ theorem prime_of_degree_eq_one (hp1 : degree p = 1) : Prime p := by classical have : Prime (normalize p) := Monic.prime_of_degree_eq_one (hp1 ▸ degree_normalize) - (monic_normalize fun hp0 => absurd hp1 (hp0.symm ▸ by simp)) + (monic_normalize fun hp0 => absurd hp1 (hp0.symm ▸ by simp only [degree_zero]; decide)) exact (normalize_associated _).prime this #align polynomial.prime_of_degree_eq_one Polynomial.prime_of_degree_eq_one diff --git a/Mathlib/Data/Polynomial/HasseDeriv.lean b/Mathlib/Data/Polynomial/HasseDeriv.lean index bcfad8d8efad0..d909e1b11a6b0 100644 --- a/Mathlib/Data/Polynomial/HasseDeriv.lean +++ b/Mathlib/Data/Polynomial/HasseDeriv.lean @@ -227,6 +227,8 @@ section open AddMonoidHom Finset.Nat +open Finset (antidiagonal mem_antidiagonal) + theorem hasseDeriv_mul (f g : R[X]) : hasseDeriv k (f * g) = ∑ ij in antidiagonal k, hasseDeriv ij.1 f * hasseDeriv ij.2 g := by let D k := (@hasseDeriv R _ k).toAddMonoidHom @@ -247,7 +249,7 @@ theorem hasseDeriv_mul (f g : R[X]) : monomial (m - x.1 + (n - x.2)) (↑(m.choose x.1) * r * (↑(n.choose x.2) * s)) = monomial (m + n - k) (↑(m.choose x.1) * ↑(n.choose x.2) * (r * s)) := by intro x hx - rw [Finset.Nat.mem_antidiagonal] at hx + rw [mem_antidiagonal] at hx subst hx by_cases hm : m < x.1 · simp only [Nat.choose_eq_zero_of_lt hm, Nat.cast_zero, zero_mul, diff --git a/Mathlib/Data/Polynomial/Module.lean b/Mathlib/Data/Polynomial/Module.lean index e891b23baf36c..117636f6122df 100644 --- a/Mathlib/Data/Polynomial/Module.lean +++ b/Mathlib/Data/Polynomial/Module.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.RingTheory.FiniteType +import Mathlib.RingTheory.Adjoin.Tower #align_import data.polynomial.module from "leanprover-community/mathlib"@"63417e01fbc711beaf25fa73b6edb395c0cfddd0" @@ -15,16 +15,105 @@ In this file, we define the polynomial module for an `R`-module `M`, i.e. the `R This is defined as a type alias `PolynomialModule R M := ℕ →₀ M`, since there might be different module structures on `ℕ →₀ M` of interest. See the docstring of `PolynomialModule` for details. +We also define, given an element `a` in an `R`-algebra `A` and an `A`-module `M`, an `R[X]`-module +`Module.AEval R M a`, which is a type synonym of `M` with the action of a polynomial `f` +given by `f • m = Polynomial.aeval a f • m`. In particular `X • m = a • m`. + +In the special case that `A = M →ₗ[R] M` and `φ : M →ₗ[R] M`, the module `Module.AEval R M a` is +abbreviated `Module.AEval' φ`. In this module we have `X • m = ↑φ m`. -/ +universe u v +open Polynomial BigOperators +namespace Module +/-- +Suppose `a` is an element of an `R`-algebra `A` and `M` is an `A`-module. +Loosely speaking, `Module.AEval R M a` is the `R[X]`-module with elements `m : M`, +where the action of a polynomial $f$ is given by $f • m = f(a) • m$. -universe u v +More precisely, `Module.AEval R M a` has elements `Module.AEval.of R M a m` for `m : M`, +and the action of `f` is `f • (of R M a m) = of R M a ((aeval a f) • m)`. +-/ +@[nolint unusedArguments] +def AEval (R M : Type*) {A : Type*} [CommSemiring R] [Semiring A] [Algebra R A] + [AddCommMonoid M] [Module A M] [Module R M] [IsScalarTower R A M] (_ : A) := M -open Polynomial +instance AEval.instAddCommGroup {R A M} [CommSemiring R] [Semiring A] (a : A) [Algebra R A] + [AddCommGroup M] [Module A M] [Module R M] [IsScalarTower R A M] : + AddCommGroup <| AEval R M a := inferInstanceAs (AddCommGroup M) -open Polynomial BigOperators +variable {R A M} [CommSemiring R] [Semiring A] (a : A) [Algebra R A] [AddCommMonoid M] [Module A M] + [Module R M] [IsScalarTower R A M] + +namespace AEval + +instance instAddCommMonoid : AddCommMonoid <| AEval R M a := inferInstanceAs (AddCommMonoid M) + +instance instModuleOrig : Module R <| AEval R M a := inferInstanceAs (Module R M) + +instance instFiniteOrig [Finite R M] : Finite R <| AEval R M a := inferInstanceAs (Finite R M) + +instance instModulePolynomial : Module R[X] <| AEval R M a := compHom M (aeval a).toRingHom + +variable (R M) +/-- +The canonical linear equivalence between `M` and `Module.AEval R M a` as an `R`-module. +-/ +def of : M ≃ₗ[R] AEval R M a := + LinearEquiv.refl _ _ + +variable {R M} +lemma of_aeval_smul (f : R[X]) (m : M) : of R M a (aeval a f • m) = f • of R M a m := rfl +@[simp] lemma of_symm_smul (f : R[X]) (m : AEval R M a) : + (of R M a).symm (f • m) = aeval a f • (of R M a).symm m := rfl + +lemma X_smul_of (m : M) : (X : R[X]) • (of R M a m) = of R M a (a • m) := by + rw [←of_aeval_smul, aeval_X] + +lemma of_symm_X_smul (m : AEval R M a) : + (of R M a).symm ((X : R[X]) • m) = a • (of R M a).symm m := by + rw [of_symm_smul, aeval_X] + +instance instIsScalarTowerOrigPolynomial : IsScalarTower R R[X] <| AEval R M a where + smul_assoc r f m := by + apply (of R M a).symm.injective + rw [of_symm_smul, map_smul, smul_assoc] + rfl + +instance instFinitePolynomial [Finite R M] : Finite R[X] <| AEval R M a := + Finite.of_restrictScalars_finite R _ _ + +end AEval + +variable (φ : M →ₗ[R] M) +/-- +Given and `R`-module `M` and a linear map `φ : M →ₗ[R] M`, `Module.AEval' φ` is loosely speaking +the `R[X]`-module with elements `m : M`, where the action of a polynomial $f$ is given by +$f • m = f(a) • m$. + +More precisely, `Module.AEval' φ` has elements `Module.AEval'.of φ m` for `m : M`, +and the action of `f` is `f • (of φ m) = of φ ((aeval φ f) • m)`. + +`Module.AEval'` is defined as a special case of `Module.AEval` in which the `R`-algebra is +`M →ₗ[R] M`. Lemmas involving `Module.AEval` may be applied to `Module.AEval'`. +-/ +abbrev AEval' := AEval R M φ +/-- +The canonical linear equivalence between `M` and `Module.AEval' φ` as an `R`-module, +where `φ : M →ₗ[R] M`. +-/ +abbrev AEval'.of : M ≃ₗ[R] AEval' φ := AEval.of R M φ +lemma AEval'_def : AEval' φ = AEval R M φ := rfl +lemma AEval'.X_smul_of (m : M) : (X : R[X]) • AEval'.of φ m = AEval'.of φ (φ m) := + AEval.X_smul_of _ _ +lemma AEval'.of_symm_X_smul (m : AEval' φ) : + (AEval'.of φ).symm ((X : R[X]) • m) = φ ((AEval'.of φ).symm m) := AEval.of_symm_X_smul _ _ + +instance [Finite R M] : Finite R[X] <| AEval' φ := inferInstance + +end Module /-- The `R[X]`-module `M[X]` for an `R`-module `M`. This is isomorphic (as an `R`-module) to `M[X]` when `M` is a ring. @@ -109,16 +198,21 @@ theorem induction_linear {P : PolynomialModule R M → Prop} (f : PolynomialModu @[semireducible] noncomputable instance polynomialModule : Module R[X] (PolynomialModule R M) := - modulePolynomialOfEndo (Finsupp.lmapDomain _ _ Nat.succ) + inferInstanceAs (Module R[X] (Module.AEval' (Finsupp.lmapDomain M R Nat.succ))) #align polynomial_module.polynomial_module PolynomialModule.polynomialModule +lemma smul_def (f : R[X]) (m : PolynomialModule R M) : + f • m = aeval (Finsupp.lmapDomain M R Nat.succ) f m := by + rfl + instance (M : Type u) [AddCommGroup M] [Module R M] [Module S M] [IsScalarTower S R M] : IsScalarTower S R (PolynomialModule R M) := Finsupp.isScalarTower _ _ instance isScalarTower' (M : Type u) [AddCommGroup M] [Module R M] [Module S M] [IsScalarTower S R M] : IsScalarTower S R[X] (PolynomialModule R M) := by - haveI : IsScalarTower R R[X] (PolynomialModule R M) := modulePolynomialOfEndo.isScalarTower _ + haveI : IsScalarTower R R[X] (PolynomialModule R M) := + inferInstanceAs <| IsScalarTower R R[X] <| Module.AEval' <| Finsupp.lmapDomain M R Nat.succ constructor intro x y z rw [← @IsScalarTower.algebraMap_smul S R, ← @IsScalarTower.algebraMap_smul S R, smul_assoc] @@ -128,7 +222,7 @@ instance isScalarTower' (M : Type u) [AddCommGroup M] [Module R M] [Module S M] theorem monomial_smul_single (i : ℕ) (r : R) (j : ℕ) (m : M) : monomial i r • single R j m = single R (i + j) (r • m) := by simp only [LinearMap.mul_apply, Polynomial.aeval_monomial, LinearMap.pow_apply, - Module.algebraMap_end_apply, modulePolynomialOfEndo_smul_def] + Module.algebraMap_end_apply, smul_def] induction i generalizing r j m with | zero => rw [Nat.zero_eq, Function.iterate_zero, zero_add] @@ -136,7 +230,7 @@ theorem monomial_smul_single (i : ℕ) (r : R) (j : ℕ) (m : M) : | succ n hn => rw [Function.iterate_succ, Function.comp_apply, Nat.succ_eq_add_one, add_assoc, ← hn] congr 2 - rw [← Nat.succ_eq_one_add] + rw [Nat.one_add] exact Finsupp.mapDomain_single #align polynomial_module.monomial_smul_single PolynomialModule.monomial_smul_single @@ -175,7 +269,7 @@ theorem smul_single_apply (i : ℕ) (f : R[X]) (m : M) (n : ℕ) : #align polynomial_module.smul_single_apply PolynomialModule.smul_single_apply theorem smul_apply (f : R[X]) (g : PolynomialModule R M) (n : ℕ) : - (f • g) n = ∑ x in Finset.Nat.antidiagonal n, f.coeff x.1 • g x.2 := by + (f • g) n = ∑ x in Finset.antidiagonal n, f.coeff x.1 • g x.2 := by induction' f using Polynomial.induction_on' with p q hp hq f_n f_a · rw [add_smul, Finsupp.add_apply, hp, hq, ← Finset.sum_add_distrib] congr @@ -207,7 +301,7 @@ noncomputable def equivPolynomialSelf : PolynomialModule R R ≃ₗ[R[X]] R[X] : · rw [Finset.sum_eq_single (i - n, n)] simp only [ite_true] · rintro ⟨p, q⟩ hpq1 hpq2 - rw [Finset.Nat.mem_antidiagonal] at hpq1 + rw [Finset.mem_antidiagonal] at hpq1 split_ifs with H · dsimp at H exfalso @@ -218,10 +312,10 @@ noncomputable def equivPolynomialSelf : PolynomialModule R R ≃ₗ[R[X]] R[X] : · intro H exfalso apply H - rw [Finset.Nat.mem_antidiagonal, tsub_add_cancel_of_le hn] + rw [Finset.mem_antidiagonal, tsub_add_cancel_of_le hn] · symm rw [Finset.sum_ite_of_false, Finset.sum_const_zero] - simp_rw [Finset.Nat.mem_antidiagonal] + simp_rw [Finset.mem_antidiagonal] intro x hx contrapose! hn rw [add_comm, ← hn] at hx diff --git a/Mathlib/Data/Polynomial/Monic.lean b/Mathlib/Data/Polynomial/Monic.lean index 700f20833f74e..765dae368d665 100644 --- a/Mathlib/Data/Polynomial/Monic.lean +++ b/Mathlib/Data/Polynomial/Monic.lean @@ -211,7 +211,7 @@ theorem nextCoeff_mul (hp : Monic p) (hq : Monic q) : nontriviality simp only [← coeff_one_reverse] rw [reverse_mul] <;> - simp [coeff_mul, Nat.antidiagonal, hp.leadingCoeff, hq.leadingCoeff, add_comm, + simp [coeff_mul, antidiagonal, hp.leadingCoeff, hq.leadingCoeff, add_comm, show Nat.succ 0 = 1 from rfl] #align polynomial.monic.next_coeff_mul Polynomial.Monic.nextCoeff_mul diff --git a/Mathlib/Data/Polynomial/RingDivision.lean b/Mathlib/Data/Polynomial/RingDivision.lean index 54660e1cba954..1af8098cc72be 100644 --- a/Mathlib/Data/Polynomial/RingDivision.lean +++ b/Mathlib/Data/Polynomial/RingDivision.lean @@ -105,6 +105,12 @@ def modByMonicHom (q : R[X]) : R[X] →ₗ[R] R[X] where map_smul' := smul_modByMonic #align polynomial.mod_by_monic_hom Polynomial.modByMonicHom +theorem neg_modByMonic (p mod : R[X]) : (-p) %ₘ mod = - (p %ₘ mod) := + (modByMonicHom mod).map_neg p + +theorem sub_modByMonic (a b mod : R[X]) : (a - b) %ₘ mod = a %ₘ mod - b %ₘ mod := + (modByMonicHom mod).map_sub a b + end section @@ -1294,7 +1300,7 @@ theorem count_map_roots_of_injective [IsDomain A] [DecidableEq B] (p : A[X]) {f (p.roots.map f).count b ≤ rootMultiplicity b (p.map f) := by by_cases hp0 : p = 0 · simp only [hp0, roots_zero, Multiset.map_zero, Multiset.count_zero, Polynomial.map_zero, - rootMultiplicity_zero] + rootMultiplicity_zero, le_refl] · exact count_map_roots ((Polynomial.map_ne_zero_iff hf).mpr hp0) b #align polynomial.count_map_roots_of_injective Polynomial.count_map_roots_of_injective diff --git a/Mathlib/Data/Polynomial/Splits.lean b/Mathlib/Data/Polynomial/Splits.lean index 0b67adb7ba7f2..d891e37afff94 100644 --- a/Mathlib/Data/Polynomial/Splits.lean +++ b/Mathlib/Data/Polynomial/Splits.lean @@ -21,12 +21,6 @@ irreducible factors over `L` have degree `1`. field and a polynomial `f` saying that `f.map i` is zero or all of its irreducible factors over `L` have degree `1`. -## Main statements - -* `lift_of_splits`: If `K` and `L` are field extensions of a field `F` and for some finite subset - `S` of `K`, the minimal polynomial of every `x ∈ K` splits as a polynomial with coefficients in - `L`, then `Algebra.adjoin F S` embeds into `L`. - -/ @@ -430,6 +424,20 @@ theorem splits_iff_exists_multiset {f : K[X]} : splits_of_exists_multiset i hs⟩ #align polynomial.splits_iff_exists_multiset Polynomial.splits_iff_exists_multiset +theorem splits_of_comp (j : L →+* F) {f : K[X]} (h : Splits (j.comp i) f) + (roots_mem_range : ∀ a ∈ (f.map (j.comp i)).roots, a ∈ j.range) : Splits i f := by + choose lift lift_eq using roots_mem_range + rw [splits_iff_exists_multiset] + refine ⟨(f.map (j.comp i)).roots.pmap lift fun _ ↦ id, map_injective _ j.injective ?_⟩ + conv_lhs => rw [Polynomial.map_map, eq_prod_roots_of_splits h] + simp_rw [Polynomial.map_mul, Polynomial.map_multiset_prod, Multiset.map_pmap, Polynomial.map_sub, + map_C, map_X, lift_eq, Multiset.pmap_eq_map] + rfl + +theorem splits_id_of_splits {f : K[X]} (h : Splits i f) + (roots_mem_range : ∀ a ∈ (f.map i).roots, a ∈ i.range) : Splits (RingHom.id K) f := + splits_of_comp (RingHom.id K) i h roots_mem_range + theorem splits_comp_of_splits (j : L →+* F) {f : K[X]} (h : Splits i f) : Splits (j.comp i) f := by -- Porting note: was -- change i with (RingHom.id _).comp i at h diff --git a/Mathlib/Data/Polynomial/UnitTrinomial.lean b/Mathlib/Data/Polynomial/UnitTrinomial.lean index f5a3dba04d03c..e91b47e0ff1dd 100644 --- a/Mathlib/Data/Polynomial/UnitTrinomial.lean +++ b/Mathlib/Data/Polynomial/UnitTrinomial.lean @@ -203,6 +203,7 @@ theorem isUnitTrinomial_iff' : sum_insert (mt mem_singleton.mp hmn.ne), sum_singleton, trinomial_leading_coeff' hkm hmn, trinomial_middle_coeff hkm hmn, trinomial_trailing_coeff' hkm hmn] simp_rw [← Units.val_pow_eq_pow_val, Int.units_sq] + decide · have key : ∀ k ∈ p.support, p.coeff k ^ 2 = 1 := fun k hk => Int.sq_eq_one_of_sq_le_three ((single_le_sum (fun k _ => sq_nonneg (p.coeff k)) hk).trans hp.le) (mem_support_iff.mp hk) diff --git a/Mathlib/Data/Prod/TProd.lean b/Mathlib/Data/Prod/TProd.lean index b80958ce652d9..3133a6173ee3f 100644 --- a/Mathlib/Data/Prod/TProd.lean +++ b/Mathlib/Data/Prod/TProd.lean @@ -177,7 +177,7 @@ theorem elim_preimage_pi [DecidableEq ι] {l : List ι} (hnd : l.Nodup) (h : ∀ simp [h] rw [← h2, ← mk_preimage_tprod, preimage_preimage] simp only [TProd.mk_elim hnd h] - dsimp; rfl + dsimp #align set.elim_preimage_pi Set.elim_preimage_pi end Set diff --git a/Mathlib/Data/QPF/Univariate/Basic.lean b/Mathlib/Data/QPF/Univariate/Basic.lean index 7b6f91f56d24b..d3bee0e10762a 100644 --- a/Mathlib/Data/QPF/Univariate/Basic.lean +++ b/Mathlib/Data/QPF/Univariate/Basic.lean @@ -13,19 +13,19 @@ import Mathlib.Data.PFunctor.Univariate.M We assume the following: -`P` : a polynomial functor -`W` : its W-type -`M` : its M-type -`F` : a functor +* `P`: a polynomial functor +* `W`: its W-type +* `M`: its M-type +* `F`: a functor We define: -`q` : `QPF` data, representing `F` as a quotient of `P` +* `q`: `QPF` data, representing `F` as a quotient of `P` The main goal is to construct: -`Fix` : the initial algebra with structure map `F Fix → Fix`. -`Cofix` : the final coalgebra with structure map `Cofix → F Cofix` +* `Fix`: the initial algebra with structure map `F Fix → Fix`. +* `Cofix`: the final coalgebra with structure map `Cofix → F Cofix` We also show that the composition of qpfs is a qpf, and that the quotient of a qpf is a qpf. @@ -53,7 +53,7 @@ class QPF (F : Type u → Type u) [Functor F] where abs : ∀ {α}, P α → F α repr : ∀ {α}, F α → P α abs_repr : ∀ {α} (x : F α), abs (repr x) = x - abs_map : ∀ {α β} (f : α → β) (p : P α), abs (f <$> p) = f <$> abs p + abs_map : ∀ {α β} (f : α → β) (p : P α), abs (P.map f p) = f <$> abs p #align qpf QPF namespace QPF @@ -167,14 +167,14 @@ set_option linter.uppercaseLean3 false in #align qpf.recF QPF.recF theorem recF_eq {α : Type _} (g : F α → α) (x : q.P.W) : - recF g x = g (abs (recF g <$> x.dest)) := by + recF g x = g (abs (q.P.map (recF g) x.dest)) := by cases x rfl set_option linter.uppercaseLean3 false in #align qpf.recF_eq QPF.recF_eq theorem recF_eq' {α : Type _} (g : F α → α) (a : q.P.A) (f : q.P.B a → q.P.W) : - recF g ⟨a, f⟩ = g (abs (recF g <$> ⟨a, f⟩)) := + recF g ⟨a, f⟩ = g (abs (q.P.map (recF g) ⟨a, f⟩)) := rfl set_option linter.uppercaseLean3 false in #align qpf.recF_eq' QPF.recF_eq' @@ -231,9 +231,9 @@ set_option linter.uppercaseLean3 false in theorem Wrepr_equiv (x : q.P.W) : Wequiv (Wrepr x) x := by induction' x with a f ih apply Wequiv.trans - · change Wequiv (Wrepr ⟨a, f⟩) (PFunctor.W.mk (Wrepr <$> ⟨a, f⟩)) + · change Wequiv (Wrepr ⟨a, f⟩) (PFunctor.W.mk (q.P.map Wrepr ⟨a, f⟩)) apply Wequiv.abs' - have : Wrepr ⟨a, f⟩ = PFunctor.W.mk (repr (abs (Wrepr <$> ⟨a, f⟩))) := rfl + have : Wrepr ⟨a, f⟩ = PFunctor.W.mk (repr (abs (q.P.map Wrepr ⟨a, f⟩))) := rfl rw [this, PFunctor.W.dest_mk, abs_repr] rfl apply Wequiv.ind; exact ih @@ -267,7 +267,7 @@ set_option linter.uppercaseLean3 false in /-- constructor of a type defined by a qpf -/ def Fix.mk (x : F (Fix F)) : Fix F := - Quot.mk _ (PFunctor.W.mk (fixToW <$> repr x)) + Quot.mk _ (PFunctor.W.mk (q.P.map fixToW (repr x))) #align qpf.fix.mk QPF.Fix.mk /-- destructor of a type defined by a qpf -/ @@ -289,7 +289,7 @@ theorem Fix.rec_eq {α : Type _} (g : F α → α) (x : F (Fix F)) : rw [Fix.rec, Fix.mk] dsimp cases' h : repr x with a f - rw [PFunctor.map_eq, recF_eq, ← PFunctor.map_eq, PFunctor.W.dest_mk, ← PFunctor.comp_map, abs_map, + rw [PFunctor.map_eq, recF_eq, ← PFunctor.map_eq, PFunctor.W.dest_mk, PFunctor.map_map, abs_map, ← h, abs_repr, this] #align qpf.fix.rec_eq QPF.Fix.rec_eq @@ -375,14 +375,16 @@ set_option linter.uppercaseLean3 false in #align qpf.corecF QPF.corecF theorem corecF_eq {α : Type _} (g : α → F α) (x : α) : - PFunctor.M.dest (corecF g x) = corecF g <$> repr (g x) := by rw [corecF, PFunctor.M.dest_corec] + PFunctor.M.dest (corecF g x) = q.P.map (corecF g) (repr (g x)) := by + rw [corecF, PFunctor.M.dest_corec] set_option linter.uppercaseLean3 false in #align qpf.corecF_eq QPF.corecF_eq -- Equivalence /-- A pre-congruence on `q.P.M` *viewed as an F-coalgebra*. Not necessarily symmetric. -/ def IsPrecongr (r : q.P.M → q.P.M → Prop) : Prop := - ∀ ⦃x y⦄, r x y → abs (Quot.mk r <$> PFunctor.M.dest x) = abs (Quot.mk r <$> PFunctor.M.dest y) + ∀ ⦃x y⦄, r x y → + abs (q.P.map (Quot.mk r) (PFunctor.M.dest x)) = abs (q.P.map (Quot.mk r) (PFunctor.M.dest y)) #align qpf.is_precongr QPF.IsPrecongr /-- The maximal congruence on `q.P.M`. -/ @@ -459,9 +461,9 @@ private theorem Cofix.bisim_aux (r : Cofix F → Cofix F → Prop) (h' : ∀ x, intro c d; apply Quot.inductionOn (motive := _) d; clear d intro d rcd; apply Quot.sound; apply rcd) have : f ∘ Quot.mk r ∘ Quot.mk Mcongr = Quot.mk r' := rfl - rw [← this, PFunctor.comp_map _ _ f, PFunctor.comp_map _ _ (Quot.mk r), abs_map, abs_map, + rw [← this, ← PFunctor.map_map _ _ f, ← PFunctor.map_map _ _ (Quot.mk r), abs_map, abs_map, abs_map, h₀] - rw [PFunctor.comp_map _ _ f, PFunctor.comp_map _ _ (Quot.mk r), abs_map, abs_map, abs_map] + rw [← PFunctor.map_map _ _ f, ← PFunctor.map_map _ _ (Quot.mk r), abs_map, abs_map, abs_map] refine' ⟨r', this, rxy⟩ theorem Cofix.bisim_rel (r : Cofix F → Cofix F → Prop) diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index ed3fff7d58693..8469f4b7101fd 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl -/ import Mathlib.Init.Data.Quot import Mathlib.Logic.Relator +import Mathlib.Mathport.Notation #align_import data.quot from "leanprover-community/mathlib"@"6ed6abbde29b8f630001a1b481603f657a3384f1" @@ -39,8 +40,8 @@ namespace Quot variable {ra : α → α → Prop} {rb : β → β → Prop} {φ : Quot ra → Quot rb → Sort*} -- mathport name: mk -@[inherit_doc] -local notation:arg "⟦" a "⟧" => Quot.mk _ a +@[inherit_doc Quot.mk] +local notation3:arg "⟦" a "⟧" => Quot.mk _ a @[elab_as_elim] protected theorem induction_on {α : Sort u} {r : α → α → Prop} {β : Quot r → Prop} (q : Quot r) @@ -220,8 +221,8 @@ variable {φ : Quotient sa → Quotient sb → Sort*} -- We have not yet decided which one works best, since the setoid instance can't always be -- reliably found but it can't always be inferred from the expected type either. -- See also: https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/confusion.20between.20equivalence.20and.20instance.20setoid/near/360822354 -@[inherit_doc] -notation:arg "⟦" a "⟧" => Quotient.mk _ a +@[inherit_doc Quotient.mk] +notation3:arg "⟦" a "⟧" => Quotient.mk _ a instance instInhabitedQuotient (s : Setoid α) [Inhabited α] : Inhabited (Quotient s) := ⟨⟦default⟧⟩ diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index 785026a02536e..10e5c3c238590 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -240,14 +240,16 @@ theorem divInt_zero_one : 0 /. 1 = 0 := theorem divInt_one_one : 1 /. 1 = 1 := show divInt _ _ = _ by rw [divInt] - simp + simp [mkRat, normalize] + rfl #align rat.mk_one_one Rat.divInt_one_one @[simp] theorem divInt_neg_one_one : -1 /. 1 = -1 := show divInt _ _ = _ by rw [divInt] - simp + simp [mkRat, normalize] + rfl #align rat.mk_neg_one_one Rat.divInt_neg_one_one theorem divInt_one (n : ℤ) : n /. 1 = n := diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index 01ab28303b775..1a09e6d8b1a18 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -233,6 +233,7 @@ theorem coe_int_div_self (n : ℤ) : ((n / n : ℤ) : ℚ) = n / n := by by_cases hn : n = 0 · subst hn simp only [Int.cast_zero, Int.zero_div, zero_div] + rfl · have : (n : ℚ) ≠ 0 := by rwa [← coe_int_inj] at hn simp only [Int.ediv_self hn, Int.cast_one, Ne.def, not_false_iff, div_self this] #align rat.coe_int_div_self Rat.coe_int_div_self diff --git a/Mathlib/Data/Rat/Sqrt.lean b/Mathlib/Data/Rat/Sqrt.lean index 15ae8a0bebf85..9e7dea9f265a3 100644 --- a/Mathlib/Data/Rat/Sqrt.lean +++ b/Mathlib/Data/Rat/Sqrt.lean @@ -42,4 +42,9 @@ theorem sqrt_nonneg (q : ℚ) : 0 ≤ Rat.sqrt q := Int.coe_nat_nonneg _ #align rat.sqrt_nonneg Rat.sqrt_nonneg +/-- `IsSquare` can be decided on `ℚ` by checking against the square root. -/ +instance : DecidablePred (IsSquare : ℚ → Prop) := + fun m => decidable_of_iff' (sqrt m * sqrt m = m) <| by + simp_rw [←exists_mul_self m, IsSquare, eq_comm] + end Rat diff --git a/Mathlib/Data/Real/ENNReal.lean b/Mathlib/Data/Real/ENNReal.lean index 80f250f87af2f..185db4b78ccbc 100644 --- a/Mathlib/Data/Real/ENNReal.lean +++ b/Mathlib/Data/Real/ENNReal.lean @@ -343,6 +343,8 @@ theorem toReal_ofReal_eq_iff {a : ℝ} : (ENNReal.ofReal a).toReal = a ↔ 0 ≤ @[simp] theorem top_ne_one : ∞ ≠ 1 := top_ne_coe #align ennreal.top_ne_one ENNReal.top_ne_one +@[simp] theorem zero_lt_top : 0 < ∞ := coe_lt_top + @[simp, norm_cast] theorem coe_eq_coe : (↑r : ℝ≥0∞) = ↑q ↔ r = q := WithTop.coe_eq_coe #align ennreal.coe_eq_coe ENNReal.coe_eq_coe @@ -428,9 +430,11 @@ nonrec theorem one_lt_two : (1 : ℝ≥0∞) < 2 := coe_one ▸ coe_two ▸ by exact_mod_cast (one_lt_two : 1 < 2) #align ennreal.one_lt_two ENNReal.one_lt_two -theorem two_ne_top : (2 : ℝ≥0∞) ≠ ∞ := coe_ne_top +@[simp] theorem two_ne_top : (2 : ℝ≥0∞) ≠ ∞ := coe_ne_top #align ennreal.two_ne_top ENNReal.two_ne_top +@[simp] theorem two_lt_top : (2 : ℝ≥0∞) < ∞ := coe_lt_top + /-- `(1 : ℝ≥0∞) ≤ 1`, recorded as a `Fact` for use with `Lp` spaces. -/ instance _root_.fact_one_le_one_ennreal : Fact ((1 : ℝ≥0∞) ≤ 1) := ⟨le_rfl⟩ @@ -635,7 +639,7 @@ theorem mul_lt_top_iff {a b : ℝ≥0∞} : a * b < ∞ ↔ a < ∞ ∧ b < ∞ theorem mul_self_lt_top_iff {a : ℝ≥0∞} : a * a < ⊤ ↔ a < ⊤ := by rw [ENNReal.mul_lt_top_iff, and_self, or_self, or_iff_left_iff_imp] rintro rfl - norm_num + decide #align ennreal.mul_self_lt_top_iff ENNReal.mul_self_lt_top_iff theorem mul_pos_iff : 0 < a * b ↔ 0 < a ∧ 0 < b := @@ -651,7 +655,7 @@ theorem mul_pos (ha : a ≠ 0) (hb : b ≠ 0) : 0 < a * b := rcases n.eq_zero_or_pos with rfl | (hn : 0 < n) · simp · induction a using recTopCoe - · simp only [Ne.def, hn.ne', top_pow hn] + · simp only [Ne.def, hn.ne', top_pow hn, not_false_eq_true, and_self] · simp only [← coe_pow, coe_ne_top, false_and] #align ennreal.pow_eq_top_iff ENNReal.pow_eq_top_iff @@ -1793,7 +1797,8 @@ theorem add_thirds (a : ℝ≥0∞) : a / 3 + a / 3 + a / 3 = a := by @[simp] theorem div_pos_iff : 0 < a / b ↔ a ≠ 0 ∧ b ≠ ∞ := by simp [pos_iff_ne_zero, not_or] #align ennreal.div_pos_iff ENNReal.div_pos_iff -protected theorem half_pos (h : a ≠ 0) : 0 < a / 2 := by simp [h] +protected theorem half_pos (h : a ≠ 0) : 0 < a / 2 := by + simp only [div_pos_iff, ne_eq, h, not_false_eq_true]; decide #align ennreal.half_pos ENNReal.half_pos protected theorem one_half_lt_one : (2⁻¹ : ℝ≥0∞) < 1 := @@ -1914,7 +1919,8 @@ theorem coe_zpow (hr : r ≠ 0) (n : ℤ) : (↑(r ^ n) : ℝ≥0∞) = (r : ℝ theorem zpow_pos (ha : a ≠ 0) (h'a : a ≠ ∞) (n : ℤ) : 0 < a ^ n := by cases n · exact ENNReal.pow_pos ha.bot_lt _ - · simp only [h'a, pow_eq_top_iff, zpow_negSucc, Ne.def, not_false, ENNReal.inv_pos, false_and] + · simp only [h'a, pow_eq_top_iff, zpow_negSucc, Ne.def, not_false, ENNReal.inv_pos, false_and, + not_false_eq_true] #align ennreal.zpow_pos ENNReal.zpow_pos theorem zpow_lt_top (ha : a ≠ 0) (h'a : a ≠ ∞) (n : ℤ) : a ^ n < ∞ := by diff --git a/Mathlib/Data/Real/NNReal.lean b/Mathlib/Data/Real/NNReal.lean index afe223186f2c5..a7fa0f834f145 100644 --- a/Mathlib/Data/Real/NNReal.lean +++ b/Mathlib/Data/Real/NNReal.lean @@ -638,11 +638,53 @@ theorem toNNReal_of_nonpos {r : ℝ} : r ≤ 0 → Real.toNNReal r = 0 := toNNReal_eq_zero.2 #align real.to_nnreal_of_nonpos Real.toNNReal_of_nonpos +lemma toNNReal_eq_iff_eq_coe {r : ℝ} {p : ℝ≥0} (hp : p ≠ 0) : r.toNNReal = p ↔ r = p := + ⟨fun h ↦ h ▸ (coe_toNNReal _ <| not_lt.1 fun hlt ↦ hp <| h ▸ toNNReal_of_nonpos hlt.le).symm, + fun h ↦ h.symm ▸ toNNReal_coe⟩ + +@[simp] +lemma toNNReal_eq_one {r : ℝ} : r.toNNReal = 1 ↔ r = 1 := toNNReal_eq_iff_eq_coe one_ne_zero + +@[simp] +lemma toNNReal_eq_nat_cast {r : ℝ} {n : ℕ} (hn : n ≠ 0) : r.toNNReal = n ↔ r = n := by + exact_mod_cast toNNReal_eq_iff_eq_coe <| Nat.cast_ne_zero.2 hn + +@[simp] +lemma toNNReal_eq_ofNat {r : ℝ} {n : ℕ} [h : n.AtLeastTwo] : + r.toNNReal = no_index (OfNat.ofNat n) ↔ r = OfNat.ofNat n := + toNNReal_eq_nat_cast h.ne_zero + @[simp] theorem toNNReal_le_toNNReal_iff {r p : ℝ} (hp : 0 ≤ p) : toNNReal r ≤ toNNReal p ↔ r ≤ p := by simp [← NNReal.coe_le_coe, hp] #align real.to_nnreal_le_to_nnreal_iff Real.toNNReal_le_toNNReal_iff +@[simp] +lemma toNNReal_le_one {r : ℝ} : r.toNNReal ≤ 1 ↔ r ≤ 1 := by + simpa using toNNReal_le_toNNReal_iff zero_le_one + +@[simp] +lemma one_lt_toNNReal {r : ℝ} : 1 < r.toNNReal ↔ 1 < r := by + simpa only [not_le] using toNNReal_le_one.not + +@[simp] +lemma toNNReal_le_nat_cast {r : ℝ} {n : ℕ} : r.toNNReal ≤ n ↔ r ≤ n := by + simpa using toNNReal_le_toNNReal_iff n.cast_nonneg + +@[simp] +lemma nat_cast_lt_toNNReal {r : ℝ} {n : ℕ} : n < r.toNNReal ↔ n < r := by + simpa only [not_le] using toNNReal_le_nat_cast.not + +@[simp] +lemma toNNReal_le_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : + r.toNNReal ≤ no_index (OfNat.ofNat n) ↔ r ≤ n := + toNNReal_le_nat_cast + +@[simp] +lemma ofNat_lt_toNNReal {r : ℝ} {n : ℕ} [n.AtLeastTwo] : + no_index (OfNat.ofNat n) < r.toNNReal ↔ n < r := + nat_cast_lt_toNNReal + @[simp] theorem toNNReal_eq_toNNReal_iff {r p : ℝ} (hr : 0 ≤ r) (hp : 0 ≤ p) : toNNReal r = toNNReal p ↔ r = p := by simp [← NNReal.coe_eq, coe_toNNReal, hr, hp] @@ -663,6 +705,41 @@ theorem toNNReal_lt_toNNReal_iff_of_nonneg {r p : ℝ} (hr : 0 ≤ r) : toNNReal_lt_toNNReal_iff'.trans ⟨And.left, fun h => ⟨h, lt_of_le_of_lt hr h⟩⟩ #align real.to_nnreal_lt_to_nnreal_iff_of_nonneg Real.toNNReal_lt_toNNReal_iff_of_nonneg +lemma toNNReal_le_toNNReal_iff' {r p : ℝ} : r.toNNReal ≤ p.toNNReal ↔ r ≤ p ∨ r ≤ 0 := by + simp_rw [← not_lt, toNNReal_lt_toNNReal_iff', not_and_or] + +lemma toNNReal_le_toNNReal_iff_of_pos {r p : ℝ} (hr : 0 < r) : r.toNNReal ≤ p.toNNReal ↔ r ≤ p := by + simp [toNNReal_le_toNNReal_iff', hr.not_le] + +@[simp] +lemma one_le_toNNReal {r : ℝ} : 1 ≤ r.toNNReal ↔ 1 ≤ r := by + simpa using toNNReal_le_toNNReal_iff_of_pos one_pos + +@[simp] +lemma toNNReal_lt_one {r : ℝ} : r.toNNReal < 1 ↔ r < 1 := by simp only [← not_le, one_le_toNNReal] + +@[simp] +lemma nat_cast_le_toNNReal' {n : ℕ} {r : ℝ} : ↑n ≤ r.toNNReal ↔ n ≤ r ∨ n = 0 := by + simpa [n.cast_nonneg.le_iff_eq] using toNNReal_le_toNNReal_iff' (r := n) + +@[simp] +lemma toNNReal_lt_nat_cast' {n : ℕ} {r : ℝ} : r.toNNReal < n ↔ r < n ∧ n ≠ 0 := by + simpa [pos_iff_ne_zero] using toNNReal_lt_toNNReal_iff' (r := r) (p := n) + +lemma nat_cast_le_toNNReal {n : ℕ} {r : ℝ} (hn : n ≠ 0) : ↑n ≤ r.toNNReal ↔ n ≤ r := by simp [hn] + +lemma toNNReal_lt_nat_cast {r : ℝ} {n : ℕ} (hn : n ≠ 0) : r.toNNReal < n ↔ r < n := by simp [hn] + +@[simp] +lemma toNNReal_lt_ofNat {r : ℝ} {n : ℕ} [h : n.AtLeastTwo] : + r.toNNReal < no_index (OfNat.ofNat n) ↔ r < OfNat.ofNat n := + toNNReal_lt_nat_cast h.ne_zero + +@[simp] +lemma ofNat_le_toNNReal {n : ℕ} {r : ℝ} [h : n.AtLeastTwo] : + no_index (OfNat.ofNat n) ≤ r.toNNReal ↔ OfNat.ofNat n ≤ r := + nat_cast_le_toNNReal h.ne_zero + @[simp] theorem toNNReal_add {r p : ℝ} (hr : 0 ≤ r) (hp : 0 ≤ p) : Real.toNNReal (r + p) = Real.toNNReal r + Real.toNNReal p := diff --git a/Mathlib/Data/Real/Pi/Wallis.lean b/Mathlib/Data/Real/Pi/Wallis.lean index 72c1ced1777b3..89b3317cb4aa1 100644 --- a/Mathlib/Data/Real/Pi/Wallis.lean +++ b/Mathlib/Data/Real/Pi/Wallis.lean @@ -95,8 +95,8 @@ theorem le_W (k : ℕ) : ((2 : ℝ) * k + 1) / (2 * k + 2) * (π / 2) ≤ W k := rw [W_eq_integral_sin_pow_div_integral_sin_pow, le_div_iff (integral_sin_pow_pos _)] convert integral_sin_pow_succ_le (2 * k + 1) rw [integral_sin_pow (2 * k)] - simp only [sin_zero, zero_pow', Ne.def, Nat.succ_ne_zero, zero_mul, sin_pi, tsub_zero, zero_div, - zero_add] + simp only [sin_zero, ne_eq, add_eq_zero, and_false, not_false_eq_true, zero_pow', cos_zero, + mul_one, sin_pi, cos_pi, mul_neg, neg_zero, sub_self, zero_div, zero_add] norm_cast #align real.wallis.le_W Real.Wallis.le_W diff --git a/Mathlib/Data/Real/Sqrt.lean b/Mathlib/Data/Real/Sqrt.lean index 970b33638e32f..ca73aab94a253 100644 --- a/Mathlib/Data/Real/Sqrt.lean +++ b/Mathlib/Data/Real/Sqrt.lean @@ -380,7 +380,7 @@ def evalSqrt : PositivityExt where eval {_ _} _zα _pα e := do let (.app _ (a : Q(Real))) ← whnfR e | throwError "not Real.sqrt" let zα' ← synthInstanceQ (q(Zero Real) : Q(Type)) let pα' ← synthInstanceQ (q(PartialOrder Real) : Q(Type)) - let ra ← core zα' pα' a + let ra ← catchNone <| core zα' pα' a assertInstancesCommute match ra with | .positive pa => pure (.positive (q(Real.sqrt_pos_of_pos $pa) : Expr)) diff --git a/Mathlib/Data/Rel.lean b/Mathlib/Data/Rel.lean index 52a5cbf219f14..347800e2c5121 100644 --- a/Mathlib/Data/Rel.lean +++ b/Mathlib/Data/Rel.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad -/ import Mathlib.Order.CompleteLattice import Mathlib.Order.GaloisConnection +import Mathlib.Order.Hom.CompleteLattice #align_import data.rel from "leanprover-community/mathlib"@"706d88f2b8fdfeb0b22796433d7a6c1a010af9f2" @@ -116,6 +117,26 @@ theorem comp_left_id (r : Rel α β) : @Eq α • r = r := by simp #align rel.comp_left_id Rel.comp_left_id +@[simp] +theorem comp_right_bot (r : Rel α β) : r • (⊥ : Rel β γ) = ⊥ := by + ext x y + simp [comp, Bot.bot] + +@[simp] +theorem comp_left_bot (r : Rel α β) : (⊥ : Rel γ α) • r = ⊥ := by + ext x y + simp [comp, Bot.bot] + +@[simp] +theorem comp_right_top (r : Rel α β) : r • (⊤ : Rel β γ) = λ x _ ↦ x ∈ r.dom := by + ext x z + simp [comp, Top.top, dom] + +@[simp] +theorem comp_left_top (r : Rel α β) : (⊤ : Rel γ α) • r = λ _ y ↦ y ∈ r.codom := by + ext x z + simp [comp, Top.top, codom] + theorem inv_id : inv (@Eq α) = @Eq α := by ext x y constructor <;> apply Eq.symm @@ -126,6 +147,12 @@ theorem inv_comp (r : Rel α β) (s : Rel β γ) : inv (r • s) = inv s • inv simp [comp, inv, flip, and_comm] #align rel.inv_comp Rel.inv_comp +@[simp] +theorem inv_bot : (⊥ : Rel α β).inv = (⊥ : Rel β α) := by simp [Bot.bot, inv, flip] + +@[simp] +theorem inv_top : (⊤ : Rel α β).inv = (⊤ : Rel β α) := by simp [Top.top, inv, flip] + /-- Image of a set under a relation -/ def image (s : Set α) : Set β := { y | ∃ x ∈ s, r x y } #align rel.image Rel.image @@ -170,6 +197,22 @@ theorem image_univ : r.image Set.univ = r.codom := by simp [mem_image, codom] #align rel.image_univ Rel.image_univ +@[simp] +theorem image_empty : r.image ∅ = ∅ := by + ext x + simp [mem_image] + +@[simp] +theorem image_bot (s : Set α) : (⊥ : Rel α β).image s = ∅ := by + rw [Set.eq_empty_iff_forall_not_mem] + intro x h + simp [mem_image, Bot.bot] at h + +@[simp] +theorem image_top {s : Set α} (h : Set.Nonempty s) : + (⊤ : Rel α β).image s = Set.univ := + Set.eq_univ_of_forall λ x ↦ ⟨h.some, by simp [h.some_mem, Top.top]⟩ + /-- Preimage of a set under a relation `r`. Same as the image of `s` under `r.inv` -/ def preimage (s : Set β) : Set α := r.inv.image s @@ -206,6 +249,60 @@ theorem preimage_comp (s : Rel β γ) (t : Set γ) : preimage (r • s) t = prei theorem preimage_univ : r.preimage Set.univ = r.dom := by rw [preimage, image_univ, codom_inv] #align rel.preimage_univ Rel.preimage_univ +@[simp] +theorem preimage_empty : r.preimage ∅ = ∅ := by rw [preimage, image_empty] + +@[simp] +theorem preimage_inv (s : Set α) : r.inv.preimage s = r.image s := by rw [preimage, inv_inv] + +@[simp] +theorem preimage_bot (s : Set β) : (⊥ : Rel α β).preimage s = ∅ := + by rw [preimage, inv_bot, image_bot] + +@[simp] +theorem preimage_top {s : Set β} (h : Set.Nonempty s) : + (⊤ : Rel α β).preimage s = Set.univ := by rwa [← inv_top, preimage, inv_inv, image_top] + +theorem image_eq_dom_of_codomain_subset {s : Set β} (h : r.codom ⊆ s) : r.preimage s = r.dom := by + rw [← preimage_univ] + apply Set.eq_of_subset_of_subset + · exact image_subset _ (Set.subset_univ _) + · intro x hx + simp only [mem_preimage, Set.mem_univ, true_and] at hx + rcases hx with ⟨y, ryx⟩ + have hy : y ∈ s := h ⟨x, ryx⟩ + exact ⟨y, ⟨hy, ryx⟩⟩ + +theorem preimage_eq_codom_of_domain_subset {s : Set α} (h : r.dom ⊆ s) : r.image s = r.codom := + by apply r.inv.image_eq_dom_of_codomain_subset (by rwa [← codom_inv] at h) + +theorem image_inter_dom_eq (s : Set α) : r.image (s ∩ r.dom) = r.image s := by + apply Set.eq_of_subset_of_subset + · apply r.image_mono (by simp) + · intro x h + rw [mem_image] at * + rcases h with ⟨y, hy, ryx⟩ + use y + suffices h : y ∈ r.dom by simp_all only [Set.mem_inter_iff, and_self] + rw [dom, Set.mem_setOf_eq] + use x + +@[simp] +theorem preimage_inter_codom_eq (s : Set β) : r.preimage (s ∩ r.codom) = r.preimage s := by + rw[←dom_inv, preimage, preimage, image_inter_dom_eq] + +theorem inter_dom_subset_preimage_image (s : Set α) : s ∩ r.dom ⊆ r.preimage (r.image s) := by + intro x hx + simp only [Set.mem_inter_iff, dom] at hx + rcases hx with ⟨hx, ⟨y, rxy⟩⟩ + use y + simp only [image, Set.mem_setOf_eq] + exact ⟨⟨x, hx, rxy⟩, rxy⟩ + +theorem image_preimage_subset_inter_codom (s : Set β) : s ∩ r.codom ⊆ r.image (r.preimage s) := by + rw [← dom_inv, ← preimage_inv] + apply inter_dom_subset_preimage_image + /-- Core of a set `s : Set β` w.r.t `r : Rel α β` is the set of `x : α` that are related *only* to elements of `s`. Other generalization of `Function.preimage`. -/ def core (s : Set β) := { x | ∀ y, r x y → y ∈ s } @@ -264,22 +361,52 @@ namespace Function def graph (f : α → β) : Rel α β := fun x y => f x = y #align function.graph Function.graph +@[simp] lemma graph_def (f : α → β) (x y) : f.graph x y ↔ (f x = y) := Iff.rfl + +theorem graph_id : graph id = @Eq α := by simp [graph] + +theorem graph_comp {f : β → γ} {g : α → β} : graph (f ∘ g) = Rel.comp (graph g) (graph f) := by + ext x y + simp [Rel.comp] + end Function +theorem Equiv.graph_inv (f : α ≃ β) : (f.symm : β → α).graph = Rel.inv (f : α → β).graph := by + ext x y + aesop (add norm Rel.inv_def) + +theorem Relation.is_graph_iff (r : Rel α β) : (∃! f, Function.graph f = r) ↔ ∀ x, ∃! y, r x y := by + unfold Function.graph + constructor + · rintro ⟨f, rfl, _⟩ x + use f x + simp only [forall_eq', and_self] + · intro h + rcases Classical.axiomOfChoice (λ x ↦ (h x).exists) with ⟨f,hf⟩ + use f + constructor + · ext x _ + constructor + · rintro rfl + exact hf x + · exact (h x).unique (hf x) + · rintro _ rfl + exact funext hf + namespace Set -- TODO: if image were defined with bounded quantification in corelib, the next two would -- be definitional theorem image_eq (f : α → β) (s : Set α) : f '' s = (Function.graph f).image s := by - simp [Set.image, Function.graph, Rel.image] + simp [Set.image, Rel.image] #align set.image_eq Set.image_eq theorem preimage_eq (f : α → β) (s : Set β) : f ⁻¹' s = (Function.graph f).preimage s := by - simp [Set.preimage, Function.graph, Rel.preimage, Rel.inv, flip, Rel.image] + simp [Set.preimage, Rel.preimage, Rel.inv, flip, Rel.image] #align set.preimage_eq Set.preimage_eq theorem preimage_eq_core (f : α → β) (s : Set β) : f ⁻¹' s = (Function.graph f).core s := by - simp [Set.preimage, Function.graph, Rel.core] + simp [Set.preimage, Rel.core] #align set.preimage_eq_core Set.preimage_eq_core end Set diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index df02455f987a3..63c7cd2efb827 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Order.SymmDiff import Mathlib.Logic.Function.Iterate import Mathlib.Tactic.Tauto import Mathlib.Tactic.ByContra +import Mathlib.Util.Delaborators #align_import data.set.basic from "leanprover-community/mathlib"@"001ffdc42920050657fd45bd2b8bfbec8eaaeb29" @@ -432,14 +433,6 @@ theorem not_not_mem : ¬a ∉ s ↔ a ∈ s := /-! ### Non-empty sets -/ - -/-- The property `s.Nonempty` expresses the fact that the set `s` is not empty. It should be used -in theorem assumptions instead of `∃ x, x ∈ s` or `s ≠ ∅` as it gives access to a nice API thanks -to the dot notation. -/ -protected def Nonempty (s : Set α) : Prop := - ∃ x, x ∈ s -#align set.nonempty Set.Nonempty - -- Porting note: we seem to need parentheses at `(↥s)`, -- even if we increase the right precedence of `↥` in `Mathlib.Tactic.Coe`. -- Porting note: removed `simp` as it is competing with `nonempty_subtype`. @@ -770,7 +763,7 @@ theorem mem_union (x : α) (a b : Set α) : x ∈ a ∪ b ↔ x ∈ a ∨ x ∈ @[simp] theorem union_self (a : Set α) : a ∪ a = a := - ext fun _ => or_self_iff _ + ext fun _ => or_self_iff #align set.union_self Set.union_self @[simp] @@ -924,7 +917,7 @@ theorem mem_of_mem_inter_right {x : α} {a b : Set α} (h : x ∈ a ∩ b) : x @[simp] theorem inter_self (a : Set α) : a ∩ a = a := - ext fun _ => and_self_iff _ + ext fun _ => and_self_iff #align set.inter_self Set.inter_self @[simp] @@ -2345,6 +2338,16 @@ theorem subset_ite {t s s' u : Set α} : u ⊆ t.ite s s' ↔ u ∩ t ⊆ s ∧ by_cases hx : x ∈ t <;> simp [*, Set.ite] #align set.subset_ite Set.subset_ite +theorem ite_eq_of_subset_left (t : Set α) {s₁ s₂ : Set α} (h : s₁ ⊆ s₂) : + t.ite s₁ s₂ = s₁ ∪ (s₂ \ t) := by + ext x + by_cases hx : x ∈ t <;> simp [*, Set.ite, or_iff_right_of_imp (@h x)] + +theorem ite_eq_of_subset_right (t : Set α) {s₁ s₂ : Set α} (h : s₂ ⊆ s₁) : + t.ite s₁ s₂ = (s₁ ∩ t) ∪ s₂ := by + ext x + by_cases hx : x ∈ t <;> simp [*, Set.ite, or_iff_left_of_imp (@h x)] + /-! ### Subsingleton -/ diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index b887fc5e48041..54be95b67b4ee 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -63,7 +63,7 @@ namespace Set variable {s t : Set α} /-- The cardinality of a set as a term in `ℕ∞` -/ -noncomputable def encard (s : Set α) := PartENat.withTopEquiv (PartENat.card s) +noncomputable def encard (s : Set α) : ℕ∞ := PartENat.withTopEquiv (PartENat.card s) @[simp] theorem encard_univ_coe (s : Set α) : encard (univ : Set s) = encard s := by rw [encard, encard, PartENat.card_congr (Equiv.Set.univ ↑s)] @@ -119,7 +119,7 @@ theorem encard_insert_of_not_mem (has : a ∉ s) : (insert a s).encard = s.encar rw [←union_singleton, encard_union_eq (by simpa), encard_singleton] theorem Finite.encard_lt_top (h : s.Finite) : s.encard < ⊤ := by - refine' h.induction_on (by simpa using WithTop.zero_lt_top) _ + refine' h.induction_on (by simp) _ rintro a t hat _ ht' rw [encard_insert_of_not_mem hat] exact lt_tsub_iff_right.1 ht' @@ -468,8 +468,7 @@ macro_rules /-- The cardinality of `s : Set α` . Has the junk value `0` if `s` is infinite -/ -noncomputable def ncard (s : Set α) := - ENat.toNat s.encard +noncomputable def ncard (s : Set α) : ℕ := ENat.toNat s.encard #align set.ncard Set.ncard theorem ncard_def (s : Set α) : s.ncard = ENat.toNat s.encard := rfl @@ -556,7 +555,7 @@ theorem nonempty_of_ncard_ne_zero (hs : s.ncard ≠ 0) : s.Nonempty := by #align set.nonempty_of_ncard_ne_zero Set.nonempty_of_ncard_ne_zero @[simp] theorem ncard_singleton (a : α) : ({a} : Set α).ncard = 1 := by - simp [ncard_eq_toFinset_card] + simp [ncard, ncard_eq_toFinset_card] #align set.ncard_singleton Set.ncard_singleton theorem ncard_singleton_inter (a : α) (s : Set α) : ({a} ∩ s).ncard ≤ 1 := by diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 1627d6772fbd2..4523a6a50fcd7 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -876,6 +876,17 @@ theorem surjOn_iff_surjective : SurjOn f s univ ↔ Surjective (s.restrict f) := ⟨a, as, e⟩⟩ #align set.surj_on_iff_surjective Set.surjOn_iff_surjective +@[simp] +theorem MapsTo.restrict_surjective_iff (h : MapsTo f s t) : + Surjective (MapsTo.restrict _ _ _ h) ↔ SurjOn f s t := by + refine ⟨fun h' b hb ↦ ?_, fun h' ⟨b, hb⟩ ↦ ?_⟩ + · obtain ⟨⟨a, ha⟩, ha'⟩ := h' ⟨b, hb⟩ + replace ha' : f a = b := by simpa [Subtype.ext_iff] using ha' + rw [← ha'] + exact mem_image_of_mem f ha + · obtain ⟨a, ha, rfl⟩ := h' hb + exact ⟨⟨a, ha⟩, rfl⟩ + theorem SurjOn.image_eq_of_mapsTo (h₁ : SurjOn f s t) (h₂ : MapsTo f s t) : f '' s = t := eq_of_subset_of_subset h₂.image_subset h₁ #align set.surj_on.image_eq_of_maps_to Set.SurjOn.image_eq_of_mapsTo diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 02fe36d71f5a2..c7d8f114341f3 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -170,7 +170,7 @@ theorem preimage_comp_eq : preimage (g ∘ f) = preimage f ∘ preimage g := theorem preimage_iterate_eq {f : α → α} {n : ℕ} : Set.preimage f^[n] = (Set.preimage f)^[n] := by induction' n with n ih; · simp - rw [iterate_succ, iterate_succ', Set.preimage_comp_eq, ih] + rw [iterate_succ, iterate_succ', preimage_comp_eq, ih] #align set.preimage_iterate_eq Set.preimage_iterate_eq theorem preimage_preimage {g : β → γ} {f : α → β} {s : Set γ} : @@ -197,7 +197,7 @@ theorem nonempty_of_nonempty_preimage {s : Set β} {f : α → β} (hf : (f ⁻ @[simp] theorem preimage_singleton_false (p : α → Prop) : p ⁻¹' {False} = {a | ¬p a} := by ext; simp #align set.preimage_singleton_false Set.preimage_singleton_false -theorem preimage_subtype_coe_eq_compl {α : Type*} {s u v : Set α} (hsuv : s ⊆ u ∪ v) +theorem preimage_subtype_coe_eq_compl {s u v : Set α} (hsuv : s ⊆ u ∪ v) (H : s ∩ (u ∩ v) = ∅) : ((↑) : s → α) ⁻¹' u = ((↑) ⁻¹' v)ᶜ := by ext ⟨x, x_in_s⟩ constructor @@ -734,7 +734,7 @@ theorem _root_.Nat.mem_range_succ (i : ℕ) : i ∈ range Nat.succ ↔ 0 < i := exact Nat.succ_pos n, fun h => ⟨_, Nat.succ_pred_eq_of_pos h⟩⟩ #align nat.mem_range_succ Nat.mem_range_succ -theorem Nonempty.preimage' {s : Set β} (hs : s.Nonempty) {f : α → β} (hf : s ⊆ Set.range f) : +theorem Nonempty.preimage' {s : Set β} (hs : s.Nonempty) {f : α → β} (hf : s ⊆ range f) : (f ⁻¹' s).Nonempty := let ⟨_, hy⟩ := hs let ⟨x, hx⟩ := hf hy @@ -889,7 +889,7 @@ theorem _root_.Prod.range_snd [Nonempty α] : range (Prod.snd : α × β → β) #align prod.range_snd Prod.range_snd @[simp] -theorem range_eval {ι : Type*} {α : ι → Sort _} [∀ i, Nonempty (α i)] (i : ι) : +theorem range_eval {α : ι → Sort _} [∀ i, Nonempty (α i)] (i : ι) : range (eval i : (∀ i, α i) → α i) = univ := (surjective_eval i).range_eq #align set.range_eval Set.range_eval @@ -1239,14 +1239,14 @@ theorem Subsingleton.preimage {s : Set β} (hs : s.Subsingleton) {f : α → β} #align set.subsingleton.preimage Set.Subsingleton.preimage /-- If the image of a set under an injective map is a subsingleton, the set is a subsingleton. -/ -theorem subsingleton_of_image {α β : Type*} {f : α → β} (hf : Function.Injective f) (s : Set α) +theorem subsingleton_of_image {f : α → β} (hf : Function.Injective f) (s : Set α) (hs : (f '' s).Subsingleton) : s.Subsingleton := (hs.preimage hf).anti <| subset_preimage_image _ _ #align set.subsingleton_of_image Set.subsingleton_of_image /-- If the preimage of a set under a surjective map is a subsingleton, the set is a subsingleton. -/ -theorem subsingleton_of_preimage {α β : Type*} {f : α → β} (hf : Function.Surjective f) (s : Set β) +theorem subsingleton_of_preimage {f : α → β} (hf : Function.Surjective f) (s : Set β) (hs : (f ⁻¹' s).Subsingleton) : s.Subsingleton := fun fx hx fy hy => by rcases hf fx, hf fy with ⟨⟨x, rfl⟩, ⟨y, rfl⟩⟩ exact congr_arg f (hs hx hy) @@ -1289,7 +1289,7 @@ end Set namespace Function -variable {ι : Sort*} {α : Type*} {β : Type*} {f : α → β} +variable {ι : Sort*} {f : α → β} open Set @@ -1402,13 +1402,13 @@ theorem coe_image {p : α → Prop} {s : Set (Subtype p)} : @[simp] theorem coe_image_of_subset {s t : Set α} (h : t ⊆ s) : (↑) '' { x : ↥s | ↑x ∈ t } = t := by ext x - rw [Set.mem_image] + rw [mem_image] exact ⟨fun ⟨_, hx', hx⟩ => hx ▸ hx', fun hx => ⟨⟨x, h hx⟩, hx, rfl⟩⟩ #align subtype.coe_image_of_subset Subtype.coe_image_of_subset theorem range_coe {s : Set α} : range ((↑) : s → α) = s := by - rw [← Set.image_univ] - simp [-Set.image_univ, coe_image] + rw [← image_univ] + simp [-image_univ, coe_image] #align subtype.range_coe Subtype.range_coe /-- A variant of `range_coe`. Try to use `range_coe` if possible. diff --git a/Mathlib/Data/Set/Intervals/Image.lean b/Mathlib/Data/Set/Intervals/Image.lean index 1e1ca90b1cb4c..6c97bec00b4d0 100644 --- a/Mathlib/Data/Set/Intervals/Image.lean +++ b/Mathlib/Data/Set/Intervals/Image.lean @@ -246,6 +246,15 @@ namespace Set variable [Preorder α] {p : α → Prop} +@[simp] lemma preimage_subtype_val_Ici (a : {x // p x}) : (↑) ⁻¹' (Ici a.1) = Ici a := rfl +@[simp] lemma preimage_subtype_val_Iic (a : {x // p x}) : (↑) ⁻¹' (Iic a.1) = Iic a := rfl +@[simp] lemma preimage_subtype_val_Ioi (a : {x // p x}) : (↑) ⁻¹' (Ioi a.1) = Ioi a := rfl +@[simp] lemma preimage_subtype_val_Iio (a : {x // p x}) : (↑) ⁻¹' (Iio a.1) = Iio a := rfl +@[simp] lemma preimage_subtype_val_Icc (a b : {x // p x}) : (↑) ⁻¹' (Icc a.1 b) = Icc a b := rfl +@[simp] lemma preimage_subtype_val_Ico (a b : {x // p x}) : (↑) ⁻¹' (Ico a.1 b) = Ico a b := rfl +@[simp] lemma preimage_subtype_val_Ioc (a b : {x // p x}) : (↑) ⁻¹' (Ioc a.1 b) = Ioc a b := rfl +@[simp] lemma preimage_subtype_val_Ioo (a b : {x // p x}) : (↑) ⁻¹' (Ioo a.1 b) = Ioo a b := rfl + theorem image_subtype_val_Icc_subset (a b : {x // p x}) : Subtype.val '' Icc a b ⊆ Icc a.val b.val := image_subset_iff.mpr fun _ m => m @@ -278,4 +287,144 @@ theorem image_subtype_val_Ioi_subset (a : {x // p x}) : Subtype.val '' Ioi a ⊆ Ioi a.val := image_subset_iff.mpr fun _ m => m +@[simp] +lemma image_subtype_val_Ici_Iic {a : α} (b : Ici a) : Subtype.val '' Iic b = Icc a b := + (Subtype.image_preimage_val (Ici a) (Iic b.1)).trans Iic_inter_Ici + +@[simp] +lemma image_subtype_val_Ici_Iio {a : α} (b : Ici a) : Subtype.val '' Iio b = Ico a b := + (Subtype.image_preimage_val (Ici a) (Iio b.1)).trans Iio_inter_Ici + +@[simp] +lemma image_subtype_val_Ici_Ici {a : α} (b : Ici a) : Subtype.val '' Ici b = Ici b.1 := + (Subtype.image_preimage_val (Ici a) (Ici b.1)).trans <| inter_eq_left.2 <| Ici_subset_Ici.2 b.2 + +@[simp] +lemma image_subtype_val_Ici_Ioi {a : α} (b : Ici a) : Subtype.val '' Ioi b = Ioi b.1 := + (Subtype.image_preimage_val (Ici a) (Ioi b.1)).trans <| inter_eq_left.2 <| Ioi_subset_Ici b.2 + +@[simp] +lemma image_subtype_val_Iic_Ici {a : α} (b : Iic a) : Subtype.val '' Ici b = Icc b.1 a := + Subtype.image_preimage_val (Iic a) (Ici b.1) + +@[simp] +lemma image_subtype_val_Iic_Ioi {a : α} (b : Iic a) : Subtype.val '' Ioi b = Ioc b.1 a := + Subtype.image_preimage_val (Iic a) (Ioi b.1) + +@[simp] +lemma image_subtype_val_Iic_Iic {a : α} (b : Iic a) : Subtype.val '' Iic b = Iic b.1 := + image_subtype_val_Ici_Ici (α := αᵒᵈ) _ + +@[simp] +lemma image_subtype_val_Iic_Iio {a : α} (b : Iic a) : Subtype.val '' Iio b = Iio b.1 := + image_subtype_val_Ici_Ioi (α := αᵒᵈ) _ + +@[simp] +lemma image_subtype_val_Ioi_Ici {a : α} (b : Ioi a) : Subtype.val '' Ici b = Ici b.1 := + (Subtype.image_preimage_val (Ioi a) (Ici b.1)).trans <| inter_eq_left.2 <| Ici_subset_Ioi.2 b.2 + +@[simp] +lemma image_subtype_val_Ioi_Iic {a : α} (b : Ioi a) : Subtype.val '' Iic b = Ioc a b := + (Subtype.image_preimage_val (Ioi a) (Iic b.1)).trans Iic_inter_Ioi + +@[simp] +lemma image_subtype_val_Ioi_Ioi {a : α} (b : Ioi a) : Subtype.val '' Ioi b = Ioi b.1 := + (Subtype.image_preimage_val (Ioi a) (Ioi b.1)).trans <| inter_eq_left.2 <| Ioi_subset_Ioi b.2.le + +@[simp] +lemma image_subtype_val_Ioi_Iio {a : α} (b : Ioi a) : Subtype.val '' Iio b = Ioo a b := + (Subtype.image_preimage_val (Ioi a) (Iio b.1)).trans Iio_inter_Ioi + +@[simp] +lemma image_subtype_val_Iio_Ici {a : α} (b : Iio a) : Subtype.val '' Ici b = Ico b.1 a := + Subtype.image_preimage_val (Iio a) (Ici b.1) + +@[simp] +lemma image_subtype_val_Iio_Iic {a : α} (b : Iio a) : Subtype.val '' Iic b = Iic b.1 := + image_subtype_val_Ioi_Ici (α := αᵒᵈ) _ + +@[simp] +lemma image_subtype_val_Iio_Ioi {a : α} (b : Iio a) : Subtype.val '' Ioi b = Ioo b.1 a := + Subtype.image_preimage_val (Iio a) (Ioi b.1) + +@[simp] +lemma image_subtype_val_Iio_Iio {a : α} (b : Iio a) : Subtype.val '' Iio b = Iio b.1 := + image_subtype_val_Ioi_Ioi (α := αᵒᵈ) _ + +private lemma image_subtype_val_Ixx_Ixi {p q r : α → α → Prop} {a b : α} (c : {x // p a x ∧ q x b}) + (h : ∀ {x}, r c x → p a x) : + Subtype.val '' {y : {x // p a x ∧ q x b} | r c.1 y.1} = {y : α | r c.1 y ∧ q y b} := + (Subtype.image_preimage_val {x | p a x ∧ q x b} {y | r c.1 y}).trans <| by + ext; simp (config := { contextual := true }) [h] + +private lemma image_subtype_val_Ixx_Iix {p q r : α → α → Prop} {a b : α} (c : {x // p a x ∧ q x b}) + (h : ∀ {x}, r x c → q x b) : + Subtype.val '' {y : {x // p a x ∧ q x b} | r y.1 c.1} = {y : α | p a y ∧ r y c.1} := + (Subtype.image_preimage_val {x | p a x ∧ q x b} {y | r y c.1}).trans <| by + ext; simp (config := { contextual := true }) [@and_comm (p _ _), h] + +@[simp] +lemma image_subtype_val_Icc_Ici {a b : α} (c : Icc a b) : Subtype.val '' Ici c = Icc c.1 b := + image_subtype_val_Ixx_Ixi c c.2.1.trans + +@[simp] +lemma image_subtype_val_Icc_Iic {a b : α} (c : Icc a b) : Subtype.val '' Iic c = Icc a c := + image_subtype_val_Ixx_Iix c (le_trans · c.2.2) + +@[simp] +lemma image_subtype_val_Icc_Ioi {a b : α} (c : Icc a b) : Subtype.val '' Ioi c = Ioc c.1 b := + image_subtype_val_Ixx_Ixi c (c.2.1.trans <| le_of_lt ·) + +@[simp] +lemma image_subtype_val_Icc_Iio {a b : α} (c : Icc a b) : Subtype.val '' Iio c = Ico a c := + image_subtype_val_Ixx_Iix c fun h ↦ (le_of_lt h).trans c.2.2 + +@[simp] +lemma image_subtype_val_Ico_Ici {a b : α} (c : Ico a b) : Subtype.val '' Ici c = Ico c.1 b := + image_subtype_val_Ixx_Ixi c c.2.1.trans + +@[simp] +lemma image_subtype_val_Ico_Iic {a b : α} (c : Ico a b) : Subtype.val '' Iic c = Icc a c := + image_subtype_val_Ixx_Iix c (lt_of_le_of_lt · c.2.2) + +@[simp] +lemma image_subtype_val_Ico_Ioi {a b : α} (c : Ico a b) : Subtype.val '' Ioi c = Ioo c.1 b := + image_subtype_val_Ixx_Ixi c (c.2.1.trans <| le_of_lt ·) + +@[simp] +lemma image_subtype_val_Ico_Iio {a b : α} (c : Ico a b) : Subtype.val '' Iio c = Ico a c := + image_subtype_val_Ixx_Iix c (lt_trans · c.2.2) + +@[simp] +lemma image_subtype_val_Ioc_Ici {a b : α} (c : Ioc a b) : Subtype.val '' Ici c = Icc c.1 b := + image_subtype_val_Ixx_Ixi c c.2.1.trans_le + +@[simp] +lemma image_subtype_val_Ioc_Iic {a b : α} (c : Ioc a b) : Subtype.val '' Iic c = Ioc a c := + image_subtype_val_Ixx_Iix c (le_trans · c.2.2) + +@[simp] +lemma image_subtype_val_Ioc_Ioi {a b : α} (c : Ioc a b) : Subtype.val '' Ioi c = Ioc c.1 b := + image_subtype_val_Ixx_Ixi c c.2.1.trans + +@[simp] +lemma image_subtype_val_Ioc_Iio {a b : α} (c : Ioc a b) : Subtype.val '' Iio c = Ioo a c := + image_subtype_val_Ixx_Iix c fun h ↦ (le_of_lt h).trans c.2.2 + +@[simp] +lemma image_subtype_val_Ioo_Ici {a b : α} (c : Ioo a b) : Subtype.val '' Ici c = Ico c.1 b := + image_subtype_val_Ixx_Ixi c c.2.1.trans_le + +@[simp] +lemma image_subtype_val_Ioo_Iic {a b : α} (c : Ioo a b) : Subtype.val '' Iic c = Ioc a c := + image_subtype_val_Ixx_Iix c (lt_of_le_of_lt · c.2.2) + +@[simp] +lemma image_subtype_val_Ioo_Ioi {a b : α} (c : Ioo a b) : Subtype.val '' Ioi c = Ioo c.1 b := + image_subtype_val_Ixx_Ixi c c.2.1.trans + +@[simp] +lemma image_subtype_val_Ioo_Iio {a b : α} (c : Ioo a b) : Subtype.val '' Iio c = Ioo a c := + image_subtype_val_Ixx_Iix c (lt_trans · c.2.2) + end Set diff --git a/Mathlib/Data/Set/Intervals/OrdConnected.lean b/Mathlib/Data/Set/Intervals/OrdConnected.lean index 956043809a71c..0f142c3b4a282 100644 --- a/Mathlib/Data/Set/Intervals/OrdConnected.lean +++ b/Mathlib/Data/Set/Intervals/OrdConnected.lean @@ -78,6 +78,29 @@ protected theorem Icc_subset (s : Set α) [hs : OrdConnected s] {x y} (hx : x hs.out hx hy #align set.Icc_subset Set.Icc_subset +@[simp] +lemma image_subtype_val_Icc {s : Set α} [OrdConnected s] (x y : s) : + Subtype.val '' Icc x y = Icc x.1 y := + (Subtype.image_preimage_val s (Icc x.1 y)).trans <| inter_eq_left.2 <| s.Icc_subset x.2 y.2 + +@[simp] +lemma image_subtype_val_Ico {s : Set α} [OrdConnected s] (x y : s) : + Subtype.val '' Ico x y = Ico x.1 y := + (Subtype.image_preimage_val s (Ico x.1 y)).trans <| inter_eq_left.2 <| + Ico_subset_Icc_self.trans <| s.Icc_subset x.2 y.2 + +@[simp] +lemma image_subtype_val_Ioc {s : Set α} [OrdConnected s] (x y : s) : + Subtype.val '' Ioc x y = Ioc x.1 y := + (Subtype.image_preimage_val s (Ioc x.1 y)).trans <| inter_eq_left.2 <| + Ioc_subset_Icc_self.trans <| s.Icc_subset x.2 y.2 + +@[simp] +lemma image_subtype_val_Ioo {s : Set α} [OrdConnected s] (x y : s) : + Subtype.val '' Ioo x y = Ioo x.1 y := + (Subtype.image_preimage_val s (Ioo x.1 y)).trans <| inter_eq_left.2 <| + Ioo_subset_Icc_self.trans <| s.Icc_subset x.2 y.2 + theorem OrdConnected.inter {s t : Set α} (hs : OrdConnected s) (ht : OrdConnected t) : OrdConnected (s ∩ t) := ⟨fun _ hx _ hy => subset_inter (hs.out hx.1 hy.1) (ht.out hx.2 hy.2)⟩ diff --git a/Mathlib/Data/Set/Pointwise/Basic.lean b/Mathlib/Data/Set/Pointwise/Basic.lean index 379bda1983bba..49e18f3e9909a 100644 --- a/Mathlib/Data/Set/Pointwise/Basic.lean +++ b/Mathlib/Data/Set/Pointwise/Basic.lean @@ -3,12 +3,12 @@ Copyright (c) 2019 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Floris van Doorn -/ +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Algebra.Group.Units.Hom import Mathlib.Algebra.GroupPower.Basic import Mathlib.Algebra.GroupWithZero.Basic -import Mathlib.Algebra.Hom.Equiv.Basic -import Mathlib.Algebra.Hom.Units -import Mathlib.Data.Set.Lattice import Mathlib.Data.Nat.Order.Basic +import Mathlib.Data.Set.Lattice import Mathlib.Tactic.Common #align_import data.set.pointwise.basic from "leanprover-community/mathlib"@"5e526d18cea33550268dcbbddcb822d5cde40654" diff --git a/Mathlib/Data/Set/Pointwise/Iterate.lean b/Mathlib/Data/Set/Pointwise/Iterate.lean index c6d58c820b04f..4e8a4fc1f41e9 100644 --- a/Mathlib/Data/Set/Pointwise/Iterate.lean +++ b/Mathlib/Data/Set/Pointwise/Iterate.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ +import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Data.Set.Pointwise.SMul -import Mathlib.Algebra.Hom.Iterate import Mathlib.Dynamics.FixedPoints.Basic #align_import data.set.pointwise.iterate from "leanprover-community/mathlib"@"9003f28797c0664a49e4179487267c494477d853" diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 73d7e068ff9e2..20221a834f8c0 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -86,7 +86,7 @@ theorem image2_smul : image2 SMul.smul s t = s • t := #align set.image2_smul Set.image2_smul #align set.image2_vadd Set.image2_vadd --- @[to_additive add_image_prod] -- Porting note: bug in mathlib3 +@[to_additive vadd_image_prod] theorem image_smul_prod : (fun x : α × β ↦ x.fst • x.snd) '' s ×ˢ t = s • t := image_prod _ #align set.image_smul_prod Set.image_smul_prod diff --git a/Mathlib/Data/Set/Prod.lean b/Mathlib/Data/Set/Prod.lean index 67b15e8c73e13..a078db8541d70 100644 --- a/Mathlib/Data/Set/Prod.lean +++ b/Mathlib/Data/Set/Prod.lean @@ -554,6 +554,89 @@ theorem diag_image (s : Set α) : (fun x => (x, x)) '' s = diagonal α ∩ s × end Diagonal +end Set + +section Pullback + +open Set + +variable {X Y Z} + +/-- The fiber product $X \times_Y Z$. -/ +abbrev Function.Pullback (f : X → Y) (g : Z → Y) := {p : X × Z // f p.1 = g p.2} + +/-- The fiber product $X \times_Y X$. -/ +abbrev Function.PullbackSelf (f : X → Y) := f.Pullback f + +/-- The projection from the fiber product to the first factor. -/ +def Function.Pullback.fst {f : X → Y} {g : Z → Y} (p : f.Pullback g) : X := p.val.1 + +/-- The projection from the fiber product to the second factor. -/ +def Function.Pullback.snd {f : X → Y} {g : Z → Y} (p : f.Pullback g) : Z := p.val.2 + +open Function.Pullback in +lemma Function.pullback_comm_sq (f : X → Y) (g : Z → Y) : + f ∘ @fst X Y Z f g = g ∘ @snd X Y Z f g := funext fun p ↦ p.2 + +/-- The diagonal map $\Delta: X \to X \times_Y X$. -/ +def toPullbackDiag (f : X → Y) (x : X) : f.Pullback f := ⟨(x, x), rfl⟩ + +/-- The diagonal $\Delta(X) \subseteq X \times_Y X$. -/ +def Function.pullbackDiagonal (f : X → Y) : Set (f.Pullback f) := {p | p.fst = p.snd} + +/-- Three functions between the three pairs of spaces $X_i, Y_i, Z_i$ that are compatible + induce a function $X_1 \times_{Y_1} Z_1 \to X_2 \times_{Y_2} Z_2$. -/ +def Function.mapPullback {X₁ X₂ Y₁ Y₂ Z₁ Z₂} + {f₁ : X₁ → Y₁} {g₁ : Z₁ → Y₁} {f₂ : X₂ → Y₂} {g₂ : Z₂ → Y₂} + (mapX : X₁ → X₂) (mapY : Y₁ → Y₂) (mapZ : Z₁ → Z₂) + (commX : f₂ ∘ mapX = mapY ∘ f₁) (commZ : g₂ ∘ mapZ = mapY ∘ g₁) + (p : f₁.Pullback g₁) : f₂.Pullback g₂ := + ⟨(mapX p.fst, mapZ p.snd), + (congr_fun commX _).trans <| (congr_arg mapY p.2).trans <| congr_fun commZ.symm _⟩ + +open Function.Pullback in +/-- The projection $(X \times_Y Z) \times_Z (X \times_Y Z) \to X \times_Y X$. -/ +def Function.PullbackSelf.map_fst {f : X → Y} {g : Z → Y} : + (@snd X Y Z f g).PullbackSelf → f.PullbackSelf := + mapPullback fst g fst (pullback_comm_sq f g) (pullback_comm_sq f g) + +open Function.Pullback in +/-- The projection $(X \times_Y Z) \times_X (X \times_Y Z) \to Z \times_Y Z$. -/ +def Function.PullbackSelf.map_snd {f : X → Y} {g : Z → Y} : + (@fst X Y Z f g).PullbackSelf → g.PullbackSelf := + mapPullback snd f snd (pullback_comm_sq f g).symm (pullback_comm_sq f g).symm + +open Function.PullbackSelf Function.Pullback +theorem preimage_map_fst_pullbackDiagonal {f : X → Y} {g : Z → Y} : + @map_fst X Y Z f g ⁻¹' pullbackDiagonal f = pullbackDiagonal (@snd X Y Z f g) := by + ext ⟨⟨p₁, p₂⟩, he⟩ + simp_rw [pullbackDiagonal, mem_setOf, Subtype.ext_iff, Prod.ext_iff] + exact (and_iff_left he).symm + +theorem Function.Injective.preimage_pullbackDiagonal {f : X → Y} {g : Z → X} (inj : g.Injective) : + mapPullback g id g (by rfl) (by rfl) ⁻¹' pullbackDiagonal f = pullbackDiagonal (f ∘ g) := + ext fun _ ↦ inj.eq_iff + +theorem image_toPullbackDiag (f : X → Y) (s : Set X) : + toPullbackDiag f '' s = pullbackDiagonal f ∩ Subtype.val ⁻¹' s ×ˢ s := by + ext x + constructor + · rintro ⟨x, hx, rfl⟩ + exact ⟨rfl, hx, hx⟩ + · obtain ⟨⟨x, y⟩, h⟩ := x + rintro ⟨rfl : x = y, h2x⟩ + exact mem_image_of_mem _ h2x.1 + +theorem range_toPullbackDiag (f : X → Y) : range (toPullbackDiag f) = pullbackDiagonal f := by + rw [← image_univ, image_toPullbackDiag, univ_prod_univ, preimage_univ, inter_univ] + +theorem injective_toPullbackDiag (f : X → Y) : (toPullbackDiag f).Injective := + fun _ _ h ↦ congr_arg Prod.fst (congr_arg Subtype.val h) + +end Pullback + +namespace Set + section OffDiag variable {α : Type*} {s t : Set α} {x : α × α} {a : α} diff --git a/Mathlib/Data/SetLike/Basic.lean b/Mathlib/Data/SetLike/Basic.lean index a92f69c877aa2..5370a3d11c39a 100644 --- a/Mathlib/Data/SetLike/Basic.lean +++ b/Mathlib/Data/SetLike/Basic.lean @@ -39,7 +39,7 @@ namespace MySubobject variables {X : Type*} [ObjectTypeclass X] {x : X} instance : SetLike (MySubobject X) X := - ⟨MySubobject.carrier, λ p q h, by cases p; cases q; congr'⟩ + ⟨MySubobject.carrier, fun p q h => by cases p; cases q; congr!⟩ @[simp] lemma mem_carrier {p : MySubobject X} : x ∈ p.carrier ↔ x ∈ (p : Set X) := Iff.rfl diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index d51cd58790179..ae98a5158f3eb 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -34,7 +34,7 @@ namespace SignType -- Porting note: Added Fintype SignType manually instance : Fintype SignType := - Fintype.ofMultiset (zero :: neg :: pos :: List.nil) (fun x ↦ by cases x <;> simp only) + Fintype.ofMultiset (zero :: neg :: pos :: List.nil) (fun x ↦ by cases x <;> decide) instance : Zero SignType := ⟨zero⟩ @@ -155,7 +155,7 @@ def fin3Equiv : SignType ≃* Fin 3 where | ⟨2, _⟩ => by simp | ⟨n + 3, h⟩ => by simp at h map_mul' a b := by - cases a <;> cases b <;> simp + cases a <;> cases b <;> rfl #align sign_type.fin3_equiv SignType.fin3Equiv section CaseBashing diff --git a/Mathlib/Data/Stream/Init.lean b/Mathlib/Data/Stream/Init.lean index 550b8f56a5712..cc6b3eff963e7 100644 --- a/Mathlib/Data/Stream/Init.lean +++ b/Mathlib/Data/Stream/Init.lean @@ -591,7 +591,7 @@ theorem length_take (n : ℕ) (s : Stream' α) : (take n s).length = n := by theorem take_take {s : Stream' α} : ∀ {m n}, (s.take n).take m = s.take (min n m) | 0, n => by rw [min_zero, List.take_zero, take_zero] | m, 0 => by rw [zero_min, take_zero, List.take_nil] - | m+1, n+1 => by rw [take_succ, List.take_cons, Nat.min_succ_succ, take_succ, take_take] + | m+1, n+1 => by rw [take_succ, List.take_cons, Nat.succ_min_succ, take_succ, take_take] @[simp] theorem concat_take_get {s : Stream' α} : s.take n ++ [s.get n] = s.take (n+1) := (take_succ' n).symm diff --git a/Mathlib/Data/String/Basic.lean b/Mathlib/Data/String/Basic.lean index 633bb00b8c2a9..e1e2b6da63107 100644 --- a/Mathlib/Data/String/Basic.lean +++ b/Mathlib/Data/String/Basic.lean @@ -75,13 +75,13 @@ theorem ltb_cons_addChar (c : Char) (cs₁ cs₂ : List Char) (i₁ i₂ : Pos) theorem lt_iff_toList_lt : ∀ {s₁ s₂ : String}, s₁ < s₂ ↔ s₁.toList < s₂.toList | ⟨s₁⟩, ⟨s₂⟩ => show ltb ⟨⟨s₁⟩, 0⟩ ⟨⟨s₂⟩, 0⟩ ↔ s₁ < s₂ by induction s₁ generalizing s₂ <;> cases s₂ - · simp + · decide · rename_i c₂ cs₂; apply iff_of_true - · rw [ltb]; simp; apply ne_false_of_eq_true; apply decide_eq_true + · rw [ltb]; simp only [Iterator.hasNext, Iterator.curr] simp [endPos, utf8ByteSize, utf8ByteSize.go, csize_pos] · apply List.nil_lt_cons · rename_i c₁ cs₁ ih; apply iff_of_false - · rw [ltb]; simp + · rw [ltb]; simp [Iterator.hasNext, Iterator.curr] · apply not_lt_of_lt; apply List.nil_lt_cons · rename_i c₁ cs₁ ih c₂ cs₂; rw [ltb] simp [Iterator.hasNext, endPos, utf8ByteSize, utf8ByteSize.go, csize_pos, Iterator.curr, get, @@ -134,7 +134,7 @@ theorem asString_inv_toList (s : String) : s.toList.asString = s := theorem toList_nonempty : ∀ {s : String}, s ≠ "" → s.toList = s.head :: (s.drop 1).toList | ⟨s⟩, h => by cases s - · simp only at h + · simp only [ne_eq, not_true_eq_false] at h · rename_i c cs simp only [toList, List.cons.injEq] constructor <;> [rfl; simp [drop_eq]] diff --git a/Mathlib/Data/Sum/Interval.lean b/Mathlib/Data/Sum/Interval.lean index 95716f7596fc3..a1186898ec78f 100644 --- a/Mathlib/Data/Sum/Interval.lean +++ b/Mathlib/Data/Sum/Interval.lean @@ -3,19 +3,17 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Data.Finset.Sum import Mathlib.Data.Sum.Order import Mathlib.Order.LocallyFinite -#align_import data.sum.interval from "leanprover-community/mathlib"@"861a26926586cd46ff80264d121cdb6fa0e35cc1" +#align_import data.sum.interval from "leanprover-community/mathlib"@"48a058d7e39a80ed56858505719a0b2197900999" /-! # Finite intervals in a disjoint union -This file provides the `LocallyFiniteOrder` instance for the disjoint sum of two orders. - -## TODO - -Do the same for the lexicographic sum of orders. +This file provides the `LocallyFiniteOrder` instance for the disjoint sum and linear sum of two +orders and calculates the cardinality of their finite intervals. -/ @@ -107,6 +105,118 @@ theorem sumLift₂_mono (h₁ : ∀ a b, f₁ a b ⊆ g₁ a b) (h₂ : ∀ a b, end SumLift₂ +section SumLexLift +variable (f₁ f₁' : α₁ → β₁ → Finset γ₁) (f₂ f₂' : α₂ → β₂ → Finset γ₂) + (g₁ g₁' : α₁ → β₂ → Finset γ₁) (g₂ g₂' : α₁ → β₂ → Finset γ₂) + +/-- Lifts maps `α₁ → β₁ → Finset γ₁`, `α₂ → β₂ → Finset γ₂`, `α₁ → β₂ → Finset γ₁`, +`α₂ → β₂ → Finset γ₂` to a map `α₁ ⊕ α₂ → β₁ ⊕ β₂ → Finset (γ₁ ⊕ γ₂)`. Could be generalized to +alternative monads if we can make sure to keep computability and universe polymorphism. -/ +def sumLexLift : Sum α₁ α₂ → Sum β₁ β₂ → Finset (Sum γ₁ γ₂) + | inl a, inl b => (f₁ a b).map Embedding.inl + | inl a, inr b => (g₁ a b).disjSum (g₂ a b) + | inr _, inl _ => ∅ + | inr a, inr b => (f₂ a b).map ⟨_, inr_injective⟩ +#align finset.sum_lex_lift Finset.sumLexLift + +@[simp] +lemma sumLexLift_inl_inl (a : α₁) (b : β₁) : + sumLexLift f₁ f₂ g₁ g₂ (inl a) (inl b) = (f₁ a b).map Embedding.inl := rfl +#align finset.sum_lex_lift_inl_inl Finset.sumLexLift_inl_inl + +@[simp] +lemma sumLexLift_inl_inr (a : α₁) (b : β₂) : + sumLexLift f₁ f₂ g₁ g₂ (inl a) (inr b) = (g₁ a b).disjSum (g₂ a b) := rfl +#align finset.sum_lex_lift_inl_inr Finset.sumLexLift_inl_inr + +@[simp] +lemma sumLexLift_inr_inl (a : α₂) (b : β₁) : sumLexLift f₁ f₂ g₁ g₂ (inr a) (inl b) = ∅ := rfl +#align finset.sum_lex_lift_inr_inl Finset.sumLexLift_inr_inl + +@[simp] +lemma sumLexLift_inr_inr (a : α₂) (b : β₂) : + sumLexLift f₁ f₂ g₁ g₂ (inr a) (inr b) = (f₂ a b).map ⟨_, inr_injective⟩ := rfl +#align finset.sum_lex_lift_inr_inr Finset.sumLexLift_inr_inr + +variable {f₁ g₁ f₂ g₂ f₁' g₁' f₂' g₂'} {a : Sum α₁ α₂} {b : Sum β₁ β₂} {c : Sum γ₁ γ₂} + +lemma mem_sumLexLift : + c ∈ sumLexLift f₁ f₂ g₁ g₂ a b ↔ + (∃ a₁ b₁ c₁, a = inl a₁ ∧ b = inl b₁ ∧ c = inl c₁ ∧ c₁ ∈ f₁ a₁ b₁) ∨ + (∃ a₁ b₂ c₁, a = inl a₁ ∧ b = inr b₂ ∧ c = inl c₁ ∧ c₁ ∈ g₁ a₁ b₂) ∨ + (∃ a₁ b₂ c₂, a = inl a₁ ∧ b = inr b₂ ∧ c = inr c₂ ∧ c₂ ∈ g₂ a₁ b₂) ∨ + ∃ a₂ b₂ c₂, a = inr a₂ ∧ b = inr b₂ ∧ c = inr c₂ ∧ c₂ ∈ f₂ a₂ b₂ := by + constructor + · obtain a | a := a <;> obtain b | b := b + · rw [sumLexLift, mem_map] + rintro ⟨c, hc, rfl⟩ + exact Or.inl ⟨a, b, c, rfl, rfl, rfl, hc⟩ + · refine' fun h ↦ (mem_disjSum.1 h).elim _ _ + · rintro ⟨c, hc, rfl⟩ + exact Or.inr (Or.inl ⟨a, b, c, rfl, rfl, rfl, hc⟩) + · rintro ⟨c, hc, rfl⟩ + exact Or.inr (Or.inr $ Or.inl ⟨a, b, c, rfl, rfl, rfl, hc⟩) + · exact fun h ↦ (not_mem_empty _ h).elim + · rw [sumLexLift, mem_map] + rintro ⟨c, hc, rfl⟩ + exact Or.inr (Or.inr $ Or.inr $ ⟨a, b, c, rfl, rfl, rfl, hc⟩) + · rintro (⟨a, b, c, rfl, rfl, rfl, hc⟩ | ⟨a, b, c, rfl, rfl, rfl, hc⟩ | + ⟨a, b, c, rfl, rfl, rfl, hc⟩ | ⟨a, b, c, rfl, rfl, rfl, hc⟩) + · exact mem_map_of_mem _ hc + · exact inl_mem_disjSum.2 hc + · exact inr_mem_disjSum.2 hc + · exact mem_map_of_mem _ hc +#align finset.mem_sum_lex_lift Finset.mem_sumLexLift + +lemma inl_mem_sumLexLift {c₁ : γ₁} : + inl c₁ ∈ sumLexLift f₁ f₂ g₁ g₂ a b ↔ + (∃ a₁ b₁, a = inl a₁ ∧ b = inl b₁ ∧ c₁ ∈ f₁ a₁ b₁) ∨ + ∃ a₁ b₂, a = inl a₁ ∧ b = inr b₂ ∧ c₁ ∈ g₁ a₁ b₂ := by + simp [mem_sumLexLift] +#align finset.inl_mem_sum_lex_lift Finset.inl_mem_sumLexLift + +lemma inr_mem_sumLexLift {c₂ : γ₂} : + inr c₂ ∈ sumLexLift f₁ f₂ g₁ g₂ a b ↔ + (∃ a₁ b₂, a = inl a₁ ∧ b = inr b₂ ∧ c₂ ∈ g₂ a₁ b₂) ∨ + ∃ a₂ b₂, a = inr a₂ ∧ b = inr b₂ ∧ c₂ ∈ f₂ a₂ b₂ := by + simp [mem_sumLexLift] +#align finset.inr_mem_sum_lex_lift Finset.inr_mem_sumLexLift + +lemma sumLexLift_mono (hf₁ : ∀ a b, f₁ a b ⊆ f₁' a b) (hf₂ : ∀ a b, f₂ a b ⊆ f₂' a b) + (hg₁ : ∀ a b, g₁ a b ⊆ g₁' a b) (hg₂ : ∀ a b, g₂ a b ⊆ g₂' a b) (a : Sum α₁ α₂) + (b : Sum β₁ β₂) : sumLexLift f₁ f₂ g₁ g₂ a b ⊆ sumLexLift f₁' f₂' g₁' g₂' a b := by + cases a <;> cases b + exacts [map_subset_map.2 (hf₁ _ _), disjSum_mono (hg₁ _ _) (hg₂ _ _), Subset.rfl, + map_subset_map.2 (hf₂ _ _)] +#align finset.sum_lex_lift_mono Finset.sumLexLift_mono + +lemma sumLexLift_eq_empty : + sumLexLift f₁ f₂ g₁ g₂ a b = ∅ ↔ + (∀ a₁ b₁, a = inl a₁ → b = inl b₁ → f₁ a₁ b₁ = ∅) ∧ + (∀ a₁ b₂, a = inl a₁ → b = inr b₂ → g₁ a₁ b₂ = ∅ ∧ g₂ a₁ b₂ = ∅) ∧ + ∀ a₂ b₂, a = inr a₂ → b = inr b₂ → f₂ a₂ b₂ = ∅ := by + refine' ⟨fun h ↦ ⟨_, _, _⟩, fun h ↦ _⟩ + any_goals rintro a b rfl rfl; exact map_eq_empty.1 h + · rintro a b rfl rfl; exact disjSum_eq_empty.1 h + cases a <;> cases b + · exact map_eq_empty.2 (h.1 _ _ rfl rfl) + · simp [h.2.1 _ _ rfl rfl] + · rfl + · exact map_eq_empty.2 (h.2.2 _ _ rfl rfl) +#align finset.sum_lex_lift_eq_empty Finset.sumLexLift_eq_empty + +lemma sumLexLift_nonempty : + (sumLexLift f₁ f₂ g₁ g₂ a b).Nonempty ↔ + (∃ a₁ b₁, a = inl a₁ ∧ b = inl b₁ ∧ (f₁ a₁ b₁).Nonempty) ∨ + (∃ a₁ b₂, a = inl a₁ ∧ b = inr b₂ ∧ ((g₁ a₁ b₂).Nonempty ∨ (g₂ a₁ b₂).Nonempty)) ∨ + ∃ a₂ b₂, a = inr a₂ ∧ b = inr b₂ ∧ (f₂ a₂ b₂).Nonempty := by + -- porting note: was `simp [nonempty_iff_ne_empty, sumLexLift_eq_empty, not_and_or]`. Could + -- add `-exists_and_left, -not_and, -exists_and_right` but easier to squeeze. + simp only [nonempty_iff_ne_empty, Ne.def, sumLexLift_eq_empty, not_and_or, exists_prop, + not_forall] +#align finset.sum_lex_lift_nonempty Finset.sumLexLift_nonempty + +end SumLexLift end Finset open Finset Function @@ -209,4 +319,115 @@ theorem Ioo_inr_inr : Ioo (inr b₁ : Sum α β) (inr b₂) = (Ioo b₁ b₂).ma end Disjoint +/-! ### Lexicographical sum of orders -/ + +namespace Lex +variable [Preorder α] [Preorder β] [OrderTop α] [OrderBot β] [LocallyFiniteOrder α] + [LocallyFiniteOrder β] + +/-- Throwaway tactic. -/ +local elab "simp_lex" : tactic => do + Lean.Elab.Tactic.evalTactic <| ← `(tactic| + refine toLex.surjective.forall₃.2 ?_; + rintro (a | a) (b | b) (c | c) <;> simp only + [sumLexLift_inl_inl, sumLexLift_inl_inr, sumLexLift_inr_inl, sumLexLift_inr_inr, + inl_le_inl_iff, inl_le_inr, not_inr_le_inl, inr_le_inr_iff, inl_lt_inl_iff, inl_lt_inr, + not_inr_lt_inl, inr_lt_inr_iff, mem_Icc, mem_Ico, mem_Ioc, mem_Ioo, mem_Ici, mem_Ioi, + mem_Iic, mem_Iio, Equiv.coe_toEmbedding, toLex_inj, exists_false, and_false, false_and, + map_empty, not_mem_empty, true_and, inl_mem_disjSum, inr_mem_disjSum, and_true, ofLex_toLex, + mem_map, Embedding.coeFn_mk, exists_prop, exists_eq_right, Embedding.inl_apply, + -- porting note: added + inl.injEq, inr.injEq] + ) + +instance locallyFiniteOrder : LocallyFiniteOrder (α ⊕ₗ β) where + finsetIcc a b := + (sumLexLift Icc Icc (fun a _ => Ici a) (fun _ => Iic) (ofLex a) (ofLex b)).map toLex.toEmbedding + finsetIco a b := + (sumLexLift Ico Ico (fun a _ => Ici a) (fun _ => Iio) (ofLex a) (ofLex b)).map toLex.toEmbedding + finsetIoc a b := + (sumLexLift Ioc Ioc (fun a _ => Ioi a) (fun _ => Iic) (ofLex a) (ofLex b)).map toLex.toEmbedding + finsetIoo a b := + (sumLexLift Ioo Ioo (fun a _ => Ioi a) (fun _ => Iio) (ofLex a) (ofLex b)).map toLex.toEmbedding + finset_mem_Icc := by simp_lex + finset_mem_Ico := by simp_lex + finset_mem_Ioc := by simp_lex + finset_mem_Ioo := by simp_lex +#align sum.lex.locally_finite_order Sum.Lex.locallyFiniteOrder + +variable (a a₁ a₂ : α) (b b₁ b₂ : β) + +lemma Icc_inl_inl : + Icc (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Icc a₁ a₂).map (Embedding.inl.trans toLex.toEmbedding) := by + rw [← Finset.map_map]; rfl +#align sum.lex.Icc_inl_inl Sum.Lex.Icc_inl_inl + +lemma Ico_inl_inl : + Ico (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Ico a₁ a₂).map (Embedding.inl.trans toLex.toEmbedding) := by + rw [← Finset.map_map]; rfl +#align sum.lex.Ico_inl_inl Sum.Lex.Ico_inl_inl + +lemma Ioc_inl_inl : + Ioc (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Ioc a₁ a₂).map (Embedding.inl.trans toLex.toEmbedding) := by + rw [← Finset.map_map]; rfl +#align sum.lex.Ioc_inl_inl Sum.Lex.Ioc_inl_inl + +lemma Ioo_inl_inl : + Ioo (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Ioo a₁ a₂).map (Embedding.inl.trans toLex.toEmbedding) := by + rw [← Finset.map_map]; rfl +#align sum.lex.Ioo_inl_inl Sum.Lex.Ioo_inl_inl + +@[simp] +lemma Icc_inl_inr : Icc (inlₗ a) (inrₗ b) = ((Ici a).disjSum (Iic b)).map toLex.toEmbedding := rfl +#align sum.lex.Icc_inl_inr Sum.Lex.Icc_inl_inr + +@[simp] +lemma Ico_inl_inr : Ico (inlₗ a) (inrₗ b) = ((Ici a).disjSum (Iio b)).map toLex.toEmbedding := rfl +#align sum.lex.Ico_inl_inr Sum.Lex.Ico_inl_inr + +@[simp] +lemma Ioc_inl_inr : Ioc (inlₗ a) (inrₗ b) = ((Ioi a).disjSum (Iic b)).map toLex.toEmbedding := rfl +#align sum.lex.Ioc_inl_inr Sum.Lex.Ioc_inl_inr + +@[simp] +lemma Ioo_inl_inr : Ioo (inlₗ a) (inrₗ b) = ((Ioi a).disjSum (Iio b)).map toLex.toEmbedding := rfl +#align sum.lex.Ioo_inl_inr Sum.Lex.Ioo_inl_inr + +@[simp, nolint simpNF] -- Porting note: dsimp cannot prove this +lemma Icc_inr_inl : Icc (inrₗ b) (inlₗ a) = ∅ := rfl +#align sum.lex.Icc_inr_inl Sum.Lex.Icc_inr_inl + +@[simp, nolint simpNF] -- Porting note: dsimp cannot prove this +lemma Ico_inr_inl : Ico (inrₗ b) (inlₗ a) = ∅ := rfl +#align sum.lex.Ico_inr_inl Sum.Lex.Ico_inr_inl + +@[simp, nolint simpNF] -- Porting note: dsimp cannot prove this +lemma Ioc_inr_inl : Ioc (inrₗ b) (inlₗ a) = ∅ := rfl +#align sum.lex.Ioc_inr_inl Sum.Lex.Ioc_inr_inl + +@[simp, nolint simpNF] -- Porting note: dsimp cannot prove this +lemma Ioo_inr_inl : Ioo (inrₗ b) (inlₗ a) = ∅ := rfl +#align sum.lex.Ioo_inr_inl Sum.Lex.Ioo_inr_inl + +lemma Icc_inr_inr : + Icc (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Icc b₁ b₂).map (Embedding.inr.trans toLex.toEmbedding) := by + rw [← Finset.map_map]; rfl +#align sum.lex.Icc_inr_inr Sum.Lex.Icc_inr_inr + +lemma Ico_inr_inr : + Ico (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Ico b₁ b₂).map (Embedding.inr.trans toLex.toEmbedding) := by + rw [← Finset.map_map]; rfl +#align sum.lex.Ico_inr_inr Sum.Lex.Ico_inr_inr + +lemma Ioc_inr_inr : + Ioc (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Ioc b₁ b₂).map (Embedding.inr.trans toLex.toEmbedding) := by + rw [← Finset.map_map]; rfl +#align sum.lex.Ioc_inr_inr Sum.Lex.Ioc_inr_inr + +lemma Ioo_inr_inr : + Ioo (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Ioo b₁ b₂).map (Embedding.inr.trans toLex.toEmbedding) := by + rw [← Finset.map_map]; rfl +#align sum.lex.Ioo_inr_inr Sum.Lex.Ioo_inr_inr + +end Lex end Sum diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 081e1cab31555..4c536fc3ab5b1 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -53,7 +53,7 @@ instance Sym.hasCoe (α : Type*) (n : ℕ) : CoeOut (Sym α n) (Multiset α) := #align sym.has_coe Sym.hasCoe -- Porting note: instance needed for Data.Finset.Sym -instance [DecidableEq α]: DecidableEq (Sym α n) := Subtype.instDecidableEqSubtype +instance [DecidableEq α] : DecidableEq (Sym α n) := Subtype.instDecidableEqSubtype /-- This is the `List.Perm` setoid lifted to `Vector`. diff --git a/Mathlib/Data/Sym/Sym2.lean b/Mathlib/Data/Sym/Sym2.lean index 41d3715a47efe..82521932fd4a9 100644 --- a/Mathlib/Data/Sym/Sym2.lean +++ b/Mathlib/Data/Sym/Sym2.lean @@ -60,7 +60,7 @@ namespace Sym2 /-- This is the relation capturing the notion of pairs equivalent up to permutations. -/ -@[aesop (rule_sets [Sym2]) [safe [constructors, cases], norm unfold]] +@[aesop (rule_sets [Sym2]) [safe [constructors, cases], norm]] inductive Rel (α : Type u) : α × α → α × α → Prop | refl (x y : α) : Rel _ (x, y) (x, y) | swap (x y : α) : Rel _ (x, y) (y, x) @@ -633,39 +633,19 @@ end SymEquiv section Decidable -/-- An algorithm for computing `Sym2.Rel`. --/ -@[aesop norm unfold (rule_sets [Sym2])] -def relBool [DecidableEq α] (x y : α × α) : Bool := - if x.1 = y.1 then x.2 = y.2 else if x.1 = y.2 then x.2 = y.1 else false -#align sym2.rel_bool Sym2.relBool - -@[aesop norm unfold (rule_sets [Sym2])] -theorem relBool_spec [DecidableEq α] (x y : α × α) : ↥(relBool x y) ↔ Rel α x y := by - cases' x with x₁ x₂; cases' y with y₁ y₂ - aesop (rule_sets [Sym2]) (add norm unfold [relBool]) -#align sym2.rel_bool_spec Sym2.relBool_spec +#noalign sym2.rel_bool +#noalign sym2.rel_bool_spec /-- Given `[DecidableEq α]` and `[Fintype α]`, the following instance gives `Fintype (Sym2 α)`. -/ -instance instRelDecidable (α : Type*) [DecidableEq α] : DecidableRel (Sym2.Rel α) := fun x y => - decidable_of_bool (relBool x y) (relBool_spec x y) --- Porting note: add this other version needed for Data.Finset.Sym -instance instRelDecidable' (α : Type*) [DecidableEq α] : - DecidableRel (· ≈ · : α × α → α × α → Prop) := instRelDecidable _ - --- porting note: extra definitions and lemmas for proving decidable equality in `Sym2` -/-- An algorithm for deciding equality in `Sym2 α`. -/ -@[aesop norm unfold (rule_sets [Sym2])] -def eqBool [DecidableEq α] : Sym2 α → Sym2 α → Bool := - Sym2.lift₂.toFun - ⟨fun x₁ x₂ y₁ y₂ => relBool (x₁, x₂) (y₁, y₂), by aesop (add norm unfold [relBool])⟩ - -@[aesop norm unfold (rule_sets [Sym2])] -theorem eqBool_spec [DecidableEq α] (a b : Sym2 α) : (eqBool a b) ↔ (a = b) := - Sym2.inductionOn₂ a b <| by aesop (rule_sets [Sym2]) +instance instDecidableRel [DecidableEq α] : DecidableRel (Rel α) := + fun _ _ => decidable_of_iff' _ rel_iff +instance instDecidableRel' [DecidableEq α] : DecidableRel (HasEquiv.Equiv (α := α × α)) := + instDecidableRel +-- the `Equiv` version above is needed for this +example [DecidableEq α] : DecidableEq (Sym2 α) := inferInstance /-! ### The other element of an element of the symmetric square -/ @@ -693,8 +673,7 @@ def Mem.other' [DecidableEq α] {a : α} {z : Sym2 α} (h : a ∈ z) : α := intro _ e _; subst e; rfl apply this · rw [mem_iff] at hy - have : relBool x y := (relBool_spec x y).mpr h - aesop (add norm unfold [pairOther, relBool])) + aesop (add norm unfold [pairOther])) z h #align sym2.mem.other' Sym2.Mem.other' diff --git a/Mathlib/Data/UInt.lean b/Mathlib/Data/UInt.lean index 81fe878fd81d4..bc17670d8751f 100644 --- a/Mathlib/Data/UInt.lean +++ b/Mathlib/Data/UInt.lean @@ -5,15 +5,20 @@ import Mathlib.Algebra.GroupWithZero.Defs import Mathlib.Algebra.Ring.Basic import Mathlib.Data.ZMod.Defs -lemma UInt8.val_eq_coe {a : Nat} : (ofNat a).val = a := rfl +lemma UInt8.val_eq_of_lt {a : Nat} : a < UInt8.size → ((ofNat a).val : Nat) = a := + Nat.mod_eq_of_lt -lemma UInt16.val_eq_coe {a : Nat} : (ofNat a).val = a := rfl +lemma UInt16.val_eq_of_lt {a : Nat} : a < UInt16.size → ((ofNat a).val : Nat) = a := + Nat.mod_eq_of_lt -lemma UInt32.val_eq_coe {a : Nat} : (ofNat a).val = a := rfl +lemma UInt32.val_eq_of_lt {a : Nat} : a < UInt32.size → ((ofNat a).val : Nat) = a := + Nat.mod_eq_of_lt -lemma UInt64.val_eq_coe {a : Nat} : (ofNat a).val = a := rfl +lemma UInt64.val_eq_of_lt {a : Nat} : a < UInt64.size → ((ofNat a).val : Nat) = a := + Nat.mod_eq_of_lt -lemma USize.val_eq_coe {a : Nat} : (ofNat a).val = a := rfl +lemma USize.val_eq_of_lt {a : Nat} : a < USize.size → ((ofNat a).val : Nat) = a := + Nat.mod_eq_of_lt instance UInt8.neZero : NeZero UInt8.size := ⟨by decide⟩ @@ -119,16 +124,9 @@ def isAlphanum (c : UInt8) : Bool := c.isAlpha || c.isDigit theorem toChar_aux (n : Nat) (h : n < size) : Nat.isValidChar (UInt32.ofNat n).1 := by - rw [UInt32.val_eq_coe] - apply Or.inl - apply Nat.lt_trans - · dsimp - rw [Nat.mod_eq n UInt32.size] - · rw [if_neg] - · exact h - · simp only [true_and, not_le] - exact Nat.lt_trans h (by decide) - · decide + rw [UInt32.val_eq_of_lt] + exact Or.inl $ Nat.lt_trans h $ by decide + exact Nat.lt_trans h $ by decide /-- The numbers from 0 to 256 are all valid UTF-8 characters, so we can embed one in the other. -/ def toChar (n : UInt8) : Char := ⟨n.toUInt32, toChar_aux n.1 n.1.2⟩ diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index a5b29bbd9a39c..c56e5ec81a1cc 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -853,7 +853,7 @@ theorem neg_one_ne_one {n : ℕ} [Fact (2 < n)] : (-1 : ZMod n) ≠ 1 := #align zmod.neg_one_ne_one ZMod.neg_one_ne_one theorem neg_eq_self_mod_two (a : ZMod 2) : -a = a := by - fin_cases a <;> apply Fin.ext <;> simp [Fin.coe_neg, Int.natMod] + fin_cases a <;> apply Fin.ext <;> simp [Fin.coe_neg, Int.natMod]; rfl #align zmod.neg_eq_self_mod_two ZMod.neg_eq_self_mod_two @[simp] diff --git a/Mathlib/Data/ZMod/IntUnitsPower.lean b/Mathlib/Data/ZMod/IntUnitsPower.lean index d6ceea3a373f6..79ff46b5386a0 100644 --- a/Mathlib/Data/ZMod/IntUnitsPower.lean +++ b/Mathlib/Data/ZMod/IntUnitsPower.lean @@ -9,56 +9,108 @@ import Mathlib.Data.Int.Order.Units import Mathlib.Data.ZMod.Basic /-! -# The power operator by `ZMod 2` on `ℤˣ` +# The power operator on `ℤˣ` by `ZMod 2`, `ℕ`, and `ℤ` See also the related `negOnePow`. ## TODO * Generalize this to `Pow G (Zmod n)` where `orderOf g = n`. -* Abstract this with a `LawfulPow` typeclass such that we can reuse the same lemmas for `z₂pow`, - `npow`, and `zpow`. + +## Implementation notes + +In future, we could consider a `LawfulPower M R` typeclass; but we can save ourselves a lot of work +by using `Module R (Additive M)` in its place, especially since this already has instances for +`R = ℕ` and `R = ℤ`. -/ local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 -/-- There is a canonical power operation by `ℤˣ` on `ZMod 2`. +instance : SMul (ZMod 2) (Additive ℤˣ) where + smul z au := .ofMul <| Additive.toMul au ^ z.val + +lemma ZMod.smul_units_def (z : ZMod 2) (au : Additive ℤˣ) : + z • au = z.val • au := rfl + +lemma ZMod.natCast_smul_units (n : ℕ) (au : Additive ℤˣ) : (n : ZMod 2) • au = n • au := + (Int.units_pow_eq_pow_mod_two au n).symm + +/-- This is an indirect way of saying that `ℤˣ` has a power operation by `ZMod 2`. -/ +instance : Module (ZMod 2) (Additive ℤˣ) where + smul z au := .ofMul <| Additive.toMul au ^ z.val + one_smul au := Additive.toMul.injective <| pow_one _ + mul_smul z₁ z₂ au := Additive.toMul.injective <| by + dsimp only [ZMod.smul_units_def, toMul_nsmul] + rw [←pow_mul, ZMod.val_mul, ←Int.units_pow_eq_pow_mod_two, mul_comm] + smul_zero z := Additive.toMul.injective <| one_pow _ + smul_add z au₁ au₂ := Additive.toMul.injective <| mul_pow _ _ _ + add_smul z₁ z₂ au := Additive.toMul.injective <| by + dsimp only [ZMod.smul_units_def, toMul_nsmul, toMul_add] + rw [←pow_add, ZMod.val_add, ←Int.units_pow_eq_pow_mod_two] + zero_smul au := Additive.toMul.injective <| pow_zero (Additive.toMul au) + +section CommSemiring +variable {R : Type*} [CommSemiring R] [Module R (Additive ℤˣ)] + +/-- There is a canonical power operation on `ℤˣ` by `R` if `Additive ℤˣ` is an `R`-module. + +In lemma names, this operations is called `uzpow` to match `zpow`. + +Notably this is satisfied by `R ∈ {ℕ, ℤ, ZMod 2}`. -/ +instance Int.instUnitsPow : Pow ℤˣ R where + pow u r := Additive.toMul (r • Additive.ofMul u) + +-- The above instance forms no typeclass diamonds with the standard power operators +example : Int.instUnitsPow = Monoid.Pow := rfl +example : Int.instUnitsPow = DivInvMonoid.Pow := rfl + +@[simp] lemma ofMul_uzpow (u : ℤˣ) (r : R) : Additive.ofMul (u ^ r) = r • Additive.ofMul u := rfl + +@[simp] lemma toMul_uzpow (u : Additive ℤˣ) (r : R) : + Additive.toMul (r • u) = Additive.toMul u ^ r := rfl + +@[norm_cast] lemma uzpow_natCast (u : ℤˣ) (n : ℕ) : u ^ (n : R) = u ^ n := by + change Additive.toMul ((n : R) • Additive.ofMul u) = _ + rw [←nsmul_eq_smul_cast, toMul_nsmul, toMul_ofMul] + +-- See note [no_index around OfNat.ofNat] +lemma uzpow_ofNat (s : ℤˣ) (n : ℕ) [n.AtLeastTwo] : + s ^ (no_index (OfNat.ofNat n : R)) = s ^ (no_index (OfNat.ofNat n : ℕ)) := + uzpow_natCast _ _ -In lemma names, this operations is called `z₂pow` to match `zpow`. -/ -instance : Pow ℤˣ (ZMod 2) where - pow s z₂ := s ^ z₂.val +@[simp] lemma one_uzpow (x : R) : (1 : ℤˣ) ^ x = 1 := + Additive.ofMul.injective <| smul_zero _ -lemma z₂pow_def (s : ℤˣ) (x : ZMod 2) : s ^ x = s ^ x.val := rfl +lemma mul_uzpow (s₁ s₂ : ℤˣ) (x : R) : (s₁ * s₂) ^ x = s₁ ^ x * s₂ ^ x := + Additive.ofMul.injective <| smul_add x (Additive.ofMul s₁) (Additive.ofMul s₂) -lemma z₂pow_natCast (s : ℤˣ) (n : ℕ) : s ^ (n : ZMod 2) = s ^ n := - (Int.units_pow_eq_pow_mod_two s n).symm +@[simp] lemma uzpow_zero (s : ℤˣ) : (s ^ (0 : R) : ℤˣ) = (1 : ℤˣ) := + Additive.ofMul.injective <| zero_smul R (Additive.ofMul s) -lemma z₂pow_ofNat (s : ℤˣ) (n : ℕ) [n.AtLeastTwo] : - s ^ (no_index (OfNat.ofNat n : ZMod 2)) = s ^ (no_index (OfNat.ofNat n : ℕ)) := - z₂pow_natCast _ _ +@[simp] lemma uzpow_one (s : ℤˣ) : (s ^ (1 : R) : ℤˣ) = s := + Additive.ofMul.injective <| one_smul R (Additive.ofMul s) -@[simp] lemma one_z₂pow (x : ZMod 2) : (1 : ℤˣ) ^ x = 1 := one_pow _ +lemma uzpow_mul (s : ℤˣ) (x y : R) : s ^ (x * y) = (s ^ x) ^ y := + Additive.ofMul.injective <| mul_comm x y ▸ mul_smul y x (Additive.ofMul s) -lemma mul_z₂pow (s₁ s₂ : ℤˣ) (x : ZMod 2) : (s₁ * s₂) ^ x = s₁ ^ x * s₂ ^ x := mul_pow _ _ _ +lemma uzpow_add (s : ℤˣ) (x y : R) : s ^ (x + y) = s ^ x * s ^ y := + Additive.ofMul.injective <| add_smul x y (Additive.ofMul s) -@[simp] lemma z₂pow_zero (s : ℤˣ) : (s ^ (0 : ZMod 2) : ℤˣ) = (1 : ℤˣ) := pow_zero _ +end CommSemiring -@[simp] lemma z₂pow_one (s : ℤˣ) : (s ^ (1 : ZMod 2) : ℤˣ) = s := pow_one _ +section CommRing -lemma z₂pow_mul (s : ℤˣ) (x y : ZMod 2) : s ^ (x * y) = (s ^ x) ^ y := by - convert pow_mul s x.val y.val - rw [←z₂pow_natCast, Nat.cast_mul, ZMod.nat_cast_zmod_val, ZMod.nat_cast_zmod_val] +section CommRing +variable {R : Type*} [CommRing R] [Module R (Additive ℤˣ)] -lemma z₂pow_add (s : ℤˣ) (x y : ZMod 2) : s ^ (x + y) = s ^ x * s ^ y := by - simp only [z₂pow_def] - rw [ZMod.val_add, ←pow_eq_pow_mod, pow_add] - fin_cases s <;> simp +lemma uzpow_sub (s : ℤˣ) (x y : R) : s ^ (x - y) = s ^ x / s ^ y := + Additive.ofMul.injective <| sub_smul x y (Additive.ofMul s) -/-! The next two lemmas are mathematically not interesting as `-` coincides with `+` and `/` with -`*`, but they match the api for powers by `ℤ`. -/ +lemma uzpow_neg (s : ℤˣ) (x : R) : s ^ (-x) = (s ^ x)⁻¹ := + Additive.ofMul.injective <| neg_smul x (Additive.ofMul s) -lemma z₂pow_sub (s : ℤˣ) (x y : ZMod 2) : s ^ (x - y) = s ^ x / s ^ y := by - simp only [CharTwo.sub_eq_add, z₂pow_add, Int.units_div_eq_mul] +@[norm_cast] lemma uzpow_intCast (u : ℤˣ) (z : ℤ) : u ^ (z : R) = u ^ z := by + change Additive.toMul ((z : R) • Additive.ofMul u) = _ + rw [←zsmul_eq_smul_cast, toMul_zsmul, toMul_ofMul] -lemma z₂pow_neg (s : ℤˣ) (x : ZMod 2) : s ^ (-x) = (s ^ x)⁻¹ := by - simp only [CharTwo.neg_eq, Int.units_inv_eq_self] +end CommRing diff --git a/Mathlib/Data/ZMod/Quotient.lean b/Mathlib/Data/ZMod/Quotient.lean index d9aeffb0d814f..f428cab83f8de 100644 --- a/Mathlib/Data/ZMod/Quotient.lean +++ b/Mathlib/Data/ZMod/Quotient.lean @@ -3,11 +3,11 @@ Copyright (c) 2021 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ +import Mathlib.Algebra.Group.Equiv.TypeTags import Mathlib.Data.ZMod.Basic import Mathlib.GroupTheory.GroupAction.Quotient -import Mathlib.RingTheory.Int.Basic import Mathlib.RingTheory.Ideal.QuotientOperations -import Mathlib.Algebra.Hom.Equiv.TypeTags +import Mathlib.RingTheory.Int.Basic #align_import data.zmod.quotient from "leanprover-community/mathlib"@"da420a8c6dd5bdfb85c4ced85c34388f633bc6ff" diff --git a/Mathlib/Deprecated/Group.lean b/Mathlib/Deprecated/Group.lean index 385d7edb5bc73..c81c50d7b6fb0 100644 --- a/Mathlib/Deprecated/Group.lean +++ b/Mathlib/Deprecated/Group.lean @@ -3,10 +3,10 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.TypeTags -import Mathlib.Algebra.Hom.Equiv.Basic -import Mathlib.Algebra.Hom.Ring.Defs -import Mathlib.Algebra.Hom.Units +import Mathlib.Algebra.Group.Units.Hom +import Mathlib.Algebra.Ring.Hom.Defs #align_import deprecated.group from "leanprover-community/mathlib"@"5a3e819569b0f12cbec59d740a2613018e7b8eec" diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index af72e3fe717e1..80ef23e7bbb42 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import Mathlib.Algebra.Hom.Iterate +import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Analysis.SpecificLimits.Basic import Mathlib.Order.Iterate import Mathlib.Order.SemiconjSup diff --git a/Mathlib/Dynamics/PeriodicPts.lean b/Mathlib/Dynamics/PeriodicPts.lean index 6af69e9059498..121bb53b38947 100644 --- a/Mathlib/Dynamics/PeriodicPts.lean +++ b/Mathlib/Dynamics/PeriodicPts.lean @@ -3,10 +3,10 @@ Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import Mathlib.Algebra.Hom.Iterate +import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Data.List.Cycle -import Mathlib.Data.PNat.Basic import Mathlib.Data.Nat.Prime +import Mathlib.Data.PNat.Basic import Mathlib.Dynamics.FixedPoints.Basic import Mathlib.GroupTheory.GroupAction.Group diff --git a/Mathlib/FieldTheory/AbelRuffini.lean b/Mathlib/FieldTheory/AbelRuffini.lean index 266de2b7c4a16..72aa3447d9a58 100644 --- a/Mathlib/FieldTheory/AbelRuffini.lean +++ b/Mathlib/FieldTheory/AbelRuffini.lean @@ -30,7 +30,7 @@ noncomputable section local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 -open scoped Classical Polynomial +open scoped Classical Polynomial IntermediateField open Polynomial IntermediateField diff --git a/Mathlib/FieldTheory/AbsoluteGaloisGroup.lean b/Mathlib/FieldTheory/AbsoluteGaloisGroup.lean new file mode 100644 index 0000000000000..16992537a9e46 --- /dev/null +++ b/Mathlib/FieldTheory/AbsoluteGaloisGroup.lean @@ -0,0 +1,60 @@ +/- +Copyright (c) 2023 María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: María Inés de Frutos-Fernández +-/ +import Mathlib.FieldTheory.KrullTopology +import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure +import Mathlib.Topology.Algebra.Group.TopologicalAbelianization + +/-! +# The topological abelianization of the absolute Galois group. + +We define the absolute Galois group of a field `K` and its topological abelianization. + +## Main definitions +- `Field.absoluteGaloisGroup` : The Galois group of the field extension `K^al/K`, where `K^al` is an + algebraic closure of `K`. +- `Field.absoluteGaloisGroupAbelianization` : The topological abelianization of + `Field.absoluteGaloisGroup K`, that is, the quotient of `Field.absoluteGaloisGroup K` by the + topological closure of its commutator subgroup. + +## Main results +- `Field.absoluteGaloisGroup.commutator_closure_isNormal` : the topological closure of the + commutator of `absoluteGaloisGroup` is a normal subgroup. + +## Tags +field, algebraic closure, galois group, abelianization + +-/ + +namespace Field + +variable (K : Type*) [Field K] + +/-! ### The absolute Galois group -/ + +/-- The absolute Galois group of `K`, defined as the Galois group of the field extension `K^al/K`, + where `K^al` is an algebraic closure of `K`. -/ +def absoluteGaloisGroup := AlgebraicClosure K ≃ₐ[K] AlgebraicClosure K + +local notation "G_K" => absoluteGaloisGroup + +noncomputable instance : Group (G_K K) := AlgEquiv.aut + +/-- `absoluteGaloisGroup` is a topological space with the Krull topology. -/ +noncomputable instance : TopologicalSpace (G_K K) := krullTopology K (AlgebraicClosure K) + +/-! ### The topological abelianization of the absolute Galois group -/ + +instance absoluteGaloisGroup.commutator_closure_isNormal : + (commutator (G_K K)).topologicalClosure.Normal := + Subgroup.is_normal_topologicalClosure (commutator (G_K K)) + +/-- The topological abelianization of `absoluteGaloisGroup`, that is, the quotient of + `absoluteGaloisGroup` by the topological closure of its commutator subgroup. -/ +abbrev absoluteGaloisGroupAbelianization := TopologicalAbelianization (G_K K) + +local notation "G_K_ab" => absoluteGaloisGroupAbelianization + +end Field diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index d9b118c78b134..f662e7e5399f5 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -25,7 +25,7 @@ For example, `Algebra.adjoin K {x}` might not include `x⁻¹`. ## Notation - - `F⟮α⟯`: adjoin a single element `α` to `F`. + - `F⟮α⟯`: adjoin a single element `α` to `F` (in scope `IntermediateField`). -/ set_option autoImplicit true @@ -75,26 +75,28 @@ def gi : GaloisInsertion (adjoin F : Set E → IntermediateField F E) choice_eq _ _ := copy_eq _ _ _ #align intermediate_field.gi IntermediateField.gi -instance : CompleteLattice (IntermediateField F E) := - GaloisInsertion.liftCompleteLattice IntermediateField.gi +instance : CompleteLattice (IntermediateField F E) where + __ := GaloisInsertion.liftCompleteLattice IntermediateField.gi + bot := + { toSubalgebra := ⊥ + inv_mem' := fun x (hx : x ∈ (⊥ : Subalgebra F E)) => show x⁻¹ ∈ (⊥ : Subalgebra F E) by + rw [Algebra.mem_bot] at hx ⊢ + obtain ⟨r, rfl⟩ := hx + exact ⟨r⁻¹, map_inv₀ _ _⟩ } + bot_le x := (bot_le : ⊥ ≤ x.toSubalgebra) instance : Inhabited (IntermediateField F E) := ⟨⊤⟩ -theorem coe_bot : ↑(⊥ : IntermediateField F E) = Set.range (algebraMap F E) := by - change ↑(Subfield.closure (Set.range (algebraMap F E) ∪ ∅)) = Set.range (algebraMap F E) - rw [Set.union_empty, ← Set.image_univ, ← RingHom.map_field_closure] - simp +theorem coe_bot : ↑(⊥ : IntermediateField F E) = Set.range (algebraMap F E) := rfl #align intermediate_field.coe_bot IntermediateField.coe_bot theorem mem_bot {x : E} : x ∈ (⊥ : IntermediateField F E) ↔ x ∈ Set.range (algebraMap F E) := - Set.ext_iff.mp coe_bot x + Iff.rfl #align intermediate_field.mem_bot IntermediateField.mem_bot @[simp] -theorem bot_toSubalgebra : (⊥ : IntermediateField F E).toSubalgebra = ⊥ := by - ext - rw [mem_toSubalgebra, Algebra.mem_bot, mem_bot] +theorem bot_toSubalgebra : (⊥ : IntermediateField F E).toSubalgebra = ⊥ := rfl #align intermediate_field.bot_to_subalgebra IntermediateField.bot_toSubalgebra @[simp] @@ -257,10 +259,8 @@ def topEquiv : (⊤ : IntermediateField F E) ≃ₐ[F] E := @[simp] theorem restrictScalars_bot_eq_self (K : IntermediateField F E) : - (⊥ : IntermediateField K E).restrictScalars _ = K := by - ext x - rw [mem_restrictScalars, mem_bot]; - exact Set.ext_iff.mp Subtype.range_coe x + (⊥ : IntermediateField K E).restrictScalars _ = K := + SetLike.coe_injective Subtype.range_coe #align intermediate_field.restrict_scalars_bot_eq_self IntermediateField.restrictScalars_bot_eq_self @[simp] @@ -269,26 +269,31 @@ theorem restrictScalars_top {K : Type*} [Field K] [Algebra K E] [Algebra K F] rfl #align intermediate_field.restrict_scalars_top IntermediateField.restrictScalars_top -theorem AlgHom.fieldRange_eq_map {K : Type*} [Field K] [Algebra F K] (f : E →ₐ[F] K) : +@[simp] +theorem map_bot {K : Type*} [Field K] [Algebra F K] (f : E →ₐ[F] K) : + IntermediateField.map f ⊥ = ⊥ := + toSubalgebra_injective <| Algebra.map_bot _ + +theorem _root_.AlgHom.fieldRange_eq_map {K : Type*} [Field K] [Algebra F K] (f : E →ₐ[F] K) : f.fieldRange = IntermediateField.map f ⊤ := SetLike.ext' Set.image_univ.symm -#align alg_hom.field_range_eq_map IntermediateField.AlgHom.fieldRange_eq_map +#align alg_hom.field_range_eq_map AlgHom.fieldRange_eq_map -theorem AlgHom.map_fieldRange {K L : Type*} [Field K] [Field L] [Algebra F K] [Algebra F L] +theorem _root_.AlgHom.map_fieldRange {K L : Type*} [Field K] [Field L] [Algebra F K] [Algebra F L] (f : E →ₐ[F] K) (g : K →ₐ[F] L) : f.fieldRange.map g = (g.comp f).fieldRange := SetLike.ext' (Set.range_comp g f).symm -#align alg_hom.map_field_range IntermediateField.AlgHom.map_fieldRange +#align alg_hom.map_field_range AlgHom.map_fieldRange -theorem AlgHom.fieldRange_eq_top {K : Type*} [Field K] [Algebra F K] {f : E →ₐ[F] K} : +theorem _root_.AlgHom.fieldRange_eq_top {K : Type*} [Field K] [Algebra F K] {f : E →ₐ[F] K} : f.fieldRange = ⊤ ↔ Function.Surjective f := SetLike.ext'_iff.trans Set.range_iff_surjective -#align alg_hom.field_range_eq_top IntermediateField.AlgHom.fieldRange_eq_top +#align alg_hom.field_range_eq_top AlgHom.fieldRange_eq_top @[simp] -theorem AlgEquiv.fieldRange_eq_top {K : Type*} [Field K] [Algebra F K] (f : E ≃ₐ[F] K) : +theorem _root_.AlgEquiv.fieldRange_eq_top {K : Type*} [Field K] [Algebra F K] (f : E ≃ₐ[F] K) : (f : E →ₐ[F] K).fieldRange = ⊤ := AlgHom.fieldRange_eq_top.mpr f.surjective -#align alg_equiv.field_range_eq_top IntermediateField.AlgEquiv.fieldRange_eq_top +#align alg_equiv.field_range_eq_top AlgEquiv.fieldRange_eq_top end Lattice @@ -469,7 +474,7 @@ where /-- If `x₁ x₂ ... xₙ : E` then `F⟮x₁,x₂,...,xₙ⟯` is the `IntermediateField F E` generated by these elements. -/ -macro:max K:term "⟮" xs:term,* "⟯" : term => do ``(adjoin $K $(← mkInsertTerm xs.getElems)) +scoped macro:max K:term "⟮" xs:term,* "⟯" : term => do ``(adjoin $K $(← mkInsertTerm xs.getElems)) open Lean PrettyPrinter.Delaborator SubExpr in @[delab app.IntermediateField.adjoin] @@ -899,6 +904,12 @@ noncomputable def algHomAdjoinIntegralEquiv (h : IsIntegral F α) : rw [adjoin.powerBasis_gen, minpoly_gen, Equiv.refl_apply]) #align intermediate_field.alg_hom_adjoin_integral_equiv IntermediateField.algHomAdjoinIntegralEquiv +lemma algHomAdjoinIntegralEquiv_symm_apply_gen (h : IsIntegral F α) + (x : { x // x ∈ (minpoly F α).aroots K }) : + (algHomAdjoinIntegralEquiv F h).symm x (AdjoinSimple.gen F α) = x := + (adjoin.powerBasis h).lift_gen x.val <| by + rw [adjoin.powerBasis_gen, minpoly_gen]; exact (mem_aroots.mp x.2).2 + /-- Fintype of algebra homomorphism `F⟮α⟯ →ₐ[F] K` -/ noncomputable def fintypeOfAlgHomAdjoinIntegral (h : IsIntegral F α) : Fintype (F⟮α⟯ →ₐ[F] K) := PowerBasis.AlgHom.fintype (adjoin.powerBasis h) diff --git a/Mathlib/FieldTheory/Fixed.lean b/Mathlib/FieldTheory/Fixed.lean index e3a524d4418f3..fc0accdeaa811 100644 --- a/Mathlib/FieldTheory/Fixed.lean +++ b/Mathlib/FieldTheory/Fixed.lean @@ -45,7 +45,7 @@ variable (F : Type v) [Field F] [MulSemiringAction M F] [MulSemiringAction G F] /-- The subfield of F fixed by the field endomorphism `m`. -/ def FixedBy.subfield : Subfield F where - carrier := fixedBy M F m + carrier := fixedBy F m zero_mem' := smul_zero m add_mem' hx hy := (smul_add m _ _).trans <| congr_arg₂ _ hx hy neg_mem' hx := (smul_neg m _).trans <| congr_arg _ hx diff --git a/Mathlib/FieldTheory/Galois.lean b/Mathlib/FieldTheory/Galois.lean index 47022d81aaa23..834fd9ebd6c3b 100644 --- a/Mathlib/FieldTheory/Galois.lean +++ b/Mathlib/FieldTheory/Galois.lean @@ -35,7 +35,7 @@ Together, these two results prove the Galois correspondence. -/ -open scoped Polynomial +open scoped Polynomial IntermediateField open FiniteDimensional AlgEquiv diff --git a/Mathlib/FieldTheory/IntermediateField.lean b/Mathlib/FieldTheory/IntermediateField.lean index abcfccec06d2b..2ef4ac58bd1f5 100644 --- a/Mathlib/FieldTheory/IntermediateField.lean +++ b/Mathlib/FieldTheory/IntermediateField.lean @@ -403,13 +403,11 @@ instance isScalarTower_mid' : IsScalarTower K S L := /-- If `f : L →+* L'` fixes `K`, `S.map f` is the intermediate field between `L'` and `K` such that `x ∈ S ↔ f x ∈ S.map f`. -/ -def map (f : L →ₐ[K] L') (S : IntermediateField K L) : IntermediateField K L' := - { - S.toSubalgebra.map - f with - inv_mem' := by - rintro _ ⟨x, hx, rfl⟩ - exact ⟨x⁻¹, S.inv_mem hx, map_inv₀ f x⟩ } +def map (f : L →ₐ[K] L') (S : IntermediateField K L) : IntermediateField K L' where + __ := S.toSubalgebra.map f + inv_mem' := by + rintro _ ⟨x, hx, rfl⟩ + exact ⟨x⁻¹, S.inv_mem hx, map_inv₀ f x⟩ #align intermediate_field.map IntermediateField.map @[simp] @@ -417,6 +415,14 @@ theorem coe_map (f : L →ₐ[K] L') : (S.map f : Set L') = f '' S := rfl #align intermediate_field.coe_map IntermediateField.coe_map +@[simp] +theorem toSubalgebra_map (f : L →ₐ[K] L') : (S.map f).toSubalgebra = S.toSubalgebra.map f := + rfl + +@[simp] +theorem toSubfield_map (f : L →ₐ[K] L') : (S.map f).toSubfield = S.toSubfield.map f := + rfl + theorem map_map {K L₁ L₂ L₃ : Type*} [Field K] [Field L₁] [Algebra K L₁] [Field L₂] [Algebra K L₂] [Field L₃] [Algebra K L₃] (E : IntermediateField K L₁) (f : L₁ →ₐ[K] L₂) (g : L₂ →ₐ[K] L₃) : (E.map f).map g = E.map (g.comp f) := @@ -567,8 +573,9 @@ theorem coe_inclusion {E F : IntermediateField K L} (hEF : E ≤ F) (e : E) : variable {S} -theorem toSubalgebra_injective {S S' : IntermediateField K L} - (h : S.toSubalgebra = S'.toSubalgebra) : S = S' := by +theorem toSubalgebra_injective : + Function.Injective (IntermediateField.toSubalgebra : IntermediateField K L → _) := by + intro S S' h ext rw [← mem_toSubalgebra, ← mem_toSubalgebra, h] #align intermediate_field.to_subalgebra_injective IntermediateField.toSubalgebra_injective diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index 9a89a9ed7c7ac..929e15ad41b03 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -306,17 +306,6 @@ theorem maximalSubfieldWithHom_is_maximal : Classical.choose_spec (exists_maximal_subfieldWithHom K L M) #align lift.subfield_with_hom.maximal_subfield_with_hom_is_maximal IsAlgClosed.lift.SubfieldWithHom.maximalSubfieldWithHom_is_maximal --- Porting note: split out this definition from `maximalSubfieldWithHom_eq_top` -/-- Produce an algebra homomorphism `Adjoin R {x} →ₐ[R] T` sending `x` to -a root of `x`'s minimal polynomial in `T`. -/ -noncomputable def _root_.Algebra.adjoin.liftSingleton (R : Type*) [Field R] {S T : Type*} - [CommRing S] [CommRing T] [Algebra R S] [Algebra R T] - (x : S) (y : T) (h : aeval y (minpoly R x) = 0) : - Algebra.adjoin R {x} →ₐ[R] T := -AlgHom.comp - (AdjoinRoot.liftHom _ y h) - (AlgEquiv.adjoinSingletonEquivAdjoinRootMinpoly R x).toAlgHom - -- porting note: this was much faster in lean 3 set_option synthInstance.maxHeartbeats 200000 in theorem maximalSubfieldWithHom_eq_top : (maximalSubfieldWithHom K L M).carrier = ⊤ := by diff --git a/Mathlib/FieldTheory/Normal.lean b/Mathlib/FieldTheory/Normal.lean index 188bce0341ebd..fc3c07f0f58ba 100644 --- a/Mathlib/FieldTheory/Normal.lean +++ b/Mathlib/FieldTheory/Normal.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Thomas Browning, Patrick Lutz -/ import Mathlib.FieldTheory.Adjoin -import Mathlib.FieldTheory.Tower +import Mathlib.FieldTheory.SplittingField.Construction import Mathlib.GroupTheory.Solvable -import Mathlib.RingTheory.PowerBasis #align_import field_theory.normal from "leanprover-community/mathlib"@"9fb8964792b4237dac6200193a0d533f1b3f7423" @@ -70,10 +69,6 @@ instance normal_self : Normal F F := (minpoly.eq_X_sub_C' x).symm ▸ splits_X_sub_C _⟩ #align normal_self normal_self -variable {K} - -variable (K) - theorem Normal.exists_isSplittingField [h : Normal F K] [FiniteDimensional F K] : ∃ p : F[X], IsSplittingField F K p := by let s := Basis.ofVectorSpace F K @@ -109,134 +104,54 @@ theorem Normal.tower_top_of_normal [h : Normal F E] : Normal K E := #align normal.tower_top_of_normal Normal.tower_top_of_normal theorem AlgHom.normal_bijective [h : Normal F E] (ϕ : E →ₐ[F] K) : Function.Bijective ϕ := - ⟨ϕ.toRingHom.injective, fun x => by - letI : Algebra E K := ϕ.toRingHom.toAlgebra - obtain ⟨h1, h2⟩ := h.out (algebraMap K E x) - cases' - minpoly.mem_range_of_degree_eq_one E x - (h2.def.resolve_left (minpoly.ne_zero h1) - (minpoly.irreducible - (isIntegral_of_isScalarTower - ((isIntegral_algebraMap_iff (algebraMap K E).injective).mp h1))) - (minpoly.dvd E x - ((algebraMap K E).injective - (by - rw [RingHom.map_zero, aeval_map_algebraMap, ← aeval_algebraMap_apply] - exact minpoly.aeval F (algebraMap K E x))))) with - y hy - exact ⟨y, hy⟩⟩ + h.isAlgebraic'.bijective_of_isScalarTower' ϕ #align alg_hom.normal_bijective AlgHom.normal_bijective -- Porting note: `[Field F] [Field E] [Algebra F E]` added by hand. -variable {F} {E} {E' : Type*} [Field F] [Field E] [Algebra F E] [Field E'] [Algebra F E'] +variable {F E} {E' : Type*} [Field F] [Field E] [Algebra F E] [Field E'] [Algebra F E'] -theorem Normal.of_algEquiv [h : Normal F E] (f : E ≃ₐ[F] E') : Normal F E' := - normal_iff.2 fun x => by - cases' h.out (f.symm x) with hx hhx - have H := map_isIntegral f.toAlgHom hx - simp [AlgEquiv.toAlgHom_eq_coe] at H - use H - apply Polynomial.splits_of_splits_of_dvd (algebraMap F E') (minpoly.ne_zero hx) - · rw [← AlgHom.comp_algebraMap f.toAlgHom] - exact Polynomial.splits_comp_of_splits (algebraMap F E) f.toAlgHom.toRingHom hhx - · apply minpoly.dvd _ _ - rw [← AddEquiv.map_eq_zero_iff f.symm.toAddEquiv] - exact - Eq.trans (Polynomial.aeval_algHom_apply f.symm.toAlgHom x (minpoly F (f.symm x))).symm - (minpoly.aeval _ _) +theorem Normal.of_algEquiv [h : Normal F E] (f : E ≃ₐ[F] E') : Normal F E' := by + rw [normal_iff] at h ⊢ + intro x; specialize h (f.symm x) + rw [← f.apply_symm_apply x, minpoly.algEquiv_eq, ← f.toAlgHom.comp_algebraMap] + exact ⟨map_isIntegral f h.1, splits_comp_of_splits _ _ h.2⟩ #align normal.of_alg_equiv Normal.of_algEquiv theorem AlgEquiv.transfer_normal (f : E ≃ₐ[F] E') : Normal F E ↔ Normal F E' := - ⟨fun _ => Normal.of_algEquiv f, fun _ => Normal.of_algEquiv f.symm⟩ + ⟨fun _ ↦ Normal.of_algEquiv f, fun _ ↦ Normal.of_algEquiv f.symm⟩ #align alg_equiv.transfer_normal AlgEquiv.transfer_normal --- seems to be causing a diamond in the below proof --- however, this may be a fluke and the proof below uses non-canonical `Algebra` instances: --- when I replaced all the instances inside the proof with the "canonical" instances we have, --- I had the (unprovable) goal (of the form) `AdjoinRoot.mk f (C x) = AdjoinRoot.mk f X` --- for some `x, f`. So maybe this is indeed the correct approach and rewriting this proof is --- salient in the future, or at least taking a closer look at the algebra instances it uses. -attribute [-instance] AdjoinRoot.instSMulAdjoinRoot +open IntermediateField theorem Normal.of_isSplittingField (p : F[X]) [hFEp : IsSplittingField F E p] : Normal F E := by rcases eq_or_ne p 0 with (rfl | hp) · have := hFEp.adjoin_rootSet - simp only [rootSet_zero, Algebra.adjoin_empty] at this - exact - Normal.of_algEquiv - (AlgEquiv.ofBijective (Algebra.ofId F E) (Algebra.bijective_algebraMap_iff.2 this.symm)) - refine' normal_iff.2 fun x => _ - have hFE : FiniteDimensional F E := IsSplittingField.finiteDimensional E p - have Hx : IsIntegral F x := isIntegral_of_noetherian (IsNoetherian.iff_fg.2 hFE) x - refine' ⟨Hx, Or.inr _⟩ - rintro q q_irred ⟨r, hr⟩ - let D := AdjoinRoot q - haveI := Fact.mk q_irred - let pbED := AdjoinRoot.powerBasis q_irred.ne_zero - haveI : FiniteDimensional E D := PowerBasis.finiteDimensional pbED - have finrankED : FiniteDimensional.finrank E D = q.natDegree := by - rw [PowerBasis.finrank pbED, AdjoinRoot.powerBasis_dim] - haveI : FiniteDimensional F D := FiniteDimensional.trans F E D - rsuffices ⟨ϕ⟩ : Nonempty (D →ₐ[F] E) - --Porting note: the `change` was `rw [← WithBot.coe_one]` - · change degree q = ↑(1 : ℕ) - rw [degree_eq_iff_natDegree_eq q_irred.ne_zero, ← finrankED] - have nat_lemma : ∀ a b c : ℕ, a * b = c → c ≤ a → 0 < c → b = 1 := by - intro a b c h1 h2 h3 - nlinarith - exact - nat_lemma _ _ _ (FiniteDimensional.finrank_mul_finrank F E D) - (LinearMap.finrank_le_finrank_of_injective - (show Function.Injective ϕ.toLinearMap from ϕ.toRingHom.injective)) - FiniteDimensional.finrank_pos - let C := AdjoinRoot (minpoly F x) - haveI Hx_irred := Fact.mk (minpoly.irreducible Hx) --- Porting note: `heval` added since now Lean wants the proof explicitly in several places. - have heval : eval₂ (algebraMap F D) (AdjoinRoot.root q) (minpoly F x) = 0 := by - rw [algebraMap_eq F E D, ← eval₂_map, hr, AdjoinRoot.algebraMap_eq, eval₂_mul, - AdjoinRoot.eval₂_root, zero_mul] - letI : Algebra C D := - RingHom.toAlgebra (AdjoinRoot.lift (algebraMap F D) (AdjoinRoot.root q) heval) - letI : Algebra C E := RingHom.toAlgebra (AdjoinRoot.lift (algebraMap F E) x (minpoly.aeval F x)) - haveI : IsScalarTower F C D := of_algebraMap_eq fun y => (AdjoinRoot.lift_of heval).symm - haveI : IsScalarTower F C E := by - refine' of_algebraMap_eq fun y => (AdjoinRoot.lift_of _).symm --- Porting note: the following proof was just `_`. - rw [← aeval_def, minpoly.aeval] - suffices Nonempty (D →ₐ[C] E) by exact Nonempty.map (AlgHom.restrictScalars F) this - let S : Set D := ((p.aroots E).map (algebraMap E D)).toFinset - suffices ⊤ ≤ IntermediateField.adjoin C S by - refine' IntermediateField.algHom_mk_adjoin_splits' (top_le_iff.mp this) fun y hy => _ - rcases Multiset.mem_map.mp (Multiset.mem_toFinset.mp hy) with ⟨z, hz1, hz2⟩ - have Hz : IsIntegral F z := isIntegral_of_noetherian (IsNoetherian.iff_fg.2 hFE) z - use - show IsIntegral C y from - isIntegral_of_noetherian (IsNoetherian.iff_fg.2 (FiniteDimensional.right F C D)) y - apply splits_of_splits_of_dvd (algebraMap C E) (map_ne_zero (minpoly.ne_zero Hz)) - · rw [splits_map_iff, ← algebraMap_eq F C E] - exact - splits_of_splits_of_dvd _ hp hFEp.splits (minpoly.dvd F z (mem_aroots.mp hz1).2) - · apply minpoly.dvd - rw [← hz2, aeval_def, eval₂_map, ← algebraMap_eq F C D, algebraMap_eq F E D, ← hom_eval₂, ← - aeval_def, minpoly.aeval F z, RingHom.map_zero] - rw [← IntermediateField.toSubalgebra_le_toSubalgebra, IntermediateField.top_toSubalgebra] - apply ge_trans (IntermediateField.algebra_adjoin_le_adjoin C S) - suffices - (Algebra.adjoin C S).restrictScalars F = - (Algebra.adjoin E {AdjoinRoot.root q}).restrictScalars F by - rw [AdjoinRoot.adjoinRoot_eq_top, Subalgebra.restrictScalars_top, ← - @Subalgebra.restrictScalars_top F C] at this - exact top_le_iff.mpr (Subalgebra.restrictScalars_injective F this) -/- Porting note: the `change` was `dsimp only [S]`. Using `set S ... with hS` doesn't work. -/ - change Subalgebra.restrictScalars F (Algebra.adjoin C - (((p.aroots E).map (algebraMap E D)).toFinset : Set D)) = _ - rw [← Finset.image_toFinset, Finset.coe_image] - apply - Eq.trans - (Algebra.adjoin_res_eq_adjoin_res F E C D hFEp.adjoin_rootSet AdjoinRoot.adjoinRoot_eq_top) - rw [Set.image_singleton, RingHom.algebraMap_toAlgebra, AdjoinRoot.lift_root] + rw [rootSet_zero, Algebra.adjoin_empty] at this + exact Normal.of_algEquiv + (AlgEquiv.ofBijective (Algebra.ofId F E) (Algebra.bijective_algebraMap_iff.2 this.symm)) + refine normal_iff.mpr fun x ↦ ?_ + haveI : FiniteDimensional F E := IsSplittingField.finiteDimensional E p + have hx := isIntegral_of_finite F x + let L := (p * minpoly F x).SplittingField + have hL := splits_of_splits_mul' _ ?_ (SplittingField.splits (p * minpoly F x)) + · let j : E →ₐ[F] L := IsSplittingField.lift E p hL.1 + refine ⟨hx, splits_of_comp _ (j : E →+* L) (j.comp_algebraMap ▸ hL.2) fun a ha ↦ ?_⟩ + rw [j.comp_algebraMap] at ha + letI : Algebra F⟮x⟯ L := ((algHomAdjoinIntegralEquiv F hx).symm ⟨a, ha⟩).toRingHom.toAlgebra + let j' : E →ₐ[F⟮x⟯] L := IsSplittingField.lift E (p.map (algebraMap F F⟮x⟯)) ?_ + · change a ∈ j.range + rw [← IsSplittingField.adjoin_rootSet_eq_range E p j, + IsSplittingField.adjoin_rootSet_eq_range E p (j'.restrictScalars F)] + exact ⟨x, (j'.commutes _).trans (algHomAdjoinIntegralEquiv_symm_apply_gen F hx _)⟩ + · rw [splits_map_iff, ← IsScalarTower.algebraMap_eq]; exact hL.1 + · rw [Polynomial.map_ne_zero_iff (algebraMap F L).injective, mul_ne_zero_iff] + exact ⟨hp, minpoly.ne_zero hx⟩ #align normal.of_is_splitting_field Normal.of_isSplittingField +instance Polynomial.SplittingField.instNormal [Field F] (p : F[X]) : Normal F p.SplittingField := + Normal.of_isSplittingField p +#align polynomial.splitting_field.normal Polynomial.SplittingField.instNormal + end NormalTower namespace IntermediateField @@ -343,9 +258,9 @@ theorem AlgHom.fieldRange_of_normal [Algebra F K] {E : IntermediateField F K} [N -- Porting note: this was `IsScalarTower F E E := by infer_instance`. letI : Algebra E E := Algebra.id E let g := f.restrictNormal' E - rw [← show E.val.comp ↑g = f from FunLike.ext_iff.mpr (f.restrictNormal_commutes E), ← - IntermediateField.AlgHom.map_fieldRange, IntermediateField.AlgEquiv.fieldRange_eq_top g, - ← IntermediateField.AlgHom.fieldRange_eq_map, IntermediateField.fieldRange_val] + rw [← show E.val.comp ↑g = f from FunLike.ext_iff.mpr (f.restrictNormal_commutes E), + ← AlgHom.map_fieldRange, AlgEquiv.fieldRange_eq_top g, ← AlgHom.fieldRange_eq_map, + IntermediateField.fieldRange_val] #align alg_hom.field_range_of_normal AlgHom.fieldRange_of_normal /-- Restrict algebra isomorphism to a normal subfield -/ @@ -455,7 +370,23 @@ theorem AlgEquiv.restrictNormalHom_surjective [Normal F K₁] [Normal F E] : ⟨χ.liftNormal E, χ.restrict_liftNormal E⟩ #align alg_equiv.restrict_normal_hom_surjective AlgEquiv.restrictNormalHom_surjective -variable (F) (K₁) +open AdjoinRoot in +theorem Normal.minpoly_eq_iff_mem_orbit [h : Normal F E] {x y : E} : + minpoly F x = minpoly F y ↔ x ∈ MulAction.orbit (E ≃ₐ[F] E) y := by + refine ⟨fun he ↦ ?_, fun ⟨f, he⟩ ↦ he ▸ minpoly.algEquiv_eq f y⟩ + let Fx := AdjoinRoot (minpoly F x) + have hx : aeval x (minpoly F x) = 0 := minpoly.aeval F x + have hy : aeval y (minpoly F x) = 0 := he ▸ minpoly.aeval F y + let Ax : Algebra Fx E := (lift (algebraMap F E) x hx).toAlgebra + have Tx : IsScalarTower F Fx E := IsScalarTower.of_ring_hom (liftHom _ x hx) + let Ay : Algebra Fx E := (lift (algebraMap F E) y hy).toAlgebra + have Ty : IsScalarTower F Fx E := IsScalarTower.of_ring_hom (liftHom _ y hy) + haveI : Fact (Irreducible <| minpoly F x) := ⟨minpoly.irreducible <| h.isIntegral x⟩ + let f : E ≃ₐ[F] E := @AlgEquiv.liftNormal F Fx Fx _ _ _ _ _ AlgEquiv.refl E _ _ Ay Ax Ty Tx _ + refine ⟨f, (congr_arg f (lift_root hy).symm).trans <| Eq.trans ?_ (lift_root hx)⟩ + exact @AlgEquiv.liftNormal_commutes F Fx Fx _ _ _ _ _ _ E _ _ Ay Ax Ty Tx _ (root _) + +variable (F K₁) theorem isSolvable_of_isScalarTower [Normal F K₁] [h1 : IsSolvable (K₁ ≃ₐ[F] K₁)] [h2 : IsSolvable (E ≃ₐ[K₁] E)] : IsSolvable (E ≃ₐ[F] E) := by diff --git a/Mathlib/FieldTheory/PolynomialGaloisGroup.lean b/Mathlib/FieldTheory/PolynomialGaloisGroup.lean index a4af395878bc8..8e16630de2371 100644 --- a/Mathlib/FieldTheory/PolynomialGaloisGroup.lean +++ b/Mathlib/FieldTheory/PolynomialGaloisGroup.lean @@ -395,6 +395,8 @@ theorem restrictComp_surjective (hq : q.natDegree ≠ 0) : variable {p q} +open scoped IntermediateField + /-- For a separable polynomial, its Galois group has cardinality equal to the dimension of its splitting field over `F`. -/ theorem card_of_separable (hp : p.Separable) : Fintype.card p.Gal = finrank F p.SplittingField := diff --git a/Mathlib/FieldTheory/RatFunc.lean b/Mathlib/FieldTheory/RatFunc.lean index 2acaf122ac270..ca5fe82a13083 100644 --- a/Mathlib/FieldTheory/RatFunc.lean +++ b/Mathlib/FieldTheory/RatFunc.lean @@ -689,6 +689,7 @@ theorem coe_mapRingHom_eq_coe_map [RingHomClass F R[X] S[X]] (φ : F) (hφ : R[X rfl #align ratfunc.coe_map_ring_hom_eq_coe_map RatFunc.coe_mapRingHom_eq_coe_map +set_option maxHeartbeats 300000 in -- TODO: Generalize to `FunLike` classes, /-- Lift a monoid with zero homomorphism `R[X] →*₀ G₀` to a `RatFunc R →*₀ G₀` on the condition that `φ` maps non zero divisors to non zero divisors, @@ -703,7 +704,7 @@ def liftMonoidWithZeroHom (φ : R[X] →*₀ G₀) (hφ : R[X]⁰ ≤ G₀⁰.co map_one' := by dsimp only -- porting note: force the function to be applied rw [← ofFractionRing_one, ← Localization.mk_one, liftOn_ofFractionRing_mk] - simp only [map_one, Submonoid.coe_one, div_one] + simp only [map_one, OneMemClass.coe_one, div_one] map_mul' x y := by cases' x with x cases' y with y diff --git a/Mathlib/FieldTheory/SplittingField/Construction.lean b/Mathlib/FieldTheory/SplittingField/Construction.lean index 72269157add18..0a6b190adcb29 100644 --- a/Mathlib/FieldTheory/SplittingField/Construction.lean +++ b/Mathlib/FieldTheory/SplittingField/Construction.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.FieldTheory.Normal +import Mathlib.FieldTheory.SplittingField.IsSplittingField #align_import field_theory.splitting_field.construction from "leanprover-community/mathlib"@"e3f4be1fcb5376c4948d7f095bec45350bfb9d1a" @@ -359,33 +359,12 @@ instance (f : K[X]) : NoZeroSMulDivisors K f.SplittingField := inferInstance /-- Any splitting field is isomorphic to `SplittingFieldAux f`. -/ -def algEquiv (f : K[X]) [IsSplittingField K L f] : L ≃ₐ[K] SplittingField f := by - refine' - AlgEquiv.ofBijective (lift L f <| splits (SplittingField f) f) - ⟨RingHom.injective (lift L f <| splits (SplittingField f) f).toRingHom, _⟩ - haveI := finiteDimensional (SplittingField f) f - haveI := finiteDimensional L f - have : FiniteDimensional.finrank K L = FiniteDimensional.finrank K (SplittingField f) := - le_antisymm - (LinearMap.finrank_le_finrank_of_injective - (show Function.Injective (lift L f <| splits (SplittingField f) f).toLinearMap from - RingHom.injective (lift L f <| splits (SplittingField f) f : L →+* f.SplittingField))) - (LinearMap.finrank_le_finrank_of_injective - (show Function.Injective (lift (SplittingField f) f <| splits L f).toLinearMap from - RingHom.injective (lift (SplittingField f) f <| splits L f : f.SplittingField →+* L))) - change Function.Surjective (lift L f <| splits (SplittingField f) f).toLinearMap - refine' (LinearMap.injective_iff_surjective_of_finrank_eq_finrank this).1 _ - exact RingHom.injective (lift L f <| splits (SplittingField f) f : L →+* f.SplittingField) +def algEquiv (f : K[X]) [h : IsSplittingField K L f] : L ≃ₐ[K] SplittingField f := + AlgEquiv.ofBijective (lift L f <| splits (SplittingField f) f) <| + have := finiteDimensional L f + ((Algebra.isAlgebraic_of_finite K L).algHom_bijective₂ _ <| lift _ f h.1).1 #align polynomial.is_splitting_field.alg_equiv Polynomial.IsSplittingField.algEquiv end IsSplittingField end Polynomial - -section Normal - -instance Polynomial.SplittingField.instNormal [Field F] (p : F[X]) : Normal F p.SplittingField := - Normal.of_isSplittingField p -#align polynomial.splitting_field.normal Polynomial.SplittingField.instNormal - -end Normal diff --git a/Mathlib/FieldTheory/SplittingField/IsSplittingField.lean b/Mathlib/FieldTheory/SplittingField/IsSplittingField.lean index 5c53ddde43eda..d1b373571a7ca 100644 --- a/Mathlib/FieldTheory/SplittingField/IsSplittingField.lean +++ b/Mathlib/FieldTheory/SplittingField/IsSplittingField.lean @@ -79,7 +79,7 @@ instance map (f : F[X]) [IsSplittingField F L f] : IsSplittingField K L (f.map < #align polynomial.is_splitting_field.map Polynomial.IsSplittingField.map theorem splits_iff (f : K[X]) [IsSplittingField K L f] : - Polynomial.Splits (RingHom.id K) f ↔ (⊤ : Subalgebra K L) = ⊥ := + Splits (RingHom.id K) f ↔ (⊤ : Subalgebra K L) = ⊥ := ⟨fun h => by -- Porting note: replaced term-mode proof rw [eq_bot_iff, ← adjoin_rootSet L f, rootSet, aroots, roots_map (algebraMap K L) h, Algebra.adjoin_le_iff] @@ -112,7 +112,7 @@ end ScalarTower /-- Splitting field of `f` embeds into any field that splits `f`. -/ def lift [Algebra K F] (f : K[X]) [IsSplittingField K L f] - (hf : Polynomial.Splits (algebraMap K F) f) : L →ₐ[K] F := + (hf : Splits (algebraMap K F) f) : L →ₐ[K] F := if hf0 : f = 0 then (Algebra.ofId K F).comp <| (Algebra.botEquiv K L : (⊥ : Subalgebra K L) →ₐ[K] K).comp <| by @@ -129,11 +129,9 @@ def lift [Algebra K F] (f : K[X]) [IsSplittingField K L f] theorem finiteDimensional (f : K[X]) [IsSplittingField K L f] : FiniteDimensional K L := ⟨@Algebra.top_toSubmodule K L _ _ _ ▸ - adjoin_rootSet L f ▸ FG_adjoin_of_finite (Finset.finite_toSet _) fun y hy => + adjoin_rootSet L f ▸ FG_adjoin_of_finite (Finset.finite_toSet _) fun y hy ↦ if hf : f = 0 then by rw [hf, rootSet_zero] at hy; cases hy - else - isAlgebraic_iff_isIntegral.1 ⟨f, hf, (eval₂_eq_eval_map _).trans <| - (mem_roots <| map_ne_zero hf).1 (Multiset.mem_toFinset.mp hy)⟩⟩ + else isAlgebraic_iff_isIntegral.1 ⟨f, hf, (mem_rootSet'.mp hy).2⟩⟩ #align polynomial.is_splitting_field.finite_dimensional Polynomial.IsSplittingField.finiteDimensional theorem of_algEquiv [Algebra K F] (p : K[X]) (f : F ≃ₐ[K] L) [IsSplittingField K F p] : @@ -145,6 +143,10 @@ theorem of_algEquiv [Algebra K F] (p : K[X]) (f : F ≃ₐ[K] L) [IsSplittingFie adjoin_rootSet_eq_range (splits F p), adjoin_rootSet F p] #align polynomial.is_splitting_field.of_alg_equiv Polynomial.IsSplittingField.of_algEquiv +theorem adjoin_rootSet_eq_range [Algebra K F] (f : K[X]) [IsSplittingField K L f] (i : L →ₐ[K] F) : + Algebra.adjoin K (rootSet f F) = i.range := + (Polynomial.adjoin_rootSet_eq_range (splits L f) i).mpr (adjoin_rootSet L f) + end IsSplittingField end Polynomial @@ -157,15 +159,8 @@ variable {K L} [Field K] [Field L] [Algebra K L] {p : K[X]} theorem splits_of_splits {F : IntermediateField K L} (h : p.Splits (algebraMap K L)) (hF : ∀ x ∈ p.rootSet L, x ∈ F) : p.Splits (algebraMap K F) := by - simp_rw [rootSet_def, Finset.mem_coe, Multiset.mem_toFinset] at hF - rw [splits_iff_exists_multiset] - refine' ⟨Multiset.pmap Subtype.mk _ hF, map_injective _ (algebraMap F L).injective _⟩ - conv_lhs => - rw [Polynomial.map_map, ← IsScalarTower.algebraMap_eq, eq_prod_roots_of_splits h, ← - Multiset.pmap_eq_map _ _ _ hF] - simp_rw [Polynomial.map_mul, Polynomial.map_multiset_prod, Multiset.map_pmap, Polynomial.map_sub, - map_C, map_X] - rfl + simp_rw [← F.fieldRange_val, rootSet_def, Finset.mem_coe, Multiset.mem_toFinset] at hF + exact splits_of_comp _ F.val.toRingHom h hF #align intermediate_field.splits_of_splits IntermediateField.splits_of_splits end IntermediateField diff --git a/Mathlib/FieldTheory/Subfield.lean b/Mathlib/FieldTheory/Subfield.lean index b1a13f486104d..19129fd415460 100644 --- a/Mathlib/FieldTheory/Subfield.lean +++ b/Mathlib/FieldTheory/Subfield.lean @@ -418,7 +418,7 @@ instance toAlgebra : Algebra s K := #align subfield.to_algebra Subfield.toAlgebra @[simp] -theorem coe_subtype : ⇑(s.subtype) = ((↑) : s → K) := +theorem coe_subtype : ⇑(s.subtype) = ((↑) : s → K) := rfl #align subfield.coe_subtype Subfield.coe_subtype diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index 9432a3a393167..ed39ed20f5d5f 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -474,16 +474,18 @@ theorem altitude_eq_mongePlane (t : Triangle ℝ P) {i₁ i₂ i₃ : Fin 3} (h (h₂₃ : i₂ ≠ i₃) : t.altitude i₁ = t.mongePlane i₂ i₃ := by have hs : ({i₂, i₃}ᶜ : Finset (Fin 3)) = {i₁} := by -- porting note: was `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> simp at h₁₂ h₁₃ h₂₃ ⊢ + fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ + <;> simp (config := {decide := true}) at h₁₂ h₁₃ h₂₃ ⊢ have he : univ.erase i₁ = {i₂, i₃} := by -- porting note: was `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> simp at h₁₂ h₁₃ h₂₃ ⊢ + fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ + <;> simp (config := {decide := true}) at h₁₂ h₁₃ h₂₃ ⊢ rw [mongePlane_def, altitude_def, direction_affineSpan, hs, he, centroid_singleton, coe_insert, coe_singleton, vectorSpan_image_eq_span_vsub_set_left_ne ℝ _ (Set.mem_insert i₂ _)] simp [h₂₃, Submodule.span_insert_eq_span] -- porting note: this didn't need the `congr` and the `fin_cases` congr - fin_cases i₂ <;> fin_cases i₃ <;> simp at h₂₃ ⊢ + fin_cases i₂ <;> fin_cases i₃ <;> simp (config := {decide := true}) at h₂₃ ⊢ #align affine.triangle.altitude_eq_monge_plane Affine.Triangle.altitude_eq_mongePlane /-- The orthocenter lies in the altitudes. -/ @@ -589,7 +591,8 @@ theorem altitude_replace_orthocenter_eq_affineSpan {t₁ t₂ : Triangle ℝ P} · have hu : (Finset.univ : Finset (Fin 3)) = {j₁, j₂, j₃} := by clear h₁ h₂ h₃ -- porting note: was `decide!` - fin_cases j₁ <;> fin_cases j₂ <;> fin_cases j₃ <;> simp at hj₁₂ hj₁₃ hj₂₃ ⊢ + fin_cases j₁ <;> fin_cases j₂ <;> fin_cases j₃ + <;> simp (config := {decide := true}) at hj₁₂ hj₁₃ hj₂₃ ⊢ rw [← Set.image_univ, ← Finset.coe_univ, hu, Finset.coe_insert, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_insert_eq, Set.image_singleton, h₁, h₂, h₃, Set.insert_subset_iff, Set.insert_subset_iff, Set.singleton_subset_iff] @@ -604,7 +607,8 @@ theorem altitude_replace_orthocenter_eq_affineSpan {t₁ t₂ : Triangle ℝ P} have hu : Finset.univ.erase j₂ = {j₁, j₃} := by clear h₁ h₂ h₃ -- porting note: was `decide!` - fin_cases j₁ <;> fin_cases j₂ <;> fin_cases j₃ <;> simp at hj₁₂ hj₁₃ hj₂₃ ⊢ + fin_cases j₁ <;> fin_cases j₂ <;> fin_cases j₃ + <;> simp (config := {decide := true}) at hj₁₂ hj₁₃ hj₂₃ ⊢ rw [hu, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_singleton, h₁, h₃] have hle : (t₁.altitude i₃).directionᗮ ≤ line[ℝ, t₁.orthocenter, t₁.points i₃].directionᗮ := Submodule.orthogonal_le (direction_le (affineSpan_orthocenter_point_le_altitude _ _)) @@ -612,7 +616,8 @@ theorem altitude_replace_orthocenter_eq_affineSpan {t₁ t₂ : Triangle ℝ P} have hui : Finset.univ.erase i₃ = {i₁, i₂} := by clear hle h₂ h₃ -- porting note: was `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> simp at hi₁₂ hi₁₃ hi₂₃ ⊢ + fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ + <;> simp (config := {decide := true}) at hi₁₂ hi₁₃ hi₂₃ ⊢ rw [hui, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_singleton] refine' vsub_mem_vectorSpan ℝ (Set.mem_insert _ _) (Set.mem_insert_of_mem _ (Set.mem_singleton _)) #align affine.triangle.altitude_replace_orthocenter_eq_affine_span Affine.Triangle.altitude_replace_orthocenter_eq_affineSpan @@ -783,7 +788,7 @@ theorem OrthocentricSystem.eq_insert_orthocenter {s : Set P} (ho : OrthocentricS ∃ j₁ : Fin 3, j₁ ≠ j₂ ∧ j₁ ≠ j₃ ∧ ∀ j : Fin 3, j = j₁ ∨ j = j₂ ∨ j = j₃ := by clear h₂ h₃ -- porting note: was `decide!` - fin_cases j₂ <;> fin_cases j₃ <;> simp at hj₂₃ ⊢ + fin_cases j₂ <;> fin_cases j₃ <;> simp (config := {decide := true}) at hj₂₃ ⊢ suffices h : t₀.points j₁ = t.orthocenter · have hui : (Set.univ : Set (Fin 3)) = {i₁, i₂, i₃} := by ext x; simpa using h₁₂₃ x have huj : (Set.univ : Set (Fin 3)) = {j₁, j₂, j₃} := by ext x; simpa using hj₁₂₃ x diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 766d53a1145ee..a1de6c986ee61 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -315,7 +315,7 @@ def idGroupoid (H : Type u) [TopologicalSpace H] : StructureGroupoid H where cases' he with he he · left have : e = e' := by - refine' eq_of_eq_on_source_univ (Setoid.symm he'e) _ _ <;> + refine' eq_of_eqOnSource_univ (Setoid.symm he'e) _ _ <;> rw [Set.mem_singleton_iff.1 he] <;> rfl rwa [← this] · right diff --git a/Mathlib/Geometry/Manifold/Metrizable.lean b/Mathlib/Geometry/Manifold/Metrizable.lean index 5c430c2a13927..81177974bc838 100644 --- a/Mathlib/Geometry/Manifold/Metrizable.lean +++ b/Mathlib/Geometry/Manifold/Metrizable.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners import Mathlib.Topology.Compactness.Paracompact -import Mathlib.Topology.MetricSpace.Metrizable +import Mathlib.Topology.Metrizable.Urysohn #align_import geometry.manifold.metrizable from "leanprover-community/mathlib"@"d1bd9c5df2867c1cb463bc6364446d57bdd9f7f1" diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 9106b44118f28..9fa6e629b9514 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -425,7 +425,7 @@ def ModelWithCorners.pi {𝕜 : Type u} [NontriviallyNormedField 𝕜] {ι : Typ [∀ i, TopologicalSpace (H i)] (I : ∀ i, ModelWithCorners 𝕜 (E i) (H i)) : ModelWithCorners 𝕜 (∀ i, E i) (ModelPi H) where toLocalEquiv := LocalEquiv.pi fun i => (I i).toLocalEquiv - source_eq := by simp only [Set.pi_univ, mfld_simps] + source_eq := by simp only [pi_univ, mfld_simps] unique_diff' := UniqueDiffOn.pi ι E _ _ fun i _ => (I i).unique_diff' continuous_toFun := continuous_pi fun i => (I i).continuous.comp (continuous_apply i) continuous_invFun := continuous_pi fun i => (I i).continuous_symm.comp (continuous_apply i) @@ -631,11 +631,11 @@ theorem contDiffGroupoid_prod {I : ModelWithCorners 𝕜 E H} {I' : ModelWithCor simp only at he he_symm he' he'_symm constructor <;> simp only [LocalEquiv.prod_source, LocalHomeomorph.prod_toLocalEquiv] · have h3 := ContDiffOn.prod_map he he' - rw [← I.image_eq, ← I'.image_eq, Set.prod_image_image_eq] at h3 + rw [← I.image_eq, ← I'.image_eq, prod_image_image_eq] at h3 rw [← (I.prod I').image_eq] exact h3 · have h3 := ContDiffOn.prod_map he_symm he'_symm - rw [← I.image_eq, ← I'.image_eq, Set.prod_image_image_eq] at h3 + rw [← I.image_eq, ← I'.image_eq, prod_image_image_eq] at h3 rw [← (I.prod I').image_eq] exact h3 #align cont_diff_groupoid_prod contDiffGroupoid_prod @@ -718,7 +718,7 @@ def analyticGroupoid : StructureGroupoid H := apply And.intro · intro x hx rcases h (I.symm x) (mem_preimage.mp hx.left) with ⟨v, hv⟩ - exact hv.right.right.left x ⟨Set.mem_preimage.mpr ⟨hx.left, hv.right.left⟩, hx.right⟩ + exact hv.right.right.left x ⟨mem_preimage.mpr ⟨hx.left, hv.right.left⟩, hx.right⟩ · apply mapsTo'.mp simp only [MapsTo] intro x hx @@ -1563,3 +1563,21 @@ theorem writtenInExtChartAt_chartAt_symm_comp [ChartedSpace H H'] (x : M') {y} simp_all only [mfld_simps, chartAt_comp] end ExtendedCharts + +section Topology +-- Let `M` be a topological manifold over the field 𝕜. +variable + {E : Type*} {𝕜 : Type*} [NontriviallyNormedField 𝕜] + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] + (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + [HasGroupoid M (contDiffGroupoid 0 I)] + +/-- A finite-dimensional manifold modelled on a locally compact field + (such as ℝ, ℂ or the `p`-adic numbers) is locally compact. -/ +lemma Manifold.locallyCompact_of_finiteDimensional [LocallyCompactSpace 𝕜] + [FiniteDimensional 𝕜 E] : LocallyCompactSpace M := by + have : ProperSpace E := FiniteDimensional.proper 𝕜 E + have : LocallyCompactSpace H := I.locallyCompactSpace + exact ChartedSpace.locallyCompactSpace H M + +end Topology diff --git a/Mathlib/GroupTheory/CommutingProbability.lean b/Mathlib/GroupTheory/CommutingProbability.lean index 26b9363e0ad6d..11b4a6ef32cc0 100644 --- a/Mathlib/GroupTheory/CommutingProbability.lean +++ b/Mathlib/GroupTheory/CommutingProbability.lean @@ -209,13 +209,13 @@ theorem commProb_reciprocal (n : ℕ) : rcases Nat.even_or_odd n with h2 | h2 · have := div_two_lt h0 rw [reciprocalFactors_even h0 h2, commProb_cons, commProb_reciprocal (n / 2), - commProb_odd (by norm_num)] + commProb_odd (by decide)] field_simp [h0, h2.two_dvd] norm_num · have := div_four_lt h0 h1 rw [reciprocalFactors_odd h1 h2, commProb_cons, commProb_reciprocal (n / 4 + 1)] have key : n % 4 = 1 ∨ n % 4 = 3 := Nat.odd_mod_four_iff.mp (Nat.odd_iff.mp h2) - have hn : Odd (n % 4) := by rcases key with h | h <;> rw [h] <;> norm_num + have hn : Odd (n % 4) := by rcases key with h | h <;> rw [h] <;> decide rw [commProb_odd (hn.mul h2), div_mul_div_comm, mul_one, div_eq_div_iff, one_mul] <;> norm_cast · have h0 : (n % 4) ^ 2 + 3 = n % 4 * 4 := by rcases key with h | h <;> rw [h] <;> norm_num have h1 := (Nat.div_add_mod n 4).symm diff --git a/Mathlib/GroupTheory/Congruence.lean b/Mathlib/GroupTheory/Congruence.lean index d3a73916d2ec1..6a6b88ae91f20 100644 --- a/Mathlib/GroupTheory/Congruence.lean +++ b/Mathlib/GroupTheory/Congruence.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston -/ import Mathlib.Algebra.Group.Prod -import Mathlib.Algebra.Hom.Equiv.Basic +import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Data.Setoid.Basic import Mathlib.GroupTheory.Submonoid.Operations diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index fdf3464e1b944..89b835c227d00 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -991,21 +991,22 @@ theorem lift_injective_of_ping_pong : Function.Injective (lift f) := by end PingPongLemma -/-- The free product of free groups is itself a free group -/ -@[simps!] --Porting note: added `!` -instance {ι : Type*} (G : ι → Type*) [∀ i, Group (G i)] [hG : ∀ i, IsFreeGroup (G i)] : - IsFreeGroup (CoprodI G) where - Generators := Σi, IsFreeGroup.Generators (G i) - MulEquiv' := - MonoidHom.toMulEquiv - (FreeGroup.lift fun x : Σi, IsFreeGroup.Generators (G i) => - CoprodI.of (IsFreeGroup.of x.2 : G x.1)) - (CoprodI.lift fun i : ι => - (IsFreeGroup.lift fun x : IsFreeGroup.Generators (G i) => - FreeGroup.of (⟨i, x⟩ : Σi, IsFreeGroup.Generators (G i)) : - G i →* FreeGroup (Σi, IsFreeGroup.Generators (G i)))) - (by ext; simp) - (by ext; simp) +/-- Given a family of free groups with distinguished bases, then their free product is free, with +a basis given by the union of the bases of the components. -/ +def FreeGroupBasis.coprodI {ι : Type*} {X : ι → Type*} {G : ι → Type*} [∀ i, Group (G i)] + (B : ∀ i, FreeGroupBasis (X i) (G i)) : + FreeGroupBasis (Σ i, X i) (CoprodI G) := + ⟨MulEquiv.symm <| MonoidHom.toMulEquiv + (FreeGroup.lift fun x : Σ i, X i => CoprodI.of (B x.1 x.2)) + (CoprodI.lift fun i : ι => (B i).lift fun x : X i => + FreeGroup.of (⟨i, x⟩ : Σ i, X i)) + (by ext; simp) + (by ext1 i; apply (B i).ext_hom; simp)⟩ + +/-- The free product of free groups is itself a free group. -/ +instance {ι : Type*} (G : ι → Type*) [∀ i, Group (G i)] [∀ i, IsFreeGroup (G i)] : + IsFreeGroup (CoprodI G) := + (FreeGroupBasis.coprodI (fun i ↦ IsFreeGroup.basis (G i))).isFreeGroup -- NB: One might expect this theorem to be phrased with ℤ, but ℤ is an additive group, -- and using `Multiplicative ℤ` runs into diamond issues. diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index a5bd4a7e4dfb2..988880b6a5f87 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -9,6 +9,7 @@ import Mathlib.GroupTheory.OrderOfElement import Mathlib.Algebra.GCDMonoid.Finset import Mathlib.Data.Nat.Factorization.Basic import Mathlib.Tactic.ByContra +import Mathlib.Tactic.Peel #align_import group_theory.exponent from "leanprover-community/mathlib"@"52fa514ec337dd970d71d8de8d0fd68b455a1e54" @@ -34,6 +35,11 @@ it is equal to the lowest common multiple of the order of all elements of the gr `Finset.lcm` of the order of its elements. * `Monoid.exponent_eq_iSup_orderOf(')`: For a commutative cancel monoid, the exponent is equal to `⨆ g : G, orderOf g` (or zero if it has any order-zero elements). +* `Monoid.exponent_pi` and `Monoid.exponent_prod`: The exponent of a finite product of monoids is + the least common multiple (`Finset.lcm` and `lcm`, respectively) of the exponents of the + constituent monoids. +* `MonoidHom.exponent_dvd`: If `f : M₁ →⋆ M₂` is surjective, then the exponent of `M₂` divides the + exponent of `M₁`. ## TODO * Refactor the characteristic of a ring to be the exponent of its underlying additive group. @@ -74,6 +80,21 @@ noncomputable def exponent := variable {G} +@[simp] +theorem _root_.AddMonoid.exponent_additive : + AddMonoid.exponent (Additive G) = exponent G := rfl + +@[simp] +theorem exponent_multiplicative {G : Type*} [AddMonoid G] : + exponent (Multiplicative G) = AddMonoid.exponent G := rfl + +open MulOpposite in +@[to_additive (attr := simp)] +theorem _root_.MulOpposite.exponent : exponent (MulOpposite G) = exponent G := by + simp only [Monoid.exponent, ExponentExists] + congr! + all_goals exact ⟨(op_injective <| · <| op ·), (unop_injective <| . <| unop .)⟩ + @[to_additive] theorem exponentExists_iff_ne_zero : ExponentExists G ↔ exponent G ≠ 0 := by rw [exponent] @@ -386,3 +407,58 @@ theorem card_dvd_exponent_pow_rank' {n : ℕ} (hG : ∀ g : G, g ^ n = 1) : #align card_dvd_exponent_nsmul_rank' card_dvd_exponent_nsmul_rank' end CommGroup + +section PiProd + +open Finset Monoid + +@[to_additive] +theorem Monoid.exponent_pi_eq_zero {ι : Type*} {M : ι → Type*} [∀ i, Monoid (M i)] {j : ι} + (hj : exponent (M j) = 0) : exponent ((i : ι) → M i) = 0 := by + rw [@exponent_eq_zero_iff, ExponentExists] at hj ⊢ + push_neg at hj ⊢ + peel hj with _ n hn + obtain ⟨m, hm⟩ := this + refine ⟨Pi.mulSingle j m, fun h ↦ hm ?_⟩ + simpa using congr_fun h j + +/-- If `f : M₁ →⋆ M₂` is surjective, then the exponent of `M₂` divides the exponent of `M₁`. -/ +@[to_additive] +theorem MonoidHom.exponent_dvd {F M₁ M₂ : Type*} [Monoid M₁] [Monoid M₂] [MonoidHomClass F M₁ M₂] + {f : F} (hf : Function.Surjective f) : exponent M₂ ∣ exponent M₁ := by + refine Monoid.exponent_dvd_of_forall_pow_eq_one M₂ _ fun m₂ ↦ ?_ + obtain ⟨m₁, rfl⟩ := hf m₂ + rw [←map_pow, pow_exponent_eq_one, map_one] + +/-- The exponent of finite product of monoids is the `Finset.lcm` of the exponents of the +constituent monoids. -/ +@[to_additive "The exponent of finite product of additive monoids is the `Finset.lcm` of the +exponents of the constituent additive monoids."] +theorem Monoid.exponent_pi {ι : Type*} [Fintype ι] {M : ι → Type*} [∀ i, Monoid (M i)] : + exponent ((i : ι) → M i) = lcm univ (exponent <| M ·) := by + refine dvd_antisymm ?_ ?_ + · refine exponent_dvd_of_forall_pow_eq_one _ _ fun m ↦ ?_ + ext i + rw [Pi.pow_apply, Pi.one_apply, ← orderOf_dvd_iff_pow_eq_one] + apply dvd_trans (Monoid.order_dvd_exponent (m i)) + exact Finset.dvd_lcm (mem_univ i) + · apply Finset.lcm_dvd fun i _ ↦ ?_ + exact MonoidHom.exponent_dvd (f := Pi.evalMonoidHom (M ·) i) (Function.surjective_eval i) + +/-- The exponent of product of two monoids is the `lcm` of the exponents of the +individuaul monoids. -/ +@[to_additive AddMonoid.exponent_prod "The exponent of product of two additive monoids is the `lcm` +of the exponents of the individuaul additive monoids."] +theorem Monoid.exponent_prod {M₁ M₂ : Type*} [Monoid M₁] [Monoid M₂] : + exponent (M₁ × M₂) = lcm (exponent M₁) (exponent M₂) := by + refine dvd_antisymm ?_ (lcm_dvd ?_ ?_) + · refine exponent_dvd_of_forall_pow_eq_one _ _ fun g ↦ ?_ + ext1 + · rw [Prod.pow_fst, Prod.fst_one, ← orderOf_dvd_iff_pow_eq_one] + exact dvd_trans (Monoid.order_dvd_exponent (g.1)) <| dvd_lcm_left _ _ + · rw [Prod.pow_snd, Prod.snd_one, ← orderOf_dvd_iff_pow_eq_one] + exact dvd_trans (Monoid.order_dvd_exponent (g.2)) <| dvd_lcm_right _ _ + · exact MonoidHom.exponent_dvd (f := MonoidHom.fst M₁ M₂) Prod.fst_surjective + · exact MonoidHom.exponent_dvd (f := MonoidHom.snd M₁ M₂) Prod.snd_surjective + +end PiProd diff --git a/Mathlib/GroupTheory/FiniteAbelian.lean b/Mathlib/GroupTheory/FiniteAbelian.lean index 9ab087f7b8a49..64243cd968c62 100644 --- a/Mathlib/GroupTheory/FiniteAbelian.lean +++ b/Mathlib/GroupTheory/FiniteAbelian.lean @@ -75,7 +75,7 @@ theorem equiv_directSum_zmod_of_fintype [Finite G] : obtain ⟨n, ι, fι, p, hp, e, ⟨f⟩⟩ := equiv_free_prod_directSum_zmod G cases' n with n · have : Unique (Fin Nat.zero →₀ ℤ) := - { uniq := by simp only [Nat.zero_eq, eq_iff_true_of_subsingleton] } + { uniq := by simp only [Nat.zero_eq, eq_iff_true_of_subsingleton]; trivial } exact ⟨ι, fι, p, hp, e, ⟨f.trans AddEquiv.uniqueProd⟩⟩ · haveI := @Fintype.prodLeft _ _ _ (Fintype.ofEquiv G f.toEquiv) _ exact diff --git a/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean b/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean index c26caff6e9de4..660245ce6458d 100644 --- a/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean +++ b/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import Mathlib.Algebra.Hom.Equiv.TypeTags +import Mathlib.Algebra.Group.Equiv.TypeTags import Mathlib.Algebra.Module.Equiv import Mathlib.Data.Finsupp.Defs import Mathlib.GroupTheory.FreeAbelianGroup diff --git a/Mathlib/GroupTheory/FreeGroup/IsFreeGroup.lean b/Mathlib/GroupTheory/FreeGroup/IsFreeGroup.lean index 0470f87a39a29..a26ba4bac968f 100644 --- a/Mathlib/GroupTheory/FreeGroup/IsFreeGroup.lean +++ b/Mathlib/GroupTheory/FreeGroup/IsFreeGroup.lean @@ -10,56 +10,175 @@ import Mathlib.GroupTheory.FreeGroup.Basic /-! # Free groups structures on arbitrary types -This file defines a type class for type that are free groups, together with the usual operations. -The type class can be instantiated by providing an isomorphism to the canonical free group, or by -proving that the universal property holds. +This file defines the notion of free basis of a group, which induces an isomorphism between the +group and the free group generated by the basis. + +It also introduced a type class for groups which are free groups, i.e., for which some free basis +exists. For the explicit construction of free groups, see `GroupTheory/FreeGroup`. ## Main definitions -* `IsFreeGroup G` - a typeclass to indicate that `G` is free over some generators -* `IsFreeGroup.of` - the canonical injection of `G`'s generators into `G` -* `IsFreeGroup.lift` - the universal property of the free group +* `FreeGroupBasis ι G` : a function from `ι` to `G` such that `G` is free over its image. + Equivalently, an isomorphism between `G` and `FreeGroup ι`. + +* `IsFreeGroup G` : a typeclass to indicate that `G` is free over some generators +* `Generators G` : given a group satisfying `IsFreeGroup G`, some indexing type over + which `G` is free. +* `IsFreeGroup.of` : the canonical injection of `G`'s generators into `G` +* `IsFreeGroup.lift` : the universal property of the free group ## Main results -* `IsFreeGroup.toFreeGroup` - any free group with generators `A` is equivalent to `FreeGroup A`. -* `IsFreeGroup.unique_lift` - the universal property of a free group -* `IsFreeGroup.ofUniqueLift` - constructing `IsFreeGroup` from the universal property +* `FreeGroupBasis.isFreeGroup`: a group admitting a free group basis is free. +* `IsFreeGroup.toFreeGroup`: any free group with generators `A` is equivalent to `FreeGroup A`. +* `IsFreeGroup.unique_lift`: the universal property of a free group. +* `FreeGroupBasis.ofUniqueLift`: a group satisfying the universal property of a free group admits + a free group basis. -/ universe u +open Function Set + +noncomputable section + +/-- A free group basis `FreeGroupBasis ι G` is a structure recording the isomorphism between a +group `G` and the free group over `ι`. One may think of such a basis as a function from `ι` to `G` +(which is registered through a `FunLike` instance) together with the fact that the morphism induced +by this function from `FreeGroup ι` to `G` is an isomorphism. -/ +structure FreeGroupBasis (ι : Type*) (G : Type*) [Group G] where + /-- `FreeGroupBasis.ofRepr` constructs a basis given an equivalence with a free group. -/ + ofRepr :: + /-- `repr` is the isomorphism between the group `G` and the free group generated by `ι`. -/ + repr : G ≃* FreeGroup ι + +/-- A group is free if it admits a free group basis. In the definition, we require the basis to +be in the same universe as `G`, although this property follows from the existence of a basis in +any universe, see `FreeGroupBasis.isFreeGroup`. -/ +class IsFreeGroup (G : Type u) [Group G] : Prop where + nonempty_basis : ∃ (ι : Type u), Nonempty (FreeGroupBasis ι G) + +namespace FreeGroupBasis + +variable {ι ι' G H : Type*} [Group G] [Group H] + +/-- A free group basis for `G` over `ι` is associated to a map `ι → G` recording the images of +the generators. -/ +instance funLike : FunLike (FreeGroupBasis ι G) ι (fun _ ↦ G) where + coe b := fun i ↦ b.repr.symm (FreeGroup.of i) + coe_injective' := by + rintro ⟨b⟩ ⟨b'⟩ hbb' + have H : (b.symm : FreeGroup ι →* G) = (b'.symm : FreeGroup ι →* G) := by + ext i; exact congr_fun hbb' i + have : b.symm = b'.symm := by ext x; exact FunLike.congr_fun H x + rw [ofRepr.injEq, ← MulEquiv.symm_symm b, ← MulEquiv.symm_symm b', this] + +@[simp] lemma repr_apply_coe (b : FreeGroupBasis ι G) (i : ι) : b.repr (b i) = FreeGroup.of i := by + change b.repr (b.repr.symm (FreeGroup.of i)) = FreeGroup.of i + simp + +/-- The canonical basis of the free group over `X`. -/ +def ofFreeGroup (X : Type*) : FreeGroupBasis X (FreeGroup X) := ofRepr (MulEquiv.refl _) + +@[simp] lemma ofFreeGroup_apply {X : Type*} (x : X) : + FreeGroupBasis.ofFreeGroup X x = FreeGroup.of x := + rfl + +/-- Reindex a free group basis through a bijection of the indexing sets. -/ +protected def reindex (b : FreeGroupBasis ι G) (e : ι ≃ ι') : FreeGroupBasis ι' G := + ofRepr (b.repr.trans (FreeGroup.freeGroupCongr e)) + +@[simp] lemma reindex_apply (b : FreeGroupBasis ι G) (e : ι ≃ ι') (x : ι') : + b.reindex e x = b (e.symm x) := rfl + +/-- Pushing a free group basis through a group isomorphism. -/ +protected def map (b : FreeGroupBasis ι G) (e : G ≃* H) : FreeGroupBasis ι H := + ofRepr (e.symm.trans b.repr) + +@[simp] lemma map_apply (b : FreeGroupBasis ι G) (e : G ≃* H) (x : ι) : + b.map e x = e (b x) := rfl + +protected lemma injective (b : FreeGroupBasis ι G) : Injective b := + b.repr.symm.injective.comp FreeGroup.of_injective + +/-- A group admitting a free group basis is a free group. -/ +lemma isFreeGroup (b : FreeGroupBasis ι G) : IsFreeGroup G := + ⟨range b, ⟨b.reindex (Equiv.ofInjective (↑b) b.injective)⟩⟩ + +instance (X : Type*) : IsFreeGroup (FreeGroup X) := + (ofFreeGroup X).isFreeGroup -/- ./././Mathport/Syntax/Translate/Command.lean:388:30: infer kinds are unsupported in Lean 4: -#[`MulEquiv] [] -/ -/- Porting Note regarding the comment above: -The mathlib3 version makes `G` explicit in `IsFreeGroup.MulEquiv`. -/ +/-- Given a free group basis of `G` over `ι`, there is a canonical bijection between maps from `ι` +to a group `H` and morphisms from `G` to `H`. -/ +@[simps!] +def lift (b : FreeGroupBasis ι G) : (ι → H) ≃ (G →* H) := + FreeGroup.lift.trans + { toFun := fun f => f.comp b.repr.toMonoidHom + invFun := fun f => f.comp b.repr.symm.toMonoidHom + left_inv := fun f => by + ext + simp + right_inv := fun f => by + ext + simp } -/-- `IsFreeGroup G` means that `G` isomorphic to a free group. -/ -class IsFreeGroup (G : Type u) [Group G] where - /-- The generators of a free group. -/ - Generators : Type u - /-- The multiplicative equivalence between "the" free group on the generators, and - the given group `G`. - Note: `IsFreeGroup.MulEquiv'` should not be used directly. - `IsFreeGroup.MulEquiv` should be used instead because it makes `G` an explicit variable.-/ - MulEquiv' : FreeGroup Generators ≃* G -#align is_free_group IsFreeGroup +/-- If two morphisms on `G` coincide on the elements of a basis, then they coincide. -/ +lemma ext_hom (b : FreeGroupBasis ι G) (f g : G →* H) (h : ∀ i, f (b i) = g (b i)) : f = g := + b.lift.symm.injective <| funext h -instance (X : Type*) : IsFreeGroup (FreeGroup X) where - Generators := X - MulEquiv' := MulEquiv.refl _ +/-- If a group satisfies the universal property of a free group with respect to a given type, then +it admits a free group basis based on this type. Here, the universal property is expressed as +in `IsFreeGroup.lift` and its properties. -/ +def ofLift {G : Type u} [Group G] (X : Type u) (of : X → G) + (lift : ∀ {H : Type u} [Group H], (X → H) ≃ (G →* H)) + (lift_of : ∀ {H : Type u} [Group H], ∀ (f : X → H) (a), lift f (of a) = f a) : + FreeGroupBasis X G where + repr := MulEquiv.symm <| MonoidHom.toMulEquiv (FreeGroup.lift of) (lift FreeGroup.of) + (by + apply FreeGroup.ext_hom; intro x + simp only [MonoidHom.coe_comp, Function.comp_apply, MonoidHom.id_apply, FreeGroup.lift.of, + lift_of]) + (by + let lift_symm_of : ∀ {H : Type u} [Group H], ∀ (f : G →* H) (a), lift.symm f a = f (of a) := + by intro H _ f a; simp [← lift_of (lift.symm f)] + apply lift.symm.injective; ext x + simp only [MonoidHom.coe_comp, Function.comp_apply, MonoidHom.id_apply, FreeGroup.lift.of, + lift_of, lift_symm_of]) + +/-- If a group satisfies the universal property of a free group with respect to a given type, then +it admits a free group basis based on this type. Here +the universal property is expressed as in `IsFreeGroup.unique_lift`. -/ +def ofUniqueLift {G : Type u} [Group G] (X : Type u) (of : X → G) + (h : ∀ {H : Type u} [Group H] (f : X → H), ∃! F : G →* H, ∀ a, F (of a) = f a) : + FreeGroupBasis X G := + let lift {H : Type u} [Group H] : (X → H) ≃ (G →* H) := + { toFun := fun f => Classical.choose (h f) + invFun := fun F => F ∘ of + left_inv := fun f => funext (Classical.choose_spec (h f)).left + right_inv := fun F => ((Classical.choose_spec (h (F ∘ of))).right F fun _ => rfl).symm } + let lift_of {H : Type u} [Group H] (f : X → H) (a : X) : lift f (of a) = f a := + congr_fun (lift.symm_apply_apply f) a + ofLift X of @lift @lift_of + +end FreeGroupBasis namespace IsFreeGroup variable (G : Type*) [Group G] [IsFreeGroup G] +/-- A set of generators of a free group, chosen arbitrarily -/ +def Generators : Type _ := (IsFreeGroup.nonempty_basis (G := G)).choose + /-- Any free group is isomorphic to "the" free group. -/ -def MulEquiv : FreeGroup (Generators G) ≃* G := IsFreeGroup.MulEquiv' +irreducible_def MulEquiv : FreeGroup (Generators G) ≃* G := + (IsFreeGroup.nonempty_basis (G := G)).choose_spec.some.repr.symm + +/-- A free group basis of a free group `G`, over the set `Generators G`. -/ +def basis : FreeGroupBasis (Generators G) G := FreeGroupBasis.ofRepr (MulEquiv G).symm /-- Any free group is isomorphic to "the" free group. -/ @[simps!] @@ -76,32 +195,14 @@ def of : Generators G → G := (MulEquiv G).toFun ∘ FreeGroup.of #align is_free_group.of IsFreeGroup.of -@[simp] -theorem of_eq_freeGroup_of {A : Type u} : @of (FreeGroup A) _ _ = FreeGroup.of := - rfl -#align is_free_group.of_eq_free_group_of IsFreeGroup.of_eq_freeGroup_of - variable {H : Type*} [Group H] /-- The equivalence between functions on the generators and group homomorphisms from a free group given by those generators. -/ def lift : (Generators G → H) ≃ (G →* H) := - FreeGroup.lift.trans - { toFun := fun f => f.comp (MulEquiv G).symm.toMonoidHom - invFun := fun f => f.comp (MulEquiv G).toMonoidHom - left_inv := fun f => by - ext - simp - right_inv := fun f => by - ext - simp } + (basis G).lift #align is_free_group.lift IsFreeGroup.lift -@[simp] -theorem lift'_eq_freeGroup_lift {A : Type u} : @lift (FreeGroup A) _ _ H _ = FreeGroup.lift := - rfl -#align is_free_group.lift'_eq_free_group_lift IsFreeGroup.lift'_eq_freeGroup_lift - @[simp] theorem lift_of (f : Generators G → H) (a : Generators G) : lift f (of a) = f a := congr_fun (lift.symm_apply_apply f) a @@ -112,8 +213,7 @@ theorem lift_symm_apply (f : G →* H) (a : Generators G) : (lift.symm f) a = f rfl #align is_free_group.lift_symm_apply IsFreeGroup.lift_symm_apply -@[ext 1050] --Porting note: increased priority, but deliberately less than for example ---`FreeProduct.ext_hom` +/- Do not register this as an ext lemma, as `Generators G` is not canonical. -/ theorem ext_hom ⦃f g : G →* H⦄ (h : ∀ a : Generators G, f (of a) = g (of a)) : f = g := lift.symm.injective (funext h) #align is_free_group.ext_hom IsFreeGroup.ext_hom @@ -127,47 +227,23 @@ theorem unique_lift (f : Generators G → H) : ∃! F : G →* H, ∀ a, F (of a simpa only [Function.funext_iff] using lift.symm.bijective.existsUnique f #align is_free_group.unique_lift IsFreeGroup.unique_lift -/-- If a group satisfies the universal property of a free group, then it is a free group, where -the universal property is expressed as in `IsFreeGroup.lift` and its properties. -/ -def ofLift {G : Type u} [Group G] (X : Type u) (of : X → G) +/-- If a group satisfies the universal property of a free group with respect to a given type, then +it is free. Here, the universal property is expressed as in `IsFreeGroup.lift` and its +properties. -/ +lemma ofLift {G : Type u} [Group G] (X : Type u) (of : X → G) (lift : ∀ {H : Type u} [Group H], (X → H) ≃ (G →* H)) - (lift_of : ∀ {H : Type u} [Group H], ∀ (f : X → H) (a), lift f (of a) = f a) : IsFreeGroup G - where - Generators := X - MulEquiv' := - MonoidHom.toMulEquiv (FreeGroup.lift of) (lift FreeGroup.of) - (by - apply FreeGroup.ext_hom; intro x - simp only [MonoidHom.coe_comp, Function.comp_apply, MonoidHom.id_apply, FreeGroup.lift.of, - lift_of]) - (by - let lift_symm_of : ∀ {H : Type u} [Group H], ∀ (f : G →* H) (a), lift.symm f a = f (of a) := - by intro H _ f a; simp [← lift_of (lift.symm f)] - apply lift.symm.injective; ext x - simp only [MonoidHom.coe_comp, Function.comp_apply, MonoidHom.id_apply, FreeGroup.lift.of, - lift_of, lift_symm_of]) -#align is_free_group.of_lift IsFreeGroup.ofLift + (lift_of : ∀ {H : Type u} [Group H], ∀ (f : X → H) (a), lift f (of a) = f a) : + IsFreeGroup G := + (FreeGroupBasis.ofLift X of lift lift_of).isFreeGroup -/-- If a group satisfies the universal property of a free group, then it is a free group, where -the universal property is expressed as in `IsFreeGroup.unique_lift`. -/ -noncomputable def ofUniqueLift {G : Type u} [Group G] (X : Type u) (of : X → G) +/-- If a group satisfies the universal property of a free group with respect to a given type, then +it is free. Here the universal property is expressed as in `IsFreeGroup.unique_lift`. -/ +lemma ofUniqueLift {G : Type u} [Group G] (X : Type u) (of : X → G) (h : ∀ {H : Type u} [Group H] (f : X → H), ∃! F : G →* H, ∀ a, F (of a) = f a) : IsFreeGroup G := - let lift {H : Type u} [Group H] : (X → H) ≃ (G →* H) := - { toFun := fun f => Classical.choose (h f) - invFun := fun F => F ∘ of - left_inv := fun f => funext (Classical.choose_spec (h f)).left - right_inv := fun F => ((Classical.choose_spec (h (F ∘ of))).right F fun _ => rfl).symm } - let lift_of {H : Type u} [Group H] (f : X → H) (a : X) : lift f (of a) = f a := - congr_fun (lift.symm_apply_apply f) a - ofLift X of @lift @lift_of -#align is_free_group.of_unique_lift IsFreeGroup.ofUniqueLift - -/-- Being a free group transports across group isomorphisms. -/ -def ofMulEquiv {H : Type _} [Group H] (h : G ≃* H) : IsFreeGroup H - where - Generators := Generators G - MulEquiv' := (MulEquiv G).trans h -#align is_free_group.of_mul_equiv IsFreeGroup.ofMulEquiv + (FreeGroupBasis.ofUniqueLift X of h).isFreeGroup + +lemma ofMulEquiv [IsFreeGroup G] (e : G ≃* H) : IsFreeGroup H := + ((basis G).map e).isFreeGroup end IsFreeGroup diff --git a/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean b/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean index b35f6179ad92c..fee519dca5465 100644 --- a/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean +++ b/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean @@ -121,7 +121,7 @@ instance actionGroupoidIsFree {G A : Type u} [Group G] [IsFreeGroup G] [MulActio · suffices SemidirectProduct.rightHom.comp F' = MonoidHom.id _ by -- Porting note: `MonoidHom.ext_iff` has been deprecated. exact FunLike.ext_iff.mp this - ext + apply IsFreeGroup.ext_hom (fun x ↦ ?_) rw [MonoidHom.comp_apply, hF'] rfl · rintro ⟨⟨⟩, a : A⟩ ⟨⟨⟩, b⟩ ⟨e, h : IsFreeGroup.of e • a = b⟩ @@ -222,7 +222,7 @@ def functorOfMonoidHom {X} [Monoid X] (f : End (root' T) →* X) : G ⥤ Categor /-- Given a free groupoid and an arborescence of its generating quiver, the vertex group at the root is freely generated by loops coming from generating arrows in the complement of the tree. -/ -def endIsFree : IsFreeGroup (End (root' T)) := +lemma endIsFree : IsFreeGroup (End (root' T)) := IsFreeGroup.ofUniqueLift ((wideSubquiverEquivSetTotal <| wideSubquiverSymmetrify T)ᶜ : Set _) (fun e => loopOfHom T (of e.val.hom)) (by diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index 2bf84ce74ce70..2045b9d4f69b3 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -37,7 +37,11 @@ open Function namespace MulAction -variable (M : Type u) {α : Type v} [Monoid M] [MulAction M α] +variable (M : Type u) [Monoid M] (α : Type v) [MulAction M α] + +section Orbit + +variable {α} /-- The orbit of an element under an action. -/ @[to_additive "The orbit of an element under an action."] @@ -55,8 +59,8 @@ theorem mem_orbit_iff {a₁ a₂ : α} : a₂ ∈ orbit M a₁ ↔ ∃ x : M, x #align add_action.mem_orbit_iff AddAction.mem_orbit_iff @[to_additive (attr := simp)] -theorem mem_orbit (a : α) (x : M) : x • a ∈ orbit M a := - ⟨x, rfl⟩ +theorem mem_orbit (a : α) (m : M) : m • a ∈ orbit M a := + ⟨m, rfl⟩ #align mul_action.mem_orbit MulAction.mem_orbit #align add_action.mem_orbit AddAction.mem_orbit @@ -102,7 +106,17 @@ theorem orbit.coe_smul {a : α} {m : M} {a' : orbit M a} : ↑(m • a') = m • #align mul_action.orbit.coe_smul MulAction.orbit.coe_smul #align add_action.orbit.coe_vadd AddAction.orbit.coe_vadd -variable (M) (α) +variable (M) + +@[to_additive] +theorem orbit_eq_univ [IsPretransitive M α] (a : α) : orbit M a = Set.univ := + (surjective_smul M a).range_eq +#align mul_action.orbit_eq_univ MulAction.orbit_eq_univ +#align add_action.orbit_eq_univ AddAction.orbit_eq_univ + +end Orbit + +section FixedPoints /-- The set of elements fixed under the whole action. -/ @[to_additive "The set of elements fixed under the whole action."] @@ -111,6 +125,8 @@ def fixedPoints : Set α := #align mul_action.fixed_points MulAction.fixedPoints #align add_action.fixed_points AddAction.fixedPoints +variable {M} + /-- `fixedBy m` is the set of elements fixed by `m`. -/ @[to_additive "`fixedBy m` is the set of elements fixed by `m`."] def fixedBy (m : M) : Set α := @@ -118,14 +134,16 @@ def fixedBy (m : M) : Set α := #align mul_action.fixed_by MulAction.fixedBy #align add_action.fixed_by AddAction.fixedBy +variable (M) + @[to_additive] -theorem fixed_eq_iInter_fixedBy : fixedPoints M α = ⋂ m : M, fixedBy M α m := +theorem fixed_eq_iInter_fixedBy : fixedPoints M α = ⋂ m : M, fixedBy α m := Set.ext fun _ => ⟨fun hx => Set.mem_iInter.2 fun m => hx m, fun hx m => (Set.mem_iInter.1 hx m : _)⟩ #align mul_action.fixed_eq_Inter_fixed_by MulAction.fixed_eq_iInter_fixedBy #align add_action.fixed_eq_Inter_fixed_by AddAction.fixed_eq_iInter_fixedBy -variable {M} +variable {M α} @[to_additive (attr := simp)] theorem mem_fixedPoints {a : α} : a ∈ fixedPoints M α ↔ ∀ m : M, m • a = a := @@ -134,7 +152,7 @@ theorem mem_fixedPoints {a : α} : a ∈ fixedPoints M α ↔ ∀ m : M, m • a #align add_action.mem_fixed_points AddAction.mem_fixedPoints @[to_additive (attr := simp)] -theorem mem_fixedBy {m : M} {a : α} : a ∈ fixedBy M α m ↔ m • a = a := +theorem mem_fixedBy {m : M} {a : α} : a ∈ fixedBy α m ↔ m • a = a := Iff.rfl #align mul_action.mem_fixed_by MulAction.mem_fixedBy #align add_action.mem_fixed_by AddAction.mem_fixedBy @@ -148,36 +166,6 @@ theorem mem_fixedPoints' {a : α} : a ∈ fixedPoints M α ↔ ∀ a', a' ∈ or #align mul_action.mem_fixed_points' MulAction.mem_fixedPoints' #align add_action.mem_fixed_points' AddAction.mem_fixedPoints' -variable (M) {α} - -/-- The stabilizer of a point `a` as a submonoid of `M`. -/ -@[to_additive "The stabilizer of m point `a` as an additive submonoid of `M`."] -def Stabilizer.submonoid (a : α) : Submonoid M where - carrier := { m | m • a = a } - one_mem' := one_smul _ a - mul_mem' {m m'} (ha : m • a = a) (hb : m' • a = a) := - show (m * m') • a = a by rw [← smul_smul, hb, ha] -#align mul_action.stabilizer.submonoid MulAction.Stabilizer.submonoid -#align add_action.stabilizer.add_submonoid AddAction.Stabilizer.addSubmonoid - -@[to_additive (attr := simp)] -theorem mem_stabilizer_submonoid_iff {a : α} {m : M} : m ∈ Stabilizer.submonoid M a ↔ m • a = a := - Iff.rfl -#align mul_action.mem_stabilizer_submonoid_iff MulAction.mem_stabilizer_submonoid_iff -#align add_action.mem_stabilizer_add_submonoid_iff AddAction.mem_stabilizer_addSubmonoid_iff - -@[to_additive] -instance [DecidableEq α] (a : α) : DecidablePred (· ∈ Stabilizer.submonoid M a) := - fun _ => inferInstanceAs <| Decidable (_ = _) - -@[to_additive] -theorem orbit_eq_univ [IsPretransitive M α] (a : α) : orbit M a = Set.univ := - (surjective_smul M a).range_eq -#align mul_action.orbit_eq_univ MulAction.orbit_eq_univ -#align add_action.orbit_eq_univ AddAction.orbit_eq_univ - -variable {M} - @[to_additive mem_fixedPoints_iff_card_orbit_eq_one] theorem mem_fixedPoints_iff_card_orbit_eq_one {a : α} [Fintype (orbit M a)] : a ∈ fixedPoints M α ↔ Fintype.card (orbit M a) = 1 := by @@ -192,34 +180,56 @@ theorem mem_fixedPoints_iff_card_orbit_eq_one {a : α} [Fintype (orbit M a)] : #align mul_action.mem_fixed_points_iff_card_orbit_eq_one MulAction.mem_fixedPoints_iff_card_orbit_eq_one #align add_action.mem_fixed_points_iff_card_orbit_eq_zero AddAction.mem_fixedPoints_iff_card_orbit_eq_one -end MulAction +end FixedPoints -namespace MulAction +section Stabilizers -variable (G : Type u) {α : Type v} [Group G] [MulAction G α] +variable {α} -/-- The stabilizer of an element under an action, i.e. what sends the element to itself. -A subgroup. -/ -@[to_additive - "The stabilizer of an element under an action, i.e. what sends the element to itself. - An additive subgroup."] -def stabilizer (a : α) : Subgroup G := - { Stabilizer.submonoid G a with - inv_mem' := fun {m} (ha : m • a = a) => show m⁻¹ • a = a by rw [inv_smul_eq_iff, ha] } -#align mul_action.stabilizer MulAction.stabilizer -#align add_action.stabilizer AddAction.stabilizer +/-- The stabilizer of a point `a` as a submonoid of `M`. -/ +@[to_additive "The stabilizer of a point `a` as an additive submonoid of `M`."] +def stabilizerSubmonoid (a : α) : Submonoid M where + carrier := { m | m • a = a } + one_mem' := one_smul _ a + mul_mem' {m m'} (ha : m • a = a) (hb : m' • a = a) := + show (m * m') • a = a by rw [← smul_smul, hb, ha] +#align mul_action.stabilizer.submonoid MulAction.stabilizerSubmonoid +#align add_action.stabilizer.add_submonoid AddAction.stabilizerAddSubmonoid -variable {G} +variable {M} + +@[to_additive] +instance [DecidableEq α] (a : α) : DecidablePred (· ∈ stabilizerSubmonoid M a) := + fun _ => inferInstanceAs <| Decidable (_ = _) @[to_additive (attr := simp)] -theorem mem_stabilizer_iff {g : G} {a : α} : g ∈ stabilizer G a ↔ g • a = a := +theorem mem_stabilizerSubmonoid_iff {a : α} {m : M} : m ∈ stabilizerSubmonoid M a ↔ m • a = a := Iff.rfl -#align mul_action.mem_stabilizer_iff MulAction.mem_stabilizer_iff -#align add_action.mem_stabilizer_iff AddAction.mem_stabilizer_iff +#align mul_action.mem_stabilizer_submonoid_iff MulAction.mem_stabilizerSubmonoid_iff +#align add_action.mem_stabilizer_add_submonoid_iff AddAction.mem_stabilizerAddSubmonoid_iff -@[to_additive] -instance [DecidableEq α] (a : α) : DecidablePred (· ∈ stabilizer G a) := - fun _ => inferInstanceAs <| Decidable (_ = _) +end Stabilizers + +end MulAction + +/-- `smul` by a `k : M` over a ring is injective, if `k` is not a zero divisor. +The general theory of such `k` is elaborated by `IsSMulRegular`. +The typeclass that restricts all terms of `M` to have this property is `NoZeroSMulDivisors`. -/ +theorem smul_cancel_of_non_zero_divisor {M R : Type*} [Monoid M] [NonUnitalNonAssocRing R] + [DistribMulAction M R] (k : M) (h : ∀ x : R, k • x = 0 → x = 0) {a b : R} (h' : k • a = k • b) : + a = b := by + rw [← sub_eq_zero] + refine' h _ _ + rw [smul_sub, h', sub_self] +#align smul_cancel_of_non_zero_divisor smul_cancel_of_non_zero_divisor + +namespace MulAction + +variable (G : Type u) [Group G] (α : Type v) [MulAction G α] + +section Orbit + +variable {G α} @[to_additive (attr := simp)] theorem smul_orbit (g : G) (a : α) : g • orbit G a = orbit G a := @@ -254,8 +264,6 @@ theorem orbit_eq_iff {a b : α} : orbit G a = orbit G b ↔ a ∈ orbit G b := #align mul_action.orbit_eq_iff MulAction.orbit_eq_iff #align add_action.orbit_eq_iff AddAction.orbit_eq_iff -variable (G) - @[to_additive] theorem mem_orbit_smul (g : G) (a : α) : a ∈ orbit G (g • a) := by simp only [orbit_smul, mem_orbit_self] @@ -268,7 +276,7 @@ theorem smul_mem_orbit_smul (g h : G) (a : α) : g • a ∈ orbit G (h • a) : #align mul_action.smul_mem_orbit_smul MulAction.smul_mem_orbit_smul #align add_action.vadd_mem_orbit_vadd AddAction.vadd_mem_orbit_vadd -variable (α) +variable (G α) /-- The relation 'in the same orbit'. -/ @[to_additive "The relation 'in the same orbit'."] @@ -420,7 +428,34 @@ def selfEquivSigmaOrbits : α ≃ Σω : Ω, orbit G ω.out' := #align mul_action.self_equiv_sigma_orbits MulAction.selfEquivSigmaOrbits #align add_action.self_equiv_sigma_orbits AddAction.selfEquivSigmaOrbits -variable {G α} +end Orbit + +section Stabilizer + +variable {α} + +/-- The stabilizer of an element under an action, i.e. what sends the element to itself. +A subgroup. -/ +@[to_additive + "The stabilizer of an element under an action, i.e. what sends the element to itself. + An additive subgroup."] +def stabilizer (a : α) : Subgroup G := + { stabilizerSubmonoid G a with + inv_mem' := fun {m} (ha : m • a = a) => show m⁻¹ • a = a by rw [inv_smul_eq_iff, ha] } +#align mul_action.stabilizer MulAction.stabilizer +#align add_action.stabilizer AddAction.stabilizer + +variable {G} + +@[to_additive] +instance [DecidableEq α] (a : α) : DecidablePred (· ∈ stabilizer G a) := + fun _ => inferInstanceAs <| Decidable (_ = _) + +@[to_additive (attr := simp)] +theorem mem_stabilizer_iff {a : α} {g : G} : g ∈ stabilizer G a ↔ g • a = a := + Iff.rfl +#align mul_action.mem_stabilizer_iff MulAction.mem_stabilizer_iff +#align add_action.mem_stabilizer_iff AddAction.mem_stabilizer_iff /-- If the stabilizer of `a` is `S`, then the stabilizer of `g • a` is `gSg⁻¹`. -/ theorem stabilizer_smul_eq_stabilizer_map_conj (g : G) (a : α) : @@ -440,11 +475,13 @@ noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : (orbitRel (MulEquiv.subgroupCongr this).trans ((MulAut.conj g).subgroupMap <| stabilizer G b).symm #align mul_action.stabilizer_equiv_stabilizer_of_orbit_rel MulAction.stabilizerEquivStabilizerOfOrbitRel +end Stabilizer + end MulAction namespace AddAction -variable (G : Type u) (α : Type v) [AddGroup G] [AddAction G α] +variable (G : Type u) [AddGroup G] (α : Type v) [AddAction G α] /-- If the stabilizer of `x` is `S`, then the stabilizer of `g +ᵥ x` is `g + S + (-g)`. -/ theorem stabilizer_vadd_eq_stabilizer_map_conj (g : G) (a : α) : @@ -466,14 +503,3 @@ noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : (orbitRel #align add_action.stabilizer_equiv_stabilizer_of_orbit_rel AddAction.stabilizerEquivStabilizerOfOrbitRel end AddAction - -/-- `smul` by a `k : M` over a ring is injective, if `k` is not a zero divisor. -The general theory of such `k` is elaborated by `IsSMulRegular`. -The typeclass that restricts all terms of `M` to have this property is `NoZeroSMulDivisors`. -/ -theorem smul_cancel_of_non_zero_divisor {M R : Type*} [Monoid M] [NonUnitalNonAssocRing R] - [DistribMulAction M R] (k : M) (h : ∀ x : R, k • x = 0 → x = 0) {a b : R} (h' : k • a = k • b) : - a = b := by - rw [← sub_eq_zero] - refine' h _ _ - rw [smul_sub, h', sub_self] -#align smul_cancel_of_non_zero_divisor smul_cancel_of_non_zero_divisor diff --git a/Mathlib/GroupTheory/GroupAction/Defs.lean b/Mathlib/GroupTheory/GroupAction/Defs.lean index 5095fe3e6e253..54e106c3b274b 100644 --- a/Mathlib/GroupTheory/GroupAction/Defs.lean +++ b/Mathlib/GroupTheory/GroupAction/Defs.lean @@ -3,9 +3,9 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yury Kudryashov -/ -import Mathlib.Algebra.Group.TypeTags import Mathlib.Algebra.Group.Commute.Defs -import Mathlib.Algebra.Hom.Group.Defs +import Mathlib.Algebra.Group.Hom.Defs +import Mathlib.Algebra.Group.TypeTags import Mathlib.Algebra.Opposites import Mathlib.Logic.Embedding.Basic diff --git a/Mathlib/GroupTheory/GroupAction/DomAct/ActionHom.lean b/Mathlib/GroupTheory/GroupAction/DomAct/ActionHom.lean index 4b1a542ad747a..caa6ab453b23c 100644 --- a/Mathlib/GroupTheory/GroupAction/DomAct/ActionHom.lean +++ b/Mathlib/GroupTheory/GroupAction/DomAct/ActionHom.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.GroupTheory.GroupAction.DomAct.Basic -import Mathlib.Algebra.Hom.GroupAction +import Mathlib.GroupTheory.GroupAction.Hom /-! # Action of `Mᵈᵐᵃ` on `α →[N] β` and `A →+[N] B` diff --git a/Mathlib/GroupTheory/GroupAction/Group.lean b/Mathlib/GroupTheory/GroupAction/Group.lean index e00ea7c6de4a9..9669ef0f0855c 100644 --- a/Mathlib/GroupTheory/GroupAction/Group.lean +++ b/Mathlib/GroupTheory/GroupAction/Group.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Hom.Aut +import Mathlib.Algebra.Group.Aut import Mathlib.GroupTheory.GroupAction.Units #align_import group_theory.group_action.group from "leanprover-community/mathlib"@"3b52265189f3fb43aa631edffce5d060fafaf82f" diff --git a/Mathlib/Algebra/Hom/GroupAction.lean b/Mathlib/GroupTheory/GroupAction/Hom.lean similarity index 100% rename from Mathlib/Algebra/Hom/GroupAction.lean rename to Mathlib/GroupTheory/GroupAction/Hom.lean diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 341ff7c09c638..17a3333557acb 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Thomas Browning -/ import Mathlib.Algebra.Group.ConjFinite -import Mathlib.Algebra.Hom.GroupAction import Mathlib.Data.Fintype.BigOperators import Mathlib.Dynamics.PeriodicPts -import Mathlib.GroupTheory.GroupAction.ConjAct import Mathlib.GroupTheory.Commutator import Mathlib.GroupTheory.Coset +import Mathlib.GroupTheory.GroupAction.ConjAct +import Mathlib.GroupTheory.GroupAction.Hom #align_import group_theory.group_action.quotient from "leanprover-community/mathlib"@"4be589053caf347b899a494da75410deb55fb3ef" @@ -307,9 +307,9 @@ theorem card_eq_sum_card_group_div_card_stabilizer [Fintype α] [Fintype β] [Fi "**Burnside's lemma** : a (noncomputable) bijection between the disjoint union of all `{x ∈ X | g • x = x}` for `g ∈ G` and the product `G × X/G`, where `G` is an additive group acting on `X` and `X/G`denotes the quotient of `X` by the relation `orbitRel G X`. "] -noncomputable def sigmaFixedByEquivOrbitsProdGroup : (Σa : α, fixedBy α β a) ≃ Ω × α := +noncomputable def sigmaFixedByEquivOrbitsProdGroup : (Σa : α, fixedBy β a) ≃ Ω × α := calc - (Σa : α, fixedBy α β a) ≃ { ab : α × β // ab.1 • ab.2 = ab.2 } := + (Σa : α, fixedBy β a) ≃ { ab : α × β // ab.1 • ab.2 = ab.2 } := (Equiv.subtypeProdEquivSigmaSubtype _).symm _ ≃ { ba : β × α // ba.2 • ba.1 = ba.1 } := (Equiv.prodComm α β).subtypeEquiv fun _ => Iff.rfl _ ≃ Σb : β, stabilizer α b := @@ -333,8 +333,8 @@ elements fixed by each `g ∈ G` is the number of orbits. -/ @[to_additive "**Burnside's lemma** : given a finite additive group `G` acting on a set `X`, the average number of elements fixed by each `g ∈ G` is the number of orbits. "] -theorem sum_card_fixedBy_eq_card_orbits_mul_card_group [Fintype α] [∀ a, Fintype <| fixedBy α β a] - [Fintype Ω] : (∑ a : α, Fintype.card (fixedBy α β a)) = Fintype.card Ω * Fintype.card α := by +theorem sum_card_fixedBy_eq_card_orbits_mul_card_group [Fintype α] [∀ a : α, Fintype <| fixedBy β a] + [Fintype Ω] : (∑ a : α, Fintype.card (fixedBy β a)) = Fintype.card Ω * Fintype.card α := by rw [← Fintype.card_prod, ← Fintype.card_sigma, Fintype.card_congr (sigmaFixedByEquivOrbitsProdGroup α β)] #align mul_action.sum_card_fixed_by_eq_card_orbits_mul_card_group MulAction.sum_card_fixedBy_eq_card_orbits_mul_card_group @@ -426,7 +426,7 @@ theorem card_comm_eq_card_conjClasses_mul_card (G : Type*) [Group G] : -- Porting note: Changed `calc` proof into a `rw` proof. rw [card_congr (Equiv.subtypeProdEquivSigmaSubtype Commute), card_sigma, sum_equiv ConjAct.toConjAct.toEquiv (fun a ↦ card { b // Commute a b }) - (fun g ↦ card (MulAction.fixedBy (ConjAct G) G g)) + (fun g ↦ card (MulAction.fixedBy G g)) fun g ↦ card_congr' <| congr_arg _ <| funext fun h ↦ mul_inv_eq_iff_eq_mul.symm.to_eq, MulAction.sum_card_fixedBy_eq_card_orbits_mul_card_group] congr 1; apply card_congr'; congr; ext; diff --git a/Mathlib/GroupTheory/GroupAction/SubMulAction.lean b/Mathlib/GroupTheory/GroupAction/SubMulAction.lean index 729bc37c12af1..4f7819ca0510e 100644 --- a/Mathlib/GroupTheory/GroupAction/SubMulAction.lean +++ b/Mathlib/GroupTheory/GroupAction/SubMulAction.lean @@ -3,10 +3,10 @@ Copyright (c) 2020 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Algebra.Hom.GroupAction import Mathlib.Algebra.Module.Basic import Mathlib.Data.SetLike.Basic import Mathlib.GroupTheory.GroupAction.Basic +import Mathlib.GroupTheory.GroupAction.Hom #align_import group_theory.group_action.sub_mul_action from "leanprover-community/mathlib"@"feb99064803fd3108e37c18b0f77d0a8344677a3" @@ -331,9 +331,9 @@ lemma orbit_of_sub_mul {p : SubMulAction R M} (m : p) : -/ /-- Stabilizers in monoid SubMulAction coincide with stabilizers in the ambient space -/ theorem stabilizer_of_subMul.submonoid {p : SubMulAction R M} (m : p) : - MulAction.Stabilizer.submonoid R m = MulAction.Stabilizer.submonoid R (m : M) := by + MulAction.stabilizerSubmonoid R m = MulAction.stabilizerSubmonoid R (m : M) := by ext - simp only [MulAction.mem_stabilizer_submonoid_iff, ← SubMulAction.val_smul, SetLike.coe_eq_coe] + simp only [MulAction.mem_stabilizerSubmonoid_iff, ← SubMulAction.val_smul, SetLike.coe_eq_coe] #align sub_mul_action.stabilizer_of_sub_mul.submonoid SubMulAction.stabilizer_of_subMul.submonoid end MulActionMonoid diff --git a/Mathlib/GroupTheory/HNNExtension.lean b/Mathlib/GroupTheory/HNNExtension.lean index 58ac00238b7eb..0350fceacee4d 100644 --- a/Mathlib/GroupTheory/HNNExtension.lean +++ b/Mathlib/GroupTheory/HNNExtension.lean @@ -400,7 +400,7 @@ theorem not_cancels_of_cons_hyp (u : ℤˣ) (w : NormalWord d) theorem unitsSMul_cancels_iff (u : ℤˣ) (w : NormalWord d) : Cancels (-u) (unitsSMul φ u w) ↔ ¬ Cancels u w := by by_cases h : Cancels u w - · simp only [unitsSMul, dif_pos trivial, h, iff_false] + · simp only [unitsSMul, h, dite_true, not_true_eq_false, iff_false] induction w using consRecOn with | ofGroup => simp [Cancels, unitsSMulWithCancel] | cons g u' w h1 h2 _ => diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 2bd30f0523b0d..a8682807bab05 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Julian Kuelshammer -/ -import Mathlib.Algebra.Hom.Iterate +import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Data.Int.ModEq import Mathlib.Data.Set.Pointwise.Basic import Mathlib.Data.Set.Intervals.Infinite @@ -679,7 +679,7 @@ theorem sum_card_orderOf_eq_card_pow_eq_one [Fintype G] [DecidableEq G] (hn : n #align sum_card_order_of_eq_card_pow_eq_one sum_card_orderOf_eq_card_pow_eq_one #align sum_card_add_order_of_eq_card_nsmul_eq_zero sum_card_addOrderOf_eq_card_nsmul_eq_zero -@[to_additive ] +@[to_additive] theorem orderOf_le_card_univ [Fintype G] : orderOf x ≤ Fintype.card G := Finset.le_card_of_inj_on_range ((· ^ ·) x) (fun _ _ => Finset.mem_univ _) fun _ hi _ hj => pow_injective_of_lt_orderOf x hi hj @@ -1129,7 +1129,7 @@ theorem LinearOrderedRing.orderOf_le_two : orderOf x ≤ 2 := by cases' ne_or_eq |x| 1 with h h · simp [orderOf_abs_ne_one h] rcases eq_or_eq_neg_of_abs_eq h with (rfl | rfl) - · simp + · simp; decide apply orderOf_le_of_pow_eq_one <;> norm_num #align linear_ordered_ring.order_of_le_two LinearOrderedRing.orderOf_le_two diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index acd152bd3e522..c110216869b65 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -198,7 +198,7 @@ theorem card_modEq_card_fixedPoints [Fintype (fixedPoints G α)] : Eq.symm (Finset.sum_bij_ne_zero (fun a _ _ => Quotient.mk'' a.1) (fun _ _ _ => Finset.mem_univ _) (fun a₁ a₂ _ _ _ _ h => - Subtype.eq ((mem_fixedPoints' α).mp a₂.2 a₁.1 (Quotient.exact' h))) + Subtype.eq (mem_fixedPoints'.mp a₂.2 a₁.1 (Quotient.exact' h))) (fun b => Quotient.inductionOn' b fun b _ hb => _) fun a ha _ => by rw [key, mem_fixedPoints_iff_card_orbit_eq_one.mp a.2]) obtain ⟨k, hk⟩ := hG.card_orbit b diff --git a/Mathlib/GroupTheory/Perm/Basic.lean b/Mathlib/GroupTheory/Perm/Basic.lean index 5e6f01f606a86..47d643e2d59a2 100644 --- a/Mathlib/GroupTheory/Perm/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Basic.lean @@ -5,7 +5,7 @@ Authors: Leonardo de Moura, Mario Carneiro -/ import Mathlib.Algebra.Group.Pi import Mathlib.Algebra.Group.Prod -import Mathlib.Algebra.Hom.Iterate +import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Logic.Equiv.Set #align_import group_theory.perm.basic from "leanprover-community/mathlib"@"b86832321b586c6ac23ef8cdef6a7a27e42b13bd" diff --git a/Mathlib/GroupTheory/Perm/Cycle/PossibleTypes.lean b/Mathlib/GroupTheory/Perm/Cycle/PossibleTypes.lean new file mode 100644 index 0000000000000..d53ced8baacd3 --- /dev/null +++ b/Mathlib/GroupTheory/Perm/Cycle/PossibleTypes.lean @@ -0,0 +1,74 @@ +/- +Copyright (c) 2023 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir +-/ + +import Mathlib.GroupTheory.Perm.Cycle.Concrete + +/-! # Possible cycle types of permutations + +* For `m : Multiset ℕ`, `Equiv.Perm.exists_with_cycleType_iff m` + proves that there are permutations with cycleType `m` if and only if + its sum is at most `Fintype.card α` and its members are at least 2. + +-/ + +variable (α : Type _) [DecidableEq α] [Fintype α] + +/-- There are permutations with cycleType `m` if and only if + its sum is at most `Fintype.card α` and its members are at least 2. -/ +theorem Equiv.Perm.exists_with_cycleType_iff {m : Multiset ℕ} : + (∃ g : Equiv.Perm α, g.cycleType = m) ↔ + (m.sum ≤ Fintype.card α ∧ ∀ a ∈ m, 2 ≤ a) := by + constructor + · -- empty case + intro h + obtain ⟨g, hg⟩ := h + constructor + · rw [← hg, Equiv.Perm.sum_cycleType] + exact (Equiv.Perm.support g).card_le_univ + · intro a + rw [← hg] + exact Equiv.Perm.two_le_of_mem_cycleType + · rintro ⟨hc, h2c⟩ + have hc' : m.toList.sum ≤ Fintype.card α + · simp only [Multiset.sum_toList] + exact hc + obtain ⟨p, hp_length, hp_nodup, hp_disj⟩ := List.exists_pw_disjoint_with_card hc' + use List.prod (List.map (fun l => List.formPerm l) p) + have hp2 : ∀ x ∈ p, 2 ≤ x.length := by + intro x hx + apply h2c x.length + rw [← Multiset.mem_toList, ← hp_length, List.mem_map] + exact ⟨x, hx, rfl⟩ + rw [Equiv.Perm.cycleType_eq _ rfl] + · -- lengths + rw [← Multiset.coe_toList m] + apply congr_arg + rw [List.map_map]; rw [← hp_length] + apply List.map_congr + intro x hx; simp only [Function.comp_apply] + rw [List.support_formPerm_of_nodup x (hp_nodup x hx)] + ·-- length + rw [List.toFinset_card_of_nodup (hp_nodup x hx)] + · -- length >= 1 + intro a h + apply Nat.not_succ_le_self 1 + conv_rhs => rw [← List.length_singleton a]; rw [← h] + exact hp2 x hx + · -- cycles + intro g + rw [List.mem_map] + rintro ⟨x, hx, rfl⟩ + have hx_nodup : x.Nodup := hp_nodup x hx + rw [← Cycle.formPerm_coe x hx_nodup] + apply Cycle.isCycle_formPerm + rw [Cycle.nontrivial_coe_nodup_iff hx_nodup] + exact hp2 x hx + · -- disjoint + rw [List.pairwise_map] + apply List.Pairwise.imp_of_mem _ hp_disj + intro a b ha hb hab + rw [List.formPerm_disjoint_iff (hp_nodup a ha) (hp_nodup b hb) (hp2 a ha) (hp2 b hb)] + exact hab diff --git a/Mathlib/GroupTheory/Perm/Fin.lean b/Mathlib/GroupTheory/Perm/Fin.lean index 7a73e7bfe41be..4a47675d6ae71 100644 --- a/Mathlib/GroupTheory/Perm/Fin.lean +++ b/Mathlib/GroupTheory/Perm/Fin.lean @@ -103,7 +103,7 @@ theorem finRotate_succ_eq_decomposeFin {n : ℕ} : @[simp] theorem sign_finRotate (n : ℕ) : Perm.sign (finRotate (n + 1)) = (-1) ^ n := by induction' n with n ih - · simp + · simp; rfl · rw [finRotate_succ_eq_decomposeFin] simp [ih, pow_succ] #align sign_fin_rotate sign_finRotate diff --git a/Mathlib/GroupTheory/Perm/Support.lean b/Mathlib/GroupTheory/Perm/Support.lean index 31d02cf7ffbed..6a5437a5ab0af 100644 --- a/Mathlib/GroupTheory/Perm/Support.lean +++ b/Mathlib/GroupTheory/Perm/Support.lean @@ -142,7 +142,7 @@ theorem nodup_of_pairwise_disjoint {l : List (Perm α)} (h1 : (1 : Perm α) ∉ suffices (σ : Perm α) = 1 by rw [this] at h_mem exact h1 h_mem - exact ext fun a => (or_self_iff _).mp (h_disjoint a) + exact ext fun a => or_self_iff.mp (h_disjoint a) #align equiv.perm.nodup_of_pairwise_disjoint Equiv.Perm.nodup_of_pairwise_disjoint theorem pow_apply_eq_self_of_apply_eq_self {x : α} (hfx : f x = x) : ∀ n : ℕ, (f ^ n) x = x diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean new file mode 100644 index 0000000000000..1b7e59e870ab7 --- /dev/null +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -0,0 +1,185 @@ +/- +Copyright (c) 2023 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ + +import Mathlib.GroupTheory.CoprodI +import Mathlib.GroupTheory.Coprod.Basic +import Mathlib.GroupTheory.QuotientGroup +import Mathlib.GroupTheory.Complement + +/-! + +## Pushouts of Monoids and Groups + +This file defines wide pushouts of monoids and groups and proves some properties +of the amalgamated product of groups (i.e. the special case where all the maps +in the diagram are injective). + +## Main definitions + +- `Monoid.PushoutI`: the pushout of a diagram of monoids indexed by a type `ι` +- `Monoid.PushoutI.base`: the map from the amalgamating monoid to the pushout +- `Monoid.PushoutI.of`: the map from each Monoid in the family to the pushout +- `Monoid.PushoutI.lift`: the universal property used to define homomorphisms out of the pushout. + +## References + +* The normal form theorem follows these [notes](https://webspace.maths.qmul.ac.uk/i.m.chiswell/ggt/lecture_notes/lecture2.pdf) +from Queen Mary University + +## Tags + +amalgamated product, pushout, group + +-/ + +namespace Monoid + +open CoprodI Subgroup Coprod Function List + +variable {ι : Type*} {G : ι → Type*} {H : Type*} {K : Type*} [Monoid K] + +/-- The relation we quotient by to form the pushout -/ +def PushoutI.con [∀ i, Monoid (G i)] [Monoid H] (φ : ∀ i, H →* G i) : + Con (Coprod (CoprodI G) H) := + conGen (fun x y : Coprod (CoprodI G) H => + ∃ i x', x = inl (of (φ i x')) ∧ y = inr x') + +/-- The indexed pushout of monoids, which is the pushout in the category of monoids, +or the category of groups. -/ +def PushoutI [∀ i, Monoid (G i)] [Monoid H] (φ : ∀ i, H →* G i) : Type _ := + (PushoutI.con φ).Quotient + +namespace PushoutI + +section Monoid + +variable [∀ i, Monoid (G i)] [Monoid H] {φ : ∀ i, H →* G i} + +protected instance mul : Mul (PushoutI φ) := by + delta PushoutI; infer_instance + +protected instance one : One (PushoutI φ) := by + delta PushoutI; infer_instance + +instance monoid : Monoid (PushoutI φ) := + { Con.monoid _ with + toMul := PushoutI.mul + toOne := PushoutI.one } + +/-- The map from each indexing group into the pushout -/ +def of (i : ι) : G i →* PushoutI φ := + (Con.mk' _).comp <| inl.comp CoprodI.of + +variable (φ) in +/-- The map from the base monoid into the pushout -/ +def base : H →* PushoutI φ := + (Con.mk' _).comp inr + +theorem of_comp_eq_base (i : ι) : (of i).comp (φ i) = (base φ) := by + ext x + apply (Con.eq _).2 + refine ConGen.Rel.of _ _ ?_ + simp only [MonoidHom.comp_apply, Set.mem_iUnion, Set.mem_range] + exact ⟨_, _, rfl, rfl⟩ + +variable (φ) in +theorem of_apply_eq_base (i : ι) (x : H) : of i (φ i x) = base φ x := by + rw [← MonoidHom.comp_apply, of_comp_eq_base] + +/-- Define a homomorphism out of the pushout of monoids be defining it on each object in the +diagram -/ +def lift (f : ∀ i, G i →* K) (k : H →* K) + (hf : ∀ i, (f i).comp (φ i) = k) : + PushoutI φ →* K := + Con.lift _ (Coprod.lift (CoprodI.lift f) k) <| by + apply Con.conGen_le <| fun x y => ?_ + rintro ⟨i, x', rfl, rfl⟩ + simp only [FunLike.ext_iff, MonoidHom.coe_comp, comp_apply] at hf + simp [hf] + +@[simp] +theorem lift_of (f : ∀ i, G i →* K) (k : H →* K) + (hf : ∀ i, (f i).comp (φ i) = k) + {i : ι} (g : G i) : (lift f k hf) (of i g : PushoutI φ) = f i g := by + delta PushoutI lift of + simp only [MonoidHom.coe_comp, Con.coe_mk', comp_apply, Con.lift_coe, + lift_apply_inl, CoprodI.lift_of] + +@[simp] +theorem lift_base (f : ∀ i, G i →* K) (k : H →* K) + (hf : ∀ i, (f i).comp (φ i) = k) + (g : H) : (lift f k hf) (base φ g : PushoutI φ) = k g := by + delta PushoutI lift base + simp only [MonoidHom.coe_comp, Con.coe_mk', comp_apply, Con.lift_coe, lift_apply_inr] + +-- `ext` attribute should be lower priority then `hom_ext_nonempty` +@[ext 1199] +theorem hom_ext {f g : PushoutI φ →* K} + (h : ∀ i, f.comp (of i : G i →* _) = g.comp (of i : G i →* _)) + (hbase : f.comp (base φ) = g.comp (base φ)) : f = g := + (MonoidHom.cancel_right Con.mk'_surjective).mp <| + Coprod.hom_ext + (CoprodI.ext_hom _ _ h) + hbase + +@[ext high] +theorem hom_ext_nonempty [hn : Nonempty ι] + {f g : PushoutI φ →* K} + (h : ∀ i, f.comp (of i : G i →* _) = g.comp (of i : G i →* _)) : f = g := + hom_ext h <| by + cases hn with + | intro i => + ext + rw [← of_comp_eq_base i, ← MonoidHom.comp_assoc, h, MonoidHom.comp_assoc] + +/-- The equivalence that is part of the universal property of the pushout. A hom out of +the pushout is just a morphism out of all groups in the pushout that satsifies a commutativity +condition. -/ +@[simps] +def homEquiv : + (PushoutI φ →* K) ≃ { f : (Π i, G i →* K) × (H →* K) // ∀ i, (f.1 i).comp (φ i) = f.2 } := + { toFun := fun f => ⟨(fun i => f.comp (of i), f.comp (base φ)), + fun i => by rw [MonoidHom.comp_assoc, of_comp_eq_base]⟩ + invFun := fun f => lift f.1.1 f.1.2 f.2, + left_inv := fun _ => hom_ext (by simp [FunLike.ext_iff]) + (by simp [FunLike.ext_iff]) + right_inv := fun ⟨⟨_, _⟩, _⟩ => by simp [FunLike.ext_iff, Function.funext_iff] } + +/-- The map from the coproduct into the pushout -/ +def ofCoprodI : CoprodI G →* PushoutI φ := + CoprodI.lift of + +@[simp] +theorem ofCoprodI_of (i : ι) (g : G i) : + (ofCoprodI (CoprodI.of g) : PushoutI φ) = of i g := by + simp [ofCoprodI] + +theorem induction_on {motive : PushoutI φ → Prop} + (x : PushoutI φ) + (of : ∀ (i : ι) (g : G i), motive (of i g)) + (base : ∀ h, motive (base φ h)) + (mul : ∀ x y, motive x → motive y → motive (x * y)) : motive x := by + delta PushoutI PushoutI.of PushoutI.base at * + induction x using Con.induction_on with + | H x => + induction x using Coprod.induction_on with + | inl g => + induction g using CoprodI.induction_on with + | h_of i g => exact of i g + | h_mul x y ihx ihy => + rw [map_mul] + exact mul _ _ ihx ihy + | h_one => simpa using base 1 + | inr h => exact base h + | mul x y ihx ihy => exact mul _ _ ihx ihy + +end Monoid + +variable [∀ i, Group (G i)] [Group H] {φ : ∀ i, H →* G i} + +instance : Group (PushoutI φ) := + { Con.group (PushoutI.con φ) with + toMonoid := PushoutI.monoid } diff --git a/Mathlib/GroupTheory/SemidirectProduct.lean b/Mathlib/GroupTheory/SemidirectProduct.lean index 0b063f4e29578..a9e412baca896 100644 --- a/Mathlib/GroupTheory/SemidirectProduct.lean +++ b/Mathlib/GroupTheory/SemidirectProduct.lean @@ -3,9 +3,9 @@ Copyright (c) 2020 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Hom.Aut -import Mathlib.Logic.Function.Basic +import Mathlib.Algebra.Group.Aut import Mathlib.GroupTheory.Subgroup.Basic +import Mathlib.Logic.Function.Basic #align_import group_theory.semidirect_product from "leanprover-community/mathlib"@"f7fc89d5d5ff1db2d1242c7bb0e9062ce47ef47c" diff --git a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean index 3dcf4048c87ee..6fd738905a9b1 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean @@ -262,6 +262,7 @@ theorem normalClosure_swap_mul_swap_five : ⟨finRotate 5, finRotate_bit1_mem_alternatingGroup (n := 2)⟩ := by rw [Subtype.ext_iff] simp only [Fin.val_mk, Subgroup.coe_mul, Subgroup.coe_inv, Fin.val_mk] + decide rw [eq_top_iff, ← normalClosure_finRotate_five] refine' normalClosure_le_normal _ rw [Set.singleton_subset_iff, SetLike.mem_coe, ← h5] @@ -280,7 +281,7 @@ theorem isConj_swap_mul_swap_of_cycleType_two {g : Perm (Fin 5)} (ha : g ∈ alt rw [← Multiset.eq_replicate_card] at h2 rw [← sum_cycleType, h2, Multiset.sum_replicate, smul_eq_mul] at h have h : Multiset.card g.cycleType ≤ 3 := - le_of_mul_le_mul_right (le_trans h (by simp only [card_fin]; ring_nf)) (by simp) + le_of_mul_le_mul_right (le_trans h (by simp only [card_fin]; ring_nf; decide)) (by simp) rw [mem_alternatingGroup, sign_of_cycleType, h2] at ha norm_num at ha rw [pow_add, pow_mul, Int.units_pow_two, one_mul, Units.ext_iff, Units.val_one, @@ -299,7 +300,7 @@ theorem isConj_swap_mul_swap_of_cycleType_two {g : Perm (Fin 5)} (ha : g ∈ alt · rw [disjoint_iff_disjoint_support, support_swap h04, support_swap h13] decide · contrapose! ha - simp [h_1] + decide #align alternating_group.is_conj_swap_mul_swap_of_cycle_type_two alternatingGroup.isConj_swap_mul_swap_of_cycleType_two /-- Shows that $A_5$ is simple by taking an arbitrary non-identity element and showing by casework diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 26c8a68654eeb..f070fecf3d77e 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -602,4 +602,13 @@ theorem IsCyclic.exponent_eq_zero_of_infinite [Group α] [IsCyclic α] [Infinite #align is_cyclic.exponent_eq_zero_of_infinite IsCyclic.exponent_eq_zero_of_infinite #align is_add_cyclic.exponent_eq_zero_of_infinite IsAddCyclic.exponent_eq_zero_of_infinite +instance ZMod.instIsAddCyclic (n : ℕ) : IsAddCyclic (ZMod n) where + exists_generator := ⟨1, fun n ↦ ⟨n, by simp⟩⟩ + +@[simp] +protected theorem ZMod.exponent (n : ℕ) : AddMonoid.exponent (ZMod n) = n := by + cases n + · rw [IsAddCyclic.exponent_eq_zero_of_infinite] + · rw [IsAddCyclic.exponent_eq_card, card] + end Exponent diff --git a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean index 962092e3c2155..df50ccc519fcc 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean @@ -118,7 +118,7 @@ instance : Infinite (DihedralGroup 0) := DihedralGroup.fintypeHelper.infinite_iff.mp inferInstance instance : Nontrivial (DihedralGroup n) := - ⟨⟨r 0, sr 0, by simp_rw [ne_eq]⟩⟩ + ⟨⟨r 0, sr 0, by simp_rw [ne_eq, not_false_eq_true]⟩⟩ /-- If `0 < n`, then `DihedralGroup n` has `2n` elements. -/ @@ -161,7 +161,7 @@ theorem orderOf_sr (i : ZMod n) : orderOf (sr i) = 2 := by · rw [sq, sr_mul_self] · -- Porting note: Previous proof was `decide` revert n - simp_rw [one_def, ne_eq, forall_const] + simp_rw [one_def, ne_eq, forall_const, not_false_eq_true] #align dihedral_group.order_of_sr DihedralGroup.orderOf_sr /-- If `0 < n`, then `r 1` has order `n`. diff --git a/Mathlib/GroupTheory/Submonoid/Basic.lean b/Mathlib/GroupTheory/Submonoid/Basic.lean index af445b913fac0..9883eb26506f9 100644 --- a/Mathlib/GroupTheory/Submonoid/Basic.lean +++ b/Mathlib/GroupTheory/Submonoid/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Kenny Lau, Johan Commelin, Mario Carneiro, Kevin Buzzard, Amelia Livingston, Yury Kudryashov -/ -import Mathlib.Algebra.Hom.Group.Defs +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Algebra.Group.Units import Mathlib.GroupTheory.Subsemigroup.Basic diff --git a/Mathlib/GroupTheory/Subsemigroup/Basic.lean b/Mathlib/GroupTheory/Subsemigroup/Basic.lean index 6b04eacf451fc..dc9ac4426d72b 100644 --- a/Mathlib/GroupTheory/Subsemigroup/Basic.lean +++ b/Mathlib/GroupTheory/Subsemigroup/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Kenny Lau, Johan Commelin, Mario Carneiro, Kevin Buzzard, Amelia Livingston, Yury Kudryashov, Yakov Pechersky -/ -import Mathlib.Algebra.Hom.Group.Defs +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Data.Set.Lattice import Mathlib.Data.SetLike.Basic diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index 81ef351ccf18b..0b32be6cc8f3f 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -503,7 +503,7 @@ namespace Sylow theorem mem_fixedPoints_mul_left_cosets_iff_mem_normalizer {H : Subgroup G} [Finite (H : Set G)] {x : G} : (x : G ⧸ H) ∈ MulAction.fixedPoints H (G ⧸ H) ↔ x ∈ normalizer H := ⟨fun hx => - have ha : ∀ {y : G ⧸ H}, y ∈ orbit H (x : G ⧸ H) → y = x := (mem_fixedPoints' _).1 hx _ + have ha : ∀ {y : G ⧸ H}, y ∈ orbit H (x : G ⧸ H) → y = x := mem_fixedPoints'.1 hx _ (inv_mem_iff (G := G)).1 (mem_normalizer_fintype fun n (hn : n ∈ H) => have : (n⁻¹ * x)⁻¹ * x ∈ H := QuotientGroup.eq.1 (ha ⟨⟨n⁻¹, inv_mem hn⟩, rfl⟩) @@ -512,7 +512,7 @@ theorem mem_fixedPoints_mul_left_cosets_iff_mem_normalizer {H : Subgroup G} [Fin convert this rw [inv_inv]), fun hx : ∀ n : G, n ∈ H ↔ x * n * x⁻¹ ∈ H => - (mem_fixedPoints' _).2 fun y => + mem_fixedPoints'.2 fun y => Quotient.inductionOn' y fun y hy => QuotientGroup.eq.2 (let ⟨⟨b, hb₁⟩, hb₂⟩ := hy diff --git a/Mathlib/Init/CCLemmas.lean b/Mathlib/Init/CCLemmas.lean index 6f63b733b01d9..22229ba15cba4 100644 --- a/Mathlib/Init/CCLemmas.lean +++ b/Mathlib/Init/CCLemmas.lean @@ -31,7 +31,7 @@ theorem and_eq_of_eq_false_right {a b : Prop} (h : b = False) : (a ∧ b) = Fals h.symm ▸ propext (and_false_iff _) theorem and_eq_of_eq {a b : Prop} (h : a = b) : (a ∧ b) = a := - h ▸ propext (and_self_iff _) + h ▸ propext and_self_iff theorem or_eq_of_eq_true_left {a b : Prop} (h : a = True) : (a ∨ b) = True := h.symm ▸ propext (true_or_iff _) @@ -46,7 +46,7 @@ theorem or_eq_of_eq_false_right {a b : Prop} (h : b = False) : (a ∨ b) = a := h.symm ▸ propext (or_false_iff _) theorem or_eq_of_eq {a b : Prop} (h : a = b) : (a ∨ b) = a := - h ▸ propext (or_self_iff _) + h ▸ propext or_self_iff theorem imp_eq_of_eq_true_left {a b : Prop} (h : a = True) : (a → b) = b := h.symm ▸ propext ⟨fun h ↦ h trivial, fun h₁ _ ↦ h₁⟩ diff --git a/Mathlib/Init/Data/Bool/Basic.lean b/Mathlib/Init/Data/Bool/Basic.lean index b5af62ef4ed96..c974d399d6d7e 100644 --- a/Mathlib/Init/Data/Bool/Basic.lean +++ b/Mathlib/Init/Data/Bool/Basic.lean @@ -4,24 +4,16 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ import Mathlib.Mathport.Rename +import Std.Data.Bool /-! # Boolean operations -In Lean 3 this file also contained the definitions of `cond`, `bor`, `band` and `bnot`, -the boolean functions. These are in Lean 4 core (as `cond`, `or`, `and` and `not`), but -apparently `xor` didn't make the cut. - +In Lean 3 this file contained the definitions of `cond`, `bor`, `band` and `bnot`, +the boolean functions. These are now in Lean 4 core or Std (as `cond`, `or`, `and`, `not`, `xor`). -/ #align bor or #align band and #align bnot not - -/-- Boolean XOR -/ -@[inline] -def xor : Bool → Bool → Bool - | true, false => true - | false, true => true - | _, _ => false #align bxor xor diff --git a/Mathlib/Init/Data/Bool/Lemmas.lean b/Mathlib/Init/Data/Bool/Lemmas.lean index ce4db27021822..bc1078774c302 100644 --- a/Mathlib/Init/Data/Bool/Lemmas.lean +++ b/Mathlib/Init/Data/Bool/Lemmas.lean @@ -47,22 +47,12 @@ namespace Bool theorem cond_self.{u} {α : Type u} (b : Bool) (a : α) : cond b a a = a := by cases b <;> rfl #align cond_a_a Bool.cond_self -@[simp] -theorem xor_self (b : Bool) : xor b b = false := by cases b <;> simp +attribute [simp] xor_self #align bxor_self Bool.xor_self -@[simp] -theorem xor_true (b : Bool) : xor b true = not b := by cases b <;> simp #align bxor_tt Bool.xor_true - -theorem xor_false (b : Bool) : xor b false = b := by cases b <;> simp #align bxor_ff Bool.xor_false - -@[simp] -theorem true_xor (b : Bool) : xor true b = not b := by cases b <;> simp #align tt_bxor Bool.true_xor - -theorem false_xor (b : Bool) : xor false b = b := by cases b <;> simp #align ff_bxor Bool.false_xor theorem true_eq_false_eq_False : ¬true = false := by decide @@ -167,7 +157,6 @@ theorem coe_or_iff (a b : Bool) : a || b ↔ a ∨ b := by simp theorem coe_and_iff (a b : Bool) : a && b ↔ a ∧ b := by simp #align band_coe_iff Bool.coe_and_iff -@[simp] theorem coe_xor_iff (a b : Bool) : xor a b ↔ Xor' (a = true) (b = true) := by cases a <;> cases b <;> exact by decide #align bxor_coe_iff Bool.coe_xor_iff diff --git a/Mathlib/Init/Data/List/Lemmas.lean b/Mathlib/Init/Data/List/Lemmas.lean index 44e71417a45ae..58423c8980808 100644 --- a/Mathlib/Init/Data/List/Lemmas.lean +++ b/Mathlib/Init/Data/List/Lemmas.lean @@ -176,7 +176,7 @@ theorem length_mapAccumr₂ : calc succ (length (mapAccumr₂ f x y c).2) = succ (min (length x) (length y)) := congr_arg succ (length_mapAccumr₂ f x y c) - _ = min (succ (length x)) (succ (length y)) := Eq.symm (min_succ_succ (length x) (length y)) + _ = min (succ (length x)) (succ (length y)) := Eq.symm (succ_min_succ (length x) (length y)) | _, _ :: _, [], _ => rfl | _, [], _ :: _, _ => rfl | _, [], [], _ => rfl diff --git a/Mathlib/Init/Data/Nat/Bitwise.lean b/Mathlib/Init/Data/Nat/Bitwise.lean index 4d12513f11fc3..bb17d49565d4b 100644 --- a/Mathlib/Init/Data/Nat/Bitwise.lean +++ b/Mathlib/Init/Data/Nat/Bitwise.lean @@ -73,10 +73,7 @@ theorem bodd_succ (n : ℕ) : bodd (succ n) = not (bodd n) := by @[simp] theorem bodd_add (m n : ℕ) : bodd (m + n) = bxor (bodd m) (bodd n) := by - induction' n with n IH - · simp - · simp [add_succ, IH] - cases bodd m <;> cases bodd n <;> rfl + induction n <;> simp_all [add_succ, Bool.xor_not] #align nat.bodd_add Nat.bodd_add @[simp] @@ -274,7 +271,7 @@ theorem binaryRec_zero {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (b theorem bodd_bit (b n) : bodd (bit b n) = b := by rw [bit_val] simp only [Nat.mul_comm, Nat.add_comm, bodd_add, bodd_mul, bodd_succ, bodd_zero, Bool.not_false, - Bool.not_true, Bool.and_false, Bool.xor_false_right] + Bool.not_true, Bool.and_false, Bool.xor_false] cases b <;> cases bodd n <;> rfl #align nat.bodd_bit Nat.bodd_bit diff --git a/Mathlib/Init/Data/Nat/Lemmas.lean b/Mathlib/Init/Data/Nat/Lemmas.lean index f6911cd52868f..eef918bc45fa2 100644 --- a/Mathlib/Init/Data/Nat/Lemmas.lean +++ b/Mathlib/Init/Data/Nat/Lemmas.lean @@ -187,9 +187,9 @@ instance linearOrder : LinearOrder ℕ where #align nat.le_succ_of_pred_le Nat.le_succ_of_pred_le -#align nat.le_lt_antisymm Nat.le_lt_antisymm +#align nat.le_lt_antisymm Nat.le_lt_asymm -#align nat.lt_le_antisymm Nat.lt_le_antisymm +#align nat.lt_le_antisymm Nat.lt_le_asymm #align nat.lt_asymm Nat.lt_asymm @@ -470,7 +470,7 @@ Many lemmas are proven more generally in mathlib `algebra/order/sub` -/ #align nat.min_zero Nat.min_zero -#align nat.min_succ_succ Nat.min_succ_succ +#align nat.succ_min_succ Nat.succ_min_succ #align nat.sub_eq_sub_min Nat.sub_eq_sub_min @@ -784,11 +784,11 @@ lemma to_digits_core_length (b : Nat) (h : 2 <= b) (f n e : Nat) to_digits_core_lens_eq b f (n / b) (Nat.digitChar $ n % b), if_false] exact Nat.succ_le_succ ih case neg => - have _ : e = 0 := Nat.eq_zero_of_nonpos e h_pred_pos + have _ : e = 0 := Nat.eq_zero_of_not_pos h_pred_pos rw [‹e = 0›] have _ : b ^ 1 = b := by simp only [pow_succ, pow_zero, Nat.one_mul] have _ : n < b := ‹b ^ 1 = b› ▸ (‹e = 0› ▸ hlt : n < b ^ Nat.succ 0) - simp only [(@Nat.div_eq_of_lt n b ‹n < b› : n / b = 0), if_true, List.length] + simp [(@Nat.div_eq_of_lt n b ‹n < b› : n / b = 0)] /-- The core implementation of `Nat.repr` returns a String with length less than or equal to the number of digits in the decimal number (represented by `e`). For example, the decimal string diff --git a/Mathlib/Init/Logic.lean b/Mathlib/Init/Logic.lean index 3d6310420a639..1df648620747d 100644 --- a/Mathlib/Init/Logic.lean +++ b/Mathlib/Init/Logic.lean @@ -152,7 +152,6 @@ theorem false_and_iff : False ∧ p ↔ False := iff_of_eq (false_and _) #align false_and false_and_iff #align not_and_self not_and_self_iff #align and_not_self and_not_self_iff -theorem and_self_iff : p ∧ p ↔ p := iff_of_eq (and_self _) #align and_self and_self_iff #align or.imp Or.impₓ -- reorder implicits @@ -187,7 +186,6 @@ theorem false_or_iff : False ∨ p ↔ p := iff_of_eq (false_or _) #align false_or false_or_iff theorem or_false_iff : p ∨ False ↔ p := iff_of_eq (or_false _) #align or_false or_false_iff -theorem or_self_iff : p ∨ p ↔ p := iff_of_eq (or_self _) #align or_self or_self_iff theorem not_or_of_not : ¬a → ¬b → ¬(a ∨ b) := fun h1 h2 ↦ not_or.2 ⟨h1, h2⟩ diff --git a/Mathlib/Init/Meta/WellFoundedTactics.lean b/Mathlib/Init/Meta/WellFoundedTactics.lean index 190c9d69b0c17..ccf9e526c6f8b 100644 --- a/Mathlib/Init/Meta/WellFoundedTactics.lean +++ b/Mathlib/Init/Meta/WellFoundedTactics.lean @@ -16,11 +16,7 @@ theorem Nat.lt_add_of_zero_lt_left (a b : Nat) (h : 0 < b) : a < a + b := assumption #align nat.lt_add_of_zero_lt_left Nat.lt_add_of_zero_lt_left -theorem Nat.zero_lt_one_add (a : Nat) : 0 < 1 + a := - suffices 0 < a + 1 by - simp only [Nat.add_comm] - assumption - Nat.zero_lt_succ _ +theorem Nat.zero_lt_one_add (a : Nat) : 0 < 1 + a := by simp [Nat.one_add] #align nat.zero_lt_one_add Nat.zero_lt_one_add #align nat.lt_add_right Nat.lt_add_right diff --git a/Mathlib/Init/Order/Defs.lean b/Mathlib/Init/Order/Defs.lean index 71354059a3a5b..b823526c32e79 100644 --- a/Mathlib/Init/Order/Defs.lean +++ b/Mathlib/Init/Order/Defs.lean @@ -262,7 +262,7 @@ implicit arguments, requires us to unfold the defs and split the `if`s in the de macro "compareOfLessAndEq_rfl" : tactic => `(tactic| (intros a b; first | rfl | (simp only [compare, compareOfLessAndEq]; split_ifs <;> rfl) | - (induction a <;> induction b <;> simp only []))) + (induction a <;> induction b <;> simp (config := {decide := true}) only []))) /-- A linear order is reflexive, transitive, antisymmetric and total relation `≤`. We assume that every linear ordered type has decidable `(≤)`, `(<)`, and `(=)`. -/ @@ -436,16 +436,16 @@ theorem compare_eq_iff_eq {a b : α} : (compare a b = .eq) ↔ a = b := by case _ _ h => exact false_iff_iff.2 h theorem compare_le_iff_le {a b : α} : (compare a b ≠ .gt) ↔ a ≤ b := by - cases h : compare a b <;> simp only [] - · exact true_iff_iff.2 <| le_of_lt <| compare_lt_iff_lt.1 h - · exact true_iff_iff.2 <| le_of_eq <| compare_eq_iff_eq.1 h - · exact false_iff_iff.2 <| not_le_of_gt <| compare_gt_iff_gt.1 h + cases h : compare a b <;> simp + · exact le_of_lt <| compare_lt_iff_lt.1 h + · exact le_of_eq <| compare_eq_iff_eq.1 h + · exact compare_gt_iff_gt.1 h theorem compare_ge_iff_ge {a b : α} : (compare a b ≠ .lt) ↔ a ≥ b := by - cases h : compare a b <;> simp only [] - · exact false_iff_iff.2 <| (lt_iff_not_ge a b).1 <| compare_lt_iff_lt.1 h - · exact true_iff_iff.2 <| le_of_eq <| (·.symm) <| compare_eq_iff_eq.1 h - · exact true_iff_iff.2 <| le_of_lt <| compare_gt_iff_gt.1 h + cases h : compare a b <;> simp + · exact compare_lt_iff_lt.1 h + · exact le_of_eq <| (·.symm) <| compare_eq_iff_eq.1 h + · exact le_of_lt <| compare_gt_iff_gt.1 h theorem compare_iff (a b : α) {o : Ordering} : compare a b = o ↔ o.toRel a b := by cases o <;> simp only [Ordering.toRel] diff --git a/Mathlib/Init/Set.lean b/Mathlib/Init/Set.lean index 9f5c2fda08f43..5eb2fd967e8dd 100644 --- a/Mathlib/Init/Set.lean +++ b/Mathlib/Init/Set.lean @@ -127,4 +127,11 @@ instance : LawfulFunctor Set where λ ⟨_, ⟨⟨a, ⟨h₁, h₂⟩⟩, h₃⟩⟩ => ⟨a, ⟨h₁, show h (g a) = c from h₂ ▸ h₃⟩⟩⟩ map_const := rfl +/-- The property `s.Nonempty` expresses the fact that the set `s` is not empty. It should be used +in theorem assumptions instead of `∃ x, x ∈ s` or `s ≠ ∅` as it gives access to a nice API thanks +to the dot notation. -/ +protected def Nonempty (s : Set α) : Prop := + ∃ x, x ∈ s +#align set.nonempty Set.Nonempty + end Set diff --git a/Mathlib/Lean/Elab/Term.lean b/Mathlib/Lean/Elab/Term.lean new file mode 100644 index 0000000000000..97faafafde5f3 --- /dev/null +++ b/Mathlib/Lean/Elab/Term.lean @@ -0,0 +1,23 @@ +/- +Copyright (c) 2023 Kyle Miller. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kyle Miller +-/ +import Lean.Elab + +/-! +# Additions to `Lean.Elab.Term` +-/ + +namespace Lean.Elab.Term + +/-- Fully elaborates the term `patt`, allowing typeclass inference failure, +but while setting `errToSorry` to false. +Typeclass failures result in plain metavariables. +Instantiates all assigned metavariables. -/ +def elabPattern (patt : Term) (expectedType? : Option Expr) : TermElabM Expr := do + withTheReader Term.Context ({ · with ignoreTCFailures := true, errToSorry := false }) <| + withSynthesizeLight do + let t ← elabTerm patt expectedType? + synthesizeSyntheticMVars (mayPostpone := false) (ignoreStuckTC := true) + instantiateMVars t diff --git a/Mathlib/Lean/Meta/DiscrTree.lean b/Mathlib/Lean/Meta/DiscrTree.lean index d41556daad7cf..90913a9df7895 100644 --- a/Mathlib/Lean/Meta/DiscrTree.lean +++ b/Mathlib/Lean/Meta/DiscrTree.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Lean.Meta.DiscrTree +import Std.Lean.Meta.DiscrTree import Mathlib.Lean.Expr.Traverse /-! @@ -14,17 +15,6 @@ set_option autoImplicit true namespace Lean.Meta.DiscrTree -/-- -Inserts a new key into a discrimination tree, -but only if it is not of the form `#[*]` or `#[=, *, *, *]`. --/ -def insertIfSpecific [BEq α] (d : DiscrTree α s) - (keys : Array (DiscrTree.Key s)) (v : α) : DiscrTree α s := - if keys == #[Key.star] || keys == #[Key.const `Eq 3, Key.star, Key.star, Key.star] then - d - else - d.insertCore keys v - /-- Find keys which match the expression, or some subexpression. @@ -36,48 +26,21 @@ Implementation: we reverse the results from `getMatch`, so that we return lemmas matching larger subexpressions first, and amongst those we return more specific lemmas first. -/ -partial def getSubexpressionMatches (d : DiscrTree α s) (e : Expr) : MetaM (Array α) := do +partial def getSubexpressionMatches (d : DiscrTree α) (e : Expr) (config : WhnfCoreConfig) : + MetaM (Array α) := do match e with | .bvar _ => return #[] | .forallE _ _ _ _ => forallTelescope e (fun args body => do args.foldlM (fun acc arg => do - pure <| acc ++ (← d.getSubexpressionMatches (← inferType arg))) - (← d.getSubexpressionMatches body).reverse) + pure <| acc ++ (← d.getSubexpressionMatches (← inferType arg) config)) + (← d.getSubexpressionMatches body config).reverse) | .lam _ _ _ _ | .letE _ _ _ _ _ => lambdaLetTelescope e (fun args body => do args.foldlM (fun acc arg => do - pure <| acc ++ (← d.getSubexpressionMatches (← inferType arg))) - (← d.getSubexpressionMatches body).reverse) + pure <| acc ++ (← d.getSubexpressionMatches (← inferType arg) config)) + (← d.getSubexpressionMatches body config).reverse) | _ => - e.foldlM (fun a f => do pure <| a ++ (← d.getSubexpressionMatches f)) (← d.getMatch e).reverse - -variable {m : Type → Type} [Monad m] - - -/-- The explicit stack of `Trie.mapArrays` -/ -private inductive Ctxt {α β s} - | empty : Ctxt - | ctxt : Array (Key s × Trie β s) → Array β → Array (Key s × Trie α s) → Key s → Ctxt → Ctxt - -/-- Apply a function to the array of values at each node in a `DiscrTree`. -/ -partial def Trie.mapArrays (t : Trie α s) (f : Array α → Array β) : Trie β s := - let .node vs0 cs0 := t - go (.mkEmpty cs0.size) (f vs0) cs0.reverse Ctxt.empty -where - /-- This implementation as a single tail-recursive function is chosen to not blow the - interpreter stack when the `Trie` is very deep -/ - go cs vs todo ps := - if todo.isEmpty then - let c := .node vs cs - match ps with - | .empty => c - | .ctxt cs' vs' todo k ps => go (cs'.push (k, c)) vs' todo ps - else - let (k, .node vs' cs') := todo.back - go (.mkEmpty cs'.size) (f vs') cs'.reverse (.ctxt cs vs todo.pop k ps) - -/-- Apply a function to the array of values at each node in a `DiscrTree`. -/ -def mapArrays (d : DiscrTree α s) (f : Array α → Array β) : DiscrTree β s := - { root := d.root.map (fun t => t.mapArrays f) } + e.foldlM (fun a f => do + pure <| a ++ (← d.getSubexpressionMatches f config)) (← d.getMatch e config).reverse end Lean.Meta.DiscrTree diff --git a/Mathlib/Lean/PrettyPrinter/Delaborator.lean b/Mathlib/Lean/PrettyPrinter/Delaborator.lean new file mode 100644 index 0000000000000..f3fe10b4c90cb --- /dev/null +++ b/Mathlib/Lean/PrettyPrinter/Delaborator.lean @@ -0,0 +1,48 @@ +/- +Copyright (c) 2023 Kyle Miller. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kyle Miller +-/ +import Lean.PrettyPrinter.Delaborator.Basic + +/-! +# Additions to the delaborator +-/ + +namespace Lean.PrettyPrinter.Delaborator + +open Lean.Meta Lean.SubExpr SubExpr + +namespace SubExpr + +variable {α : Type} [Inhabited α] +variable {m : Type → Type} [Monad m] + +section Descend + +variable [MonadReaderOf SubExpr m] [MonadWithReaderOf SubExpr m] +variable [MonadLiftT MetaM m] [MonadControlT MetaM m] +variable [MonadLiftT IO m] + +/-- Assuming the current expression is a lambda or pi, +descend into the body using the given name `n` for the username of the fvar. +Provides `x` with the fresh fvar for the bound variable. +See also `Lean.PrettyPrinter.Delaborator.SubExpr.withBindingBody`. -/ +def withBindingBody' (n : Name) (x : Expr → m α) : m α := do + let e ← getExpr + Meta.withLocalDecl n e.binderInfo e.bindingDomain! fun fvar => + descend (e.bindingBody!.instantiate1 fvar) 1 (x fvar) + +end Descend + +end SubExpr + +/-- Assuming the current expression in a lambda or pi, +descend into the body using an unused name generated from the binder's name. +Provides `d` with both `Syntax` for the bound name as an identifier +as well as the fresh fvar for the bound variable. +See also `Lean.PrettyPrinter.Delaborator.withBindingBodyUnusedName`. -/ +def withBindingBodyUnusedName' {α} (d : Syntax → Expr → DelabM α) : DelabM α := do + let n ← getUnusedName (← getExpr).bindingName! (← getExpr).bindingBody! + let stxN ← annotateCurPos (mkIdent n) + withBindingBody' n $ d stxN diff --git a/Mathlib/Lean/Thunk.lean b/Mathlib/Lean/Thunk.lean index 2a314f96db090..5da21d33be8e6 100644 --- a/Mathlib/Lean/Thunk.lean +++ b/Mathlib/Lean/Thunk.lean @@ -30,10 +30,17 @@ instance {α : Type u} [DecidableEq α] : DecidableEq (Thunk α) := by rw [this] infer_instance -/-- The product of two thunks. -/ +/-- The cartesian product of two thunks. -/ def prod (a : Thunk α) (b : Thunk β) : Thunk (α × β) := Thunk.mk fun _ => (a.get, b.get) @[simp] theorem prod_get_fst : (prod a b).get.1 = a.get := rfl @[simp] theorem prod_get_snd : (prod a b).get.2 = b.get := rfl +/-- The sum of two thunks. -/ +def add [Add α] (a b : Thunk α) : Thunk α := Thunk.mk fun _ => a.get + b.get + +instance [Add α] : Add (Thunk α) := ⟨add⟩ + +@[simp] theorem add_get [Add α] {a b : Thunk α} : (a + b).get = a.get + b.get := rfl + end Thunk diff --git a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean index 14819d2553db9..381e4f20e2fe0 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean @@ -479,7 +479,8 @@ theorem affineIndependent_iff_not_collinear_of_ne {p : Fin 3 → P} {i₁ i₂ i AffineIndependent k p ↔ ¬Collinear k ({p i₁, p i₂, p i₃} : Set P) := by have hu : (Finset.univ : Finset (Fin 3)) = {i₁, i₂, i₃} := by -- Porting note: Originally `by decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> simp only at h₁₂ h₁₃ h₂₃ ⊢ + fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ + <;> simp (config := {decide := true}) only at h₁₂ h₁₃ h₂₃ ⊢ rw [affineIndependent_iff_not_collinear, ← Set.image_univ, ← Finset.coe_univ, hu, Finset.coe_insert, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_pair] #align affine_independent_iff_not_collinear_of_ne affineIndependent_iff_not_collinear_of_ne diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 3b32e603f4928..f162ad0dbda54 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -688,7 +688,7 @@ theorem affineIndependent_of_ne_of_mem_of_mem_of_not_mem {s : AffineSubspace k P refine' hp₃ ((AffineSubspace.le_def' _ s).1 _ p₃ h) simp_rw [affineSpan_le, Set.image_subset_iff, Set.subset_def, Set.mem_preimage] intro x - fin_cases x <;> simp [hp₁, hp₂] + fin_cases x <;> simp (config := {decide := true}) [hp₁, hp₂] #align affine_independent_of_ne_of_mem_of_mem_of_not_mem affineIndependent_of_ne_of_mem_of_mem_of_not_mem /-- If distinct points `p₁` and `p₃` lie in `s` but `p₂` does not, the three points are affinely diff --git a/Mathlib/LinearAlgebra/Basic.lean b/Mathlib/LinearAlgebra/Basic.lean index d6af9170d4bac..6896c11de78b5 100644 --- a/Mathlib/LinearAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/Basic.lean @@ -7,7 +7,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov, Fréd import Mathlib.Algebra.BigOperators.Pi import Mathlib.Algebra.Module.Hom import Mathlib.Algebra.Module.Prod -import Mathlib.Algebra.Module.Submodule.Lattice +import Mathlib.Algebra.Module.Submodule.Map import Mathlib.Algebra.Module.LinearMap import Mathlib.Data.DFinsupp.Basic import Mathlib.Data.Finsupp.Basic @@ -549,391 +549,6 @@ theorem subtype_comp_ofLe (p q : Submodule R M) (h : p ≤ q) : rfl #align submodule.subtype_comp_of_le Submodule.subtype_comp_ofLe -section - -variable [RingHomSurjective σ₁₂] {F : Type*} [sc : SemilinearMapClass F σ₁₂ M M₂] - -/-- The pushforward of a submodule `p ⊆ M` by `f : M → M₂` -/ -def map (f : F) (p : Submodule R M) : Submodule R₂ M₂ := - { p.toAddSubmonoid.map f with - carrier := f '' p - smul_mem' := by - rintro c x ⟨y, hy, rfl⟩ - obtain ⟨a, rfl⟩ := σ₁₂.surjective c - exact ⟨_, p.smul_mem a hy, map_smulₛₗ f _ _⟩ } -#align submodule.map Submodule.map - -@[simp] -theorem map_coe (f : F) (p : Submodule R M) : (map f p : Set M₂) = f '' p := - rfl -#align submodule.map_coe Submodule.map_coe - -theorem map_toAddSubmonoid (f : M →ₛₗ[σ₁₂] M₂) (p : Submodule R M) : - (p.map f).toAddSubmonoid = p.toAddSubmonoid.map (f : M →+ M₂) := - SetLike.coe_injective rfl -#align submodule.map_to_add_submonoid Submodule.map_toAddSubmonoid - -theorem map_toAddSubmonoid' (f : M →ₛₗ[σ₁₂] M₂) (p : Submodule R M) : - (p.map f).toAddSubmonoid = p.toAddSubmonoid.map f := - SetLike.coe_injective rfl -#align submodule.map_to_add_submonoid' Submodule.map_toAddSubmonoid' - -@[simp] -theorem _root_.AddMonoidHom.coe_toIntLinearMap_map {A A₂ : Type*} [AddCommGroup A] [AddCommGroup A₂] - (f : A →+ A₂) (s : AddSubgroup A) : - (AddSubgroup.toIntSubmodule s).map f.toIntLinearMap = - AddSubgroup.toIntSubmodule (s.map f) := rfl - -@[simp] -theorem _root_.MonoidHom.coe_toAdditive_map {G G₂ : Type*} [Group G] [Group G₂] (f : G →* G₂) - (s : Subgroup G) : - s.toAddSubgroup.map (MonoidHom.toAdditive f) = Subgroup.toAddSubgroup (s.map f) := rfl - -@[simp] -theorem _root_.AddMonoidHom.coe_toMultiplicative_map {G G₂ : Type*} [AddGroup G] [AddGroup G₂] - (f : G →+ G₂) (s : AddSubgroup G) : - s.toSubgroup.map (AddMonoidHom.toMultiplicative f) = AddSubgroup.toSubgroup (s.map f) := rfl - -@[simp] -theorem mem_map {f : F} {p : Submodule R M} {x : M₂} : x ∈ map f p ↔ ∃ y, y ∈ p ∧ f y = x := - Iff.rfl -#align submodule.mem_map Submodule.mem_map - -theorem mem_map_of_mem {f : F} {p : Submodule R M} {r} (h : r ∈ p) : f r ∈ map f p := - Set.mem_image_of_mem _ h -#align submodule.mem_map_of_mem Submodule.mem_map_of_mem - -theorem apply_coe_mem_map (f : F) {p : Submodule R M} (r : p) : f r ∈ map f p := - mem_map_of_mem r.prop -#align submodule.apply_coe_mem_map Submodule.apply_coe_mem_map - -@[simp] -theorem map_id : map (LinearMap.id : M →ₗ[R] M) p = p := - Submodule.ext fun a => by simp -#align submodule.map_id Submodule.map_id - -theorem map_comp [RingHomSurjective σ₂₃] [RingHomSurjective σ₁₃] (f : M →ₛₗ[σ₁₂] M₂) - (g : M₂ →ₛₗ[σ₂₃] M₃) (p : Submodule R M) : map (g.comp f : M →ₛₗ[σ₁₃] M₃) p = map g (map f p) := - SetLike.coe_injective <| by simp only [← image_comp, map_coe, LinearMap.coe_comp, comp_apply] -#align submodule.map_comp Submodule.map_comp - -theorem map_mono {f : F} {p p' : Submodule R M} : p ≤ p' → map f p ≤ map f p' := - image_subset _ -#align submodule.map_mono Submodule.map_mono - -@[simp] -theorem map_zero : map (0 : M →ₛₗ[σ₁₂] M₂) p = ⊥ := - have : ∃ x : M, x ∈ p := ⟨0, p.zero_mem⟩ - ext <| by simp [this, eq_comm] -#align submodule.map_zero Submodule.map_zero - -theorem map_add_le (f g : M →ₛₗ[σ₁₂] M₂) : map (f + g) p ≤ map f p ⊔ map g p := by - rintro x ⟨m, hm, rfl⟩ - exact add_mem_sup (mem_map_of_mem hm) (mem_map_of_mem hm) -#align submodule.map_add_le Submodule.map_add_le - -theorem map_inf_le (f : F) {p q : Submodule R M} : - (p ⊓ q).map f ≤ p.map f ⊓ q.map f := - image_inter_subset f p q - -theorem map_inf (f : F) {p q : Submodule R M} (hf : Injective f) : - (p ⊓ q).map f = p.map f ⊓ q.map f := - SetLike.coe_injective <| Set.image_inter hf - -theorem range_map_nonempty (N : Submodule R M) : - (Set.range (fun ϕ => Submodule.map ϕ N : (M →ₛₗ[σ₁₂] M₂) → Submodule R₂ M₂)).Nonempty := - ⟨_, Set.mem_range.mpr ⟨0, rfl⟩⟩ -#align submodule.range_map_nonempty Submodule.range_map_nonempty - -end - -section SemilinearMap - -variable {F : Type*} [sc : SemilinearMapClass F σ₁₂ M M₂] - -/-- The pushforward of a submodule by an injective linear map is -linearly equivalent to the original submodule. See also `LinearEquiv.submoduleMap` for a -computable version when `f` has an explicit inverse. -/ -noncomputable def equivMapOfInjective (f : F) (i : Injective f) (p : Submodule R M) : - p ≃ₛₗ[σ₁₂] p.map f := - { Equiv.Set.image f p i with - map_add' := by - intros - simp only [coe_add, map_add, Equiv.toFun_as_coe, Equiv.Set.image_apply] - rfl - map_smul' := by - intros - simp only [coe_smul_of_tower, map_smulₛₗ, Equiv.toFun_as_coe, Equiv.Set.image_apply] - rfl } -#align submodule.equiv_map_of_injective Submodule.equivMapOfInjective - -@[simp] -theorem coe_equivMapOfInjective_apply (f : F) (i : Injective f) (p : Submodule R M) (x : p) : - (equivMapOfInjective f i p x : M₂) = f x := - rfl -#align submodule.coe_equiv_map_of_injective_apply Submodule.coe_equivMapOfInjective_apply - -/-- The pullback of a submodule `p ⊆ M₂` along `f : M → M₂` -/ -def comap (f : F) (p : Submodule R₂ M₂) : Submodule R M := - { p.toAddSubmonoid.comap f with - carrier := f ⁻¹' p - smul_mem' := fun a x h => by simp [p.smul_mem (σ₁₂ a) h] } -#align submodule.comap Submodule.comap - -@[simp] -theorem comap_coe (f : F) (p : Submodule R₂ M₂) : (comap f p : Set M) = f ⁻¹' p := - rfl -#align submodule.comap_coe Submodule.comap_coe - -@[simp] -theorem AddMonoidHom.coe_toIntLinearMap_comap {A A₂ : Type*} [AddCommGroup A] [AddCommGroup A₂] - (f : A →+ A₂) (s : AddSubgroup A₂) : - (AddSubgroup.toIntSubmodule s).comap f.toIntLinearMap = - AddSubgroup.toIntSubmodule (s.comap f) := rfl - -@[simp] -theorem mem_comap {f : F} {p : Submodule R₂ M₂} : x ∈ comap f p ↔ f x ∈ p := - Iff.rfl -#align submodule.mem_comap Submodule.mem_comap - -@[simp] -theorem comap_id : comap (LinearMap.id : M →ₗ[R] M) p = p := - SetLike.coe_injective rfl -#align submodule.comap_id Submodule.comap_id - -theorem comap_comp (f : M →ₛₗ[σ₁₂] M₂) (g : M₂ →ₛₗ[σ₂₃] M₃) (p : Submodule R₃ M₃) : - comap (g.comp f : M →ₛₗ[σ₁₃] M₃) p = comap f (comap g p) := - rfl -#align submodule.comap_comp Submodule.comap_comp - -theorem comap_mono {f : F} {q q' : Submodule R₂ M₂} : q ≤ q' → comap f q ≤ comap f q' := - preimage_mono -#align submodule.comap_mono Submodule.comap_mono - -theorem le_comap_pow_of_le_comap (p : Submodule R M) {f : M →ₗ[R] M} (h : p ≤ p.comap f) (k : ℕ) : - p ≤ p.comap (f ^ k) := by - induction' k with k ih - · simp [LinearMap.one_eq_id] - · simp [LinearMap.iterate_succ, comap_comp, h.trans (comap_mono ih)] -#align submodule.le_comap_pow_of_le_comap Submodule.le_comap_pow_of_le_comap - -section - -variable [RingHomSurjective σ₁₂] - -theorem map_le_iff_le_comap {f : F} {p : Submodule R M} {q : Submodule R₂ M₂} : - map f p ≤ q ↔ p ≤ comap f q := - image_subset_iff -#align submodule.map_le_iff_le_comap Submodule.map_le_iff_le_comap - -theorem gc_map_comap (f : F) : GaloisConnection (map f) (comap f) - | _, _ => map_le_iff_le_comap -#align submodule.gc_map_comap Submodule.gc_map_comap - -@[simp] -theorem map_bot (f : F) : map f ⊥ = ⊥ := - (gc_map_comap f).l_bot -#align submodule.map_bot Submodule.map_bot - -@[simp] -theorem map_sup (f : F) : map f (p ⊔ p') = map f p ⊔ map f p' := - (gc_map_comap f : GaloisConnection (map f) (comap f)).l_sup -#align submodule.map_sup Submodule.map_sup - -@[simp] -theorem map_iSup {ι : Sort*} (f : F) (p : ι → Submodule R M) : - map f (⨆ i, p i) = ⨆ i, map f (p i) := - (gc_map_comap f : GaloisConnection (map f) (comap f)).l_iSup -#align submodule.map_supr Submodule.map_iSup - -end - -@[simp] -theorem comap_top (f : F) : comap f ⊤ = ⊤ := - rfl -#align submodule.comap_top Submodule.comap_top - -@[simp] -theorem comap_inf (f : F) : comap f (q ⊓ q') = comap f q ⊓ comap f q' := - rfl -#align submodule.comap_inf Submodule.comap_inf - -@[simp] -theorem comap_iInf [RingHomSurjective σ₁₂] {ι : Sort*} (f : F) (p : ι → Submodule R₂ M₂) : - comap f (⨅ i, p i) = ⨅ i, comap f (p i) := - (gc_map_comap f : GaloisConnection (map f) (comap f)).u_iInf -#align submodule.comap_infi Submodule.comap_iInf - -@[simp] -theorem comap_zero : comap (0 : M →ₛₗ[σ₁₂] M₂) q = ⊤ := - ext <| by simp -#align submodule.comap_zero Submodule.comap_zero - -theorem map_comap_le [RingHomSurjective σ₁₂] (f : F) (q : Submodule R₂ M₂) : - map f (comap f q) ≤ q := - (gc_map_comap f).l_u_le _ -#align submodule.map_comap_le Submodule.map_comap_le - -theorem le_comap_map [RingHomSurjective σ₁₂] (f : F) (p : Submodule R M) : p ≤ comap f (map f p) := - (gc_map_comap f).le_u_l _ -#align submodule.le_comap_map Submodule.le_comap_map - -section GaloisInsertion - -variable {f : F} (hf : Surjective f) - -variable [RingHomSurjective σ₁₂] - -/-- `map f` and `comap f` form a `GaloisInsertion` when `f` is surjective. -/ -def giMapComap : GaloisInsertion (map f) (comap f) := - (gc_map_comap f).toGaloisInsertion fun S x hx => by - rcases hf x with ⟨y, rfl⟩ - simp only [mem_map, mem_comap] - exact ⟨y, hx, rfl⟩ -#align submodule.gi_map_comap Submodule.giMapComap - -theorem map_comap_eq_of_surjective (p : Submodule R₂ M₂) : (p.comap f).map f = p := - (giMapComap hf).l_u_eq _ -#align submodule.map_comap_eq_of_surjective Submodule.map_comap_eq_of_surjective - -theorem map_surjective_of_surjective : Function.Surjective (map f) := - (giMapComap hf).l_surjective -#align submodule.map_surjective_of_surjective Submodule.map_surjective_of_surjective - -theorem comap_injective_of_surjective : Function.Injective (comap f) := - (giMapComap hf).u_injective -#align submodule.comap_injective_of_surjective Submodule.comap_injective_of_surjective - -theorem map_sup_comap_of_surjective (p q : Submodule R₂ M₂) : - (p.comap f ⊔ q.comap f).map f = p ⊔ q := - (giMapComap hf).l_sup_u _ _ -#align submodule.map_sup_comap_of_surjective Submodule.map_sup_comap_of_surjective - -theorem map_iSup_comap_of_sujective {ι : Sort*} (S : ι → Submodule R₂ M₂) : - (⨆ i, (S i).comap f).map f = iSup S := - (giMapComap hf).l_iSup_u _ -#align submodule.map_supr_comap_of_sujective Submodule.map_iSup_comap_of_sujective - -theorem map_inf_comap_of_surjective (p q : Submodule R₂ M₂) : - (p.comap f ⊓ q.comap f).map f = p ⊓ q := - (giMapComap hf).l_inf_u _ _ -#align submodule.map_inf_comap_of_surjective Submodule.map_inf_comap_of_surjective - -theorem map_iInf_comap_of_surjective {ι : Sort*} (S : ι → Submodule R₂ M₂) : - (⨅ i, (S i).comap f).map f = iInf S := - (giMapComap hf).l_iInf_u _ -#align submodule.map_infi_comap_of_surjective Submodule.map_iInf_comap_of_surjective - -theorem comap_le_comap_iff_of_surjective (p q : Submodule R₂ M₂) : p.comap f ≤ q.comap f ↔ p ≤ q := - (giMapComap hf).u_le_u_iff -#align submodule.comap_le_comap_iff_of_surjective Submodule.comap_le_comap_iff_of_surjective - -theorem comap_strictMono_of_surjective : StrictMono (comap f) := - (giMapComap hf).strictMono_u -#align submodule.comap_strict_mono_of_surjective Submodule.comap_strictMono_of_surjective - -end GaloisInsertion - -section GaloisCoinsertion - -variable [RingHomSurjective σ₁₂] {f : F} (hf : Injective f) - -/-- `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. -/ -def gciMapComap : GaloisCoinsertion (map f) (comap f) := - (gc_map_comap f).toGaloisCoinsertion fun S x => by - simp [mem_comap, mem_map, forall_exists_index, and_imp] - intro y hy hxy - rw [hf.eq_iff] at hxy - rwa [← hxy] -#align submodule.gci_map_comap Submodule.gciMapComap - -theorem comap_map_eq_of_injective (p : Submodule R M) : (p.map f).comap f = p := - (gciMapComap hf).u_l_eq _ -#align submodule.comap_map_eq_of_injective Submodule.comap_map_eq_of_injective - -theorem comap_surjective_of_injective : Function.Surjective (comap f) := - (gciMapComap hf).u_surjective -#align submodule.comap_surjective_of_injective Submodule.comap_surjective_of_injective - -theorem map_injective_of_injective : Function.Injective (map f) := - (gciMapComap hf).l_injective -#align submodule.map_injective_of_injective Submodule.map_injective_of_injective - -theorem comap_inf_map_of_injective (p q : Submodule R M) : (p.map f ⊓ q.map f).comap f = p ⊓ q := - (gciMapComap hf).u_inf_l _ _ -#align submodule.comap_inf_map_of_injective Submodule.comap_inf_map_of_injective - -theorem comap_iInf_map_of_injective {ι : Sort*} (S : ι → Submodule R M) : - (⨅ i, (S i).map f).comap f = iInf S := - (gciMapComap hf).u_iInf_l _ -#align submodule.comap_infi_map_of_injective Submodule.comap_iInf_map_of_injective - -theorem comap_sup_map_of_injective (p q : Submodule R M) : (p.map f ⊔ q.map f).comap f = p ⊔ q := - (gciMapComap hf).u_sup_l _ _ -#align submodule.comap_sup_map_of_injective Submodule.comap_sup_map_of_injective - -theorem comap_iSup_map_of_injective {ι : Sort*} (S : ι → Submodule R M) : - (⨆ i, (S i).map f).comap f = iSup S := - (gciMapComap hf).u_iSup_l _ -#align submodule.comap_supr_map_of_injective Submodule.comap_iSup_map_of_injective - -theorem map_le_map_iff_of_injective (p q : Submodule R M) : p.map f ≤ q.map f ↔ p ≤ q := - (gciMapComap hf).l_le_l_iff -#align submodule.map_le_map_iff_of_injective Submodule.map_le_map_iff_of_injective - -theorem map_strictMono_of_injective : StrictMono (map f) := - (gciMapComap hf).strictMono_l -#align submodule.map_strict_mono_of_injective Submodule.map_strictMono_of_injective - -end GaloisCoinsertion - -end SemilinearMap - -section OrderIso - -variable {F : Type*} [SemilinearEquivClass F σ₁₂ M M₂] - -/-- A linear isomorphism induces an order isomorphism of submodules. -/ -@[simps symm_apply apply] -def orderIsoMapComap (f : F) : Submodule R M ≃o Submodule R₂ M₂ where - toFun := map f - invFun := comap f - left_inv := comap_map_eq_of_injective (EquivLike.injective f) - right_inv := map_comap_eq_of_surjective (EquivLike.surjective f) - map_rel_iff' := map_le_map_iff_of_injective (EquivLike.injective f) _ _ -#align submodule.order_iso_map_comap Submodule.orderIsoMapComap - -end OrderIso - -variable {F : Type*} [sc : SemilinearMapClass F σ₁₂ M M₂] - ---TODO(Mario): is there a way to prove this from order properties? -theorem map_inf_eq_map_inf_comap [RingHomSurjective σ₁₂] {f : F} {p : Submodule R M} - {p' : Submodule R₂ M₂} : map f p ⊓ p' = map f (p ⊓ comap f p') := - le_antisymm (by rintro _ ⟨⟨x, h₁, rfl⟩, h₂⟩; exact ⟨_, ⟨h₁, h₂⟩, rfl⟩) - (le_inf (map_mono inf_le_left) (map_le_iff_le_comap.2 inf_le_right)) -#align submodule.map_inf_eq_map_inf_comap Submodule.map_inf_eq_map_inf_comap - -theorem map_comap_subtype : map p.subtype (comap p.subtype p') = p ⊓ p' := - ext fun x => ⟨by rintro ⟨⟨_, h₁⟩, h₂, rfl⟩; exact ⟨h₁, h₂⟩, fun ⟨h₁, h₂⟩ => ⟨⟨_, h₁⟩, h₂, rfl⟩⟩ -#align submodule.map_comap_subtype Submodule.map_comap_subtype - -theorem eq_zero_of_bot_submodule : ∀ b : (⊥ : Submodule R M), b = 0 - | ⟨b', hb⟩ => Subtype.eq <| show b' = 0 from (mem_bot R).1 hb -#align submodule.eq_zero_of_bot_submodule Submodule.eq_zero_of_bot_submodule - -/-- The infimum of a family of invariant submodule of an endomorphism is also an invariant -submodule. -/ -theorem _root_.LinearMap.iInf_invariant {σ : R →+* R} [RingHomSurjective σ] {ι : Sort*} - (f : M →ₛₗ[σ] M) {p : ι → Submodule R M} (hf : ∀ i, ∀ v ∈ p i, f v ∈ p i) : - ∀ v ∈ iInf p, f v ∈ iInf p := by - have : ∀ i, (p i).map f ≤ p i := by - rintro i - ⟨v, hv, rfl⟩ - exact hf i v hv - suffices (iInf p).map f ≤ iInf p by exact fun v hv => this ⟨v, hv, rfl⟩ - exact le_iInf fun i => (Submodule.map_mono (iInf_le p i)).trans (this i) -#align linear_map.infi_invariant LinearMap.iInf_invariant - end AddCommMonoid section AddCommGroup @@ -946,46 +561,10 @@ theorem neg_coe : -(p : Set M) = p := Set.ext fun _ => p.neg_mem_iff #align submodule.neg_coe Submodule.neg_coe -@[simp] -protected theorem map_neg (f : M →ₗ[R] M₂) : map (-f) p = map f p := - ext fun _ => - ⟨fun ⟨x, hx, hy⟩ => hy ▸ ⟨-x, show -x ∈ p from neg_mem hx, map_neg f x⟩, fun ⟨x, hx, hy⟩ => - hy ▸ ⟨-x, show -x ∈ p from neg_mem hx, (map_neg (-f) _).trans (neg_neg (f x))⟩⟩ -#align submodule.map_neg Submodule.map_neg - end AddCommGroup end Submodule -namespace Submodule - -variable [Semifield K] -variable [AddCommMonoid V] [Module K V] -variable [AddCommMonoid V₂] [Module K V₂] - -theorem comap_smul (f : V →ₗ[K] V₂) (p : Submodule K V₂) (a : K) (h : a ≠ 0) : - p.comap (a • f) = p.comap f := by - ext b; simp only [Submodule.mem_comap, p.smul_mem_iff h, LinearMap.smul_apply] -#align submodule.comap_smul Submodule.comap_smul - -protected theorem map_smul (f : V →ₗ[K] V₂) (p : Submodule K V) (a : K) (h : a ≠ 0) : - p.map (a • f) = p.map f := - le_antisymm (by rw [map_le_iff_le_comap, comap_smul f _ a h, ← map_le_iff_le_comap]) - (by rw [map_le_iff_le_comap, ← comap_smul f _ a h, ← map_le_iff_le_comap]) -#align submodule.map_smul Submodule.map_smul - -theorem comap_smul' (f : V →ₗ[K] V₂) (p : Submodule K V₂) (a : K) : - p.comap (a • f) = ⨅ _ : a ≠ 0, p.comap f := by - classical by_cases h : a = 0 <;> simp [h, comap_smul] -#align submodule.comap_smul' Submodule.comap_smul' - -theorem map_smul' (f : V →ₗ[K] V₂) (p : Submodule K V) (a : K) : - p.map (a • f) = ⨆ _ : a ≠ 0, map f p := by - classical by_cases h : a = 0 <;> simp [h, Submodule.map_smul] -#align submodule.map_smul' Submodule.map_smul' - -end Submodule - /-! ### Properties of linear maps -/ @@ -1258,6 +837,17 @@ theorem ker_le_ker_comp (f : M →ₛₗ[τ₁₂] M₂) (g : M₂ →ₛₗ[τ ker f ≤ ker (g.comp f : M →ₛₗ[τ₁₃] M₃) := by rw [ker_comp]; exact comap_mono bot_le #align linear_map.ker_le_ker_comp LinearMap.ker_le_ker_comp +theorem ker_sup_ker_le_ker_comp_of_commute {f g : M →ₗ[R] M} (h : Commute f g) : + ker f ⊔ ker g ≤ ker (f ∘ₗ g) := by + refine sup_le_iff.mpr ⟨?_, ker_le_ker_comp g f⟩ + rw [← mul_eq_comp, h.eq, mul_eq_comp] + exact ker_le_ker_comp f g + +@[simp] +theorem ker_le_comap {p : Submodule R₂ M₂} (f : M →ₛₗ[τ₁₂] M₂) : + ker f ≤ p.comap f := + fun x hx ↦ by simp [mem_ker.mp hx] + theorem disjoint_ker {f : F} {p : Submodule R M} : Disjoint p (ker f) ↔ ∀ x ∈ p, f x = 0 → x = 0 := by simp [disjoint_def] #align linear_map.disjoint_ker LinearMap.disjoint_ker @@ -1456,6 +1046,11 @@ theorem ker_le_iff [RingHomSurjective τ₁₂] {p : Submodule R M} : rw [h] at this simpa using this +@[simp] theorem injective_restrict_iff_disjoint {p : Submodule R M} {f : M →ₗ[R] M} + (hf : ∀ x ∈ p, f x ∈ p) : + Injective (f.restrict hf) ↔ Disjoint p (ker f) := by + rw [← ker_eq_bot, ker_restrict hf, ker_eq_bot, injective_domRestrict_iff, disjoint_iff] + end Ring section Semifield @@ -2345,145 +1940,10 @@ theorem equivSubtypeMap_symm_apply {p : Submodule R M} {q : Submodule R p} (x : rfl #align submodule.equiv_subtype_map_symm_apply Submodule.equivSubtypeMap_symm_apply -/-- If `s ≤ t`, then we can view `s` as a submodule of `t` by taking the comap -of `t.subtype`. -/ -@[simps symm_apply] -def comapSubtypeEquivOfLe {p q : Submodule R M} (hpq : p ≤ q) : comap q.subtype p ≃ₗ[R] p where - toFun x := ⟨x, x.2⟩ - invFun x := ⟨⟨x, hpq x.2⟩, x.2⟩ - left_inv x := by simp only [coe_mk, SetLike.eta, LinearEquiv.coe_coe] - right_inv x := by simp only [Subtype.coe_mk, SetLike.eta, LinearEquiv.coe_coe] - map_add' x y := rfl - map_smul' c x := rfl -#align submodule.comap_subtype_equiv_of_le Submodule.comapSubtypeEquivOfLe -#align submodule.comap_subtype_equiv_of_le_symm_apply_coe_coe Submodule.comapSubtypeEquivOfLe_symm_apply - --- Porting note: The original theorem generated by `simps` was using `LinearEquiv.toLinearMap`, --- different from the theorem on Lean 3, and not simp-normal form. -@[simp] -theorem comapSubtypeEquivOfLe_apply_coe {p q : Submodule R M} (hpq : p ≤ q) - (x : comap q.subtype p) : - (comapSubtypeEquivOfLe hpq x : M) = (x : M) := - rfl -#align submodule.comap_subtype_equiv_of_le_apply_coe Submodule.comapSubtypeEquivOfLe_apply_coe - end Module end Submodule -namespace Submodule - -variable [Semiring R] [Semiring R₂] - -variable [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] - -variable [AddCommMonoid N] [AddCommMonoid N₂] [Module R N] [Module R N₂] - -variable {τ₁₂ : R →+* R₂} {τ₂₁ : R₂ →+* R} - -variable [RingHomInvPair τ₁₂ τ₂₁] [RingHomInvPair τ₂₁ τ₁₂] - -variable (p : Submodule R M) (q : Submodule R₂ M₂) - -variable (pₗ : Submodule R N) (qₗ : Submodule R N₂) - --- Porting note: Was `@[simp]`. -@[simp high] -theorem mem_map_equiv {e : M ≃ₛₗ[τ₁₂] M₂} {x : M₂} : - x ∈ p.map (e : M →ₛₗ[τ₁₂] M₂) ↔ e.symm x ∈ p := by - rw [Submodule.mem_map]; constructor - · rintro ⟨y, hy, hx⟩ - simp [← hx, hy] - · intro hx - refine' ⟨e.symm x, hx, by simp⟩ -#align submodule.mem_map_equiv Submodule.mem_map_equiv - -theorem map_equiv_eq_comap_symm (e : M ≃ₛₗ[τ₁₂] M₂) (K : Submodule R M) : - K.map (e : M →ₛₗ[τ₁₂] M₂) = K.comap (e.symm : M₂ →ₛₗ[τ₂₁] M) := - Submodule.ext fun _ => by rw [mem_map_equiv, mem_comap, LinearEquiv.coe_coe] -#align submodule.map_equiv_eq_comap_symm Submodule.map_equiv_eq_comap_symm - -theorem comap_equiv_eq_map_symm (e : M ≃ₛₗ[τ₁₂] M₂) (K : Submodule R₂ M₂) : - K.comap (e : M →ₛₗ[τ₁₂] M₂) = K.map (e.symm : M₂ →ₛₗ[τ₂₁] M) := - (map_equiv_eq_comap_symm e.symm K).symm -#align submodule.comap_equiv_eq_map_symm Submodule.comap_equiv_eq_map_symm - -variable {p} - -theorem map_symm_eq_iff (e : M ≃ₛₗ[τ₁₂] M₂) {K : Submodule R₂ M₂} : - K.map e.symm = p ↔ p.map e = K := by - constructor <;> rintro rfl - · calc - map e (map e.symm K) = comap e.symm (map e.symm K) := map_equiv_eq_comap_symm _ _ - _ = K := comap_map_eq_of_injective e.symm.injective _ - · calc - map e.symm (map e p) = comap e (map e p) := (comap_equiv_eq_map_symm _ _).symm - _ = p := comap_map_eq_of_injective e.injective _ -#align submodule.map_symm_eq_iff Submodule.map_symm_eq_iff - -theorem orderIsoMapComap_apply' (e : M ≃ₛₗ[τ₁₂] M₂) (p : Submodule R M) : - orderIsoMapComap e p = comap e.symm p := - p.map_equiv_eq_comap_symm _ -#align submodule.order_iso_map_comap_apply' Submodule.orderIsoMapComap_apply' - -theorem orderIsoMapComap_symm_apply' (e : M ≃ₛₗ[τ₁₂] M₂) (p : Submodule R₂ M₂) : - (orderIsoMapComap e).symm p = map e.symm p := - p.comap_equiv_eq_map_symm _ -#align submodule.order_iso_map_comap_symm_apply' Submodule.orderIsoMapComap_symm_apply' - -theorem inf_comap_le_comap_add (f₁ f₂ : M →ₛₗ[τ₁₂] M₂) : - comap f₁ q ⊓ comap f₂ q ≤ comap (f₁ + f₂) q := by - rw [SetLike.le_def] - intro m h - change f₁ m + f₂ m ∈ q - change f₁ m ∈ q ∧ f₂ m ∈ q at h - apply q.add_mem h.1 h.2 -#align submodule.inf_comap_le_comap_add Submodule.inf_comap_le_comap_add - -end Submodule - -namespace Submodule - -variable [CommSemiring R] [CommSemiring R₂] - -variable [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] - -variable [AddCommMonoid N] [AddCommMonoid N₂] [Module R N] [Module R N₂] - -variable {τ₁₂ : R →+* R₂} {τ₂₁ : R₂ →+* R} - -variable [RingHomInvPair τ₁₂ τ₂₁] [RingHomInvPair τ₂₁ τ₁₂] - -variable (p : Submodule R M) (q : Submodule R₂ M₂) - -variable (pₗ : Submodule R N) (qₗ : Submodule R N₂) - -theorem comap_le_comap_smul (fₗ : N →ₗ[R] N₂) (c : R) : comap fₗ qₗ ≤ comap (c • fₗ) qₗ := by - rw [SetLike.le_def] - intro m h - change c • fₗ m ∈ qₗ - change fₗ m ∈ qₗ at h - apply qₗ.smul_mem _ h -#align submodule.comap_le_comap_smul Submodule.comap_le_comap_smul - -/-- Given modules `M`, `M₂` over a commutative ring, together with submodules `p ⊆ M`, `q ⊆ M₂`, -the set of maps $\{f ∈ Hom(M, M₂) | f(p) ⊆ q \}$ is a submodule of `Hom(M, M₂)`. -/ -def compatibleMaps : Submodule R (N →ₗ[R] N₂) where - carrier := { fₗ | pₗ ≤ comap fₗ qₗ } - zero_mem' := by - change pₗ ≤ comap (0 : N →ₗ[R] N₂) qₗ - rw [comap_zero] - refine' le_top - add_mem' {f₁ f₂} h₁ h₂ := by - apply le_trans _ (inf_comap_le_comap_add qₗ f₁ f₂) - rw [le_inf_iff] - exact ⟨h₁, h₂⟩ - smul_mem' c fₗ h := by - dsimp at h - exact le_trans h (comap_le_comap_smul qₗ fₗ c) -#align submodule.compatible_maps Submodule.compatibleMaps - -end Submodule namespace Equiv diff --git a/Mathlib/LinearAlgebra/Basis.lean b/Mathlib/LinearAlgebra/Basis.lean index 69348f183579f..1d01f3470e227 100644 --- a/Mathlib/LinearAlgebra/Basis.lean +++ b/Mathlib/LinearAlgebra/Basis.lean @@ -1393,6 +1393,19 @@ def Submodule.inductionOnRankAux (b : Basis ι R M) (P : Submodule R M → Sort* end Induction +/-- An element of a non-unital-non-associative algebra is in the center exactly when it commutes +with the basis elements. -/ +lemma Basis.mem_center_iff {A} + [Semiring R] [NonUnitalNonAssocSemiring A] + [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] + (b : Basis ι R A) {x : A} : + x ∈ Set.center A ↔ ∀ i, Commute (b i) x := by + constructor + · intro h i; apply h + · intro h y + rw [← b.total_repr y, Finsupp.total_apply, Finsupp.sum, Finset.sum_mul, Finset.mul_sum] + simp_rw [mul_smul_comm, smul_mul_assoc, (h _).eq] + section RestrictScalars variable {S : Type*} [CommRing R] [Ring S] [Nontrivial S] [AddCommGroup M] diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean index d354eca2d3092..5ce702ad50ec2 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean @@ -399,4 +399,8 @@ def equivExterior [Invertible (2 : R)] : CliffordAlgebra Q ≃ₗ[R] ExteriorAlg changeFormEquiv changeForm.associated_neg_proof #align clifford_algebra.equiv_exterior CliffordAlgebra.equivExterior +/-- A `CliffordAlgebra` over a nontrivial ring is nontrivial, in characteristic not two. -/ +instance [Nontrivial R] [Invertible (2 : R)] : + Nontrivial (CliffordAlgebra Q) := (equivExterior Q).symm.injective.nontrivial + end CliffordAlgebra diff --git a/Mathlib/LinearAlgebra/DFinsupp.lean b/Mathlib/LinearAlgebra/DFinsupp.lean index d61b12907a3ed..64c769c096168 100644 --- a/Mathlib/LinearAlgebra/DFinsupp.lean +++ b/Mathlib/LinearAlgebra/DFinsupp.lean @@ -337,13 +337,18 @@ theorem biSup_eq_range_dfinsupp_lsum (p : ι → Prop) [DecidablePred p] (S : ι · simp [hp] #align submodule.bsupr_eq_range_dfinsupp_lsum Submodule.biSup_eq_range_dfinsupp_lsum +/-- A characterisation of the span of a family of submodules. + +See also `Submodule.mem_iSup_iff_exists_finsupp`. -/ theorem mem_iSup_iff_exists_dfinsupp (p : ι → Submodule R N) (x : N) : x ∈ iSup p ↔ ∃ f : Π₀ i, p i, DFinsupp.lsum ℕ (M := fun i ↦ ↥(p i)) (fun i => (p i).subtype) f = x := SetLike.ext_iff.mp (iSup_eq_range_dfinsupp_lsum p) x #align submodule.mem_supr_iff_exists_dfinsupp Submodule.mem_iSup_iff_exists_dfinsupp -/-- A variant of `Submodule.mem_iSup_iff_exists_dfinsupp` with the RHS fully unfolded. -/ +/-- A variant of `Submodule.mem_iSup_iff_exists_dfinsupp` with the RHS fully unfolded. + +See also `Submodule.mem_iSup_iff_exists_finsupp`. -/ theorem mem_iSup_iff_exists_dfinsupp' (p : ι → Submodule R N) [∀ (i) (x : p i), Decidable (x ≠ 0)] (x : N) : x ∈ iSup p ↔ ∃ f : Π₀ i, p i, (f.sum fun i xi => ↑xi) = x := by rw [mem_iSup_iff_exists_dfinsupp] @@ -351,6 +356,16 @@ theorem mem_iSup_iff_exists_dfinsupp' (p : ι → Submodule R N) [∀ (i) (x : p LinearMap.toAddMonoidHom_coe, coeSubtype] #align submodule.mem_supr_iff_exists_dfinsupp' Submodule.mem_iSup_iff_exists_dfinsupp' +lemma mem_iSup_iff_exists_finsupp (p : ι → Submodule R N) (x : N) : + x ∈ iSup p ↔ ∃ (f : ι →₀ N), (∀ i, f i ∈ p i) ∧ (f.sum fun _i xi ↦ xi) = x := by + classical + rw [mem_iSup_iff_exists_dfinsupp'] + refine ⟨fun ⟨f, hf⟩ ↦ ⟨⟨f.support, fun i ↦ (f i : N), by simp⟩, by simp, hf⟩, ?_⟩ + rintro ⟨f, hf, rfl⟩ + refine ⟨DFinsupp.mk f.support <| fun i ↦ ⟨f i, hf i⟩, Finset.sum_congr ?_ fun i hi ↦ ?_⟩ + · ext; simp + · simp [Finsupp.mem_support_iff.mp hi] + theorem mem_biSup_iff_exists_dfinsupp (p : ι → Prop) [DecidablePred p] (S : ι → Submodule R N) (x : N) : (x ∈ ⨆ (i) (_ : p i), S i) ↔ diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index 18dd3b59db798..b8890e6f4f830 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -322,7 +322,7 @@ theorem det_conj {N : Type*} [AddCommGroup N] [Module A N] (f : M →ₗ[A] M) ( contrapose! H rcases H with ⟨s, ⟨b⟩⟩ exact ⟨_, ⟨(b.map e.symm).reindexFinsetRange⟩⟩ - simp only [coe_det, H, H', MonoidHom.one_apply, dif_neg] + simp only [coe_det, H, H', MonoidHom.one_apply, dif_neg, not_false_eq_true] #align linear_map.det_conj LinearMap.det_conj /-- If a linear map is invertible, so is its determinant. -/ diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 4f93fa4cac23b..64c4eb7d3005a 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -659,7 +659,7 @@ lemma equiv [IsReflexive R M] (e : M ≃ₗ[R] N) : IsReflexive R N where exact FunLike.congr_arg f (e.apply_symm_apply m).symm simp only [this, LinearEquiv.trans_symm, LinearEquiv.symm_symm, LinearEquiv.dualMap_symm, coe_comp, LinearEquiv.coe_coe, EquivLike.comp_bijective] - refine Bijective.comp (bijective_dual_eval R M) (LinearEquiv.bijective _) + exact Bijective.comp (bijective_dual_eval R M) (LinearEquiv.bijective _) instance _root_.MulOpposite.instModuleIsReflexive : IsReflexive R (MulOpposite M) := equiv <| MulOpposite.opLinearEquiv _ diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean index f126bfc12a2ab..224094e5a2c73 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean @@ -333,6 +333,10 @@ theorem ιMulti_succ_curryLeft {n : ℕ} (m : M) : rfl #align exterior_algebra.ι_multi_succ_curry_left ExteriorAlgebra.ιMulti_succ_curryLeft +/-- An `ExteriorAlgebra` over a nontrivial ring is nontrivial. -/ +instance [Nontrivial R] : Nontrivial (ExteriorAlgebra R M) := + (algebraMap_leftInverse M).injective.nontrivial + end ExteriorAlgebra namespace TensorAlgebra diff --git a/Mathlib/LinearAlgebra/FiniteDimensional.lean b/Mathlib/LinearAlgebra/FiniteDimensional.lean index 2b61ea5d9154c..4a2fc7503f817 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional.lean @@ -922,6 +922,13 @@ theorem injective_iff_surjective [FiniteDimensional K V] {f : V →ₗ[K] V} : this).injective⟩ #align linear_map.injective_iff_surjective LinearMap.injective_iff_surjective +lemma injOn_iff_surjOn {p : Submodule K V} [FiniteDimensional K p] + {f : V →ₗ[K] V} (h : ∀ x ∈ p, f x ∈ p) : + Set.InjOn f p ↔ Set.SurjOn f p p := by + rw [Set.injOn_iff_injective, ← Set.MapsTo.restrict_surjective_iff h] + change Injective (f.domRestrict p) ↔ Surjective (f.restrict h) + simp [disjoint_iff, ← injective_iff_surjective] + theorem ker_eq_bot_iff_range_eq_top [FiniteDimensional K V] {f : V →ₗ[K] V} : LinearMap.ker f = ⊥ ↔ LinearMap.range f = ⊤ := by rw [range_eq_top, ker_eq_bot, injective_iff_surjective] @@ -959,6 +966,37 @@ theorem finrank_range_add_finrank_ker [FiniteDimensional K V] (f : V →ₗ[K] V exact Submodule.finrank_quotient_add_finrank _ #align linear_map.finrank_range_add_finrank_ker LinearMap.finrank_range_add_finrank_ker +theorem comap_eq_sup_ker_of_disjoint {p : Submodule K V} [FiniteDimensional K p] {f : V →ₗ[K] V} + (h : ∀ x ∈ p, f x ∈ p) (h' : Disjoint p (ker f)) : + p.comap f = p ⊔ ker f := by + refine le_antisymm (fun x hx ↦ ?_) (sup_le_iff.mpr ⟨h, ker_le_comap _⟩) + obtain ⟨⟨y, hy⟩, hxy⟩ := + surjective_of_injective ((injective_restrict_iff_disjoint h).mpr h') ⟨f x, hx⟩ + replace hxy : f y = f x := by simpa [Subtype.ext_iff] using hxy + exact Submodule.mem_sup.mpr ⟨y, hy, x - y, by simp [hxy], add_sub_cancel'_right y x⟩ + +theorem ker_comp_eq_of_commute_of_disjoint_ker [FiniteDimensional K V] {f g : V →ₗ[K] V} + (h : Commute f g) (h' : Disjoint (ker f) (ker g)) : + ker (f ∘ₗ g) = ker f ⊔ ker g := by + suffices ∀ x, f x = 0 → f (g x) = 0 by rw [ker_comp, comap_eq_sup_ker_of_disjoint _ h']; simpa + intro x hx + rw [← comp_apply, ← mul_eq_comp, h.eq, mul_apply, hx, _root_.map_zero] + +theorem ker_noncommProd_eq_of_supIndep_ker [FiniteDimensional K V] {ι : Type*} {f : ι → V →ₗ[K] V} + (s : Finset ι) (comm) (h : s.SupIndep fun i ↦ ker (f i)) : + ker (s.noncommProd f comm) = ⨆ i ∈ s, ker (f i) := by + classical + induction' s using Finset.induction_on with i s hi ih; simpa using LinearMap.ker_id + replace ih : ker (Finset.noncommProd s f <| Set.Pairwise.mono (s.subset_insert i) comm) = + ⨆ x ∈ s, ker (f x) := ih _ (h.subset (s.subset_insert i)) + rw [Finset.noncommProd_insert_of_not_mem _ _ _ _ hi, mul_eq_comp, + ker_comp_eq_of_commute_of_disjoint_ker] + · simp_rw [Finset.mem_insert_coe, iSup_insert, Finset.mem_coe, ih] + · exact s.noncommProd_commute _ _ _ fun j hj ↦ + comm (s.mem_insert_self i) (Finset.mem_insert_of_mem hj) (by aesop) + · replace h := Finset.supIndep_iff_disjoint_erase.mp h i (s.mem_insert_self i) + simpa [ih, hi, Finset.sup_eq_iSup] using h + end DivisionRing end LinearMap diff --git a/Mathlib/LinearAlgebra/Finrank.lean b/Mathlib/LinearAlgebra/Finrank.lean index b06a9e20c8194..b1187d4359813 100644 --- a/Mathlib/LinearAlgebra/Finrank.lean +++ b/Mathlib/LinearAlgebra/Finrank.lean @@ -435,7 +435,7 @@ theorem linearIndependent_of_top_le_span_of_card_eq_finrank {ι : Type*} [Fintyp · refine' neg_mem (smul_mem _ _ (sum_mem fun k hk => _)) obtain ⟨k_ne_i, _⟩ := Finset.mem_erase.mp hk refine' smul_mem _ _ (subset_span ⟨k, _, rfl⟩) - simp_all only [Set.mem_univ, Set.mem_diff, Set.mem_singleton_iff] + simp_all only [Set.mem_univ, Set.mem_diff, Set.mem_singleton_iff, and_self, not_false_eq_true] -- To show `b i` is a weighted sum of the other `b j`s, we'll rewrite this sum -- to have the form of the assumption `dependent`. apply eq_neg_of_add_eq_zero_left @@ -497,15 +497,12 @@ theorem coe_basisOfTopLeSpanOfCardEqFinrank {ι : Type*} [Fintype ι] (b : ι /-- A finset of `finrank K V` vectors forms a basis if they span the whole space. -/ @[simps! repr_apply] noncomputable def finsetBasisOfTopLeSpanOfCardEqFinrank {s : Finset V} - (le_span : ⊤ ≤ span K (s : Set V)) (card_eq : s.card = finrank K V) : Basis (s : Set V) K V := + (le_span : ⊤ ≤ span K (s : Set V)) (card_eq : s.card = finrank K V) : Basis {x // x ∈ s} K V := basisOfTopLeSpanOfCardEqFinrank ((↑) : ↥(s : Set V) → V) ((@Subtype.range_coe_subtype _ fun x => x ∈ s).symm ▸ le_span) (_root_.trans (Fintype.card_coe _) card_eq) #align finset_basis_of_top_le_span_of_card_eq_finrank finsetBasisOfTopLeSpanOfCardEqFinrank --- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing -attribute [nolint simpNF] finsetBasisOfTopLeSpanOfCardEqFinrank_repr_apply - /-- A set of `finrank K V` vectors forms a basis if they span the whole space. -/ @[simps! repr_apply] noncomputable def setBasisOfTopLeSpanOfCardEqFinrank {s : Set V} [Fintype s] @@ -594,11 +591,10 @@ variable [StrongRankCondition F] [NoZeroSMulDivisors F E] [Nontrivial E] @[simp] theorem Subalgebra.rank_bot : Module.rank F (⊥ : Subalgebra F E) = 1 := - ((Subalgebra.toSubmoduleEquiv (⊥ : Subalgebra F E)).symm.trans <| - LinearEquiv.ofEq _ _ Algebra.toSubmodule_bot).rank_eq.trans <| by + (Subalgebra.toSubmoduleEquiv (⊥ : Subalgebra F E)).symm.rank_eq.trans <| by + rw [Algebra.toSubmodule_bot, one_eq_span, rank_span_set, mk_singleton _] letI := Module.nontrivial F E - rw [rank_span_set] - exacts [mk_singleton _, linearIndependent_singleton one_ne_zero] + exact linearIndependent_singleton one_ne_zero #align subalgebra.rank_bot Subalgebra.rank_bot @[simp] diff --git a/Mathlib/LinearAlgebra/FreeModule/Norm.lean b/Mathlib/LinearAlgebra/FreeModule/Norm.lean index 4533553bd1613..9da4d463fb12e 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Norm.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Norm.lean @@ -67,7 +67,7 @@ instance (b : Basis ι F[X] S) {I : Ideal S} (hI : I ≠ ⊥) (i : ι) : -- irreducible so that they don't expose `Quotient.lift` accidentally. refine PowerBasis.finiteDimensional ?_ refine AdjoinRoot.powerBasis ?_ - refine I.smithCoeffs_ne_zero b hI i + exact I.smithCoeffs_ne_zero b hI i /-- For a nonzero element `f` in a `F[X]`-module `S`, the dimension of $S/\langle f \rangle$ as an `F`-vector space is the degree of the norm of `f` relative to `F[X]`. -/ diff --git a/Mathlib/LinearAlgebra/LinearPMap.lean b/Mathlib/LinearAlgebra/LinearPMap.lean index bb182f960c90a..917ea90d6ca8e 100644 --- a/Mathlib/LinearAlgebra/LinearPMap.lean +++ b/Mathlib/LinearAlgebra/LinearPMap.lean @@ -792,7 +792,7 @@ theorem graph_map_fst_eq_domain (f : E →ₗ.[R] F) : · rcases h with ⟨x, hx, _⟩ exact hx · use f ⟨x, h⟩ - simp only [h, exists_prop] + simp only [h, exists_const] theorem graph_map_snd_eq_range (f : E →ₗ.[R] F) : f.graph.map (LinearMap.snd R E F) = LinearMap.range f.toFun := by ext; simp diff --git a/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean index fef72ee7fe08a..8cd1a3a02f0eb 100644 --- a/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean @@ -568,7 +568,8 @@ theorem _root_.Matrix.nondegenerate_toBilin_iff {M : Matrix ι ι R₃} (b : Bas rw [← Matrix.nondegenerate_toBilin'_iff_nondegenerate_toBilin, Matrix.nondegenerate_toBilin'_iff] #align matrix.nondegenerate_to_bilin_iff Matrix.nondegenerate_toBilin_iff --- Lemmas transferring nondegeneracy between a bilinear form and its associated matrix +/-! Lemmas transferring nondegeneracy between a bilinear form and its associated matrix -/ + @[simp] theorem nondegenerate_toMatrix'_iff {B : BilinForm R₃ (ι → R₃)} : B.toMatrix'.Nondegenerate ↔ B.Nondegenerate := @@ -591,7 +592,8 @@ theorem Nondegenerate.toMatrix {B : BilinForm R₃ M₃} (h : B.Nondegenerate) ( (nondegenerate_toMatrix_iff b).mpr h #align bilin_form.nondegenerate.to_matrix BilinForm.Nondegenerate.toMatrix --- Some shorthands for combining the above with `Matrix.nondegenerate_of_det_ne_zero` +/-! Some shorthands for combining the above with `Matrix.nondegenerate_of_det_ne_zero` -/ + theorem nondegenerate_toBilin'_iff_det_ne_zero {M : Matrix ι ι A} : M.toBilin'.Nondegenerate ↔ M.det ≠ 0 := by rw [Matrix.nondegenerate_toBilin'_iff, Matrix.nondegenerate_iff_det_ne_zero] diff --git a/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean b/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean index 893fe143f0dd4..a837d7ffb842e 100644 --- a/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean +++ b/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean @@ -29,7 +29,7 @@ theorem eigenvalue_mem_ball {μ : K} (hμ : Module.End.HasEigenvalue (Matrix.toL ∃ k, μ ∈ Metric.closedBall (A k k) (∑ j in Finset.univ.erase k, ‖A k j‖) := by cases isEmpty_or_nonempty n · exfalso - exact hμ (Submodule.eq_bot_of_subsingleton _) + exact hμ Submodule.eq_bot_of_subsingleton · obtain ⟨v, h_eg, h_nz⟩ := hμ.exists_hasEigenvector obtain ⟨i, -, h_i⟩ := Finset.exists_mem_eq_sup' Finset.univ_nonempty (fun i => ‖v i‖) have h_nz : v i ≠ 0 := by diff --git a/Mathlib/LinearAlgebra/Matrix/Trace.lean b/Mathlib/LinearAlgebra/Matrix/Trace.lean index 17fbe513bf476..9924d417f21e7 100644 --- a/Mathlib/LinearAlgebra/Matrix/Trace.lean +++ b/Mathlib/LinearAlgebra/Matrix/Trace.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ import Mathlib.Data.Matrix.Basic +import Mathlib.Data.Matrix.RowCol #align_import linear_algebra.matrix.trace from "leanprover-community/mathlib"@"32b08ef840dd25ca2e47e035c5da03ce16d2dc3c" diff --git a/Mathlib/LinearAlgebra/Matrix/Transvection.lean b/Mathlib/LinearAlgebra/Matrix/Transvection.lean index 48ba5b08e9e15..21d209adf9735 100644 --- a/Mathlib/LinearAlgebra/Matrix/Transvection.lean +++ b/Mathlib/LinearAlgebra/Matrix/Transvection.lean @@ -484,7 +484,8 @@ theorem mul_listTransvecRow_last_row (hM : M (inr unit) (inr unit) ≠ 0) (i : F rintro rfl cases i tauto - simp only [IH hnr.le, Ne.def, mul_transvection_apply_of_ne, Ne.symm h, inl.injEq] + simp only [IH hnr.le, Ne.def, mul_transvection_apply_of_ne, Ne.symm h, inl.injEq, + not_false_eq_true] rcases le_or_lt (n + 1) i with (hi | hi) · simp [hi, n.le_succ.trans hi, if_true] · rw [if_neg, if_neg] diff --git a/Mathlib/LinearAlgebra/Orientation.lean b/Mathlib/LinearAlgebra/Orientation.lean index e4f16b826870b..824632e4bc289 100644 --- a/Mathlib/LinearAlgebra/Orientation.lean +++ b/Mathlib/LinearAlgebra/Orientation.lean @@ -115,14 +115,14 @@ theorem Orientation.reindex_symm (e : ι ≃ ι') : end Reindex /-- A module is canonically oriented with respect to an empty index type. -/ -instance (priority := 100) IsEmpty.oriented [Nontrivial R] [IsEmpty ι] : Module.Oriented R M ι where +instance (priority := 100) IsEmpty.oriented [IsEmpty ι] : Module.Oriented R M ι where positiveOrientation := rayOfNeZero R (AlternatingMap.constLinearEquivOfIsEmpty 1) <| AlternatingMap.constLinearEquivOfIsEmpty.injective.ne (by exact one_ne_zero) #align is_empty.oriented IsEmpty.oriented @[simp] -theorem Orientation.map_positiveOrientation_of_isEmpty [Nontrivial R] [IsEmpty ι] (f : M ≃ₗ[R] N) : +theorem Orientation.map_positiveOrientation_of_isEmpty [IsEmpty ι] (f : M ≃ₗ[R] N) : Orientation.map ι f positiveOrientation = positiveOrientation := rfl #align orientation.map_positive_orientation_of_is_empty Orientation.map_positiveOrientation_of_isEmpty @@ -180,23 +180,23 @@ theorem map_orientation_eq_det_inv_smul [Finite ι] (e : Basis ι R M) (x : Orie variable [Fintype ι] [DecidableEq ι] [Fintype ι'] [DecidableEq ι'] /-- The orientation given by a basis. -/ -protected def orientation [Nontrivial R] (e : Basis ι R M) : Orientation R M ι := +protected def orientation (e : Basis ι R M) : Orientation R M ι := rayOfNeZero R _ e.det_ne_zero #align basis.orientation Basis.orientation -theorem orientation_map [Nontrivial R] (e : Basis ι R M) (f : M ≃ₗ[R] N) : +theorem orientation_map (e : Basis ι R M) (f : M ≃ₗ[R] N) : (e.map f).orientation = Orientation.map ι f e.orientation := by simp_rw [Basis.orientation, Orientation.map_apply, Basis.det_map'] #align basis.orientation_map Basis.orientation_map -theorem orientation_reindex [Nontrivial R] (e : Basis ι R M) (eι : ι ≃ ι') : +theorem orientation_reindex (e : Basis ι R M) (eι : ι ≃ ι') : (e.reindex eι).orientation = Orientation.reindex R M eι e.orientation := by simp_rw [Basis.orientation, Orientation.reindex_apply, Basis.det_reindex'] #align basis.orientation_reindex Basis.orientation_reindex /-- The orientation given by a basis derived using `units_smul`, in terms of the product of those units. -/ -theorem orientation_unitsSMul [Nontrivial R] (e : Basis ι R M) (w : ι → Units R) : +theorem orientation_unitsSMul (e : Basis ι R M) (w : ι → Units R) : (e.unitsSMul w).orientation = (∏ i, w i)⁻¹ • e.orientation := by rw [Basis.orientation, Basis.orientation, smul_rayOfNeZero, ray_eq_iff, e.det.eq_smul_basis_det (e.unitsSMul w), det_unitsSMul_self, Units.smul_def, smul_smul] @@ -206,7 +206,7 @@ theorem orientation_unitsSMul [Nontrivial R] (e : Basis ι R M) (w : ι → Unit #align basis.orientation_units_smul Basis.orientation_unitsSMul @[simp] -theorem orientation_isEmpty [Nontrivial R] [IsEmpty ι] (b : Basis ι R M) : +theorem orientation_isEmpty [IsEmpty ι] (b : Basis ι R M) : b.orientation = positiveOrientation := by rw [Basis.orientation] congr @@ -230,7 +230,7 @@ namespace Orientation /-- A module `M` over a linearly ordered commutative ring has precisely two "orientations" with respect to an empty index type. (Note that these are only orientations of `M` of in the conventional mathematical sense if `M` is zero-dimensional.) -/ -theorem eq_or_eq_neg_of_isEmpty [Nontrivial R] [IsEmpty ι] (o : Orientation R M ι) : +theorem eq_or_eq_neg_of_isEmpty [IsEmpty ι] (o : Orientation R M ι) : o = positiveOrientation ∨ o = -positiveOrientation := by induction' o using Module.Ray.ind with x hx dsimp [positiveOrientation] @@ -300,7 +300,7 @@ theorem orientation_comp_linearEquiv_eq_neg_iff_det_neg (e : Basis ι R M) (f : /-- Negating a single basis vector (represented using `units_smul`) negates the corresponding orientation. -/ @[simp] -theorem orientation_neg_single [Nontrivial R] (e : Basis ι R M) (i : ι) : +theorem orientation_neg_single (e : Basis ι R M) (i : ι) : (e.unitsSMul (Function.update 1 i (-1))).orientation = -e.orientation := by rw [orientation_unitsSMul, Finset.prod_update_of_mem (Finset.mem_univ _)] simp @@ -308,7 +308,7 @@ theorem orientation_neg_single [Nontrivial R] (e : Basis ι R M) (i : ι) : /-- Given a basis and an orientation, return a basis giving that orientation: either the original basis, or one constructed by negating a single (arbitrary) basis vector. -/ -def adjustToOrientation [Nontrivial R] [Nonempty ι] (e : Basis ι R M) (x : Orientation R M ι) : +def adjustToOrientation [Nonempty ι] (e : Basis ι R M) (x : Orientation R M ι) : Basis ι R M := haveI := Classical.decEq (Orientation R M ι) if e.orientation = x then e else e.unitsSMul (Function.update 1 (Classical.arbitrary ι) (-1)) @@ -316,7 +316,7 @@ def adjustToOrientation [Nontrivial R] [Nonempty ι] (e : Basis ι R M) (x : Ori /-- `adjust_to_orientation` gives a basis with the required orientation. -/ @[simp] -theorem orientation_adjustToOrientation [Nontrivial R] [Nonempty ι] (e : Basis ι R M) +theorem orientation_adjustToOrientation [Nonempty ι] (e : Basis ι R M) (x : Orientation R M ι) : (e.adjustToOrientation x).orientation = x := by rw [adjustToOrientation] split_ifs with h @@ -327,7 +327,7 @@ theorem orientation_adjustToOrientation [Nontrivial R] [Nonempty ι] (e : Basis /-- Every basis vector from `adjust_to_orientation` is either that from the original basis or its negation. -/ -theorem adjustToOrientation_apply_eq_or_eq_neg [Nontrivial R] [Nonempty ι] (e : Basis ι R M) +theorem adjustToOrientation_apply_eq_or_eq_neg [Nonempty ι] (e : Basis ι R M) (x : Orientation R M ι) (i : ι) : e.adjustToOrientation x i = e i ∨ e.adjustToOrientation x i = -e i := by rw [adjustToOrientation] @@ -336,7 +336,7 @@ theorem adjustToOrientation_apply_eq_or_eq_neg [Nontrivial R] [Nonempty ι] (e : · by_cases hi : i = Classical.arbitrary ι <;> simp [unitsSMul_apply, hi] #align basis.adjust_to_orientation_apply_eq_or_eq_neg Basis.adjustToOrientation_apply_eq_or_eq_neg -theorem det_adjustToOrientation [Nontrivial R] [Nonempty ι] (e : Basis ι R M) +theorem det_adjustToOrientation [Nonempty ι] (e : Basis ι R M) (x : Orientation R M ι) : (e.adjustToOrientation x).det = e.det ∨ (e.adjustToOrientation x).det = -e.det := by dsimp [Basis.adjustToOrientation] @@ -351,7 +351,7 @@ theorem det_adjustToOrientation [Nontrivial R] [Nonempty ι] (e : Basis ι R M) #align basis.det_adjust_to_orientation Basis.det_adjustToOrientation @[simp] -theorem abs_det_adjustToOrientation [Nontrivial R] [Nonempty ι] (e : Basis ι R M) +theorem abs_det_adjustToOrientation [Nonempty ι] (e : Basis ι R M) (x : Orientation R M ι) (v : ι → M) : |(e.adjustToOrientation x).det v| = |e.det v| := by cases' e.det_adjustToOrientation x with h h <;> simp [h] #align basis.abs_det_adjust_to_orientation Basis.abs_det_adjustToOrientation diff --git a/Mathlib/LinearAlgebra/PiTensorProduct.lean b/Mathlib/LinearAlgebra/PiTensorProduct.lean index a3b16d0d5b2bf..7554ee8a356c7 100644 --- a/Mathlib/LinearAlgebra/PiTensorProduct.lean +++ b/Mathlib/LinearAlgebra/PiTensorProduct.lean @@ -114,10 +114,9 @@ def PiTensorProduct : Type _ := variable {R} unsuppress_compilation in --- This enables the notation `⨂[R] i : ι, s i` for the pi tensor product, given `s : ι → Type*`. ---scoped[TensorProduct] -- Porting note: `scoped` caused an error, so I commented it out. -/-- notation for tensor product over some indexed type -/ -notation3:100"⨂["R"] "(...)", "r:(scoped f => PiTensorProduct R f) => r +/-- This enables the notation `⨂[R] i : ι, s i` for the pi tensor product `PiTensorProduct`, +given an indexed family of types `s : ι → Type*`. -/ +scoped[TensorProduct] notation3:100"⨂["R"] "(...)", "r:(scoped f => PiTensorProduct R f) => r open TensorProduct @@ -311,9 +310,8 @@ variable {R} unsuppress_compilation in /-- pure tensor in tensor product over some index type -/ --- Porting note: use `FunLike.coe` as an explicit coercion to help `notation3` pretty print, --- was just `tprod R f`. -notation3:100 "⨂ₜ["R"] "(...)", "r:(scoped f => FunLike.coe (tprod R) f) => r +-- TODO(kmill) The generated delaborator never applies; figure out why this doesn't pretty print. +notation3:100 "⨂ₜ["R"] "(...)", "r:(scoped f => tprod R f) => r --Porting note: new theorem theorem tprod_eq_tprodCoeff_one : diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 0bfc666f5c2cb..6ade80c39662c 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.LinearAlgebra.Matrix.Determinant import Mathlib.LinearAlgebra.Matrix.BilinearForm import Mathlib.LinearAlgebra.Matrix.Symmetric -#align_import linear_algebra.quadratic_form.basic from "leanprover-community/mathlib"@"11b92770e4d49ff3982504c4dab918ac0887fe33" +#align_import linear_algebra.quadratic_form.basic from "leanprover-community/mathlib"@"d11f435d4e34a6cea0a1797d6b625b0c170be845" /-! # Quadratic forms diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean b/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean index 997dd5e7d6123..9a09c469912d6 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean @@ -54,7 +54,7 @@ noncomputable def isometryEquivSumSquares [DecidableEq ι] (w' : ι → ℂ) : split_ifs with h · simp only [h, zero_smul, zero_mul] have hww' : w' j = w j := by simp only [dif_neg h, Units.val_mk0] - simp only [one_mul, Units.val_mk0, smul_eq_mul] + simp (config := {zeta := false}) only [one_mul, Units.val_mk0, smul_eq_mul] rw [hww'] suffices v j * v j = w j ^ (-(1 / 2 : ℂ)) * w j ^ (-(1 / 2 : ℂ)) * w j * v j * v j by rw [this]; ring diff --git a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean index 32f31c6e273d4..a82a1a4c39d7e 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean @@ -71,7 +71,7 @@ abbrev ofHom {X : Type v} [AddCommGroup X] [Module R X] (f ≫ g).toIsometry = g.toIsometry.comp f.toIsometry := rfl -@[simp] theorem toIsometry_id {M : QuadraticModuleCat.{v} R} : +@[simp] theorem toIsometry_id {M : QuadraticModuleCat.{v} R} : Hom.toIsometry (𝟙 M) = Isometry.id _ := rfl diff --git a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat/Monoidal.lean b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat/Monoidal.lean index 974d928136643..a064029c7d484 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat/Monoidal.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat/Monoidal.lean @@ -80,13 +80,13 @@ theorem forget₂_map_associator_inv (X Y Z : QuadraticModuleCat.{u} R) : noncomputable instance instMonoidalCategory : MonoidalCategory (QuadraticModuleCat.{u} R) := Monoidal.induced (forget₂ (QuadraticModuleCat R) (ModuleCat R)) - { μIsoSymm := fun X Y => Iso.refl _ - εIsoSymm := Iso.refl _ + { μIso := fun X Y => Iso.refl _ + εIso := Iso.refl _ associator_eq := fun X Y Z => by dsimp only [forget₂_obj, forget₂_map_associator_hom] simp only [eqToIso_refl, Iso.refl_trans, Iso.refl_symm, Iso.trans_hom, tensorIso_hom, Iso.refl_hom, MonoidalCategory.tensor_id] - erw [Category.id_comp, Category.comp_id, MonoidalCategory.tensor_id, Category.comp_id] + erw [Category.id_comp, Category.comp_id, MonoidalCategory.tensor_id, Category.id_comp] rfl } variable (R) in diff --git a/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean b/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean index adf045c42c8a4..50606d04b806b 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean @@ -90,12 +90,12 @@ theorem baseChange_tmul (Q : QuadraticForm R M₂) (a : A) (m₂ : M₂) : Q.baseChange A (a ⊗ₜ m₂) = Q m₂ • (a * a) := tensorDistrib_tmul _ _ _ _ -theorem associated_baseChange [Invertible (2 : A)] (Q : QuadraticForm R M₂) : +theorem associated_baseChange [Invertible (2 : A)] (Q : QuadraticForm R M₂) : associated (R := A) (Q.baseChange A) = (associated (R := R) Q).baseChange A := by dsimp only [QuadraticForm.baseChange, BilinForm.baseChange] rw [associated_tmul (QuadraticForm.sq (R := A)) Q, associated_sq] -theorem polarBilin_baseChange [Invertible (2 : A)] (Q : QuadraticForm R M₂) : +theorem polarBilin_baseChange [Invertible (2 : A)] (Q : QuadraticForm R M₂) : polarBilin (Q.baseChange A) = (polarBilin Q).baseChange A := by rw [QuadraticForm.baseChange, BilinForm.baseChange, polarBilin_tmul, BilinForm.tmul, ←LinearMap.map_smul, smul_tmul', ←two_nsmul_associated R, coe_associatedHom, associated_sq, diff --git a/Mathlib/LinearAlgebra/TensorProduct/Tower.lean b/Mathlib/LinearAlgebra/TensorProduct/Tower.lean index 2ef90c06d2c2a..7bb911dafa240 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Tower.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Tower.lean @@ -308,11 +308,14 @@ protected def rid : M ⊗[R] R ≃ₗ[A] M := theorem rid_eq_rid : AlgebraTensorModule.rid R R M = TensorProduct.rid R M := LinearEquiv.toLinearMap_injective <| TensorProduct.ext' fun _ _ => rfl -variable {R M} - +variable {R M} in @[simp] theorem rid_tmul (r : R) (m : M) : AlgebraTensorModule.rid R A M (m ⊗ₜ r) = r • m := rfl +variable {M} in +@[simp] +theorem rid_symm_apply (m : M) : (AlgebraTensorModule.rid R A M).symm m = m ⊗ₜ 1 := rfl + end Semiring section CommSemiring diff --git a/Mathlib/LinearAlgebra/Vandermonde.lean b/Mathlib/LinearAlgebra/Vandermonde.lean index b70dce683f097..806fe0359cd12 100644 --- a/Mathlib/LinearAlgebra/Vandermonde.lean +++ b/Mathlib/LinearAlgebra/Vandermonde.lean @@ -73,7 +73,7 @@ theorem vandermonde_transpose_mul_vandermonde {n : ℕ} (v : Fin n → R) (i j) #align matrix.vandermonde_transpose_mul_vandermonde Matrix.vandermonde_transpose_mul_vandermonde theorem det_vandermonde {n : ℕ} (v : Fin n → R) : - det (vandermonde v) = ∏ i : Fin n, Finset.prod (Ioi i) (fun j => v j - v i) := by + det (vandermonde v) = ∏ i : Fin n, ∏ j in Ioi i, (v j - v i) := by unfold vandermonde induction' n with n ih · exact det_eq_one_of_card_eq_zero (Fintype.card_fin 0) @@ -103,13 +103,13 @@ theorem det_vandermonde {n : ℕ} (v : Fin n → R) : rw [Fin.succAbove_zero, Matrix.cons_val_succ, Fin.val_succ, mul_comm] exact (geom_sum₂_mul (v i.succ) (v 0) (j + 1 : ℕ)).symm _ = - (Finset.prod Finset.univ (fun i => v (Fin.succ i) - v 0)) * + (∏ i in Finset.univ, (v (Fin.succ i) - v 0)) * det fun i j : Fin n => ∑ k in Finset.range (j + 1 : ℕ), v i.succ ^ k * v 0 ^ (j - k : ℕ) := (det_mul_column (fun i => v (Fin.succ i) - v 0) _) - _ = (Finset.prod Finset.univ (fun i => v (Fin.succ i) - v 0)) * + _ = (∏ i in Finset.univ, (v (Fin.succ i) - v 0)) * det fun i j : Fin n => v (Fin.succ i) ^ (j : ℕ) := congr_arg _ ?_ - _ = ∏ i : Fin n.succ, Finset.prod (Ioi i) (fun j => v j - v i) := by + _ = ∏ i : Fin n.succ, ∏ j in Ioi i, (v j - v i) := by simp_rw [Fin.prod_univ_succ, Fin.prod_Ioi_zero, Fin.prod_Ioi_succ] have h := ih (v ∘ Fin.succ) unfold Function.comp at h diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 9d1e0a2840801..0bad09637471d 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -7,7 +7,6 @@ import Mathlib.Init.Logic import Mathlib.Init.Function import Mathlib.Init.Algebra.Classes import Mathlib.Tactic.Basic -import Mathlib.Tactic.LeftRight import Std.Util.LibraryNote import Std.Tactic.Lint.Basic @@ -357,11 +356,9 @@ theorem Or.imp3 (had : a → d) (hbe : b → e) (hcf : c → f) : a ∨ b ∨ c #align or_imp_distrib or_imp -theorem or_iff_not_imp_left : a ∨ b ↔ ¬a → b := Decidable.or_iff_not_imp_left -#align or_iff_not_imp_left or_iff_not_imp_left - -theorem or_iff_not_imp_right : a ∨ b ↔ ¬b → a := Decidable.or_iff_not_imp_right -#align or_iff_not_imp_right or_iff_not_imp_right +export Classical (or_iff_not_imp_left or_iff_not_imp_right) +#align or_iff_not_imp_left Classical.or_iff_not_imp_left +#align or_iff_not_imp_right Classical.or_iff_not_imp_right theorem not_or_of_imp : (a → b) → ¬a ∨ b := Decidable.not_or_of_imp #align not_or_of_imp not_or_of_imp diff --git a/Mathlib/Logic/Embedding/Basic.lean b/Mathlib/Logic/Embedding/Basic.lean index 977d82afec1b7..5aac7a405c639 100644 --- a/Mathlib/Logic/Embedding/Basic.lean +++ b/Mathlib/Logic/Embedding/Basic.lean @@ -192,7 +192,7 @@ def setValue {α β} (f : α ↪ β) (a : α) (b : β) [∀ a', Decidable (a' = -- split_ifs at h <;> (try subst b) <;> (try simp only [f.injective.eq_iff] at *) <;> cc split_ifs at h with h₁ h₂ _ _ h₅ h₆ <;> (try subst b) <;> - (try simp only [f.injective.eq_iff] at *) + (try simp only [f.injective.eq_iff, not_true_eq_false] at *) · rw[h₁,h₂] · rw[h₁,h] · rw[h₅,←h] diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 90af77c06f775..34fabdcd679ac 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -70,6 +70,7 @@ structure Equiv (α : Sort*) (β : Sort _) where protected right_inv : RightInverse invFun toFun #align equiv Equiv +@[inherit_doc] infixl:25 " ≃ " => Equiv /-- Turn an element of a type `F` satisfying `EquivLike F α β` into an actual diff --git a/Mathlib/Logic/Equiv/LocalEquiv.lean b/Mathlib/Logic/Equiv/LocalEquiv.lean index cc691e71b646e..bc1260ec6c5f5 100644 --- a/Mathlib/Logic/Equiv/LocalEquiv.lean +++ b/Mathlib/Logic/Equiv/LocalEquiv.lean @@ -28,13 +28,13 @@ As for equivs, we register a coercion to functions and use it in our simp normal ## Main definitions -`Equiv.toLocalEquiv`: associating a local equiv to an equiv, with source = target = univ -`LocalEquiv.symm` : the inverse of a local equiv -`LocalEquiv.trans` : the composition of two local equivs -`LocalEquiv.refl` : the identity local equiv -`LocalEquiv.ofSet` : the identity on a set `s` -`EqOnSource` : equivalence relation describing the "right" notion of equality for local - equivs (see below in implementation notes) +* `Equiv.toLocalEquiv`: associating a local equiv to an equiv, with source = target = univ +* `LocalEquiv.symm`: the inverse of a local equiv +* `LocalEquiv.trans`: the composition of two local equivs +* `LocalEquiv.refl`: the identity local equiv +* `LocalEquiv.ofSet`: the identity on a set `s` +* `EqOnSource`: equivalence relation describing the "right" notion of equality for local + equivs (see below in implementation notes) ## Implementation notes @@ -827,19 +827,19 @@ theorem eqOnSource_refl : e ≈ e := Setoid.refl _ #align local_equiv.eq_on_source_refl LocalEquiv.eqOnSource_refl -/-- Two equivalent local equivs have the same source -/ +/-- Two equivalent local equivs have the same source. -/ theorem EqOnSource.source_eq {e e' : LocalEquiv α β} (h : e ≈ e') : e.source = e'.source := h.1 #align local_equiv.eq_on_source.source_eq LocalEquiv.EqOnSource.source_eq -/-- Two equivalent local equivs coincide on the source -/ +/-- Two equivalent local equivs coincide on the source. -/ theorem EqOnSource.eqOn {e e' : LocalEquiv α β} (h : e ≈ e') : e.source.EqOn e e' := h.2 #align local_equiv.eq_on_source.eq_on LocalEquiv.EqOnSource.eqOn --Porting note: A lot of dot notation failures here. Maybe we should not use `≈` -/-- Two equivalent local equivs have the same target -/ +/-- Two equivalent local equivs have the same target. -/ theorem EqOnSource.target_eq {e e' : LocalEquiv α β} (h : e ≈ e') : e.target = e'.target := by simp only [← image_source_eq_target, ← source_eq h, h.2.image_eq] #align local_equiv.eq_on_source.target_eq LocalEquiv.EqOnSource.target_eq @@ -851,14 +851,14 @@ theorem EqOnSource.symm' {e e' : LocalEquiv α β} (h : e ≈ e') : e.symm ≈ e exact e'.rightInvOn.congr_right e'.symm_mapsTo (source_eq h ▸ h.eqOn.symm) #align local_equiv.eq_on_source.symm' LocalEquiv.EqOnSource.symm' -/-- Two equivalent local equivs have coinciding inverses on the target -/ +/-- Two equivalent local equivs have coinciding inverses on the target. -/ theorem EqOnSource.symm_eqOn {e e' : LocalEquiv α β} (h : e ≈ e') : EqOn e.symm e'.symm e.target := -- Porting note: `h.symm'` dot notation doesn't work anymore because `h` is not recognised as -- `LocalEquiv.EqOnSource` for some reason. eqOn (symm' h) #align local_equiv.eq_on_source.symm_eq_on LocalEquiv.EqOnSource.symm_eqOn -/-- Composition of local equivs respects equivalence -/ +/-- Composition of local equivs respects equivalence. -/ theorem EqOnSource.trans' {e e' : LocalEquiv α β} {f f' : LocalEquiv β γ} (he : e ≈ e') (hf : f ≈ f') : e.trans f ≈ e'.trans f' := by constructor @@ -869,7 +869,7 @@ theorem EqOnSource.trans' {e e' : LocalEquiv α β} {f f' : LocalEquiv β γ} (h simp [Function.comp_apply, LocalEquiv.coe_trans, (he.2 hx.1).symm, hf.2 hx.2] #align local_equiv.eq_on_source.trans' LocalEquiv.EqOnSource.trans' -/-- Restriction of local equivs respects equivalence -/ +/-- Restriction of local equivs respects equivalence. -/ theorem EqOnSource.restr {e e' : LocalEquiv α β} (he : e ≈ e') (s : Set α) : e.restr s ≈ e'.restr s := by constructor @@ -879,13 +879,13 @@ theorem EqOnSource.restr {e e' : LocalEquiv α β} (he : e ≈ e') (s : Set α) exact he.2 hx.1 #align local_equiv.eq_on_source.restr LocalEquiv.EqOnSource.restr -/-- Preimages are respected by equivalence -/ +/-- Preimages are respected by equivalence. -/ theorem EqOnSource.source_inter_preimage_eq {e e' : LocalEquiv α β} (he : e ≈ e') (s : Set β) : e.source ∩ e ⁻¹' s = e'.source ∩ e' ⁻¹' s := by rw [he.eqOn.inter_preimage_eq, source_eq he] #align local_equiv.eq_on_source.source_inter_preimage_eq LocalEquiv.EqOnSource.source_inter_preimage_eq /-- Composition of a local equiv and its inverse is equivalent to the restriction of the identity -to the source -/ +to the source. -/ theorem trans_self_symm : e.trans e.symm ≈ ofSet e.source := by have A : (e.trans e.symm).source = e.source := by mfld_set_tac refine' ⟨by rw [A, ofSet_source], fun x hx => _⟩ @@ -894,13 +894,13 @@ theorem trans_self_symm : e.trans e.symm ≈ ofSet e.source := by #align local_equiv.trans_self_symm LocalEquiv.trans_self_symm /-- Composition of the inverse of a local equiv and this local equiv is equivalent to the -restriction of the identity to the target -/ +restriction of the identity to the target. -/ theorem trans_symm_self : e.symm.trans e ≈ LocalEquiv.ofSet e.target := trans_self_symm e.symm #align local_equiv.trans_symm_self LocalEquiv.trans_symm_self -/-- Two equivalent local equivs are equal when the source and target are univ -/ -theorem eq_of_eq_on_source_univ (e e' : LocalEquiv α β) (h : e ≈ e') (s : e.source = univ) +/-- Two equivalent local equivs are equal when the source and target are `univ`. -/ +theorem eq_of_eqOnSource_univ (e e' : LocalEquiv α β) (h : e ≈ e') (s : e.source = univ) (t : e.target = univ) : e = e' := by refine LocalEquiv.ext (fun x => ?_) (fun x => ?_) h.1 · apply h.2 @@ -909,7 +909,7 @@ theorem eq_of_eq_on_source_univ (e e' : LocalEquiv α β) (h : e ≈ e') (s : e. · apply h.symm'.2 rw [symm_source, t] exact mem_univ _ -#align local_equiv.eq_of_eq_on_source_univ LocalEquiv.eq_of_eq_on_source_univ +#align local_equiv.eq_of_eq_on_source_univ LocalEquiv.eq_of_eqOnSource_univ section Prod diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index 51a44a1749748..d1aa8c202f9b5 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -122,23 +122,30 @@ protected def Injective.decidableEq [DecidableEq β] (I : Injective f) : Decidab fun _ _ ↦ decidable_of_iff _ I.eq_iff #align function.injective.decidable_eq Function.Injective.decidableEq -theorem Injective.of_comp {g : γ → α} (I : Injective (f ∘ g)) : Injective g := fun x y h ↦ - I <| show f (g x) = f (g y) from congr_arg f h +theorem Injective.of_comp {g : γ → α} (I : Injective (f ∘ g)) : Injective g := + fun _ _ h ↦ I <| congr_arg f h #align function.injective.of_comp Function.Injective.of_comp @[simp] -theorem Injective.of_comp_iff {f : α → β} (hf : Injective f) (g : γ → α) : +theorem Injective.of_comp_iff (hf : Injective f) (g : γ → α) : Injective (f ∘ g) ↔ Injective g := ⟨Injective.of_comp, hf.comp⟩ #align function.injective.of_comp_iff Function.Injective.of_comp_iff +theorem Injective.of_comp_right {g : γ → α} (I : Injective (f ∘ g)) (hg : Surjective g) : + Injective f := fun x y h ↦ by + obtain ⟨x, rfl⟩ := hg x + obtain ⟨y, rfl⟩ := hg y + exact congr_arg g (I h) + +theorem Surjective.bijective₂_of_injective {g : γ → α} (hf : Surjective f) (hg : Surjective g) + (I : Injective (f ∘ g)) : Bijective f ∧ Bijective g := + ⟨⟨I.of_comp_right hg, hf⟩, I.of_comp, hg⟩ + @[simp] theorem Injective.of_comp_iff' (f : α → β) {g : γ → α} (hg : Bijective g) : Injective (f ∘ g) ↔ Injective f := -⟨ λ h x y => let ⟨_, hx⟩ := hg.surjective x - let ⟨_, hy⟩ := hg.surjective y - hx ▸ hy ▸ λ hf => h hf ▸ rfl, - λ h => h.comp hg.injective⟩ + ⟨fun I ↦ I.of_comp_right hg.2, fun h ↦ h.comp hg.injective⟩ #align function.injective.of_comp_iff' Function.Injective.of_comp_iff' /-- Composition by an injective function on the left is itself injective. -/ @@ -176,13 +183,17 @@ theorem Surjective.of_comp_iff (f : α → β) {g : γ → α} (hg : Surjective ⟨Surjective.of_comp, fun h ↦ h.comp hg⟩ #align function.surjective.of_comp_iff Function.Surjective.of_comp_iff +theorem Surjective.of_comp_left {g : γ → α} (S : Surjective (f ∘ g)) (hf : Injective f) : + Surjective g := fun a ↦ let ⟨c, hc⟩ := S (f a); ⟨c, hf hc⟩ + +theorem Injective.bijective₂_of_surjective {g : γ → α} (hf : Injective f) (hg : Injective g) + (S : Surjective (f ∘ g)) : Bijective f ∧ Bijective g := + ⟨⟨hf, S.of_comp⟩, hg, S.of_comp_left hf⟩ + @[simp] theorem Surjective.of_comp_iff' (hf : Bijective f) (g : γ → α) : Surjective (f ∘ g) ↔ Surjective g := - ⟨fun h x ↦ - let ⟨x', hx'⟩ := h (f x) - ⟨x', hf.injective hx'⟩, - hf.surjective.comp⟩ + ⟨fun S ↦ S.of_comp_left hf.1, hf.surjective.comp⟩ #align function.surjective.of_comp_iff' Function.Surjective.of_comp_iff' instance decidableEqPfun (p : Prop) [Decidable p] (α : p → Type*) [∀ hp, DecidableEq (α hp)] : diff --git a/Mathlib/Logic/Hydra.lean b/Mathlib/Logic/Hydra.lean index c0cb9ad258153..9130d740a1e83 100644 --- a/Mathlib/Logic/Hydra.lean +++ b/Mathlib/Logic/Hydra.lean @@ -110,10 +110,7 @@ theorem cutExpand_fibration (r : α → α → Prop) : Fibration (GameAdd (CutExpand r) (CutExpand r)) (CutExpand r) fun s ↦ s.1 + s.2 := by rintro ⟨s₁, s₂⟩ s ⟨t, a, hr, he⟩; dsimp at he ⊢ classical - -- Porting note: Originally `obtain ⟨ha, rfl⟩` - -- This is https://github.com/leanprover/std4/issues/62 - obtain ⟨ha, hb⟩ := add_singleton_eq_iff.1 he - rw [hb] + obtain ⟨ha, rfl⟩ := add_singleton_eq_iff.1 he rw [add_assoc, mem_add] at ha obtain h | h := ha · refine' ⟨(s₁.erase a + t, s₂), GameAdd.fst ⟨t, a, hr, _⟩, _⟩ diff --git a/Mathlib/Mathport/Notation.lean b/Mathlib/Mathport/Notation.lean index e508f51f3c577..c92dae794875f 100644 --- a/Mathlib/Mathport/Notation.lean +++ b/Mathlib/Mathport/Notation.lean @@ -3,7 +3,10 @@ Copyright (c) 2021 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kyle Miller -/ +import Mathlib.Lean.Elab.Term import Mathlib.Lean.Expr +import Mathlib.Lean.PrettyPrinter.Delaborator +import Mathlib.Tactic.ScopedNS import Mathlib.Util.Syntax import Std.Data.Option.Basic @@ -158,63 +161,119 @@ def matchVar (c : Name) : Matcher := fun s => do else s.captureSubexpr c -/-- Matcher for `Expr.const`. -/ -def matchConst (c : Name) : Matcher := fun s => do - guard <| (← getExpr).isConstOf c +/-- Matcher for an expression satisfying a given predicate. -/ +def matchExpr (p : Expr → Bool) : Matcher := fun s => do + guard <| p (← getExpr) return s -/-- Matcher for `Expr.fvar`. This is only used for `local notation3`. -It checks that the user name agrees, which isn't completely accurate, but probably sufficient. -/ -def matchFVar (userName : Name) : Matcher := fun s => do +/-- Matcher for `Expr.fvar`. +It checks that the user name agrees and that the type of the expression is matched by `matchTy`. -/ +def matchFVar (userName : Name) (matchTy : Matcher) : Matcher := fun s => do let .fvar fvarId ← getExpr | failure guard <| userName == (← fvarId.getUserName) - return s + withType (matchTy s) + +/-- Matcher that checks that the type of the expression is matched by `matchTy`. -/ +def matchTypeOf (matchTy : Matcher) : Matcher := fun s => do + withType (matchTy s) -/-- Matches raw nat lits and `OfNat.ofNat` expressions. -/ +/-- Matches raw nat lits. -/ def natLitMatcher (n : Nat) : Matcher := fun s => do - let mut e ← getExpr - if e.isAppOfArity ``OfNat.ofNat 3 then - e := e.getArg! 1 - guard <| e.natLit? == n + guard <| (← getExpr).natLit? == n return s -/-- Given an identifier `f`, returns -(1) the resolved constant (if it's not an fvar) -(2) a Term for a matcher for the function -(3) the arity -(4) the positions at which the function takes an explicit argument -/ -def getExplicitArgIndices (f : Syntax) : - OptionT TermElabM (Option Name × Term × Nat × Array Nat) := do - let fe? ← try liftM <| Term.resolveId? f catch _ => pure none - match fe? with - | some fe@(.const f _) => - return (some f, ← ``(matchConst $(quote f)), ← collectIdxs (← inferType fe)) - | some fe@(.fvar fvarId) => - let userName ← fvarId.getUserName - return (none, ← ``(matchFVar $(quote userName)), ← collectIdxs (← inferType fe)) +/-- Matches applications. -/ +def matchApp (matchFun matchArg : Matcher) : Matcher := fun s => do + guard <| (← getExpr).isApp + let s ← withAppFn <| matchFun s + let s ← withAppArg <| matchArg s + return s + +/-- Matches pi types. The name `n` should be unique, and `matchBody` should use `n` +as the `userName` of its fvar. -/ +def matchForall (matchDom : Matcher) (matchBody : Expr → Matcher) : Matcher := fun s => do + guard <| (← getExpr).isForall + let s ← withBindingDomain <| matchDom s + let s ← withBindingBodyUnusedName' fun _ arg => matchBody arg s + return s + +/-- Matches lambdas. The `matchBody` takes the fvar introduced when visiting the body. -/ +def matchLambda (matchDom : Matcher) (matchBody : Expr → Matcher) : Matcher := fun s => do + guard <| (← getExpr).isLambda + let s ← withBindingDomain <| matchDom s + let s ← withBindingBodyUnusedName' fun _ arg => matchBody arg s + return s + +/-- Adds all the names in `boundNames` to the local context +with types that are fresh metavariables. +This is used for example when initializing `p` in `(scoped p => ...)` when elaborating `...`. -/ +def setupLCtx (lctx : LocalContext) (boundNames : Array Name) : + MetaM (LocalContext × HashMap FVarId Name) := do + let mut lctx := lctx + let mut boundFVars := {} + for name in boundNames do + let fvarId ← mkFreshFVarId + lctx := lctx.mkLocalDecl fvarId name (← withLCtx lctx (← getLocalInstances) mkFreshTypeMVar) + boundFVars := boundFVars.insert fvarId name + return (lctx, boundFVars) + +/-- Given an expression, generate a matcher for it. +The `boundFVars` hash map records which state variables certain fvars correspond to. +The `localFVars` hash map records which local variable the matcher should use for an exact +expression match. + +If it succeeds generating a matcher, returns +1. a list of keys that should be used for the `delab` attribute + when defining the elaborator +2. a `Term` that represents a `Matcher` for the given expression `e`. -/ +partial def exprToMatcher (boundFVars : HashMap FVarId Name) (localFVars : HashMap FVarId Term) + (e : Expr) : + OptionT TermElabM (List Name × Term) := do + match e with + | .mvar .. => return ([], ← `(pure)) + | .const n _ => return ([`app ++ n], ← ``(matchExpr (Expr.isConstOf · $(quote n)))) + | .sort .. => return ([`sort], ← ``(matchExpr Expr.isSort)) + | .fvar fvarId => + if let some n := boundFVars.find? fvarId then + -- This fvar is a pattern variable. + return ([], ← ``(matchVar $(quote n))) + else if let some s := localFVars.find? fvarId then + -- This fvar is bound by a lambda or forall expression in the pattern itself + return ([], ← ``(matchExpr (· == $s))) + else + let n ← fvarId.getUserName + if n.hasMacroScopes then + -- Match by just the type; this is likely an unnamed instance for example + let (_, m) ← exprToMatcher boundFVars localFVars (← instantiateMVars (← inferType e)) + return ([`fvar], ← ``(matchTypeOf $m)) + else + -- This is an fvar from a `variable`. Match by name and type. + let (_, m) ← exprToMatcher boundFVars localFVars (← instantiateMVars (← inferType e)) + return ([`fvar], ← ``(matchFVar $(quote n) $m)) + | .app f arg => + let (keys, matchF) ← exprToMatcher boundFVars localFVars f + let (_, matchArg) ← exprToMatcher boundFVars localFVars arg + return (if keys.isEmpty then [`app] else keys, ← ``(matchApp $matchF $matchArg)) + | .lit (.natVal n) => return ([`lit], ← ``(natLitMatcher $(quote n))) + | .forallE n t b bi => + let (_, matchDom) ← exprToMatcher boundFVars localFVars t + withLocalDecl n bi t fun arg => withFreshMacroScope do + let n' ← `(n) + let body := b.instantiate1 arg + let localFVars' := localFVars.insert arg.fvarId! n' + let (_, matchBody) ← exprToMatcher boundFVars localFVars' body + return ([`forallE], ← ``(matchForall $matchDom (fun $n' => $matchBody))) + | .lam n t b bi => + let (_, matchDom) ← exprToMatcher boundFVars localFVars t + withLocalDecl n bi t fun arg => withFreshMacroScope do + let n' ← `(n) + let body := b.instantiate1 arg + let localFVars' := localFVars.insert arg.fvarId! n' + let (_, matchBody) ← exprToMatcher boundFVars localFVars' body + return ([`lam], ← ``(matchLambda $matchDom (fun $n' => $matchBody))) | _ => - trace[notation3] "could not resolve name {f}" + trace[notation3] "can't generate matcher for {e}" failure -where - collectIdxs (ty : Expr) : MetaM (Nat × Array Nat) := do - let (_, binderInfos, _) ← Meta.forallMetaTelescope ty - let mut idxs := #[] - for bi in binderInfos, i in [0:binderInfos.size] do - if bi.isExplicit then - idxs := idxs.push i - return (binderInfos.size, idxs) - -/-- A matcher that runs `matchf` for the function and the `matchers` for the associated -argument indices. Fails if the function doesn't have exactly `arity` arguments. -/ -def fnArgMatcher (arity : Nat) (matchf : Matcher) (matchers : Array (Nat × Matcher)) : - Matcher := fun s => do - let mut s := s - let nargs := (← getExpr).getAppNumArgs - guard <| nargs == arity - s ← withNaryFn <| matchf s - for (i, matcher) in matchers do - s ← withNaryArg i <| matcher s - return s /-- Returns a `Term` that represents a `Matcher` for the given pattern `stx`. The `boundNames` set determines which identifiers are variables in the pattern. @@ -222,44 +281,21 @@ Fails in the `OptionT` sense if it comes across something it's unable to handle. Also returns constant names that could serve as a key for a delaborator. For example, if it's for a function `f`, then `app.f`. -/ -partial def mkExprMatcher (stx : Term) (boundNames : HashSet Name) : +partial def mkExprMatcher (stx : Term) (boundNames : Array Name) : OptionT TermElabM (List Name × Term) := do - let stx'? ← liftM <| (liftMacroM <| expandMacro? stx : TermElabM (Option Syntax)) - match stx'? with - | some stx' => mkExprMatcher ⟨stx'⟩ boundNames - | none => - match stx with - | `(_) => return ([], ← `(pure)) - | `($n:ident) => - if boundNames.contains n.getId then - return ([], ← ``(matchVar $(quote n.getId))) - else - processFn n #[] - | `($f:ident $args*) => processFn f args - | `(term| $n:num) => return ([], ← ``(natLitMatcher $n)) - | `(($stx)) => - if Term.hasCDot stx then + let (lctx, boundFVars) ← setupLCtx (← getLCtx) boundNames + withLCtx lctx (← getLocalInstances) do + let patt ← + try + Term.elabPattern stx none + catch e => + logException e + trace[notation3] "Could not elaborate pattern{indentD stx}\nError: {e.toMessageData}" + -- Convert the exception into an `OptionT` failure so that the `(prettyPrint := false)` + -- suggestion appears. failure - else - mkExprMatcher stx boundNames - | _ => - trace[notation3] "mkExprMatcher can't handle {stx}" - failure -where - processFn (f : Term) (args : TSyntaxArray `term) : OptionT TermElabM (List Name × Term) := do - let (name?, matchf, arity, idxs) ← getExplicitArgIndices f - unless args.size ≤ idxs.size do - trace[notation3] "Function {f} has been given more explicit arguments than expected" - failure - let mut matchers := #[] - for i in idxs, arg in args do - let (_, matcher) ← mkExprMatcher arg boundNames - matchers := matchers.push <| ← `(($(quote i), $matcher)) - -- `arity'` is the number of arguments (including trailing implicits) given these - -- explicit arguments. This reflects how the function would be elaborated. - let arity' := if _ : args.size < idxs.size then idxs[args.size] else arity - let key? := name?.map (`app ++ ·) - return (key?.toList, ← ``(fnArgMatcher $(quote arity') $matchf #[$matchers,*])) + trace[notation3] "Generating matcher for pattern {patt}" + exprToMatcher boundFVars {} patt /-- Matcher for processing `scoped` syntax. Assumes the expression to be matched against is in the `lit` variable. @@ -308,10 +344,10 @@ partial def matchScoped (lit scopeId : Name) (smatcher : Matcher) : Matcher := g Fails in the `OptionT` sense if a matcher couldn't be constructed. Also returns a delaborator key like in `mkExprMatcher`. Reminder: `$lit:ident : (scoped $scopedId:ident => $scopedTerm:Term)` -/ -partial def mkScopedMatcher (lit scopeId : Name) (scopedTerm : Term) (boundNames : HashSet Name) : +partial def mkScopedMatcher (lit scopeId : Name) (scopedTerm : Term) (boundNames : Array Name) : OptionT TermElabM (List Name × Term) := do -- Build the matcher for `scopedTerm` with `scopeId` as an additional variable - let (keys, smatcher) ← mkExprMatcher scopedTerm (boundNames.insert scopeId) + let (keys, smatcher) ← mkExprMatcher scopedTerm (boundNames.push scopeId) return (keys, ← ``(matchScoped $(quote lit) $(quote scopeId) $smatcher)) /-- Matcher for expressions produced by `foldl`. -/ @@ -337,20 +373,20 @@ partial def matchFoldl (lit x y : Name) (smatcher : Matcher) (sinit : Matcher) : /-- Create a `Term` that represents a matcher for `foldl` notation. Reminder: `( lit ","* => foldl (x y => scopedTerm) init)` -/ -partial def mkFoldlMatcher (lit x y : Name) (scopedTerm init : Term) (boundNames : HashSet Name) : +partial def mkFoldlMatcher (lit x y : Name) (scopedTerm init : Term) (boundNames : Array Name) : OptionT TermElabM (List Name × Term) := do -- Build the `scopedTerm` matcher with `x` and `y` as additional variables - let boundNames' := boundNames |>.insert x |>.insert y + let boundNames' := boundNames |>.push x |>.push y let (keys, smatcher) ← mkExprMatcher scopedTerm boundNames' let (keys', sinit) ← mkExprMatcher init boundNames return (keys ++ keys', ← ``(matchFoldl $(quote lit) $(quote x) $(quote y) $smatcher $sinit)) /-- Create a `Term` that represents a matcher for `foldr` notation. Reminder: `( lit ","* => foldr (x y => scopedTerm) init)` -/ -partial def mkFoldrMatcher (lit x y : Name) (scopedTerm init : Term) (boundNames : HashSet Name) : +partial def mkFoldrMatcher (lit x y : Name) (scopedTerm init : Term) (boundNames : Array Name) : OptionT TermElabM (List Name × Term) := do -- Build the `scopedTerm` matcher with `x` and `y` as additional variables - let boundNames' := boundNames |>.insert x |>.insert y + let boundNames' := boundNames |>.push x |>.push y let (keys, smatcher) ← mkExprMatcher scopedTerm boundNames' let (keys', sinit) ← mkExprMatcher init boundNames -- N.B. by swapping `x` and `y` we can just use the foldl matcher @@ -416,12 +452,10 @@ elab (name := notation3) doc:(docComment)? attrs?:(Parser.Term.attributes)? attr let mut boundIdents : HashMap Name Ident := {} -- Replacements to use for the `macro` let mut boundValues : HashMap Name Syntax := {} + -- The names of the bound names in order, used when constructing patterns for delaboration. + let mut boundNames : Array Name := #[] -- The normal/foldl/foldr type of each variable (for delaborator) let mut boundType : HashMap Name BoundValueType := {} - -- Function to get the keys of `boundValues`. This set is used when constructing - -- patterns for delaboration. - let getBoundNames (boundValues : HashMap Name Syntax) : HashSet Name := - HashSet.empty.insertMany <| boundValues.toArray.map Prod.fst -- Function to update `syntaxArgs` and `pattArgs` using `macroArg` syntax let pushMacro (syntaxArgs : Array (TSyntax `stx)) (pattArgs : Array Syntax) (mac : TSyntax ``macroArg) := do @@ -466,32 +500,36 @@ elab (name := notation3) doc:(docComment)? attrs?:(Parser.Term.attributes)? attr | `(foldKind| foldl) => boundValues := boundValues.insert id.getId <| ← `(expand_foldl% ($x $y => $scopedTerm') $init' [$$(.ofElems $id),*]) + boundNames := boundNames.push id.getId boundType := boundType.insert id.getId .foldl matchers := matchers.push <| - mkFoldlMatcher id.getId x.getId y.getId scopedTerm init (getBoundNames boundValues) + mkFoldlMatcher id.getId x.getId y.getId scopedTerm init boundNames | `(foldKind| foldr) => boundValues := boundValues.insert id.getId <| ← `(expand_foldr% ($x $y => $scopedTerm') $init' [$$(.ofElems $id),*]) + boundNames := boundNames.push id.getId boundType := boundType.insert id.getId .foldr matchers := matchers.push <| - mkFoldrMatcher id.getId x.getId y.getId scopedTerm init (getBoundNames boundValues) + mkFoldrMatcher id.getId x.getId y.getId scopedTerm init boundNames | _ => throwUnsupportedSyntax | `(notation3Item| $lit:ident $(prec?)? : (scoped $scopedId:ident => $scopedTerm)) => hasScoped := true (syntaxArgs, pattArgs) ← pushMacro syntaxArgs pattArgs <|← `(macroArg| $lit:ident:term $(prec?)?) matchers := matchers.push <| - mkScopedMatcher lit.getId scopedId.getId scopedTerm (getBoundNames boundValues) + mkScopedMatcher lit.getId scopedId.getId scopedTerm boundNames let scopedTerm' ← scopedTerm.replaceM fun s => pure (boundValues.find? s.getId) boundIdents := boundIdents.insert lit.getId lit boundValues := boundValues.insert lit.getId <| ← `(expand_binders% ($scopedId => $scopedTerm') $$binders:extBinders, $(⟨lit.1.mkAntiquotNode `term⟩):term) + boundNames := boundNames.push lit.getId | `(notation3Item| $lit:ident $(prec?)?) => (syntaxArgs, pattArgs) ← pushMacro syntaxArgs pattArgs <|← `(macroArg| $lit:ident:term $(prec?)?) boundIdents := boundIdents.insert lit.getId lit boundValues := boundValues.insert lit.getId <| lit.1.mkAntiquotNode `term + boundNames := boundNames.push lit.getId | stx => throwUnsupportedSyntax if hasScoped && !hasBindersItem then throwError "If there is a `scoped` item then there must be a `(...)` item for binders." @@ -518,7 +556,7 @@ elab (name := notation3) doc:(docComment)? attrs?:(Parser.Term.attributes)? attr -- 3. Create a delaborator if getPrettyPrintOpt pp? then - matchers := matchers.push <| Mathlib.Notation3.mkExprMatcher val (getBoundNames boundValues) + matchers := matchers.push <| Mathlib.Notation3.mkExprMatcher val boundNames -- The matchers need to run in reverse order, so may as well reverse them here. let matchersM? := (matchers.reverse.mapM id).run -- We let local notations have access to `variable` declarations @@ -560,3 +598,10 @@ elab (name := notation3) doc:(docComment)? attrs?:(Parser.Term.attributes)? attr ""} (Use `set_option trace.notation3 true` to get some debug information.)" initialize Std.Linter.UnreachableTactic.addIgnoreTacticKind ``«notation3» + +/-! `scoped[ns]` support -/ + +macro_rules + | `($[$doc]? $(attr)? scoped[$ns] notation3 $(prec)? $(n)? $(prio)? $(pp)? $items* => $t) => + `(with_weak_namespace $(mkIdentFrom ns <| rootNamespace ++ ns.getId) + $[$doc]? $(attr)? scoped notation3 $(prec)? $(n)? $(prio)? $(pp)? $items* => $t) diff --git a/Mathlib/Mathport/Syntax.lean b/Mathlib/Mathport/Syntax.lean index 085ec006637ca..142a42abf886e 100644 --- a/Mathlib/Mathport/Syntax.lean +++ b/Mathlib/Mathport/Syntax.lean @@ -52,7 +52,6 @@ import Mathlib.Tactic.InferParam import Mathlib.Tactic.IntervalCases import Mathlib.Tactic.Inhabit import Mathlib.Tactic.IrreducibleDef -import Mathlib.Tactic.LeftRight import Mathlib.Tactic.LibrarySearch import Mathlib.Tactic.Lift import Mathlib.Tactic.Linarith diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 9a66690bd4259..6a6beae60542b 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -1645,8 +1645,7 @@ theorem measure_eq_measure_preimage_add_measure_tsum_Ico_zpow [MeasurableSpace simp only [mem_singleton_iff, mem_union, mem_Ioo, mem_Ioi, mem_preimage] have H : f x = ∞ ∨ f x < ∞ := eq_or_lt_of_le le_top cases' H with H H - · simp only [H, eq_self_iff_true, or_false_iff, WithTop.zero_lt_top, not_top_lt, - and_false_iff] + · simp only [H, eq_self_iff_true, or_false_iff, zero_lt_top, not_top_lt, and_false_iff] · simp only [H, H.ne, and_true_iff, false_or_iff] · refine disjoint_left.2 fun x hx h'x => ?_ have : f x < ∞ := h'x.2.2 @@ -2101,6 +2100,14 @@ instance instMeasurableInv : MeasurableInv ℝ≥0∞ := ⟨continuous_inv.measurable⟩ #align ennreal.has_measurable_inv ENNReal.instMeasurableInv +instance : MeasurableSMul ℝ≥0 ℝ≥0∞ where + measurable_const_smul := by + simp_rw [ENNReal.smul_def] + exact fun _ ↦ MeasurableSMul.measurable_const_smul _ + measurable_smul_const := fun x ↦ by + simp_rw [ENNReal.smul_def] + exact measurable_coe_nnreal_ennreal.mul_const _ + end ENNReal @[measurability] diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean index 959f9fbaf9731..5875920ed8162 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic -import Mathlib.Topology.MetricSpace.Metrizable +import Mathlib.Topology.Metrizable.Basic #align_import measure_theory.constructions.borel_space.metrizable from "leanprover-community/mathlib"@"bf6a01357ff5684b1ebcd0f1a13be314fc82c0bf" diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean b/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean index aced4bf3ce654..dbcb52ffd3970 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean @@ -24,6 +24,9 @@ In this file we prove Fubini's theorem. Tonelli's theorem (see `MeasureTheory.lintegral_prod`). The lemma `MeasureTheory.Integrable.integral_prod_right` states that the inner integral of the right-hand side is integrable. +* `MeasureTheory.integral_integral_swap_of_hasCompactSupport`: a version of Fubini theorem for + continuous functions with compact support, which does not assume that the measures are σ-finite + contrary to all the usual versions of Fubini. ## Tags @@ -540,4 +543,59 @@ theorem integral_fun_fst (f : α → E) : ∫ z, f z.1 ∂μ.prod ν = (ν univ) rw [← integral_prod_swap] apply integral_fun_snd +section + +variable {X Y : Type*} + [TopologicalSpace X] [TopologicalSpace Y] [MeasurableSpace X] [MeasurableSpace Y] + [OpensMeasurableSpace X] [OpensMeasurableSpace Y] + +/-- A version of *Fubini theorem* for continuous functions with compact support: one may swap +the order of integration with respect to locally finite measures. One does not assume that the +measures are σ-finite, contrary to the usual Fubini theorem. -/ +lemma integral_integral_swap_of_hasCompactSupport + {f : X → Y → E} (hf : Continuous f.uncurry) (h'f : HasCompactSupport f.uncurry) + {μ : Measure X} {ν : Measure Y} [IsFiniteMeasureOnCompacts μ] [IsFiniteMeasureOnCompacts ν] : + ∫ x, (∫ y, f x y ∂ν) ∂μ = ∫ y, (∫ x, f x y ∂μ) ∂ν := by + let U := Prod.fst '' (tsupport f.uncurry) + have : Fact (μ U < ∞) := ⟨(IsCompact.image h'f continuous_fst).measure_lt_top⟩ + let V := Prod.snd '' (tsupport f.uncurry) + have : Fact (ν V < ∞) := ⟨(IsCompact.image h'f continuous_snd).measure_lt_top⟩ + calc + ∫ x, (∫ y, f x y ∂ν) ∂μ = ∫ x, (∫ y in V, f x y ∂ν) ∂μ := by + congr 1 with x + apply (set_integral_eq_integral_of_forall_compl_eq_zero (fun y hy ↦ ?_)).symm + contrapose! hy + have : (x, y) ∈ Function.support f.uncurry := hy + exact mem_image_of_mem _ (subset_tsupport _ this) + _ = ∫ x in U, (∫ y in V, f x y ∂ν) ∂μ := by + apply (set_integral_eq_integral_of_forall_compl_eq_zero (fun x hx ↦ ?_)).symm + have : ∀ y, f x y = 0 := by + intro y + contrapose! hx + have : (x, y) ∈ Function.support f.uncurry := hx + exact mem_image_of_mem _ (subset_tsupport _ this) + simp [this] + _ = ∫ y in V, (∫ x in U, f x y ∂μ) ∂ν := by + apply integral_integral_swap + apply (integrableOn_iff_integrable_of_support_subset (subset_tsupport f.uncurry)).mp + refine ⟨(h'f.stronglyMeasurable_of_prod hf).aestronglyMeasurable, ?_⟩ + obtain ⟨C, hC⟩ : ∃ C, ∀ p, ‖f.uncurry p‖ ≤ C := hf.bounded_above_of_compact_support h'f + exact hasFiniteIntegral_of_bounded (C := C) (eventually_of_forall hC) + _ = ∫ y, (∫ x in U, f x y ∂μ) ∂ν := by + apply set_integral_eq_integral_of_forall_compl_eq_zero (fun y hy ↦ ?_) + have : ∀ x, f x y = 0 := by + intro x + contrapose! hy + have : (x, y) ∈ Function.support f.uncurry := hy + exact mem_image_of_mem _ (subset_tsupport _ this) + simp [this] + _ = ∫ y, (∫ x, f x y ∂μ) ∂ν := by + congr 1 with y + apply set_integral_eq_integral_of_forall_compl_eq_zero (fun x hx ↦ ?_) + contrapose! hx + have : (x, y) ∈ Function.support f.uncurry := hx + exact mem_image_of_mem _ (subset_tsupport _ this) + +end + end MeasureTheory diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index 4044536644405..31b309490ea9e 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -928,7 +928,7 @@ theorem exists_closedBall_covering_tsum_measure_le (μ : Measure α) [SigmaFinit obtain ⟨u, su, u_open, μu⟩ : ∃ U, U ⊇ s ∧ IsOpen U ∧ μ U ≤ μ s + ε / 2 := Set.exists_isOpen_le_add _ _ (by - simpa only [or_false_iff, Ne.def, ENNReal.div_eq_zero_iff, ENNReal.one_ne_top] using hε) + simpa only [or_false, Ne.def, ENNReal.div_eq_zero_iff, ENNReal.two_ne_top] using hε) have : ∀ x ∈ s, ∃ R > 0, ball x R ⊆ u := fun x hx => Metric.mem_nhds_iff.1 (u_open.mem_nhds (su hx)) choose! R hR using this @@ -946,8 +946,8 @@ theorem exists_closedBall_covering_tsum_measure_le (μ : Measure α) [SigmaFinit obtain ⟨v, s'v, v_open, μv⟩ : ∃ v, v ⊇ s' ∧ IsOpen v ∧ μ v ≤ μ s' + ε / 2 / N := Set.exists_isOpen_le_add _ _ (by - simp only [hε, ENNReal.nat_ne_top, WithTop.mul_eq_top_iff, Ne.def, ENNReal.div_eq_zero_iff, - ENNReal.one_ne_top, not_false_iff, and_false_iff, false_and_iff, or_self_iff]) + simp only [ne_eq, ENNReal.div_eq_zero_iff, hε, ENNReal.two_ne_top, or_self, + ENNReal.nat_ne_top, not_false_eq_true]) have : ∀ x ∈ s', ∃ r1 ∈ f x ∩ Ioo (0 : ℝ) 1, closedBall x r1 ⊆ v := by intro x hx rcases Metric.mem_nhds_iff.1 (v_open.mem_nhds (s'v hx)) with ⟨r, rpos, hr⟩ diff --git a/Mathlib/MeasureTheory/Covering/Vitali.lean b/Mathlib/MeasureTheory/Covering/Vitali.lean index ed88b9983ea26..30c8262af514e 100644 --- a/Mathlib/MeasureTheory/Covering/Vitali.lean +++ b/Mathlib/MeasureTheory/Covering/Vitali.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Topology.MetricSpace.Basic import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic import Mathlib.MeasureTheory.Covering.VitaliFamily @@ -143,7 +142,7 @@ theorem exists_disjoint_subfamily_covering_enlargment (B : ι → Set α) (t : S -- Moreover, `δ c` is smaller than the maximum `m` of `δ` over `A`, which is `≤ δ a' / τ` -- thanks to the good choice of `a'`. This is the desired inequality. push_neg at H - simp only [← not_disjoint_iff_nonempty_inter, Classical.not_not] at H + simp only [← disjoint_iff_inter_eq_empty] at H rcases mem_insert_iff.1 ba'u with (rfl | H') · refine' ⟨b, mem_insert _ _, hcb, _⟩ calc diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index 566c66a8c4a97..3278e26efbd76 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -96,6 +96,16 @@ irreducible_def rnDeriv (μ ν : Measure α) : α → ℝ≥0∞ := if h : HaveLebesgueDecomposition μ ν then (Classical.choose h.lebesgue_decomposition).2 else 0 #align measure_theory.measure.rn_deriv MeasureTheory.Measure.rnDeriv +lemma singularPart_of_not_haveLebesgueDecomposition {μ ν : Measure α} + (h : ¬ HaveLebesgueDecomposition μ ν) : + μ.singularPart ν = 0 := by + rw [singularPart]; exact dif_neg h + +lemma rnDeriv_of_not_haveLebesgueDecomposition {μ ν : Measure α} + (h : ¬ HaveLebesgueDecomposition μ ν) : + μ.rnDeriv ν = 0 := by + rw [rnDeriv]; exact dif_neg h + theorem haveLebesgueDecomposition_spec (μ ν : Measure α) [h : HaveLebesgueDecomposition μ ν] : Measurable (μ.rnDeriv ν) ∧ μ.singularPart ν ⟂ₘ ν ∧ μ = μ.singularPart ν + ν.withDensity (μ.rnDeriv ν) := by @@ -120,6 +130,25 @@ instance haveLebesgueDecomposition_smul (μ ν : Measure α) [HaveLebesgueDecomp rfl #align measure_theory.measure.have_lebesgue_decomposition_smul MeasureTheory.Measure.haveLebesgueDecomposition_smul +instance haveLebesgueDecomposition_smul_right (μ ν : Measure α) [HaveLebesgueDecomposition μ ν] + (r : ℝ≥0) : + μ.HaveLebesgueDecomposition (r • ν) where + lebesgue_decomposition := by + obtain ⟨hmeas, hsing, hadd⟩ := haveLebesgueDecomposition_spec μ ν + by_cases hr : r = 0 + · exact ⟨⟨μ, 0⟩, measurable_const, by simp [hr], by simp⟩ + refine ⟨⟨μ.singularPart ν, r⁻¹ • μ.rnDeriv ν⟩, ?_, ?_, ?_⟩ + · change Measurable (r⁻¹ • μ.rnDeriv ν) + exact hmeas.const_smul _ + · refine MutuallySingular.mono_ac hsing AbsolutelyContinuous.rfl ?_ + exact absolutelyContinuous_of_le_smul le_rfl + · have : r⁻¹ • rnDeriv μ ν = ((r⁻¹ : ℝ≥0) : ℝ≥0∞) • rnDeriv μ ν := by simp [ENNReal.smul_def] + rw [this, withDensity_smul _ hmeas, ENNReal.smul_def r, withDensity_smul_measure, + ← smul_assoc, smul_eq_mul, ENNReal.coe_inv hr, ENNReal.inv_mul_cancel, one_smul] + · exact hadd + · simp [hr] + · exact ENNReal.coe_ne_top + @[measurability] theorem measurable_rnDeriv (μ ν : Measure α) : Measurable <| μ.rnDeriv ν := by by_cases h : HaveLebesgueDecomposition μ ν @@ -153,6 +182,32 @@ theorem withDensity_rnDeriv_le (μ ν : Measure α) : ν.withDensity (μ.rnDeriv exact Measure.zero_le μ #align measure_theory.measure.with_density_rn_deriv_le MeasureTheory.Measure.withDensity_rnDeriv_le +@[simp] +lemma withDensity_rnDeriv_eq_zero (μ ν : Measure α) [μ.HaveLebesgueDecomposition ν] : + ν.withDensity (μ.rnDeriv ν) = 0 ↔ μ ⟂ₘ ν := by + have h_dec := haveLebesgueDecomposition_add μ ν + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rw [h, add_zero] at h_dec + rw [h_dec] + exact mutuallySingular_singularPart μ ν + · rw [← MutuallySingular.self_iff] + rw [h_dec, MutuallySingular.add_left_iff] at h + refine MutuallySingular.mono_ac h.2 AbsolutelyContinuous.rfl ?_ + exact withDensity_absolutelyContinuous _ _ + +@[simp] +lemma rnDeriv_eq_zero (μ ν : Measure α) [μ.HaveLebesgueDecomposition ν] : + μ.rnDeriv ν =ᵐ[ν] 0 ↔ μ ⟂ₘ ν := by + rw [← withDensity_rnDeriv_eq_zero, + withDensity_eq_zero_iff (measurable_rnDeriv _ _).aemeasurable] + +lemma MutuallySingular.rnDeriv_ae_eq_zero {μ ν : Measure α} (hμν : μ ⟂ₘ ν) : + μ.rnDeriv ν =ᵐ[ν] 0 := by + by_cases h : μ.HaveLebesgueDecomposition ν + · rw [rnDeriv_eq_zero] + exact hμν + · rw [rnDeriv_of_not_haveLebesgueDecomposition h] + instance singularPart.instIsFiniteMeasure [IsFiniteMeasure μ] : IsFiniteMeasure (μ.singularPart ν) := isFiniteMeasure_of_le μ <| singularPart_le μ ν @@ -183,7 +238,7 @@ instance withDensity.instIsLocallyFiniteMeasure [TopologicalSpace α] [IsLocally #align measure_theory.measure.with_density.measure_theory.is_locally_finite_measure MeasureTheory.Measure.withDensity.instIsLocallyFiniteMeasure theorem lintegral_rnDeriv_lt_top_of_measure_ne_top {μ : Measure α} (ν : Measure α) {s : Set α} - (hs : μ s ≠ ∞) : (∫⁻ x in s, μ.rnDeriv ν x ∂ν) < ∞ := by + (hs : μ s ≠ ∞) : ∫⁻ x in s, μ.rnDeriv ν x ∂ν < ∞ := by by_cases hl : HaveLebesgueDecomposition μ ν · haveI := hl obtain ⟨-, -, hadd⟩ := haveLebesgueDecomposition_spec μ ν @@ -207,6 +262,11 @@ theorem lintegral_rnDeriv_lt_top (μ ν : Measure α) [IsFiniteMeasure μ] : exact lintegral_rnDeriv_lt_top_of_measure_ne_top _ (measure_lt_top _ _).ne #align measure_theory.measure.lintegral_rn_deriv_lt_top MeasureTheory.Measure.lintegral_rnDeriv_lt_top +lemma integrable_toReal_rnDeriv {μ ν : Measure α} [IsFiniteMeasure μ] : + Integrable (fun x ↦ (μ.rnDeriv ν x).toReal) ν := + integrable_toReal_of_lintegral_ne_top (Measure.measurable_rnDeriv _ _).aemeasurable + (Measure.lintegral_rnDeriv_lt_top _ _).ne + /-- The Radon-Nikodym derivative of a sigma-finite measure `μ` with respect to another measure `ν` is `ν`-almost everywhere finite. -/ theorem rnDeriv_lt_top (μ ν : Measure α) [SigmaFinite μ] : ∀ᵐ x ∂ν, μ.rnDeriv ν x < ∞ := by @@ -219,6 +279,9 @@ theorem rnDeriv_lt_top (μ ν : Measure α) [SigmaFinite μ] : ∀ᵐ x ∂ν, exact (measure_spanningSets_lt_top _ _).ne #align measure_theory.measure.rn_deriv_lt_top MeasureTheory.Measure.rnDeriv_lt_top +lemma rnDeriv_ne_top (μ ν : Measure α) [SigmaFinite μ] : ∀ᵐ x ∂ν, μ.rnDeriv ν x ≠ ∞ := by + filter_upwards [Measure.rnDeriv_lt_top μ ν] with x hx using hx.ne + /-- Given measures `μ` and `ν`, if `s` is a measure mutually singular to `ν` and `f` is a measurable function such that `μ = s + fν`, then `s = μ.singularPart μ`. @@ -270,8 +333,7 @@ theorem singularPart_smul (μ ν : Measure α) (r : ℝ≥0) : by_cases hr : r = 0 · rw [hr, zero_smul, zero_smul, singularPart_zero] by_cases hl : HaveLebesgueDecomposition μ ν - · haveI := hl - refine' + · refine' (eq_singularPart ((measurable_rnDeriv μ ν).const_smul (r : ℝ≥0∞)) (MutuallySingular.smul r (haveLebesgueDecomposition_spec _ _).2.1) _).symm rw [withDensity_smul _ (measurable_rnDeriv _ _), ← smul_add, @@ -282,6 +344,23 @@ theorem singularPart_smul (μ ν : Measure α) (r : ℝ≥0) : exact @Measure.haveLebesgueDecomposition_smul _ _ _ _ hl' _ #align measure_theory.measure.singular_part_smul MeasureTheory.Measure.singularPart_smul +theorem singularPart_smul_right (μ ν : Measure α) (r : ℝ≥0) (hr : r ≠ 0) : + μ.singularPart (r • ν) = μ.singularPart ν := by + by_cases hl : HaveLebesgueDecomposition μ ν + · refine (eq_singularPart ((measurable_rnDeriv μ ν).const_smul r⁻¹) ?_ ?_).symm + · refine (mutuallySingular_singularPart μ ν).mono_ac AbsolutelyContinuous.rfl ?_ + exact absolutelyContinuous_of_le_smul le_rfl + · rw [ENNReal.smul_def r, withDensity_smul_measure, ← withDensity_smul] + swap; · exact (measurable_rnDeriv _ _).const_smul _ + convert haveLebesgueDecomposition_add μ ν + ext x + simp only [Pi.smul_apply, ne_eq] + rw [← ENNReal.smul_def, ← smul_assoc, smul_eq_mul, mul_inv_cancel hr, one_smul] + · rw [singularPart, singularPart, dif_neg hl, dif_neg] + refine fun hl' => hl ?_ + rw [← inv_smul_smul₀ hr ν] + infer_instance + theorem singularPart_add (μ₁ μ₂ ν : Measure α) [HaveLebesgueDecomposition μ₁ ν] [HaveLebesgueDecomposition μ₂ ν] : (μ₁ + μ₂).singularPart ν = μ₁.singularPart ν + μ₂.singularPart ν := by @@ -348,6 +427,22 @@ theorem eq_withDensity_rnDeriv {s : Measure α} {f : α → ℝ≥0∞} (hf : Me restrict_apply hA, ← diff_eq, measure_inter_add_diff _ (hS₁.inter hT₁)] #align measure_theory.measure.eq_with_density_rn_deriv MeasureTheory.Measure.eq_withDensity_rnDeriv +theorem eq_withDensity_rnDeriv₀ {μ ν : Measure α} {s : Measure α} {f : α → ℝ≥0∞} + (hf : AEMeasurable f ν) (hs : s ⟂ₘ ν) (hadd : μ = s + ν.withDensity f) : + ν.withDensity f = ν.withDensity (μ.rnDeriv ν) := by + rw [withDensity_congr_ae hf.ae_eq_mk] at hadd ⊢ + exact eq_withDensity_rnDeriv hf.measurable_mk hs hadd + +theorem eq_rnDeriv₀ {μ ν : Measure α} [SigmaFinite ν] {s : Measure α} {f : α → ℝ≥0∞} + (hf : AEMeasurable f ν) (hs : s ⟂ₘ ν) (hadd : μ = s + ν.withDensity f) : + f =ᵐ[ν] μ.rnDeriv ν := by + refine' ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite₀ hf + (measurable_rnDeriv μ ν).aemeasurable _ + intro a ha _ + calc ∫⁻ x : α in a, f x ∂ν = ν.withDensity f a := (withDensity_apply f ha).symm + _ = ν.withDensity (μ.rnDeriv ν) a := by rw [eq_withDensity_rnDeriv₀ hf hs hadd] + _ = ∫⁻ x : α in a, μ.rnDeriv ν x ∂ν := withDensity_apply _ ha + /-- Given measures `μ` and `ν`, if `s` is a measure mutually singular to `ν` and `f` is a measurable function such that `μ = s + fν`, then `f = μ.rnDeriv ν`. @@ -356,13 +451,8 @@ theorem, while `MeasureTheory.Measure.eq_singularPart` provides the uniqueness o `singularPart`. Here, the uniqueness is given in terms of the functions, while the uniqueness in terms of the functions is given in `eq_withDensity_rnDeriv`. -/ theorem eq_rnDeriv [SigmaFinite ν] {s : Measure α} {f : α → ℝ≥0∞} (hf : Measurable f) (hs : s ⟂ₘ ν) - (hadd : μ = s + ν.withDensity f) : f =ᵐ[ν] μ.rnDeriv ν := by - refine' ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite hf (measurable_rnDeriv μ ν) _ - intro a ha _ - calc - ∫⁻ x : α in a, f x ∂ν = ν.withDensity f a := (withDensity_apply f ha).symm - _ = ν.withDensity (μ.rnDeriv ν) a := by rw [eq_withDensity_rnDeriv hf hs hadd] - _ = ∫⁻ x : α in a, μ.rnDeriv ν x ∂ν := withDensity_apply _ ha + (hadd : μ = s + ν.withDensity f) : f =ᵐ[ν] μ.rnDeriv ν := + eq_rnDeriv₀ hf.aemeasurable hs hadd #align measure_theory.measure.eq_rn_deriv MeasureTheory.Measure.eq_rnDeriv /-- The Radon-Nikodym derivative of `f ν` with respect to `ν` is `f`. -/ @@ -380,6 +470,93 @@ theorem rnDeriv_restrict (ν : Measure α) [SigmaFinite ν] {s : Set α} (hs : M exact rnDeriv_withDensity _ (measurable_one.indicator hs) #align measure_theory.measure.rn_deriv_restrict MeasureTheory.Measure.rnDeriv_restrict +/-- Radon-Nikodym derivative of the scalar multiple of a measure. +See also `rnDeriv_smul_left'`, which requires sigma-finite `ν` and `μ`. -/ +theorem rnDeriv_smul_left (ν μ : Measure α) [IsFiniteMeasure ν] + [ν.HaveLebesgueDecomposition μ] (r : ℝ≥0) : + (r • ν).rnDeriv μ =ᵐ[μ] r • ν.rnDeriv μ := by + rw [← withDensity_eq_iff] + · simp_rw [ENNReal.smul_def] + rw [withDensity_smul _ (measurable_rnDeriv _ _)] + suffices (r • ν).singularPart μ + withDensity μ (rnDeriv (r • ν) μ) + = (r • ν).singularPart μ + r • withDensity μ (rnDeriv ν μ) by + rwa [Measure.add_right_inj] at this + rw [← (r • ν).haveLebesgueDecomposition_add μ, singularPart_smul, ← smul_add, + ← ν.haveLebesgueDecomposition_add μ] + · exact (measurable_rnDeriv _ _).aemeasurable + · exact (measurable_rnDeriv _ _).aemeasurable.const_smul _ + · exact (lintegral_rnDeriv_lt_top (r • ν) μ).ne + +/-- Radon-Nikodym derivative of the scalar multiple of a measure. +See also `rnDeriv_smul_left_of_ne_top'`, which requires sigma-finite `ν` and `μ`. -/ +theorem rnDeriv_smul_left_of_ne_top (ν μ : Measure α) [IsFiniteMeasure ν] + [ν.HaveLebesgueDecomposition μ] {r : ℝ≥0∞} (hr : r ≠ ∞) : + (r • ν).rnDeriv μ =ᵐ[μ] r • ν.rnDeriv μ := by + have h : (r.toNNReal • ν).rnDeriv μ =ᵐ[μ] r.toNNReal • ν.rnDeriv μ := + rnDeriv_smul_left ν μ r.toNNReal + simpa [ENNReal.smul_def, ENNReal.coe_toNNReal hr] using h + +/-- Radon-Nikodym derivative with respect to the scalar multiple of a measure. +See also `rnDeriv_smul_right'`, which requires sigma-finite `ν` and `μ`. -/ +theorem rnDeriv_smul_right (ν μ : Measure α) [IsFiniteMeasure ν] + [ν.HaveLebesgueDecomposition μ] {r : ℝ≥0} (hr : r ≠ 0) : + ν.rnDeriv (r • μ) =ᵐ[μ] r⁻¹ • ν.rnDeriv μ := by + suffices ν.rnDeriv (r • μ) =ᵐ[r • μ] r⁻¹ • ν.rnDeriv μ by + suffices hμ : μ ≪ r • μ by exact hμ.ae_le this + refine absolutelyContinuous_of_le_smul (c := r⁻¹) ?_ + rw [← ENNReal.coe_inv hr, ← ENNReal.smul_def, ← smul_assoc, smul_eq_mul, + inv_mul_cancel hr, one_smul] + rw [← withDensity_eq_iff] + rotate_left + · exact (measurable_rnDeriv _ _).aemeasurable + · exact (measurable_rnDeriv _ _).aemeasurable.const_smul _ + · exact (lintegral_rnDeriv_lt_top ν _).ne + · simp_rw [ENNReal.smul_def] + rw [withDensity_smul _ (measurable_rnDeriv _ _)] + suffices ν.singularPart (r • μ) + withDensity (r • μ) (rnDeriv ν (r • μ)) + = ν.singularPart (r • μ) + r⁻¹ • withDensity (r • μ) (rnDeriv ν μ) by + rwa [add_right_inj] at this + rw [← ν.haveLebesgueDecomposition_add (r • μ), singularPart_smul_right _ _ _ hr, + ENNReal.smul_def r, withDensity_smul_measure, ← ENNReal.smul_def, ← smul_assoc, + smul_eq_mul, inv_mul_cancel hr, one_smul] + exact ν.haveLebesgueDecomposition_add μ + +/-- Radon-Nikodym derivative with respect to the scalar multiple of a measure. +See also `rnDeriv_smul_right_of_ne_top'`, which requires sigma-finite `ν` and `μ`. -/ +theorem rnDeriv_smul_right_of_ne_top (ν μ : Measure α) [IsFiniteMeasure ν] + [ν.HaveLebesgueDecomposition μ] {r : ℝ≥0∞} (hr : r ≠ 0) (hr_ne_top : r ≠ ∞) : + ν.rnDeriv (r • μ) =ᵐ[μ] r⁻¹ • ν.rnDeriv μ := by + have h : ν.rnDeriv (r.toNNReal • μ) =ᵐ[μ] r.toNNReal⁻¹ • ν.rnDeriv μ := by + refine rnDeriv_smul_right ν μ ?_ + rw [ne_eq, ENNReal.toNNReal_eq_zero_iff] + simp [hr, hr_ne_top] + have : (r.toNNReal)⁻¹ • rnDeriv ν μ = r⁻¹ • rnDeriv ν μ := by + ext x + simp only [Pi.smul_apply, ENNReal.smul_def, ne_eq, smul_eq_mul] + rw [ENNReal.coe_inv, ENNReal.coe_toNNReal hr_ne_top] + rw [ne_eq, ENNReal.toNNReal_eq_zero_iff] + simp [hr, hr_ne_top] + simp_rw [this, ENNReal.smul_def, ENNReal.coe_toNNReal hr_ne_top] at h + exact h + +/-- Radon-Nikodym derivative of a sum of two measures. +See also `rnDeriv_add'`, which requires sigma-finite `ν₁`, `ν₂` and `μ`. -/ +lemma rnDeriv_add (ν₁ ν₂ μ : Measure α) [IsFiniteMeasure ν₁] [IsFiniteMeasure ν₂] + [ν₁.HaveLebesgueDecomposition μ] [ν₂.HaveLebesgueDecomposition μ] + [(ν₁ + ν₂).HaveLebesgueDecomposition μ] : + (ν₁ + ν₂).rnDeriv μ =ᵐ[μ] ν₁.rnDeriv μ + ν₂.rnDeriv μ := by + rw [← withDensity_eq_iff] + · suffices (ν₁ + ν₂).singularPart μ + μ.withDensity ((ν₁ + ν₂).rnDeriv μ) + = (ν₁ + ν₂).singularPart μ + μ.withDensity (ν₁.rnDeriv μ + ν₂.rnDeriv μ) by + rwa [add_right_inj] at this + rw [← (ν₁ + ν₂).haveLebesgueDecomposition_add μ, singularPart_add, + withDensity_add_left (measurable_rnDeriv _ _), add_assoc, + add_comm (ν₂.singularPart μ), add_assoc, add_comm _ (ν₂.singularPart μ), + ← ν₂.haveLebesgueDecomposition_add μ, ← add_assoc, ← ν₁.haveLebesgueDecomposition_add μ] + · exact (measurable_rnDeriv _ _).aemeasurable + · exact ((measurable_rnDeriv _ _).add (measurable_rnDeriv _ _)).aemeasurable + · exact (lintegral_rnDeriv_lt_top (ν₁ + ν₂) μ).ne + open VectorMeasure SignedMeasure /-- If two finite measures `μ` and `ν` are not mutually singular, there exists some `ε > 0` and @@ -763,6 +940,90 @@ instance (priority := 100) haveLebesgueDecomposition_of_sigmaFinite (μ ν : Mea · exact fun n => Measurable.indicator (measurable_rnDeriv _ _) (S.set_mem n)⟩ #align measure_theory.measure.have_lebesgue_decomposition_of_sigma_finite MeasureTheory.Measure.haveLebesgueDecomposition_of_sigmaFinite +section rnDeriv + +/-- Radon-Nikodym derivative of the scalar multiple of a measure. +See also `rnDeriv_smul_left`, which has no hypothesis on `μ` but requires finite `ν`. -/ +theorem rnDeriv_smul_left' (ν μ : Measure α) [SigmaFinite ν] [SigmaFinite μ] (r : ℝ≥0) : + (r • ν).rnDeriv μ =ᵐ[μ] r • ν.rnDeriv μ := by + rw [← withDensity_eq_iff_of_sigmaFinite] + · simp_rw [ENNReal.smul_def] + rw [withDensity_smul _ (measurable_rnDeriv _ _)] + suffices (r • ν).singularPart μ + withDensity μ (rnDeriv (r • ν) μ) + = (r • ν).singularPart μ + r • withDensity μ (rnDeriv ν μ) by + rwa [Measure.add_right_inj] at this + rw [← (r • ν).haveLebesgueDecomposition_add μ, singularPart_smul, ← smul_add, + ← ν.haveLebesgueDecomposition_add μ] + · exact (measurable_rnDeriv _ _).aemeasurable + · exact (measurable_rnDeriv _ _).aemeasurable.const_smul _ + +/-- Radon-Nikodym derivative of the scalar multiple of a measure. +See also `rnDeriv_smul_left_of_ne_top`, which has no hypothesis on `μ` but requires finite `ν`. -/ +theorem rnDeriv_smul_left_of_ne_top' (ν μ : Measure α) [SigmaFinite ν] [SigmaFinite μ] + {r : ℝ≥0∞} (hr : r ≠ ∞) : + (r • ν).rnDeriv μ =ᵐ[μ] r • ν.rnDeriv μ := by + have h : (r.toNNReal • ν).rnDeriv μ =ᵐ[μ] r.toNNReal • ν.rnDeriv μ := + rnDeriv_smul_left' ν μ r.toNNReal + simpa [ENNReal.smul_def, ENNReal.coe_toNNReal hr] using h + +/-- Radon-Nikodym derivative with respect to the scalar multiple of a measure. +See also `rnDeriv_smul_right`, which has no hypothesis on `μ` but requires finite `ν`. -/ +theorem rnDeriv_smul_right' (ν μ : Measure α) [SigmaFinite ν] [SigmaFinite μ] + {r : ℝ≥0} (hr : r ≠ 0) : + ν.rnDeriv (r • μ) =ᵐ[μ] r⁻¹ • ν.rnDeriv μ := by + suffices ν.rnDeriv (r • μ) =ᵐ[r • μ] r⁻¹ • ν.rnDeriv μ by + suffices hμ : μ ≪ r • μ by exact hμ.ae_le this + refine absolutelyContinuous_of_le_smul (c := r⁻¹) ?_ + rw [← ENNReal.coe_inv hr, ← ENNReal.smul_def, ← smul_assoc, smul_eq_mul, + inv_mul_cancel hr, one_smul] + rw [← withDensity_eq_iff_of_sigmaFinite] + · simp_rw [ENNReal.smul_def] + rw [withDensity_smul _ (measurable_rnDeriv _ _)] + suffices ν.singularPart (r • μ) + withDensity (r • μ) (rnDeriv ν (r • μ)) + = ν.singularPart (r • μ) + r⁻¹ • withDensity (r • μ) (rnDeriv ν μ) by + rwa [add_right_inj] at this + rw [← ν.haveLebesgueDecomposition_add (r • μ), singularPart_smul_right _ _ _ hr, + ENNReal.smul_def r, withDensity_smul_measure, ← ENNReal.smul_def, ← smul_assoc, + smul_eq_mul, inv_mul_cancel hr, one_smul] + exact ν.haveLebesgueDecomposition_add μ + · exact (measurable_rnDeriv _ _).aemeasurable + · exact (measurable_rnDeriv _ _).aemeasurable.const_smul _ + +/-- Radon-Nikodym derivative with respect to the scalar multiple of a measure. +See also `rnDeriv_smul_right_of_ne_top`, which has no hypothesis on `μ` but requires finite `ν`. -/ +theorem rnDeriv_smul_right_of_ne_top' (ν μ : Measure α) [SigmaFinite ν] [SigmaFinite μ] + {r : ℝ≥0∞} (hr : r ≠ 0) (hr_ne_top : r ≠ ∞) : + ν.rnDeriv (r • μ) =ᵐ[μ] r⁻¹ • ν.rnDeriv μ := by + have h : ν.rnDeriv (r.toNNReal • μ) =ᵐ[μ] r.toNNReal⁻¹ • ν.rnDeriv μ := by + refine rnDeriv_smul_right' ν μ ?_ + rw [ne_eq, ENNReal.toNNReal_eq_zero_iff] + simp [hr, hr_ne_top] + have : (r.toNNReal)⁻¹ • rnDeriv ν μ = r⁻¹ • rnDeriv ν μ := by + ext x + simp only [Pi.smul_apply, ENNReal.smul_def, ne_eq, smul_eq_mul] + rw [ENNReal.coe_inv, ENNReal.coe_toNNReal hr_ne_top] + rw [ne_eq, ENNReal.toNNReal_eq_zero_iff] + simp [hr, hr_ne_top] + simp_rw [this, ENNReal.smul_def, ENNReal.coe_toNNReal hr_ne_top] at h + exact h + +/-- Radon-Nikodym derivative of a sum of two measures. +See also `rnDeriv_add`, which has no hypothesis on `μ` but requires finite `ν₁` and `ν₂`. -/ +lemma rnDeriv_add' (ν₁ ν₂ μ : Measure α) [SigmaFinite ν₁] [SigmaFinite ν₂] [SigmaFinite μ] : + (ν₁ + ν₂).rnDeriv μ =ᵐ[μ] ν₁.rnDeriv μ + ν₂.rnDeriv μ := by + rw [← withDensity_eq_iff_of_sigmaFinite] + · suffices (ν₁ + ν₂).singularPart μ + μ.withDensity ((ν₁ + ν₂).rnDeriv μ) + = (ν₁ + ν₂).singularPart μ + μ.withDensity (ν₁.rnDeriv μ + ν₂.rnDeriv μ) by + rwa [add_right_inj] at this + rw [← (ν₁ + ν₂).haveLebesgueDecomposition_add μ, singularPart_add, + withDensity_add_left (measurable_rnDeriv _ _), add_assoc, + add_comm (ν₂.singularPart μ), add_assoc, add_comm _ (ν₂.singularPart μ), + ← ν₂.haveLebesgueDecomposition_add μ, ← add_assoc, ← ν₁.haveLebesgueDecomposition_add μ] + · exact (measurable_rnDeriv _ _).aemeasurable + · exact ((measurable_rnDeriv _ _).add (measurable_rnDeriv _ _)).aemeasurable + +end rnDeriv + end Measure namespace SignedMeasure diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index 5e1db4a11a807..b52a8bac954b0 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -59,16 +59,18 @@ theorem withDensity_rnDeriv_eq (μ ν : Measure α) [HaveLebesgueDecomposition exact hadd.symm #align measure_theory.measure.with_density_rn_deriv_eq MeasureTheory.Measure.withDensity_rnDeriv_eq +variable {μ ν : Measure α} + /-- **The Radon-Nikodym theorem**: Given two measures `μ` and `ν`, if `HaveLebesgueDecomposition μ ν`, then `μ` is absolutely continuous to `ν` if and only if `ν.withDensity (rnDeriv μ ν) = μ`. -/ -theorem absolutelyContinuous_iff_withDensity_rnDeriv_eq {μ ν : Measure α} +theorem absolutelyContinuous_iff_withDensity_rnDeriv_eq [HaveLebesgueDecomposition μ ν] : μ ≪ ν ↔ ν.withDensity (rnDeriv μ ν) = μ := ⟨withDensity_rnDeriv_eq μ ν, fun h => h ▸ withDensity_absolutelyContinuous _ _⟩ #align measure_theory.measure.absolutely_continuous_iff_with_density_rn_deriv_eq MeasureTheory.Measure.absolutelyContinuous_iff_withDensity_rnDeriv_eq -theorem withDensity_rnDeriv_toReal_eq {μ ν : Measure α} [IsFiniteMeasure μ] - [HaveLebesgueDecomposition μ ν] (h : μ ≪ ν) {i : Set α} (hi : MeasurableSet i) : +theorem withDensity_rnDeriv_toReal_eq [IsFiniteMeasure μ] [HaveLebesgueDecomposition μ ν] + (h : μ ≪ ν) {i : Set α} (hi : MeasurableSet i) : (∫ x in i, (μ.rnDeriv ν x).toReal ∂ν) = (μ i).toReal := by rw [integral_toReal, ← withDensity_apply _ hi, withDensity_rnDeriv_eq μ ν h] · measurability @@ -78,6 +80,47 @@ theorem withDensity_rnDeriv_toReal_eq {μ ν : Measure α} [IsFiniteMeasure μ] exact measure_lt_top _ _ #align measure_theory.measure.with_density_rn_deriv_to_real_eq MeasureTheory.Measure.withDensity_rnDeriv_toReal_eq +lemma rnDeriv_pos [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ ν) : + ∀ᵐ x ∂μ, 0 < μ.rnDeriv ν x := by + rw [← Measure.withDensity_rnDeriv_eq _ _ hμν, + ae_withDensity_iff (Measure.measurable_rnDeriv _ _), Measure.withDensity_rnDeriv_eq _ _ hμν] + exact ae_of_all _ (fun x hx ↦ lt_of_le_of_ne (zero_le _) hx.symm) + +lemma inv_rnDeriv [SigmaFinite μ] [SigmaFinite ν] (hμν : μ ≪ ν) (hνμ : ν ≪ μ) : + (μ.rnDeriv ν)⁻¹ =ᵐ[μ] ν.rnDeriv μ := by + suffices μ.withDensity (μ.rnDeriv ν)⁻¹ = μ.withDensity (ν.rnDeriv μ) by + calc (μ.rnDeriv ν)⁻¹ =ᵐ[μ] (μ.withDensity (μ.rnDeriv ν)⁻¹).rnDeriv μ := + (rnDeriv_withDensity _ (measurable_rnDeriv _ _).inv).symm + _ = (μ.withDensity (ν.rnDeriv μ)).rnDeriv μ := by rw [this] + _ =ᵐ[μ] ν.rnDeriv μ := rnDeriv_withDensity _ (measurable_rnDeriv _ _) + rw [withDensity_rnDeriv_eq _ _ hνμ, ← withDensity_rnDeriv_eq _ _ hμν] + conv in ((ν.withDensity (μ.rnDeriv ν)).rnDeriv ν)⁻¹ => rw [withDensity_rnDeriv_eq _ _ hμν] + change (ν.withDensity (μ.rnDeriv ν)).withDensity (fun x ↦ (μ.rnDeriv ν x)⁻¹) = ν + rw [withDensity_inv_same (measurable_rnDeriv _ _) + (by filter_upwards [hνμ.ae_le (rnDeriv_pos hμν)] with x hx using hx.ne') + (rnDeriv_ne_top _ _)] + +lemma set_lintegral_rnDeriv [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ ν) {s : Set α} + (hs : MeasurableSet s) : + ∫⁻ x in s, μ.rnDeriv ν x ∂ν = μ s := by + conv_rhs => rw [← Measure.withDensity_rnDeriv_eq _ _ hμν, withDensity_apply _ hs] + +lemma lintegral_rnDeriv [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ ν) : + ∫⁻ x, μ.rnDeriv ν x ∂ν = μ Set.univ := by + rw [← set_lintegral_univ, set_lintegral_rnDeriv hμν MeasurableSet.univ] + +lemma set_integral_toReal_rnDeriv {μ ν : Measure α} [SigmaFinite μ] [SigmaFinite ν] + (hμν : μ ≪ ν) {s : Set α} (hs : MeasurableSet s) : + ∫ x in s, (μ.rnDeriv ν x).toReal ∂ν = (μ s).toReal := by + rw [integral_toReal (Measure.measurable_rnDeriv _ _).aemeasurable] + · rw [ENNReal.toReal_eq_toReal_iff, set_lintegral_rnDeriv hμν hs] + simp + · exact ae_restrict_of_ae (Measure.rnDeriv_lt_top _ _) + +lemma integral_toReal_rnDeriv {μ ν : Measure α} [SigmaFinite μ] [SigmaFinite ν] (hμν : μ ≪ ν) : + ∫ x, (μ.rnDeriv ν x).toReal ∂ν = (μ Set.univ).toReal := by + rw [← integral_univ, set_integral_toReal_rnDeriv hμν MeasurableSet.univ] + end Measure namespace SignedMeasure diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index e7f5c68466e33..6e4dd29317b30 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -222,14 +222,33 @@ theorem ae_le_of_forall_set_lintegral_le_of_sigmaFinite [SigmaFinite μ] {f g : _ = 0 := by simp only [μs, tsum_zero] #align measure_theory.ae_le_of_forall_set_lintegral_le_of_sigma_finite MeasureTheory.ae_le_of_forall_set_lintegral_le_of_sigmaFinite -theorem ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} - (hf : Measurable f) (hg : Measurable g) +theorem ae_le_of_forall_set_lintegral_le_of_sigmaFinite₀ [SigmaFinite μ] + {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) + (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ ≤ ∫⁻ x in s, g x ∂μ) : + f ≤ᵐ[μ] g := by + have h' : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, hf.mk f x ∂μ ≤ ∫⁻ x in s, hg.mk g x ∂μ := by + refine fun s hs hμs ↦ (set_lintegral_congr_fun hs ?_).trans_le + ((h s hs hμs).trans_eq (set_lintegral_congr_fun hs ?_)) + · filter_upwards [hf.ae_eq_mk] with a ha using fun _ ↦ ha.symm + · filter_upwards [hg.ae_eq_mk] with a ha using fun _ ↦ ha + filter_upwards [hf.ae_eq_mk, hg.ae_eq_mk, + ae_le_of_forall_set_lintegral_le_of_sigmaFinite hf.measurable_mk hg.measurable_mk h'] + with a haf hag ha + rwa [haf, hag] + +theorem ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite₀ [SigmaFinite μ] + {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : f =ᵐ[μ] g := by have A : f ≤ᵐ[μ] g := - ae_le_of_forall_set_lintegral_le_of_sigmaFinite hf hg fun s hs h's => le_of_eq (h s hs h's) + ae_le_of_forall_set_lintegral_le_of_sigmaFinite₀ hf hg fun s hs h's => le_of_eq (h s hs h's) have B : g ≤ᵐ[μ] f := - ae_le_of_forall_set_lintegral_le_of_sigmaFinite hg hf fun s hs h's => ge_of_eq (h s hs h's) + ae_le_of_forall_set_lintegral_le_of_sigmaFinite₀ hg hf fun s hs h's => ge_of_eq (h s hs h's) filter_upwards [A, B] with x using le_antisymm + +theorem ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} + (hf : Measurable f) (hg : Measurable g) + (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : f =ᵐ[μ] g := + ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite₀ hf.aemeasurable hg.aemeasurable h #align measure_theory.ae_eq_of_forall_set_lintegral_eq_of_sigma_finite MeasureTheory.ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite end ENNReal @@ -604,7 +623,7 @@ end AeEqOfForallSetIntegralEq section Lintegral theorem AEMeasurable.ae_eq_of_forall_set_lintegral_eq {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) - (hg : AEMeasurable g μ) (hfi : (∫⁻ x, f x ∂μ) ≠ ∞) (hgi : (∫⁻ x, g x ∂μ) ≠ ∞) + (hg : AEMeasurable g μ) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) (hgi : ∫⁻ x, g x ∂μ ≠ ∞) (hfg : ∀ ⦃s⦄, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : f =ᵐ[μ] g := by refine' @@ -629,4 +648,25 @@ theorem AEMeasurable.ae_eq_of_forall_set_lintegral_eq {f g : α → ℝ≥0∞} end Lintegral +section WithDensity + +variable {m : MeasurableSpace α} {μ : Measure α} + +theorem withDensity_eq_iff_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) + (hg : AEMeasurable g μ) : μ.withDensity f = μ.withDensity g ↔ f =ᵐ[μ] g := + ⟨fun hfg ↦ by + refine ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite₀ hf hg fun s hs _ ↦ ?_ + rw [← withDensity_apply f hs, ← withDensity_apply g hs, ← hfg], withDensity_congr_ae⟩ + +theorem withDensity_eq_iff {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) + (hg : AEMeasurable g μ) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) : + μ.withDensity f = μ.withDensity g ↔ f =ᵐ[μ] g := + ⟨fun hfg ↦ by + refine AEMeasurable.ae_eq_of_forall_set_lintegral_eq hf hg hfi ?_ fun s hs _ ↦ ?_ + · rwa [← set_lintegral_univ, ← withDensity_apply g MeasurableSet.univ, ← hfg, + withDensity_apply f MeasurableSet.univ, set_lintegral_univ] + · rw [← withDensity_apply f hs, ← withDensity_apply g hs, ← hfg], withDensity_congr_ae⟩ + +end WithDensity + end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean index cf9189f17055a..fe2e0cf08bec2 100644 --- a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean +++ b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean @@ -211,7 +211,7 @@ theorem TendstoInMeasure.exists_seq_tendsto_ae (hfg : TendstoInMeasure μ f atTo set s := Filter.atTop.limsup S with hs have hμs : μ s = 0 := by refine' measure_limsup_eq_zero (ne_of_lt <| lt_of_le_of_lt (ENNReal.tsum_le_tsum hμS_le) _) - simp only [ENNReal.tsum_geometric, ENNReal.one_sub_inv_two, inv_inv] + simp only [ENNReal.tsum_geometric, ENNReal.one_sub_inv_two, ENNReal.two_lt_top, inv_inv] have h_tendsto : ∀ x ∈ sᶜ, Tendsto (fun i => f (ns i) x) atTop (𝓝 (g x)) := by refine' fun x hx => Metric.tendsto_atTop.mpr fun ε hε => _ rw [hs, limsup_eq_iInf_iSup_of_nat] at hx diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index bc29e25cd7181..d2f3721665fee 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -243,7 +243,8 @@ theorem exists_closed_cover_approximatesLinearOn_of_hasFDerivWithinAt [SecondCou obtain ⟨q, hq⟩ : ∃ q, F q = (n, z, p) := hF _ -- then `x` belongs to `t q`. apply mem_iUnion.2 ⟨q, _⟩ - simp (config := { zeta := false }) only [hq, subset_closure hnz, hp, mem_inter_iff, and_true, hnz] + simp (config := { zeta := false }) only [K, hq, mem_inter_iff, hp, and_true] + exact subset_closure hnz #align exists_closed_cover_approximates_linear_on_of_has_fderiv_within_at exists_closed_cover_approximatesLinearOn_of_hasFDerivWithinAt variable [MeasurableSpace E] [BorelSpace E] (μ : Measure E) [IsAddHaarMeasure μ] @@ -1263,4 +1264,72 @@ theorem integral_target_eq_integral_abs_det_fderiv_smul {f : LocalHomeomorph E E exact (hf' x hx).hasFDerivWithinAt #align measure_theory.integral_target_eq_integral_abs_det_fderiv_smul MeasureTheory.integral_target_eq_integral_abs_det_fderiv_smul +section withDensity + +lemma _root_.MeasurableEmbedding.withDensity_ofReal_comap_apply_eq_integral_abs_det_fderiv_mul + (hs : MeasurableSet s) (hf : MeasurableEmbedding f) + {g : E → ℝ} (hg : ∀ᵐ x ∂μ, x ∈ f '' s → 0 ≤ g x) (hg_int : IntegrableOn g (f '' s) μ) + (hf' : ∀ x ∈ s, HasFDerivWithinAt f (f' x) s x) : + (μ.withDensity (fun x ↦ ENNReal.ofReal (g x))).comap f s + = ENNReal.ofReal (∫ x in s, |(f' x).det| * g (f x) ∂μ) := by + rw [Measure.comap_apply f hf.injective (fun t ht ↦ hf.measurableSet_image' ht) _ hs, + withDensity_apply _ (hf.measurableSet_image' hs), + ← ofReal_integral_eq_lintegral_ofReal hg_int + ((ae_restrict_iff' (hf.measurableSet_image' hs)).mpr hg), + integral_image_eq_integral_abs_det_fderiv_smul μ hs hf' (hf.injective.injOn _)] + simp_rw [smul_eq_mul] + +lemma _root_.MeasurableEquiv.withDensity_ofReal_map_symm_apply_eq_integral_abs_det_fderiv_mul + (hs : MeasurableSet s) (f : E ≃ᵐ E) + {g : E → ℝ} (hg : ∀ᵐ x ∂μ, x ∈ f '' s → 0 ≤ g x) (hg_int : IntegrableOn g (f '' s) μ) + (hf' : ∀ x ∈ s, HasFDerivWithinAt f (f' x) s x) : + (μ.withDensity (fun x ↦ ENNReal.ofReal (g x))).map f.symm s + = ENNReal.ofReal (∫ x in s, |(f' x).det| * g (f x) ∂μ) := by + rw [MeasurableEquiv.map_symm, + MeasurableEmbedding.withDensity_ofReal_comap_apply_eq_integral_abs_det_fderiv_mul μ hs + f.measurableEmbedding hg hg_int hf'] + +lemma _root_.MeasurableEmbedding.withDensity_ofReal_comap_apply_eq_integral_abs_deriv_mul + {f : ℝ → ℝ} (hf : MeasurableEmbedding f) {s : Set ℝ} (hs : MeasurableSet s) + {g : ℝ → ℝ} (hg : ∀ᵐ x, x ∈ f '' s → 0 ≤ g x) (hg_int : IntegrableOn g (f '' s)) + {f' : ℝ → ℝ} (hf' : ∀ x ∈ s, HasDerivWithinAt f (f' x) s x) : + (volume.withDensity (fun x ↦ ENNReal.ofReal (g x))).comap f s + = ENNReal.ofReal (∫ x in s, |f' x| * g (f x)) := by + rw [hf.withDensity_ofReal_comap_apply_eq_integral_abs_det_fderiv_mul volume hs + hg hg_int hf'] + simp only [det_one_smulRight] + +lemma _root_.MeasurableEquiv.withDensity_ofReal_map_symm_apply_eq_integral_abs_deriv_mul + (f : ℝ ≃ᵐ ℝ) {s : Set ℝ} (hs : MeasurableSet s) + {g : ℝ → ℝ} (hg : ∀ᵐ x, x ∈ f '' s → 0 ≤ g x) (hg_int : IntegrableOn g (f '' s)) + {f' : ℝ → ℝ} (hf' : ∀ x ∈ s, HasDerivWithinAt f (f' x) s x) : + (volume.withDensity (fun x ↦ ENNReal.ofReal (g x))).map f.symm s + = ENNReal.ofReal (∫ x in s, |f' x| * g (f x)) := by + rw [MeasurableEquiv.withDensity_ofReal_map_symm_apply_eq_integral_abs_det_fderiv_mul volume hs + f hg hg_int hf'] + simp only [det_one_smulRight] + +lemma _root_.MeasurableEmbedding.withDensity_ofReal_comap_apply_eq_integral_abs_deriv_mul' + {f : ℝ → ℝ} (hf : MeasurableEmbedding f) {s : Set ℝ} (hs : MeasurableSet s) + {f' : ℝ → ℝ} (hf' : ∀ x, HasDerivAt f (f' x) x) + {g : ℝ → ℝ} (hg : 0 ≤ᵐ[volume] g) (hg_int : Integrable g) : + (volume.withDensity (fun x ↦ ENNReal.ofReal (g x))).comap f s + = ENNReal.ofReal (∫ x in s, |f' x| * g (f x)) := + hf.withDensity_ofReal_comap_apply_eq_integral_abs_deriv_mul hs + (by filter_upwards [hg] with x hx using fun _ ↦ hx) hg_int.integrableOn + (fun x _ => (hf' x).hasDerivWithinAt) + +lemma _root_.MeasurableEquiv.withDensity_ofReal_map_symm_apply_eq_integral_abs_deriv_mul' + (f : ℝ ≃ᵐ ℝ) {s : Set ℝ} (hs : MeasurableSet s) + {f' : ℝ → ℝ} (hf' : ∀ x, HasDerivAt f (f' x) x) + {g : ℝ → ℝ} (hg : 0 ≤ᵐ[volume] g) (hg_int : Integrable g) : + (volume.withDensity (fun x ↦ ENNReal.ofReal (g x))).map f.symm s + = ENNReal.ofReal (∫ x in s, |f' x| * g (f x)) := by + rw [MeasurableEquiv.withDensity_ofReal_map_symm_apply_eq_integral_abs_det_fderiv_mul volume hs + f (by filter_upwards [hg] with x hx using fun _ ↦ hx) hg_int.integrableOn + (fun x _ => (hf' x).hasDerivWithinAt)] + simp only [det_one_smulRight] + +end withDensity + end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 8e270be8df3e9..69cb75949f123 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -229,7 +229,7 @@ theorem HasFiniteIntegral.smul_measure {f : α → β} (h : HasFiniteIntegral f @[simp] theorem hasFiniteIntegral_zero_measure {m : MeasurableSpace α} (f : α → β) : HasFiniteIntegral f (0 : Measure α) := by - simp only [HasFiniteIntegral, lintegral_zero_measure, WithTop.zero_lt_top] + simp only [HasFiniteIntegral, lintegral_zero_measure, zero_lt_top] #align measure_theory.has_finite_integral_zero_measure MeasureTheory.hasFiniteIntegral_zero_measure variable (α β μ) @@ -388,9 +388,7 @@ theorem HasFiniteIntegral.max_zero {f : α → ℝ} (hf : HasFiniteIntegral f μ theorem HasFiniteIntegral.min_zero {f : α → ℝ} (hf : HasFiniteIntegral f μ) : HasFiniteIntegral (fun a => min (f a) 0) μ := - hf.mono <| - eventually_of_forall fun x => by - simp [abs_le, neg_le, neg_le_abs_self, abs_eq_max_neg, le_total] + hf.mono <| eventually_of_forall fun x => by simpa [abs_le] using neg_abs_le_self _ #align measure_theory.has_finite_integral.min_zero MeasureTheory.HasFiniteIntegral.min_zero end PosPart @@ -406,6 +404,8 @@ theorem HasFiniteIntegral.smul [NormedAddCommGroup 𝕜] [SMulZeroClass 𝕜 β] (∫⁻ a : α, ‖c • f a‖₊ ∂μ) ≤ ∫⁻ a : α, ‖c‖₊ * ‖f a‖₊ ∂μ := by refine' lintegral_mono _ intro i + -- After leanprover/lean4#2734, we need to do beta reduction `exact_mod_cast` + beta_reduce exact_mod_cast (nnnorm_smul_le c (f i)) _ < ∞ := by rw [lintegral_const_mul'] @@ -669,7 +669,10 @@ theorem Integrable.add' {f g : α → β} (hf : Integrable f μ) (hg : Integrabl HasFiniteIntegral (f + g) μ := calc (∫⁻ a, ‖f a + g a‖₊ ∂μ) ≤ ∫⁻ a, ‖f a‖₊ + ‖g a‖₊ ∂μ := - lintegral_mono fun a => by exact_mod_cast nnnorm_add_le _ _ + lintegral_mono fun a => by + -- After leanprover/lean4#2734, we need to do beta reduction before `exact_mod_cast` + beta_reduce + exact_mod_cast nnnorm_add_le _ _ _ = _ := (lintegral_nnnorm_add_left hf.aestronglyMeasurable _) _ < ∞ := add_lt_top.2 ⟨hf.hasFiniteIntegral, hg.hasFiniteIntegral⟩ #align measure_theory.integrable.add' MeasureTheory.Integrable.add' diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm.lean b/Mathlib/MeasureTheory/Function/LpSeminorm.lean index d598845f2bd33..e6b3d6bc3d584 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm.lean @@ -326,10 +326,10 @@ theorem snorm_const_lt_top_iff {p : ℝ≥0∞} {c : F} (hp_ne_zero : p ≠ 0) ( snorm (fun _ : α => c) p μ < ∞ ↔ c = 0 ∨ μ Set.univ < ∞ := by have hp : 0 < p.toReal := ENNReal.toReal_pos hp_ne_zero hp_ne_top by_cases hμ : μ = 0 - · simp only [hμ, Measure.coe_zero, Pi.zero_apply, or_true_iff, WithTop.zero_lt_top, + · simp only [hμ, Measure.coe_zero, Pi.zero_apply, or_true_iff, zero_lt_top, snorm_measure_zero] by_cases hc : c = 0 - · simp only [hc, true_or_iff, eq_self_iff_true, WithTop.zero_lt_top, snorm_zero'] + · simp only [hc, true_or_iff, eq_self_iff_true, zero_lt_top, snorm_zero'] rw [snorm_const' c hp_ne_zero hp_ne_top] by_cases hμ_top : μ Set.univ = ∞ · simp [hc, hμ_top, hp] @@ -356,7 +356,7 @@ theorem memℒp_const (c : E) [IsFiniteMeasure μ] : Memℒp (fun _ : α => c) p theorem memℒp_top_const (c : E) : Memℒp (fun _ : α => c) ∞ μ := by refine' ⟨aestronglyMeasurable_const, _⟩ by_cases h : μ = 0 - · simp only [h, snorm_measure_zero, WithTop.zero_lt_top] + · simp only [h, snorm_measure_zero, zero_lt_top] · rw [snorm_const _ ENNReal.top_ne_zero h] simp only [ENNReal.top_toReal, _root_.div_zero, ENNReal.rpow_zero, mul_one, ENNReal.coe_lt_top] #align measure_theory.mem_ℒp_top_const MeasureTheory.memℒp_top_const diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 43c2b5da816c3..4a79b034f928b 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -253,6 +253,11 @@ theorem piecewise_empty (f g : α →ₛ β) : piecewise ∅ MeasurableSet.empty coe_injective <| by simp; convert Set.piecewise_empty f g #align measure_theory.simple_func.piecewise_empty MeasureTheory.SimpleFunc.piecewise_empty +@[simp] +theorem piecewise_same (f : α →ₛ β) {s : Set α} (hs : MeasurableSet s) : + piecewise s hs f f = f := + coe_injective <| Set.piecewise_same _ _ + theorem support_indicator [Zero β] {s : Set α} (hs : MeasurableSet s) (f : α →ₛ β) : Function.support (f.piecewise s hs (SimpleFunc.const α 0)) = s ∩ Function.support f := Set.support_indicator @@ -1312,7 +1317,7 @@ theorem _root_.Measurable.add_simpleFunc · simp only [SimpleFunc.const_zero, SimpleFunc.coe_piecewise, SimpleFunc.coe_const, SimpleFunc.coe_zero] change Measurable (g + s.piecewise (Function.const α c) (0 : α → E)) - rw [← piecewise_same s g, ← piecewise_add] + rw [← s.piecewise_same g, ← piecewise_add] exact Measurable.piecewise hs (hg.add_const _) (hg.add_const _) · have : (g + ↑(f + f')) = (Function.support f).piecewise (g + (f : α → E)) (g + f') := by @@ -1337,7 +1342,7 @@ theorem _root_.Measurable.simpleFunc_add · simp only [SimpleFunc.const_zero, SimpleFunc.coe_piecewise, SimpleFunc.coe_const, SimpleFunc.coe_zero] change Measurable (s.piecewise (Function.const α c) (0 : α → E) + g) - rw [← piecewise_same s g, ← piecewise_add] + rw [← s.piecewise_same g, ← piecewise_add] exact Measurable.piecewise hs (hg.const_add _) (hg.const_add _) · have : (↑(f + f') + g) = (Function.support f).piecewise ((f : α → E) + g) (f' + g) := by diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean index 9d38865f554b7..d6354f5cdac21 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Zhouhang Zhou, Yury Kudryashov, Heather Macbeth -/ import Mathlib.MeasureTheory.Function.SimpleFunc +import Mathlib.MeasureTheory.Constructions.BorelSpace.Metrizable #align_import measure_theory.function.simple_func_dense from "leanprover-community/mathlib"@"7317149f12f55affbc900fc873d0d422485122b9" @@ -190,3 +191,77 @@ theorem edist_approxOn_y0_le {f : β → α} (hf : Measurable f) {s : Set α} {y end SimpleFunc end MeasureTheory + +section CompactSupport + +variable {X Y α : Type*} [Zero α] + [TopologicalSpace X] [TopologicalSpace Y] [MeasurableSpace X] [MeasurableSpace Y] + [OpensMeasurableSpace X] [OpensMeasurableSpace Y] + +/-- A continuous function with compact support on a product space can be uniformly approximated by +simple functions. The subtlety is that we do not assume that the spaces are separable, so the +product of the Borel sigma algebras might not contain all open sets, but still it contains enough +of them to approximate compactly supported continuous functions. -/ +lemma HasCompactSupport.exists_simpleFunc_approx_of_prod [PseudoMetricSpace α] + {f : X × Y → α} (hf : Continuous f) (h'f : HasCompactSupport f) + {ε : ℝ} (hε : 0 < ε) : + ∃ (g : SimpleFunc (X × Y) α), ∀ x, dist (f x) (g x) < ε := by + have M : ∀ (K : Set (X × Y)), IsCompact K → + ∃ (g : SimpleFunc (X × Y) α), ∃ (s : Set (X × Y)), MeasurableSet s ∧ K ⊆ s ∧ + ∀ x ∈ s, dist (f x) (g x) < ε := by + intro K hK + apply IsCompact.induction_on + (p := fun t ↦ ∃ (g : SimpleFunc (X × Y) α), ∃ (s : Set (X × Y)), MeasurableSet s ∧ t ⊆ s ∧ + ∀ x ∈ s, dist (f x) (g x) < ε) hK + · exact ⟨0, ∅, by simp⟩ + · intro t t' htt' ⟨g, s, s_meas, ts, hg⟩ + exact ⟨g, s, s_meas, htt'.trans ts, hg⟩ + · intro t t' ⟨g, s, s_meas, ts, hg⟩ ⟨g', s', s'_meas, t's', hg'⟩ + refine ⟨g.piecewise s s_meas g', s ∪ s', s_meas.union s'_meas, + union_subset_union ts t's', fun p hp ↦ ?_⟩ + by_cases H : p ∈ s + · simpa [H, SimpleFunc.piecewise_apply] using hg p H + · simp only [SimpleFunc.piecewise_apply, H, ite_false] + apply hg' + simpa [H] using (mem_union _ _ _).1 hp + · rintro ⟨x, y⟩ - + obtain ⟨u, v, hu, xu, hv, yv, huv⟩ : ∃ u v, IsOpen u ∧ x ∈ u ∧ IsOpen v ∧ y ∈ v ∧ + u ×ˢ v ⊆ {z | dist (f z) (f (x, y)) < ε} := + mem_nhds_prod_iff'.1 <| Metric.continuousAt_iff'.1 hf.continuousAt ε hε + refine ⟨u ×ˢ v, nhdsWithin_le_nhds <| (hu.prod hv).mem_nhds (mk_mem_prod xu yv), ?_⟩ + exact ⟨SimpleFunc.const _ (f (x, y)), u ×ˢ v, hu.measurableSet.prod hv.measurableSet, + Subset.rfl, fun z hz ↦ huv hz⟩ + obtain ⟨g, s, s_meas, fs, hg⟩ : ∃ g s, MeasurableSet s ∧ tsupport f ⊆ s ∧ + ∀ (x : X × Y), x ∈ s → dist (f x) (g x) < ε := M _ h'f + refine ⟨g.piecewise s s_meas 0, fun p ↦ ?_⟩ + by_cases H : p ∈ s + · simpa [H, SimpleFunc.piecewise_apply] using hg p H + · have : f p = 0 := by + contrapose! H + rw [← Function.mem_support] at H + exact fs (subset_tsupport _ H) + simp [SimpleFunc.piecewise_apply, H, ite_false, this, hε] + +/-- A continuous function with compact support on a product space is measurable for the product +sigma-algebra. The subtlety is that we do not assume that the spaces are separable, so the +product of the Borel sigma algebras might not contain all open sets, but still it contains enough +of them to approximate compactly supported continuous functions. -/ +lemma HasCompactSupport.measurable_of_prod + [TopologicalSpace α] [PseudoMetrizableSpace α] [MeasurableSpace α] [BorelSpace α] + {f : X × Y → α} (hf : Continuous f) (h'f : HasCompactSupport f) : + Measurable f := by + letI : PseudoMetricSpace α := TopologicalSpace.pseudoMetrizableSpacePseudoMetric α + obtain ⟨u, -, u_pos, u_lim⟩ : ∃ u, StrictAnti u ∧ (∀ (n : ℕ), 0 < u n) ∧ Tendsto u atTop (𝓝 0) := + exists_seq_strictAnti_tendsto (0 : ℝ) + have : ∀ n, ∃ (g : SimpleFunc (X × Y) α), ∀ x, dist (f x) (g x) < u n := + fun n ↦ h'f.exists_simpleFunc_approx_of_prod hf (u_pos n) + choose g hg using this + have A : ∀ x, Tendsto (fun n ↦ g n x) atTop (𝓝 (f x)) := by + intro x + rw [tendsto_iff_dist_tendsto_zero] + apply squeeze_zero (fun n ↦ dist_nonneg) (fun n ↦ ?_) u_lim + rw [dist_comm] + exact (hg n x).le + apply measurable_of_tendsto_metrizable (fun n ↦ (g n).measurable) (tendsto_pi_nhds.2 A) + +end CompactSupport diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index a6a68b4cc15bf..3b1595ea6cbaf 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -986,7 +986,7 @@ theorem Memℒp.induction_dense (hp_ne_top : p ≠ ∞) (P : (α → E) → Prop (h1P : ∀ f g, P f → P g → P (f + g)) (h2P : ∀ f, P f → AEStronglyMeasurable f μ) {f : α → E} (hf : Memℒp f p μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : ∃ g : α → E, snorm (f - g) p μ ≤ ε ∧ P g := by rcases eq_or_ne p 0 with (rfl | hp_pos) - · rcases h0P (0 : E) MeasurableSet.empty (by simp only [measure_empty, WithTop.zero_lt_top]) + · rcases h0P (0 : E) MeasurableSet.empty (by simp only [measure_empty, zero_lt_top]) hε with ⟨g, _, Pg⟩ exact ⟨g, by simp only [snorm_exponent_zero, zero_le'], Pg⟩ suffices H : @@ -1001,7 +1001,7 @@ theorem Memℒp.induction_dense (hp_ne_top : p ≠ ∞) (P : (α → E) → Prop apply SimpleFunc.induction · intro c s hs ε εpos Hs rcases eq_or_ne c 0 with (rfl | hc) - · rcases h0P (0 : E) MeasurableSet.empty (by simp only [measure_empty, WithTop.zero_lt_top]) + · rcases h0P (0 : E) MeasurableSet.empty (by simp only [measure_empty, zero_lt_top]) εpos with ⟨g, hg, Pg⟩ rw [← snorm_neg, neg_sub] at hg refine' ⟨g, _, Pg⟩ diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 2637161541db3..3e62c6c19166b 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -5,7 +5,6 @@ Authors: Rémy Degenne, Sébastien Gouëzel -/ import Mathlib.Analysis.NormedSpace.FiniteDimension import Mathlib.Analysis.NormedSpace.BoundedLinearMaps -import Mathlib.MeasureTheory.Constructions.BorelSpace.Metrizable import Mathlib.MeasureTheory.Measure.WithDensity import Mathlib.MeasureTheory.Function.SimpleFuncDense @@ -699,14 +698,38 @@ theorem _root_.Continuous.stronglyMeasurable [MeasurableSpace α] [TopologicalSp · exact hf.measurable.stronglyMeasurable #align continuous.strongly_measurable Continuous.stronglyMeasurable -/-- A continuous function with compact support is strongly measurable. -/ -theorem _root_.Continuous.stronglyMeasurable_of_hasCompactSupport +/-- A continuous function whose support is contained in a compact set is strongly measurable. -/ +@[to_additive] +theorem _root_.Continuous.stronglyMeasurable_of_mulSupport_subset_isCompact [MeasurableSpace α] [TopologicalSpace α] [OpensMeasurableSpace α] [MeasurableSpace β] - [TopologicalSpace β] [PseudoMetrizableSpace β] [BorelSpace β] [Zero β] {f : α → β} - (hf : Continuous f) (h'f : HasCompactSupport f) : StronglyMeasurable f := by + [TopologicalSpace β] [PseudoMetrizableSpace β] [BorelSpace β] [One β] {f : α → β} + (hf : Continuous f) {k : Set α} (hk : IsCompact k) + (h'f : mulSupport f ⊆ k) : StronglyMeasurable f := by letI : PseudoMetricSpace β := pseudoMetrizableSpacePseudoMetric β rw [stronglyMeasurable_iff_measurable_separable] - refine ⟨hf.measurable, IsCompact.isSeparable (s := range f) (h'f.isCompact_range hf)⟩ + exact ⟨hf.measurable, (isCompact_range_of_mulSupport_subset_isCompact hf hk h'f).isSeparable⟩ + +/-- A continuous function with compact support is strongly measurable. -/ +@[to_additive] +theorem _root_.Continuous.stronglyMeasurable_of_hasCompactMulSupport + [MeasurableSpace α] [TopologicalSpace α] [OpensMeasurableSpace α] [MeasurableSpace β] + [TopologicalSpace β] [PseudoMetrizableSpace β] [BorelSpace β] [One β] {f : α → β} + (hf : Continuous f) (h'f : HasCompactMulSupport f) : StronglyMeasurable f := + hf.stronglyMeasurable_of_mulSupport_subset_isCompact h'f (subset_mulTSupport f) + +/-- A continuous function with compact support on a product space is strongly measurable for the +product sigma-algebra. The subtlety is that we do not assume that the spaces are separable, so the +product of the Borel sigma algebras might not contain all open sets, but still it contains enough +of them to approximate compactly supported continuous functions. -/ +lemma _root_.HasCompactSupport.stronglyMeasurable_of_prod {X Y : Type*} [Zero α] + [TopologicalSpace X] [TopologicalSpace Y] [MeasurableSpace X] [MeasurableSpace Y] + [OpensMeasurableSpace X] [OpensMeasurableSpace Y] [TopologicalSpace α] [PseudoMetrizableSpace α] + {f : X × Y → α} (hf : Continuous f) (h'f : HasCompactSupport f) : + StronglyMeasurable f := by + borelize α + apply stronglyMeasurable_iff_measurable_separable.2 ⟨h'f.measurable_of_prod hf, ?_⟩ + letI : PseudoMetricSpace α := pseudoMetrizableSpacePseudoMetric α + exact IsCompact.isSeparable (s := range f) (h'f.isCompact_range hf) /-- If `g` is a topological embedding, then `f` is strongly measurable iff `g ∘ f` is. -/ theorem _root_.Embedding.comp_stronglyMeasurable_iff {m : MeasurableSpace α} [TopologicalSpace β] @@ -982,7 +1005,7 @@ theorem finStronglyMeasurable_zero {α β} {m : MeasurableSpace α} {μ : Measur [TopologicalSpace β] : FinStronglyMeasurable (0 : α → β) μ := ⟨0, by simp only [Pi.zero_apply, SimpleFunc.coe_zero, support_zero', measure_empty, - WithTop.zero_lt_top, forall_const], + zero_lt_top, forall_const], fun _ => tendsto_const_nhds⟩ #align measure_theory.fin_strongly_measurable_zero MeasureTheory.finStronglyMeasurable_zero diff --git a/Mathlib/MeasureTheory/Group/Action.lean b/Mathlib/MeasureTheory/Group/Action.lean index 4ec1718a98a7c..f7a5d5cee19a7 100644 --- a/Mathlib/MeasureTheory/Group/Action.lean +++ b/Mathlib/MeasureTheory/Group/Action.lean @@ -3,11 +3,11 @@ Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import Mathlib.MeasureTheory.Group.MeasurableEquiv -import Mathlib.MeasureTheory.Measure.Regular import Mathlib.Dynamics.Ergodic.MeasurePreserving import Mathlib.Dynamics.Minimal -import Mathlib.Algebra.Hom.GroupAction +import Mathlib.GroupTheory.GroupAction.Hom +import Mathlib.MeasureTheory.Group.MeasurableEquiv +import Mathlib.MeasureTheory.Measure.Regular #align_import measure_theory.group.action from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 29950005d2dd6..17dc60696913c 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -3,13 +3,13 @@ Copyright (c) 2020 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Mathlib.Algebra.Hom.GroupAction import Mathlib.Dynamics.Ergodic.MeasurePreserving -import Mathlib.MeasureTheory.Measure.Regular +import Mathlib.GroupTheory.GroupAction.Hom +import Mathlib.MeasureTheory.Constructions.Prod.Basic +import Mathlib.MeasureTheory.Group.Action import Mathlib.MeasureTheory.Group.MeasurableEquiv import Mathlib.MeasureTheory.Measure.OpenPos -import Mathlib.MeasureTheory.Group.Action -import Mathlib.MeasureTheory.Constructions.Prod.Basic +import Mathlib.MeasureTheory.Measure.Regular import Mathlib.Topology.ContinuousFunction.CocompactMap import Mathlib.Topology.Homeomorph diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 612837bc868b7..e1678df86b8cf 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1265,7 +1265,7 @@ theorem integral_nonpos {f : α → ℝ} (hf : f ≤ 0) : ∫ a, f a ∂μ ≤ 0 theorem integral_eq_zero_iff_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f) (hfi : Integrable f μ) : ∫ x, f x ∂μ = 0 ↔ f =ᵐ[μ] 0 := by simp_rw [integral_eq_lintegral_of_nonneg_ae hf hfi.1, ENNReal.toReal_eq_zero_iff, - ← ENNReal.not_lt_top, ← hasFiniteIntegral_iff_ofReal hf, hfi.2, or_false_iff] + ← ENNReal.not_lt_top, ← hasFiniteIntegral_iff_ofReal hf, hfi.2, not_true_eq_false, or_false_iff] -- Porting note: split into parts, to make `rw` and `simp` work rw [lintegral_eq_zero_iff'] · rw [← hf.le_iff_eq, Filter.EventuallyEq, Filter.EventuallyLE] @@ -1645,15 +1645,15 @@ theorem MeasurePreserving.integral_comp {β} {_ : MeasurableSpace β} {f : α h₁.map_eq ▸ (h₂.integral_map g).symm #align measure_theory.measure_preserving.integral_comp MeasureTheory.MeasurePreserving.integral_comp -theorem set_integral_eq_subtype' {α} [MeasurableSpace α] {μ : Measure α} {s : Set α} +theorem integral_subtype_comap {α} [MeasurableSpace α] {μ : Measure α} {s : Set α} (hs : MeasurableSet s) (f : α → G) : - ∫ x in s, f x ∂μ = ∫ x : s, f (x : α) ∂(Measure.comap Subtype.val μ) := by + ∫ x : s, f (x : α) ∂(Measure.comap Subtype.val μ) = ∫ x in s, f x ∂μ := by rw [← map_comap_subtype_coe hs] - exact (MeasurableEmbedding.subtype_coe hs).integral_map _ + exact ((MeasurableEmbedding.subtype_coe hs).integral_map _).symm -theorem set_integral_eq_subtype {α} [MeasureSpace α] {s : Set α} (hs : MeasurableSet s) - (f : α → G) : ∫ x in s, f x = ∫ x : s, f x := set_integral_eq_subtype' hs f -#align measure_theory.set_integral_eq_subtype MeasureTheory.set_integral_eq_subtype +theorem integral_subtype {α} [MeasureSpace α] {s : Set α} (hs : MeasurableSet s) (f : α → G) : + ∫ x : s, f x = ∫ x in s, f x := integral_subtype_comap hs f +#align measure_theory.set_integral_eq_subtype MeasureTheory.integral_subtype @[simp] theorem integral_dirac' [MeasurableSpace α] (f : α → E) (a : α) (hfm : StronglyMeasurable f) : @@ -1806,7 +1806,7 @@ theorem integral_countable [MeasurableSingletonClass α] (f : α → E) {s : Set · exact Integrable.aestronglyMeasurable hf · exact Measurable.aemeasurable measurable_subtype_coe · exact Countable.measurableSet hs - rw [set_integral_eq_subtype' hs.measurableSet, integral_countable' hf'] + rw [← integral_subtype_comap hs.measurableSet, integral_countable' hf'] congr 1 with a : 1 rw [Measure.comap_apply Subtype.val Subtype.coe_injective (fun s' hs' => MeasurableSet.subtype_image (Countable.measurableSet hs) hs') _ diff --git a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean index 56519e2710549..3cbe2a779d118 100644 --- a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean +++ b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean @@ -423,7 +423,7 @@ protected theorem IntegrableAtFilter.smul {𝕜 : Type*} [NormedAddCommGroup [BoundedSMul 𝕜 E] {f : α → E} (hf : IntegrableAtFilter f l μ) (c : 𝕜) : IntegrableAtFilter (c • f) l μ := by rcases hf with ⟨s, sl, hs⟩ - refine ⟨s, sl, hs.smul c⟩ + exact ⟨s, sl, hs.smul c⟩ theorem IntegrableAtFilter.filter_mono (hl : l ≤ l') (hl' : IntegrableAtFilter f l' μ) : IntegrableAtFilter f l μ := diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index 714090c74e6c8..c4d9f4f6c91a0 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker, Bhavik Mehta -/ +import Mathlib.Analysis.Calculus.Deriv.Support import Mathlib.Analysis.SpecialFunctions.Pow.Deriv import Mathlib.MeasureTheory.Integral.FundThmCalculus import Mathlib.Order.Filter.AtTopBot @@ -665,10 +666,15 @@ variable {E : Type*} {f f' : ℝ → E} {g g' : ℝ → ℝ} {a b l : ℝ} {m : /-- **Fundamental theorem of calculus-2**, on semi-infinite intervals `(a, +∞)`. When a function has a limit at infinity `m`, and its derivative is integrable, then the integral of the derivative on `(a, +∞)` is `m - f a`. Version assuming differentiability -on `(a, +∞)` and continuity on `[a, +∞)`.-/ -theorem integral_Ioi_of_hasDerivAt_of_tendsto (hcont : ContinuousOn f (Ici a)) +on `(a, +∞)` and continuity at `a⁺`.-/ +theorem integral_Ioi_of_hasDerivAt_of_tendsto (hcont : ContinuousWithinAt f (Ici a) a) (hderiv : ∀ x ∈ Ioi a, HasDerivAt f (f' x) x) (f'int : IntegrableOn f' (Ioi a)) (hf : Tendsto f atTop (𝓝 m)) : ∫ x in Ioi a, f' x = m - f a := by + have hcont : ContinuousOn f (Ici a) := by + intro x hx + rcases hx.out.eq_or_lt with rfl|hx + · exact hcont + · exact (hderiv x hx).continuousAt.continuousWithinAt refine' tendsto_nhds_unique (intervalIntegral_tendsto_integral_Ioi a f'int tendsto_id) _ apply Tendsto.congr' _ (hf.sub_const _) filter_upwards [Ioi_mem_atTop a] with x hx @@ -688,17 +694,32 @@ on `[a, +∞)`. -/ theorem integral_Ioi_of_hasDerivAt_of_tendsto' (hderiv : ∀ x ∈ Ici a, HasDerivAt f (f' x) x) (f'int : IntegrableOn f' (Ioi a)) (hf : Tendsto f atTop (𝓝 m)) : ∫ x in Ioi a, f' x = m - f a := by - refine integral_Ioi_of_hasDerivAt_of_tendsto (fun x hx ↦ ?_) (fun x hx => hderiv x hx.out.le) + refine integral_Ioi_of_hasDerivAt_of_tendsto ?_ (fun x hx => hderiv x hx.out.le) f'int hf - exact (hderiv x hx).continuousAt.continuousWithinAt + exact (hderiv a left_mem_Ici).continuousAt.continuousWithinAt #align measure_theory.integral_Ioi_of_has_deriv_at_of_tendsto' MeasureTheory.integral_Ioi_of_hasDerivAt_of_tendsto' +/-- A special case of `integral_Ioi_of_hasDerivAt_of_tendsto` where we assume that `f` is C^1 with +compact support. -/ +theorem HasCompactSupport.integral_Ioi_deriv_eq (hf : ContDiff ℝ 1 f) (h2f : HasCompactSupport f) + (b : ℝ) : ∫ x in Ioi b, deriv f x = - f b := by + have := fun x (_ : x ∈ Ioi b) ↦ hf.differentiable le_rfl x |>.hasDerivAt + rw [integral_Ioi_of_hasDerivAt_of_tendsto hf.continuous.continuousWithinAt this, zero_sub] + refine hf.continuous_deriv le_rfl |>.integrable_of_hasCompactSupport h2f.deriv |>.integrableOn + rw [hasCompactSupport_iff_eventuallyEq, Filter.coclosedCompact_eq_cocompact] at h2f + exact h2f.filter_mono atTop_le_cocompact |>.tendsto + /-- When a function has a limit at infinity, and its derivative is nonnegative, then the derivative is automatically integrable on `(a, +∞)`. Version assuming differentiability -on `(a, +∞)` and continuity on `[a, +∞)`. -/ -theorem integrableOn_Ioi_deriv_of_nonneg (hcont : ContinuousOn g (Ici a)) +on `(a, +∞)` and continuity at `a⁺`. -/ +theorem integrableOn_Ioi_deriv_of_nonneg (hcont : ContinuousWithinAt g (Ici a) a) (hderiv : ∀ x ∈ Ioi a, HasDerivAt g (g' x) x) (g'pos : ∀ x ∈ Ioi a, 0 ≤ g' x) (hg : Tendsto g atTop (𝓝 l)) : IntegrableOn g' (Ioi a) := by + have hcont : ContinuousOn g (Ici a) := by + intro x hx + rcases hx.out.eq_or_lt with rfl|hx + · exact hcont + · exact (hderiv x hx).continuousAt.continuousWithinAt refine integrableOn_Ioi_of_intervalIntegral_norm_tendsto (l - g a) a (fun x => ?_) tendsto_id ?_ · exact intervalIntegral.integrableOn_deriv_of_nonneg (hcont.mono Icc_subset_Ici_self) (fun y hy => hderiv y hy.1) fun y hy => g'pos y hy.1 @@ -726,15 +747,15 @@ is automatically integrable on `(a, +∞)`. Version assuming differentiability on `[a, +∞)`. -/ theorem integrableOn_Ioi_deriv_of_nonneg' (hderiv : ∀ x ∈ Ici a, HasDerivAt g (g' x) x) (g'pos : ∀ x ∈ Ioi a, 0 ≤ g' x) (hg : Tendsto g atTop (𝓝 l)) : IntegrableOn g' (Ioi a) := by - refine integrableOn_Ioi_deriv_of_nonneg (fun x hx ↦ ?_) (fun x hx => hderiv x hx.out.le) g'pos hg - exact (hderiv x hx).continuousAt.continuousWithinAt + refine integrableOn_Ioi_deriv_of_nonneg ?_ (fun x hx => hderiv x hx.out.le) g'pos hg + exact (hderiv a left_mem_Ici).continuousAt.continuousWithinAt #align measure_theory.integrable_on_Ioi_deriv_of_nonneg' MeasureTheory.integrableOn_Ioi_deriv_of_nonneg' /-- When a function has a limit at infinity `l`, and its derivative is nonnegative, then the integral of the derivative on `(a, +∞)` is `l - g a` (and the derivative is integrable, see `integrable_on_Ioi_deriv_of_nonneg`). Version assuming differentiability on `(a, +∞)` and -continuity on `[a, +∞)`. -/ -theorem integral_Ioi_of_hasDerivAt_of_nonneg (hcont : ContinuousOn g (Ici a)) +continuity at `a⁺`. -/ +theorem integral_Ioi_of_hasDerivAt_of_nonneg (hcont : ContinuousWithinAt g (Ici a) a) (hderiv : ∀ x ∈ Ioi a, HasDerivAt g (g' x) x) (g'pos : ∀ x ∈ Ioi a, 0 ≤ g' x) (hg : Tendsto g atTop (𝓝 l)) : ∫ x in Ioi a, g' x = l - g a := integral_Ioi_of_hasDerivAt_of_tendsto hcont hderiv @@ -752,8 +773,8 @@ theorem integral_Ioi_of_hasDerivAt_of_nonneg' (hderiv : ∀ x ∈ Ici a, HasDeri /-- When a function has a limit at infinity, and its derivative is nonpositive, then the derivative is automatically integrable on `(a, +∞)`. Version assuming differentiability -on `(a, +∞)` and continuity on `[a, +∞)`. -/ -theorem integrableOn_Ioi_deriv_of_nonpos (hcont : ContinuousOn g (Ici a)) +on `(a, +∞)` and continuity at `a⁺`. -/ +theorem integrableOn_Ioi_deriv_of_nonpos (hcont : ContinuousWithinAt g (Ici a) a) (hderiv : ∀ x ∈ Ioi a, HasDerivAt g (g' x) x) (g'neg : ∀ x ∈ Ioi a, g' x ≤ 0) (hg : Tendsto g atTop (𝓝 l)) : IntegrableOn g' (Ioi a) := by apply integrable_neg_iff.1 @@ -766,15 +787,15 @@ is automatically integrable on `(a, +∞)`. Version assuming differentiability on `[a, +∞)`. -/ theorem integrableOn_Ioi_deriv_of_nonpos' (hderiv : ∀ x ∈ Ici a, HasDerivAt g (g' x) x) (g'neg : ∀ x ∈ Ioi a, g' x ≤ 0) (hg : Tendsto g atTop (𝓝 l)) : IntegrableOn g' (Ioi a) := by - refine integrableOn_Ioi_deriv_of_nonpos (fun x hx ↦ ?_) (fun x hx ↦ hderiv x hx.out.le) g'neg hg - exact (hderiv x hx).continuousAt.continuousWithinAt + refine integrableOn_Ioi_deriv_of_nonpos ?_ (fun x hx ↦ hderiv x hx.out.le) g'neg hg + exact (hderiv a left_mem_Ici).continuousAt.continuousWithinAt #align measure_theory.integrable_on_Ioi_deriv_of_nonpos' MeasureTheory.integrableOn_Ioi_deriv_of_nonpos' /-- When a function has a limit at infinity `l`, and its derivative is nonpositive, then the integral of the derivative on `(a, +∞)` is `l - g a` (and the derivative is integrable, see `integrable_on_Ioi_deriv_of_nonneg`). Version assuming differentiability on `(a, +∞)` and -continuity on `[a, +∞)`. -/ -theorem integral_Ioi_of_hasDerivAt_of_nonpos (hcont : ContinuousOn g (Ici a)) +continuity at `a⁺`. -/ +theorem integral_Ioi_of_hasDerivAt_of_nonpos (hcont : ContinuousWithinAt g (Ici a) a) (hderiv : ∀ x ∈ Ioi a, HasDerivAt g (g' x) x) (g'neg : ∀ x ∈ Ioi a, g' x ≤ 0) (hg : Tendsto g atTop (𝓝 l)) : ∫ x in Ioi a, g' x = l - g a := integral_Ioi_of_hasDerivAt_of_tendsto hcont hderiv @@ -792,6 +813,54 @@ theorem integral_Ioi_of_hasDerivAt_of_nonpos' (hderiv : ∀ x ∈ Ici a, HasDeri end IoiFTC +section IicFTC + +variable {E : Type*} {f f' : ℝ → E} {g g' : ℝ → ℝ} {a b l : ℝ} {m : E} [NormedAddCommGroup E] + [NormedSpace ℝ E] [CompleteSpace E] +/-- **Fundamental theorem of calculus-2**, on semi-infinite intervals `(-∞, a)`. +When a function has a limit `m` at `-∞`, and its derivative is integrable, then the +integral of the derivative on `(-∞, a)` is `f a - m`. Version assuming differentiability +on `(-∞, a)` and continuity at `a⁻`. -/ +theorem integral_Iic_of_hasDerivAt_of_tendsto (hcont : ContinuousWithinAt f (Iic a) a) + (hderiv : ∀ x ∈ Iio a, HasDerivAt f (f' x) x) (f'int : IntegrableOn f' (Iic a)) + (hf : Tendsto f atBot (𝓝 m)) : ∫ x in Iic a, f' x = f a - m := by + have hcont : ContinuousOn f (Iic a) := by + intro x hx + rcases hx.out.eq_or_lt with rfl|hx + · exact hcont + · exact (hderiv x hx).continuousAt.continuousWithinAt + refine' tendsto_nhds_unique (intervalIntegral_tendsto_integral_Iic a f'int tendsto_id) _ + apply Tendsto.congr' _ (hf.const_sub _) + filter_upwards [Iic_mem_atBot a] with x hx + symm + apply intervalIntegral.integral_eq_sub_of_hasDerivAt_of_le hx + (hcont.mono Icc_subset_Iic_self) fun y hy => hderiv y hy.2 + rw [intervalIntegrable_iff_integrable_Ioc_of_le hx] + exact f'int.mono (fun y hy => hy.2) le_rfl + +/-- **Fundamental theorem of calculus-2**, on semi-infinite intervals `(-∞, a)`. +When a function has a limit `m` at `-∞`, and its derivative is integrable, then the +integral of the derivative on `(-∞, a)` is `f a - m`. Version assuming differentiability +on `(-∞, a]`. -/ +theorem integral_Iic_of_hasDerivAt_of_tendsto' + (hderiv : ∀ x ∈ Iic a, HasDerivAt f (f' x) x) (f'int : IntegrableOn f' (Iic a)) + (hf : Tendsto f atBot (𝓝 m)) : ∫ x in Iic a, f' x = f a - m := by + refine integral_Iic_of_hasDerivAt_of_tendsto ?_ (fun x hx => hderiv x hx.out.le) + f'int hf + exact (hderiv a right_mem_Iic).continuousAt.continuousWithinAt + +/-- A special case of `integral_Iic_of_hasDerivAt_of_tendsto` where we assume that `f` is C^1 with +compact support. -/ +theorem HasCompactSupport.integral_Iic_deriv_eq (hf : ContDiff ℝ 1 f) (h2f : HasCompactSupport f) + (b : ℝ) : ∫ x in Iic b, deriv f x = f b := by + have := fun x (_ : x ∈ Iio b) ↦ hf.differentiable le_rfl x |>.hasDerivAt + rw [integral_Iic_of_hasDerivAt_of_tendsto hf.continuous.continuousWithinAt this, sub_zero] + refine hf.continuous_deriv le_rfl |>.integrable_of_hasCompactSupport h2f.deriv |>.integrableOn + rw [hasCompactSupport_iff_eventuallyEq, Filter.coclosedCompact_eq_cocompact] at h2f + exact h2f.filter_mono atBot_le_cocompact |>.tendsto + +end IicFTC + section IoiChangeVariables open Real diff --git a/Mathlib/MeasureTheory/Integral/Layercake.lean b/Mathlib/MeasureTheory/Integral/Layercake.lean index 6b563d7fbfed1..41b54e7c3a31d 100644 --- a/Mathlib/MeasureTheory/Integral/Layercake.lean +++ b/Mathlib/MeasureTheory/Integral/Layercake.lean @@ -79,7 +79,7 @@ theorem countable_meas_le_ne_meas_lt (g : α → R) : intro t ht have : μ {a | t < g a} < μ {a | t ≤ g a} := lt_of_le_of_ne (measure_mono (fun a ha ↦ le_of_lt ha)) (Ne.symm ht) - refine ⟨μ {a | t < g a}, this, fun s hs ↦ measure_mono (fun a ha ↦ hs.trans_le ha)⟩ + exact ⟨μ {a | t < g a}, this, fun s hs ↦ measure_mono (fun a ha ↦ hs.trans_le ha)⟩ theorem meas_le_ae_eq_meas_lt {R : Type*} [LinearOrder R] [MeasurableSpace R] (ν : Measure R) [NoAtoms ν] (g : α → R) : @@ -178,7 +178,7 @@ theorem lintegral_comp_eq_lintegral_meas_le_mul_of_measurable_of_sigmaFinite rw [aux₂] have mble₀ : MeasurableSet {p : α × ℝ | p.snd ∈ Ioc 0 (f p.fst)} := by simpa only [mem_univ, Pi.zero_apply, gt_iff_lt, not_lt, ge_iff_le, true_and] using - measurableSet_region_between_oc measurable_zero f_mble MeasurableSet.univ + measurableSet_region_between_oc measurable_zero f_mble MeasurableSet.univ exact (ENNReal.measurable_ofReal.comp (g_mble.comp measurable_snd)).aemeasurable.indicator₀ mble₀.nullMeasurableSet #align measure_theory.lintegral_comp_eq_lintegral_meas_le_mul_of_measurable MeasureTheory.lintegral_comp_eq_lintegral_meas_le_mul_of_measurable_of_sigmaFinite diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 9618c9bd395a3..044298d8be7fd 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -1383,6 +1383,15 @@ theorem MeasurePreserving.set_lintegral_comp_emb {mb : MeasurableSpace β} {ν : rw [← hg.set_lintegral_comp_preimage_emb hge, preimage_image_eq _ hge.injective] #align measure_theory.measure_preserving.set_lintegral_comp_emb MeasureTheory.MeasurePreserving.set_lintegral_comp_emb +theorem lintegral_subtype_comap {s : Set α} (hs : MeasurableSet s) (f : α → ℝ≥0∞) : + ∫⁻ x : s, f x ∂(μ.comap (↑)) = ∫⁻ x in s, f x ∂μ := by + rw [← (MeasurableEmbedding.subtype_coe hs).lintegral_map, map_comap_subtype_coe hs] + +theorem set_lintegral_subtype {s : Set α} (hs : MeasurableSet s) (t : Set s) (f : α → ℝ≥0∞) : + ∫⁻ x in t, f x ∂(μ.comap (↑)) = ∫⁻ x in (↑) '' t, f x ∂μ := by + rw [(MeasurableEmbedding.subtype_coe hs).restrict_comap, lintegral_subtype_comap hs, + restrict_restrict hs, inter_eq_right.2 (Subtype.coe_image_subset _ _)] + section DiracAndCount instance (priority := 10) _root_.MeasurableSpace.Top.measurableSingletonClass {α : Type*} : diff --git a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean index c018b986246e2..b5a9582183f7b 100644 --- a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean +++ b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean @@ -24,6 +24,12 @@ and `α → (E)NNReal` functions in two cases, * `ENNReal.lintegral_mul_le_Lp_mul_Lq` : ℝ≥0∞ functions, * `NNReal.lintegral_mul_le_Lp_mul_Lq` : ℝ≥0 functions. +`ENNReal.lintegral_mul_norm_pow_le` is a variant where the exponents are not reciprocals: +`∫ (f ^ p * g ^ q) ∂μ ≤ (∫ f ∂μ) ^ p * (∫ g ∂μ) ^ q` where `p, q ≥ 0` and `p + q = 1`. +`ENNReal.lintegral_prod_norm_pow_le` generalizes this to a finite family of functions: +`∫ (∏ i, f i ^ p i) ∂μ ≤ ∏ i, (∫ f i ∂μ) ^ p i` when the `p` is a collection +of nonnegative weights with sum 1. + Minkowski's inequality for the Lebesgue integral of measurable functions with `ℝ≥0∞` values: we prove `(∫ (f + g)^p ∂μ) ^ (1/p) ≤ (∫ f^p ∂μ) ^ (1/p) + (∫ g^p ∂μ) ^ (1/p)` for `1 ≤ p`. -/ @@ -48,7 +54,7 @@ only to prove the more general results: noncomputable section -open Classical BigOperators NNReal ENNReal MeasureTheory +open Classical BigOperators NNReal ENNReal MeasureTheory Finset set_option linter.uppercaseLean3 false @@ -175,6 +181,101 @@ theorem lintegral_mul_le_Lp_mul_Lq (μ : Measure α) {p q : ℝ} (hpq : p.IsConj exact ENNReal.lintegral_mul_le_Lp_mul_Lq_of_ne_zero_of_ne_top hpq hf hf_top hg_top hf_zero hg_zero #align ennreal.lintegral_mul_le_Lp_mul_Lq ENNReal.lintegral_mul_le_Lp_mul_Lq +/-- A different formulation of Hölder's inequality for two functions, with two exponents that sum to +1, instead of reciprocals of -/ +theorem lintegral_mul_norm_pow_le {α} [MeasurableSpace α] {μ : Measure α} + {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) + {p q : ℝ} (hp : 0 ≤ p) (hq : 0 ≤ q) (hpq : p + q = 1) : + ∫⁻ a, f a ^ p * g a ^ q ∂μ ≤ (∫⁻ a, f a ∂μ) ^ p * (∫⁻ a, g a ∂μ) ^ q := by + rcases hp.eq_or_lt with rfl|hp + · rw [zero_add] at hpq + simp [hpq] + rcases hq.eq_or_lt with rfl|hq + · rw [add_zero] at hpq + simp [hpq] + have h2p : 1 < 1 / p := by + rw [one_div] + apply one_lt_inv hp + linarith + have h2pq : 1 / (1 / p) + 1 / (1 / q) = 1 := by + simp [hp.ne', hq.ne', hpq] + have := ENNReal.lintegral_mul_le_Lp_mul_Lq μ ⟨h2p, h2pq⟩ (hf.pow_const p) (hg.pow_const q) + simpa [← ENNReal.rpow_mul, hp.ne', hq.ne'] using this + +/-- A version of Hölder with multiple arguments -/ +theorem lintegral_prod_norm_pow_le {α ι : Type*} [MeasurableSpace α] {μ : Measure α} + (s : Finset ι) {f : ι → α → ℝ≥0∞} (hf : ∀ i ∈ s, AEMeasurable (f i) μ) + {p : ι → ℝ} (hp : ∑ i in s, p i = 1) (h2p : ∀ i ∈ s, 0 ≤ p i) : + ∫⁻ a, ∏ i in s, f i a ^ p i ∂μ ≤ ∏ i in s, (∫⁻ a, f i a ∂μ) ^ p i := by + induction s using Finset.induction generalizing p + case empty => + simp at hp + case insert i₀ s hi₀ ih => + rcases eq_or_ne (p i₀) 1 with h2i₀|h2i₀ + · simp [hi₀] + have h2p : ∀ i ∈ s, p i = 0 := by + simpa [hi₀, h2i₀, sum_eq_zero_iff_of_nonneg (fun i hi ↦ h2p i <| mem_insert_of_mem hi)] + using hp + calc ∫⁻ a, f i₀ a ^ p i₀ * ∏ i in s, f i a ^ p i ∂μ + = ∫⁻ a, f i₀ a ^ p i₀ * ∏ i in s, 1 ∂μ := by + congr! 3 with x + apply prod_congr rfl fun i hi ↦ by rw [h2p i hi, ENNReal.rpow_zero] + _ ≤ (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * ∏ i in s, 1 := by simp [h2i₀] + _ = (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * ∏ i in s, (∫⁻ a, f i a ∂μ) ^ p i := by + congr 1 + apply prod_congr rfl fun i hi ↦ by rw [h2p i hi, ENNReal.rpow_zero] + · have hpi₀ : 0 ≤ 1 - p i₀ := by + simp_rw [sub_nonneg, ← hp, single_le_sum h2p (mem_insert_self ..)] + have h2pi₀ : 1 - p i₀ ≠ 0 := by + rwa [sub_ne_zero, ne_comm] + let q := fun i ↦ p i / (1 - p i₀) + have hq : ∑ i in s, q i = 1 := by + rw [← Finset.sum_div, ← sum_insert_sub hi₀, hp, div_self h2pi₀] + have h2q : ∀ i ∈ s, 0 ≤ q i := + fun i hi ↦ div_nonneg (h2p i <| mem_insert_of_mem hi) hpi₀ + calc ∫⁻ a, ∏ i in insert i₀ s, f i a ^ p i ∂μ + = ∫⁻ a, f i₀ a ^ p i₀ * ∏ i in s, f i a ^ p i ∂μ := by simp [hi₀] + _ = ∫⁻ a, f i₀ a ^ p i₀ * (∏ i in s, f i a ^ q i) ^ (1 - p i₀) ∂μ := by + simp [← ENNReal.prod_rpow_of_nonneg hpi₀, ← ENNReal.rpow_mul, + div_mul_cancel (h := h2pi₀)] + _ ≤ (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * (∫⁻ a, ∏ i in s, f i a ^ q i ∂μ) ^ (1 - p i₀) := by + apply ENNReal.lintegral_mul_norm_pow_le + · exact hf i₀ <| mem_insert_self .. + · exact s.aemeasurable_prod <| fun i hi ↦ (hf i <| mem_insert_of_mem hi).pow_const _ + · exact h2p i₀ <| mem_insert_self .. + · exact hpi₀ + · apply add_sub_cancel'_right + _ ≤ (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * (∏ i in s, (∫⁻ a, f i a ∂μ) ^ q i) ^ (1 - p i₀) := by + gcongr -- behavior of gcongr is heartbeat-dependent, which makes code really fragile... + exact ih (fun i hi ↦ hf i <| mem_insert_of_mem hi) hq h2q + _ = (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * ∏ i in s, (∫⁻ a, f i a ∂μ) ^ p i := by + simp [← ENNReal.prod_rpow_of_nonneg hpi₀, ← ENNReal.rpow_mul, + div_mul_cancel (h := h2pi₀)] + _ = ∏ i in insert i₀ s, (∫⁻ a, f i a ∂μ) ^ p i := by simp [hi₀] + +/-- A version of Hölder with multiple arguments, one of which plays a distinguished role. -/ +theorem lintegral_mul_prod_norm_pow_le {α ι : Type*} [MeasurableSpace α] {μ : Measure α} + (s : Finset ι) {g : α → ℝ≥0∞} {f : ι → α → ℝ≥0∞} (hg : AEMeasurable g μ) + (hf : ∀ i ∈ s, AEMeasurable (f i) μ) (q : ℝ) {p : ι → ℝ} (hpq : q + ∑ i in s, p i = 1) + (hq : 0 ≤ q) (hp : ∀ i ∈ s, 0 ≤ p i) : + ∫⁻ a, g a ^ q * ∏ i in s, f i a ^ p i ∂μ ≤ + (∫⁻ a, g a ∂μ) ^ q * ∏ i in s, (∫⁻ a, f i a ∂μ) ^ p i := by + suffices + ∫⁻ t, ∏ j in insertNone s, Option.elim j (g t) (fun j ↦ f j t) ^ Option.elim j q p ∂μ + ≤ ∏ j in insertNone s, (∫⁻ t, Option.elim j (g t) (fun j ↦ f j t) ∂μ) ^ Option.elim j q p by + simpa using this + refine ENNReal.lintegral_prod_norm_pow_le _ ?_ ?_ ?_ + · rintro (_|i) hi + · exact hg + · refine hf i ?_ + simpa using hi + · simp_rw [sum_insertNone, Option.elim] + exact hpq + · rintro (_|i) hi + · exact hq + · refine hp i ?_ + simpa using hi + theorem lintegral_rpow_add_lt_top_of_lintegral_rpow_lt_top {p : ℝ} {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hf_top : (∫⁻ a, f a ^ p ∂μ) < ⊤) (hg_top : (∫⁻ a, g a ^ p ∂μ) < ⊤) (hp1 : 1 ≤ p) : (∫⁻ a, (f + g) a ^ p ∂μ) < ⊤ := by diff --git a/Mathlib/MeasureTheory/Integral/Periodic.lean b/Mathlib/MeasureTheory/Integral/Periodic.lean index e2e06207beda2..de9e66bcd7044 100644 --- a/Mathlib/MeasureTheory/Integral/Periodic.lean +++ b/Mathlib/MeasureTheory/Integral/Periodic.lean @@ -172,7 +172,7 @@ protected theorem integral_preimage (t : ℝ) (f : AddCircle T → E) : have := integral_map_equiv (μ := volume) (measurableEquivIoc T t).symm f simp only [measurableEquivIoc, equivIoc, QuotientAddGroup.equivIocMod, MeasurableEquiv.symm_mk, MeasurableEquiv.coe_mk, Equiv.coe_fn_symm_mk] at this - rw [← (AddCircle.measurePreserving_mk T t).map_eq, set_integral_eq_subtype m, ← this] + rw [← (AddCircle.measurePreserving_mk T t).map_eq, ← integral_subtype m, ← this] have : ((↑) : Ioc t (t + T) → AddCircle T) = ((↑) : ℝ → AddCircle T) ∘ ((↑) : _ → ℝ) := by ext1 x; rfl simp_rw [this] diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 3967342254245..1343285c8ca35 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -1260,10 +1260,10 @@ theorem integral_smul_const {𝕜 : Type*} [IsROrC 𝕜] [NormedSpace 𝕜 E] [C by_cases hf : Integrable f μ · exact ((1 : 𝕜 →L[𝕜] 𝕜).smulRight c).integral_comp_comm hf · by_cases hc : c = 0 - · simp only [hc, integral_zero, smul_zero] + · simp [hc, integral_zero, smul_zero] rw [integral_undef hf, integral_undef, zero_smul] rw [integrable_smul_const hc] - simp_rw [hf] + simp_rw [hf, not_false_eq_true] #align integral_smul_const integral_smul_const theorem integral_withDensity_eq_integral_smul {f : α → ℝ≥0} (f_meas : Measurable f) (g : α → E) : diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index 3329c2114b21a..6ad05c2e2f51a 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -44,6 +44,12 @@ theorem aemeasurable_id'' (μ : Measure α) {m : MeasurableSpace α} (hm : m ≤ @Measurable.aemeasurable α α m0 m id μ (measurable_id'' hm) #align probability_theory.ae_measurable_id'' aemeasurable_id'' +lemma aemeasurable_of_map_neZero {mβ : MeasurableSpace β} {μ : Measure α} + {f : α → β} (h : NeZero (μ.map f)) : + AEMeasurable f μ := by + by_contra h' + simp [h'] at h + namespace AEMeasurable lemma mono_ac (hf : AEMeasurable f ν) (hμν : μ ≪ ν) : AEMeasurable f μ := diff --git a/Mathlib/MeasureTheory/Measure/Dirac.lean b/Mathlib/MeasureTheory/Measure/Dirac.lean index da27cf72f3951..b1cec0a7a859e 100644 --- a/Mathlib/MeasureTheory/Measure/Dirac.lean +++ b/Mathlib/MeasureTheory/Measure/Dirac.lean @@ -63,6 +63,16 @@ theorem map_dirac {f : α → β} (hf : Measurable f) (a : α) : (dirac a).map f ext fun s hs => by simp [hs, map_apply hf hs, hf hs, indicator_apply] #align measure_theory.measure.map_dirac MeasureTheory.Measure.map_dirac +lemma map_const (μ : Measure α) (c : β) : μ.map (fun _ ↦ c) = (μ Set.univ) • dirac c := by + ext s hs + simp only [aemeasurable_const, measurable_const, smul_toOuterMeasure, OuterMeasure.coe_smul, + Pi.smul_apply, dirac_apply' _ hs, smul_eq_mul] + classical + rw [Measure.map_apply measurable_const hs, Set.preimage_const] + by_cases hsc : c ∈ s + · rw [(Set.indicator_eq_one_iff_mem _).mpr hsc, mul_one, if_pos hsc] + · rw [if_neg hsc, (Set.indicator_eq_zero_iff_not_mem _).mpr hsc, measure_empty, mul_zero] + @[simp] theorem restrict_singleton (μ : Measure α) (a : α) : μ.restrict {a} = μ {a} • dirac a := by ext1 s hs diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean index 2b1b3e34432bf..738b49f5f5172 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean @@ -732,12 +732,12 @@ lemma map_apply_of_aemeasurable (ν : FiniteMeasure Ω) {f : Ω → Ω'} (f_aemb ν.map f A = ν (f ⁻¹' A) := map_apply_of_aemeasurable ν f_mble.aemeasurable A_mble -@[simp] lemma map_add {f : Ω → Ω'} (f_mble : Measurable f) (ν₁ ν₂ : FiniteMeasure Ω) : +@[simp] lemma map_add {f : Ω → Ω'} (f_mble : Measurable f) (ν₁ ν₂ : FiniteMeasure Ω) : (ν₁ + ν₂).map f = ν₁.map f + ν₂.map f := by ext s s_mble simp [map_apply' _ f_mble.aemeasurable s_mble, toMeasure_add] -@[simp] lemma map_smul {f : Ω → Ω'} (f_mble : Measurable f) (c : ℝ≥0) (ν : FiniteMeasure Ω) : +@[simp] lemma map_smul {f : Ω → Ω'} (f_mble : Measurable f) (c : ℝ≥0) (ν : FiniteMeasure Ω) : (c • ν).map f = c • (ν.map f) := by ext s s_mble simp [map_apply' _ f_mble.aemeasurable s_mble, toMeasure_smul] diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index 49ee1df879122..6858344d54368 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -665,7 +665,7 @@ instance isHaarMeasure_haarMeasure (K₀ : PositiveCompacts G) : IsHaarMeasure ( isHaarMeasure_of_isCompact_nonempty_interior (haarMeasure K₀) K₀ K₀.isCompact K₀.interior_nonempty · simp only [haarMeasure_self]; exact one_ne_zero - · simp only [haarMeasure_self] + · simp only [haarMeasure_self, ne_eq, ENNReal.one_ne_top, not_false_eq_true] #align measure_theory.measure.is_haar_measure_haar_measure MeasureTheory.Measure.isHaarMeasure_haarMeasure #align measure_theory.measure.is_add_haar_measure_add_haar_measure MeasureTheory.Measure.isAddHaarMeasure_addHaarMeasure diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean index 1779f92122ed7..5a4ccf2b9cf7f 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean @@ -104,3 +104,30 @@ theorem integral_comp_neg_Ioi {E : Type*} [NormedAddCommGroup E] [NormedSpace rw [← neg_neg c, ← integral_comp_neg_Iic] simp only [neg_neg] #align integral_comp_neg_Ioi integral_comp_neg_Ioi + +theorem integral_comp_abs {f : ℝ → ℝ} : + ∫ x, f |x| = 2 * ∫ x in Ioi (0:ℝ), f x := by + have eq : ∫ (x : ℝ) in Ioi 0, f |x| = ∫ (x : ℝ) in Ioi 0, f x := by + refine set_integral_congr measurableSet_Ioi (fun _ hx => ?_) + rw [abs_eq_self.mpr (le_of_lt (by exact hx))] + by_cases hf : IntegrableOn (fun x => f |x|) (Ioi 0) + · have int_Iic : IntegrableOn (fun x ↦ f |x|) (Iic 0) := by + rw [← Measure.map_neg_eq_self (volume : Measure ℝ)] + let m : MeasurableEmbedding fun x : ℝ => -x := (Homeomorph.neg ℝ).measurableEmbedding + rw [m.integrableOn_map_iff] + simp_rw [Function.comp, abs_neg, neg_preimage, preimage_neg_Iic, neg_zero] + exact integrableOn_Ici_iff_integrableOn_Ioi.mpr hf + calc + _ = (∫ x in Iic 0, f |x|) + ∫ x in Ioi 0, f |x| := by + rw [← integral_union (Iic_disjoint_Ioi le_rfl) measurableSet_Ioi int_Iic hf, + Iic_union_Ioi, restrict_univ] + _ = 2 * ∫ x in Ioi 0, f x := by + rw [two_mul, eq] + congr! 1 + rw [← neg_zero, ← integral_comp_neg_Iic, neg_zero] + refine set_integral_congr measurableSet_Iic (fun _ hx => ?_) + rw [abs_eq_neg_self.mpr (by exact hx)] + · have : ¬ Integrable (fun x => f |x|) := by + contrapose! hf + exact hf.integrableOn + rw [← eq, integral_undef hf, integral_undef this, mul_zero] diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 5cba3addb58eb..82c7cf8d226fa 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -418,8 +418,7 @@ theorem exists_nonempty_inter_of_measure_univ_lt_tsum_measure {m : MeasurableSpa contrapose! H apply tsum_measure_le_measure_univ hs intro i j hij - rw [Function.onFun, disjoint_iff_inf_le] - exact fun x hx => H i j hij ⟨x, hx⟩ + exact disjoint_iff_inter_eq_empty.mpr (H i j hij) #align measure_theory.exists_nonempty_inter_of_measure_univ_lt_tsum_measure MeasureTheory.exists_nonempty_inter_of_measure_univ_lt_tsum_measure /-- Pigeonhole principle for measure spaces: if `s` is a `Finset` and @@ -431,8 +430,7 @@ theorem exists_nonempty_inter_of_measure_univ_lt_sum_measure {m : MeasurableSpac contrapose! H apply sum_measure_le_measure_univ h intro i hi j hj hij - rw [Function.onFun, disjoint_iff_inf_le] - exact fun x hx => H i hi j hj hij ⟨x, hx⟩ + exact disjoint_iff_inter_eq_empty.mpr (H i hi j hj hij) #align measure_theory.exists_nonempty_inter_of_measure_univ_lt_sum_measure MeasureTheory.exists_nonempty_inter_of_measure_univ_lt_sum_measure /-- If two sets `s` and `t` are included in a set `u`, and `μ s + μ t > μ u`, @@ -1184,6 +1182,7 @@ theorem map_zero (f : α → β) : (0 : Measure α).map f = 0 := by by_cases hf : AEMeasurable f (0 : Measure α) <;> simp [map, hf] #align measure_theory.measure.map_zero MeasureTheory.Measure.map_zero +@[simp] theorem map_of_not_aemeasurable {f : α → β} {μ : Measure α} (hf : ¬AEMeasurable f μ) : μ.map f = 0 := by simp [map, hf] #align measure_theory.measure.map_of_not_ae_measurable MeasureTheory.Measure.map_of_not_aemeasurable @@ -1579,7 +1578,6 @@ theorem restrict_le_self : μ.restrict s ≤ μ := fun t ht => calc μ.restrict s t = μ (t ∩ s) := restrict_apply ht _ ≤ μ t := measure_mono <| inter_subset_left t s - #align measure_theory.measure.restrict_le_self MeasureTheory.Measure.restrict_le_self variable (μ) @@ -1591,7 +1589,6 @@ theorem restrict_eq_self (h : s ⊆ t) : μ.restrict t s = μ s := measure_mono (subset_inter (subset_toMeasurable _ _) h) _ = μ.restrict t s := by rw [← restrict_apply (measurableSet_toMeasurable _ _), measure_toMeasurable] - #align measure_theory.measure.restrict_eq_self MeasureTheory.Measure.restrict_eq_self @[simp] @@ -1609,7 +1606,6 @@ theorem le_restrict_apply (s t : Set α) : μ (t ∩ s) ≤ μ.restrict s t := calc μ (t ∩ s) = μ.restrict s (t ∩ s) := (restrict_eq_self μ (inter_subset_right _ _)).symm _ ≤ μ.restrict s t := measure_mono (inter_subset_left _ _) - #align measure_theory.measure.le_restrict_apply MeasureTheory.Measure.le_restrict_apply theorem restrict_apply_superset (h : s ⊆ t) : μ.restrict s t = μ s := @@ -1803,7 +1799,6 @@ theorem restrict_eq_self_of_ae_mem {_m0 : MeasurableSpace α} ⦃s : Set α⦄ calc μ.restrict s = μ.restrict univ := restrict_congr_set (eventuallyEq_univ.mpr hs) _ = μ := restrict_univ - #align measure_theory.measure.restrict_eq_self_of_ae_mem MeasureTheory.Measure.restrict_eq_self_of_ae_mem theorem restrict_congr_meas (hs : MeasurableSet s) : @@ -3110,8 +3105,8 @@ attribute [simp] measure_singleton variable [NoAtoms μ] -theorem _root_.Set.Subsingleton.measure_zero {α : Type*} {_m : MeasurableSpace α} {s : Set α} - (hs : s.Subsingleton) (μ : Measure α) [NoAtoms μ] : μ s = 0 := +theorem _root_.Set.Subsingleton.measure_zero (hs : s.Subsingleton) (μ : Measure α) [NoAtoms μ] : + μ s = 0 := hs.induction_on (p := fun s => μ s = 0) measure_empty measure_singleton #align set.subsingleton.measure_zero Set.Subsingleton.measure_zero @@ -3127,25 +3122,31 @@ instance Measure.restrict.instNoAtoms (s : Set α) : NoAtoms (μ.restrict s) := apply measure_mono_null (inter_subset_left t s) ht2 #align measure_theory.measure.restrict.has_no_atoms MeasureTheory.Measure.restrict.instNoAtoms -theorem _root_.Set.Countable.measure_zero {α : Type*} {m : MeasurableSpace α} {s : Set α} - (h : s.Countable) (μ : Measure α) [NoAtoms μ] : μ s = 0 := by +theorem _root_.Set.Countable.measure_zero (h : s.Countable) (μ : Measure α) [NoAtoms μ] : + μ s = 0 := by rw [← biUnion_of_singleton s, ← nonpos_iff_eq_zero] refine' le_trans (measure_biUnion_le h _) _ simp #align set.countable.measure_zero Set.Countable.measure_zero -theorem _root_.Set.Countable.ae_not_mem {α : Type*} {m : MeasurableSpace α} {s : Set α} - (h : s.Countable) (μ : Measure α) [NoAtoms μ] : ∀ᵐ x ∂μ, x ∉ s := by +theorem _root_.Set.Countable.ae_not_mem (h : s.Countable) (μ : Measure α) [NoAtoms μ] : + ∀ᵐ x ∂μ, x ∉ s := by simpa only [ae_iff, Classical.not_not] using h.measure_zero μ #align set.countable.ae_not_mem Set.Countable.ae_not_mem -theorem _root_.Set.Finite.measure_zero {α : Type*} {_m : MeasurableSpace α} {s : Set α} - (h : s.Finite) (μ : Measure α) [NoAtoms μ] : μ s = 0 := +lemma _root_.Set.Countable.measure_restrict_compl (h : s.Countable) (μ : Measure α) [NoAtoms μ] : + μ.restrict sᶜ = μ := + restrict_eq_self_of_ae_mem <| h.ae_not_mem μ + +@[simp] +lemma restrict_compl_singleton (a : α) : μ.restrict ({a}ᶜ) = μ := + (countable_singleton _).measure_restrict_compl μ + +theorem _root_.Set.Finite.measure_zero (h : s.Finite) (μ : Measure α) [NoAtoms μ] : μ s = 0 := h.countable.measure_zero μ #align set.finite.measure_zero Set.Finite.measure_zero -theorem _root_.Finset.measure_zero {α : Type*} {_m : MeasurableSpace α} (s : Finset α) - (μ : Measure α) [NoAtoms μ] : μ s = 0 := +theorem _root_.Finset.measure_zero (s : Finset α) (μ : Measure α) [NoAtoms μ] : μ s = 0 := s.finite_toSet.measure_zero μ #align finset.measure_zero Finset.measure_zero @@ -3260,7 +3261,7 @@ theorem FiniteAtFilter.exists_mem_basis {f : Filter α} (hμ : FiniteAtFilter μ #align measure_theory.measure.finite_at_filter.exists_mem_basis MeasureTheory.Measure.FiniteAtFilter.exists_mem_basis theorem finiteAtBot {m0 : MeasurableSpace α} (μ : Measure α) : μ.FiniteAtFilter ⊥ := - ⟨∅, mem_bot, by simp only [measure_empty, WithTop.zero_lt_top]⟩ + ⟨∅, mem_bot, by simp only [measure_empty, zero_lt_top]⟩ #align measure_theory.measure.finite_at_bot MeasureTheory.Measure.finiteAtBot /-- `μ` has finite spanning sets in `C` if there is a countable sequence of sets in `C` that have @@ -3824,8 +3825,7 @@ theorem Measure.smul_finite (μ : Measure α) [IsFiniteMeasure μ] {c : ℝ≥0 theorem Measure.exists_isOpen_measure_lt_top [TopologicalSpace α] (μ : Measure α) [IsLocallyFiniteMeasure μ] (x : α) : ∃ s : Set α, x ∈ s ∧ IsOpen s ∧ μ s < ∞ := by - simpa only [exists_prop, and_assoc] using - (μ.finiteAt_nhds x).exists_mem_basis (nhds_basis_opens x) + simpa only [and_assoc] using (μ.finiteAt_nhds x).exists_mem_basis (nhds_basis_opens x) #align measure_theory.measure.exists_is_open_measure_lt_top MeasureTheory.Measure.exists_isOpen_measure_lt_top instance isLocallyFiniteMeasureSMulNNReal [TopologicalSpace α] (μ : Measure α) @@ -4128,9 +4128,12 @@ theorem comap_apply (μ : Measure β) (s : Set α) : comap f μ s = μ (f '' s) _ = μ (f '' s) := by rw [hf.map_comap, restrict_apply' hf.measurableSet_range, inter_eq_self_of_subset_left (image_subset_range _ _)] - #align measurable_embedding.comap_apply MeasurableEmbedding.comap_apply +theorem comap_map (μ : Measure α) : (map f μ).comap f = μ := by + ext t _ + rw [hf.comap_apply, hf.map_apply, preimage_image_eq _ hf.injective] + theorem ae_map_iff {p : β → Prop} {μ : Measure α} : (∀ᵐ x ∂μ.map f, p x) ↔ ∀ᵐ x ∂μ, p (f x) := by simp only [ae_iff, hf.map_apply, preimage_setOf_eq] #align measurable_embedding.ae_map_iff MeasurableEmbedding.ae_map_iff @@ -4140,12 +4143,21 @@ theorem restrict_map (μ : Measure α) (s : Set β) : Measure.ext fun t ht => by simp [hf.map_apply, ht, hf.measurable ht] #align measurable_embedding.restrict_map MeasurableEmbedding.restrict_map -protected theorem comap_preimage (μ : Measure β) {s : Set β} (hs : MeasurableSet s) : - μ.comap f (f ⁻¹' s) = μ (s ∩ range f) := - comap_preimage _ _ hf.injective hf.measurable - (fun _t ht => (hf.measurableSet_image' ht).nullMeasurableSet) hs +protected theorem comap_preimage (μ : Measure β) (s : Set β) : + μ.comap f (f ⁻¹' s) = μ (s ∩ range f) := by + rw [← hf.map_apply, hf.map_comap, restrict_apply' hf.measurableSet_range] #align measurable_embedding.comap_preimage MeasurableEmbedding.comap_preimage +lemma comap_restrict (μ : Measure β) (s : Set β) : + (μ.restrict s).comap f = (μ.comap f).restrict (f ⁻¹' s) := by + ext t ht + rw [Measure.restrict_apply ht, comap_apply hf, comap_apply hf, + Measure.restrict_apply (hf.measurableSet_image.2 ht), image_inter_preimage] + +lemma restrict_comap (μ : Measure β) (s : Set α) : + (μ.comap f).restrict s = (μ.restrict (f '' s)).comap f := by + rw [comap_restrict hf, preimage_image_eq _ hf.injective] + end MeasurableEmbedding section Subtype diff --git a/Mathlib/MeasureTheory/Measure/MutuallySingular.lean b/Mathlib/MeasureTheory/Measure/MutuallySingular.lean index b834dca224cd2..37efb5e7a5e65 100644 --- a/Mathlib/MeasureTheory/Measure/MutuallySingular.lean +++ b/Mathlib/MeasureTheory/Measure/MutuallySingular.lean @@ -82,6 +82,14 @@ theorem mono (h : μ₁ ⟂ₘ ν₁) (hμ : μ₂ ≤ μ₁) (hν : ν₂ ≤ h.mono_ac hμ.absolutelyContinuous hν.absolutelyContinuous #align measure_theory.measure.mutually_singular.mono MeasureTheory.Measure.MutuallySingular.mono +@[simp] +lemma self_iff (μ : Measure α) : μ ⟂ₘ μ ↔ μ = 0 := by + refine ⟨?_, fun h ↦ by (rw [h]; exact zero_left)⟩ + rintro ⟨s, hs, hμs, hμs_compl⟩ + suffices μ Set.univ = 0 by rwa [measure_univ_eq_zero] at this + rw [← Set.union_compl_self s, measure_union disjoint_compl_right hs.compl, hμs, hμs_compl, + add_zero] + @[simp] theorem sum_left {ι : Type*} [Countable ι] {μ : ι → Measure α} : sum μ ⟂ₘ ν ↔ ∀ i, μ i ⟂ₘ ν := by refine' ⟨fun h i => h.mono (le_sum _ _) le_rfl, fun H => _⟩ diff --git a/Mathlib/MeasureTheory/Measure/WithDensity.lean b/Mathlib/MeasureTheory/Measure/WithDensity.lean index 85d551c2f5108..fb140a8f2c598 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensity.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensity.lean @@ -186,6 +186,11 @@ theorem withDensity_eq_zero {f : α → ℝ≥0∞} (hf : AEMeasurable f μ) (h h, Measure.coe_zero, Pi.zero_apply] #align measure_theory.with_density_eq_zero MeasureTheory.withDensity_eq_zero +@[simp] +theorem withDensity_eq_zero_iff {f : α → ℝ≥0∞} (hf : AEMeasurable f μ) : + μ.withDensity f = 0 ↔ f =ᵐ[μ] 0 := + ⟨withDensity_eq_zero hf, fun h => withDensity_zero (μ := μ) ▸ withDensity_congr_ae h⟩ + theorem withDensity_apply_eq_zero {f : α → ℝ≥0∞} {s : Set α} (hf : Measurable f) : μ.withDensity f s = 0 ↔ μ ({ x | f x ≠ 0 } ∩ s) = 0 := by constructor diff --git a/Mathlib/NumberTheory/Bernoulli.lean b/Mathlib/NumberTheory/Bernoulli.lean index 4184598955f54..8de2a7d4d5ae6 100644 --- a/Mathlib/NumberTheory/Bernoulli.lean +++ b/Mathlib/NumberTheory/Bernoulli.lean @@ -125,7 +125,7 @@ theorem bernoulli'_three : bernoulli' 3 = 0 := by @[simp] theorem bernoulli'_four : bernoulli' 4 = -1 / 30 := by - have : Nat.choose 4 2 = 6 := by norm_num -- shrug + have : Nat.choose 4 2 = 6 := by decide -- shrug rw [bernoulli'_def] norm_num [sum_range_succ, sum_range_succ, sum_range_zero, this] #align bernoulli'_four bernoulli'_four @@ -185,7 +185,7 @@ theorem bernoulli'_odd_eq_zero {n : ℕ} (h_odd : Odd n) (hlt : 1 < n) : bernoul · apply eq_zero_of_neg_eq specialize h n split_ifs at h <;> simp_all [h_odd.neg_one_pow, factorial_ne_zero] - · simpa [Nat.factorial] using h 1 + · simpa (config := {decide := true}) [Nat.factorial] using h 1 have h : B * (exp ℚ - 1) = X * exp ℚ := by simpa [bernoulli'PowerSeries] using bernoulli'PowerSeries_mul_exp_sub_one ℚ rw [sub_mul, h, mul_sub X, sub_right_inj, ← neg_sub, mul_neg, neg_eq_iff_eq_neg] @@ -349,7 +349,7 @@ theorem sum_range_pow (n p : ℕ) : have hexp : exp ℚ - 1 ≠ 0 := by simp only [exp, PowerSeries.ext_iff, Ne, not_forall] use 1 - simp + simp [factorial_ne_zero] have h_r : exp ℚ ^ n - 1 = X * mk fun p => coeff ℚ (p + 1) (exp ℚ ^ n) := by have h_const : C ℚ (constantCoeff ℚ (exp ℚ ^ n)) = 1 := by simp rw [← h_const, sub_const_eq_X_mul_shift] diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index dfbad84e0bf80..72a54250182cf 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -328,7 +328,7 @@ theorem finite [IsDomain B] [h₁ : Finite S] [h₂ : IsCyclotomicExtension S A refine' Set.Finite.induction_on (Set.Finite.intro h) (fun A B => _) @fun n S _ _ H A B => _ · intro _ _ _ _ _ refine' Module.finite_def.2 ⟨({1} : Finset B), _⟩ - simp [← top_toSubmodule, ← empty, toSubmodule_bot] + simp [← top_toSubmodule, ← empty, toSubmodule_bot, Submodule.one_eq_span] · intro _ _ _ _ h haveI : IsCyclotomicExtension S A (adjoin A {b : B | ∃ n : ℕ+, n ∈ S ∧ b ^ (n : ℕ) = 1}) := union_left _ (insert n S) _ _ (subset_insert n S) diff --git a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean index 83c3817c8cbaf..17d08ac9eb5f7 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean @@ -176,7 +176,7 @@ theorem discr_prime_pow [hcycl : IsCyclotomicExtension {p ^ k} K L] [hp : Fact ( simp only [discr, traceMatrix_apply, Matrix.det_unique, Fin.default_eq_zero, Fin.val_zero, _root_.pow_zero, traceForm_apply, mul_one] rw [← (algebraMap K L).map_one, trace_algebraMap, finrank _ hirr, hp, hk]; norm_num - simp [← coe_two] + simp [← coe_two, Even.neg_pow (by decide : Even (1 / 2))] · exact discr_prime_pow_ne_two hζ hirr hk #align is_cyclotomic_extension.discr_prime_pow IsCyclotomicExtension.discr_prime_pow diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index c1560a04efc68..5b2b78cc50727 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -64,6 +64,7 @@ and only at the "final step", when we need to provide an "explicit" primitive ro open Polynomial Algebra Finset FiniteDimensional IsCyclotomicExtension Nat PNat Set +open scoped IntermediateField universe u v w z @@ -468,7 +469,7 @@ theorem pow_sub_one_norm_prime_pow_of_ne_zero {k s : ℕ} (hζ : IsPrimitiveRoot rw [← PNat.coe_inj, PNat.pow_coe, ← pow_one 2] at htwo replace htwo := eq_of_prime_pow_eq (prime_iff.1 hpri.out) (prime_iff.1 Nat.prime_two) (succ_pos _) htwo - rwa [show 2 = ((2 : ℕ+) : ℕ) by simp, PNat.coe_inj] at htwo + rwa [show 2 = ((2 : ℕ+) : ℕ) by decide, PNat.coe_inj] at htwo replace hs : s = k · rw [hp, ← PNat.coe_inj, PNat.pow_coe] at htwo nth_rw 2 [← pow_one 2] at htwo diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 6ab57f89f54c3..7a90aaed0a596 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -404,7 +404,8 @@ theorem properDivisors_eq_singleton_one_iff_prime : n.properDivisors = {1} ↔ n · simp [hdvd, this] exact (le_iff_eq_or_lt.mp this).symm · by_contra' - simp [nonpos_iff_eq_zero.mp this, this] at h + simp only [nonpos_iff_eq_zero.mp this, this] at h + contradiction · exact fun h => Prime.properDivisors h #align nat.proper_divisors_eq_singleton_one_iff_prime Nat.properDivisors_eq_singleton_one_iff_prime diff --git a/Mathlib/NumberTheory/FLT/Four.lean b/Mathlib/NumberTheory/FLT/Four.lean index ab76f84eefc1b..20cde9384f650 100644 --- a/Mathlib/NumberTheory/FLT/Four.lean +++ b/Mathlib/NumberTheory/FLT/Four.lean @@ -99,7 +99,7 @@ theorem coprime_of_minimal {a b c : ℤ} (h : Minimal a b c) : IsCoprime a b := obtain ⟨c1, rfl⟩ := hpc have hf : Fermat42 a1 b1 c1 := (Fermat42.mul (Int.coe_nat_ne_zero.mpr (Nat.Prime.ne_zero hp))).mpr h.1 - apply Nat.le_lt_antisymm (h.2 _ _ _ hf) + apply Nat.le_lt_asymm (h.2 _ _ _ hf) rw [Int.natAbs_mul, lt_mul_iff_one_lt_left, Int.natAbs_pow, Int.natAbs_ofNat] · exact Nat.one_lt_pow _ _ zero_lt_two (Nat.Prime.one_lt hp) · exact Nat.pos_of_ne_zero (Int.natAbs_ne_zero.2 (ne_zero hf)) @@ -131,7 +131,7 @@ theorem exists_odd_minimal {a b c : ℤ} (h : Fermat42 a b c) : Int.dvd_gcd (Int.dvd_of_emod_eq_zero hap) (Int.dvd_of_emod_eq_zero hbp) rw [Int.gcd_eq_one_iff_coprime.mpr (coprime_of_minimal hf)] at h1 revert h1 - norm_num + decide · exact ⟨b0, ⟨a0, ⟨c0, minimal_comm hf, hbp⟩⟩⟩ exact ⟨a0, ⟨b0, ⟨c0, hf, hap⟩⟩⟩ #align fermat_42.exists_odd_minimal Fermat42.exists_odd_minimal @@ -179,7 +179,7 @@ theorem not_minimal {a b c : ℤ} (h : Minimal a b c) (ha2 : a % 2 = 1) (hc : 0 -- it helps if we know the parity of a ^ 2 (and the sign of c): have ha22 : a ^ 2 % 2 = 1 := by rw [sq, Int.mul_emod, ha2] - norm_num + decide obtain ⟨m, n, ht1, ht2, ht3, ht4, ht5, ht6⟩ := ht.coprime_classification' h2 ha22 hc -- Now a, n, m form a pythagorean triple and so we can obtain r and s such that -- a = r ^ 2 - s ^ 2, n = 2 * r * s and m = r ^ 2 + s ^ 2 diff --git a/Mathlib/NumberTheory/LSeries.lean b/Mathlib/NumberTheory/LSeries.lean index 3fedfcbe0a775..34ac5a05414e7 100644 --- a/Mathlib/NumberTheory/LSeries.lean +++ b/Mathlib/NumberTheory/LSeries.lean @@ -117,7 +117,7 @@ theorem zeta_LSeriesSummable_iff_one_lt_re {z : ℂ} : LSeriesSummable ζ z ↔ · simp [h0] simp only [cast_zero, natCoe_apply, zeta_apply, succ_ne_zero, if_false, cast_succ, one_div, Complex.norm_eq_abs, map_inv₀, Complex.abs_cpow_real, inv_inj, zero_add] - rw [← cast_one, ← cast_add, Complex.abs_of_nat, cast_add, cast_one] + rw [← cast_one, ← cast_add, Complex.abs_natCast, cast_add, cast_one] #align nat.arithmetic_function.zeta_l_series_summable_iff_one_lt_re Nat.ArithmeticFunction.zeta_LSeriesSummable_iff_one_lt_re @[simp] diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index f008ccd93a3a8..ffd1d3f8e681d 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -5,7 +5,7 @@ Authors: Chris Hughes -/ import Mathlib.NumberTheory.LegendreSymbol.QuadraticReciprocity -#align_import number_theory.legendre_symbol.gauss_eisenstein_lemmas from "leanprover-community/mathlib"@"442a83d738cb208d3600056c489be16900ba701d" +#align_import number_theory.legendre_symbol.gauss_eisenstein_lemmas from "leanprover-community/mathlib"@"8818fdefc78642a7e6afcd20be5c184f3c7d9699" /-! # Lemmas of Gauss and Eisenstein diff --git a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean index 29e832b2e7e19..7b15dceee2531 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean @@ -468,7 +468,7 @@ theorem quadratic_reciprocity_three_mod_four {a b : ℕ} (ha : a % 4 = 3) (hb : theorem mod_right' (a : ℕ) {b : ℕ} (hb : Odd b) : J(a | b) = J(a | b % (4 * a)) := by rcases eq_or_ne a 0 with (rfl | ha₀) · rw [mul_zero, mod_zero] - have hb' : Odd (b % (4 * a)) := hb.mod_even (Even.mul_right (by norm_num) _) + have hb' : Odd (b % (4 * a)) := hb.mod_even (Even.mul_right (by decide) _) rcases exists_eq_pow_mul_and_not_dvd ha₀ 2 (by norm_num) with ⟨e, a', ha₁', ha₂⟩ have ha₁ := odd_iff.mpr (two_dvd_ne_zero.mp ha₁') nth_rw 2 [ha₂]; nth_rw 1 [ha₂] @@ -493,7 +493,7 @@ theorem mod_right' (a : ℕ) {b : ℕ} (hb : Odd b) : J(a | b) = J(a | b % (4 * theorem mod_right (a : ℤ) {b : ℕ} (hb : Odd b) : J(a | b) = J(a | b % (4 * a.natAbs)) := by cases' Int.natAbs_eq a with ha ha <;> nth_rw 2 [ha] <;> nth_rw 1 [ha] · exact mod_right' a.natAbs hb - · have hb' : Odd (b % (4 * a.natAbs)) := hb.mod_even (Even.mul_right (by norm_num) _) + · have hb' : Odd (b % (4 * a.natAbs)) := hb.mod_even (Even.mul_right (by decide) _) rw [jacobiSym.neg _ hb, jacobiSym.neg _ hb', mod_right' _ hb, χ₄_nat_mod_four, χ₄_nat_mod_four (b % (4 * _)), mod_mod_of_dvd b (dvd_mul_right 4 _)] #align jacobi_sym.mod_right jacobiSym.mod_right diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index 4a608d0f1874e..aa50a12f53dbb 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -85,7 +85,7 @@ theorem quadraticCharFun_one : quadraticCharFun F 1 = 1 := by theorem quadraticCharFun_eq_one_of_char_two (hF : ringChar F = 2) {a : F} (ha : a ≠ 0) : quadraticCharFun F a = 1 := by simp only [quadraticCharFun, ha, if_false, ite_eq_left_iff] - exact fun h => h (FiniteField.isSquare_of_char_two hF a) + exact fun h => (h (FiniteField.isSquare_of_char_two hF a)).elim #align quadratic_char_fun_eq_one_of_char_two quadraticCharFun_eq_one_of_char_two /-- If `ringChar F` is odd, then `quadraticCharFun F a` can be computed in diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean index f7d3bddc79562..3fe7b1f41e066 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean @@ -54,7 +54,7 @@ theorem FiniteField.isSquare_two_iff : simp only [h, Nat.one_ne_zero, if_false, ite_eq_left_iff, Ne.def, (by decide : (-1 : ℤ) ≠ 1), imp_false, Classical.not_not] all_goals - rw [← Nat.mod_mod_of_dvd _ (by norm_num : 2 ∣ 8)] at h + rw [← Nat.mod_mod_of_dvd _ (by decide : 2 ∣ 8)] at h have h₁ := Nat.mod_lt (Fintype.card F) (by decide : 0 < 8) revert h₁ h generalize Fintype.card F % 8 = n @@ -65,7 +65,7 @@ theorem FiniteField.isSquare_two_iff : theorem quadraticChar_neg_two [DecidableEq F] (hF : ringChar F ≠ 2) : quadraticChar F (-2) = χ₈' (Fintype.card F) := by rw [(by norm_num : (-2 : F) = -1 * 2), map_mul, χ₈'_eq_χ₄_mul_χ₈, quadraticChar_neg_one hF, - quadraticChar_two hF, @cast_nat_cast _ (ZMod 4) _ _ _ (by norm_num : 4 ∣ 8)] + quadraticChar_two hF, @cast_nat_cast _ (ZMod 4) _ _ _ (by decide : 4 ∣ 8)] #align quadratic_char_neg_two quadraticChar_neg_two /-- `-2` is a square in `F` iff `#F` is not congruent to `5` or `7` mod `8`. -/ @@ -84,7 +84,7 @@ theorem FiniteField.isSquare_neg_two_iff : simp only [h, Nat.one_ne_zero, if_false, ite_eq_left_iff, Ne.def, (by decide : (-1 : ℤ) ≠ 1), imp_false, Classical.not_not] all_goals - rw [← Nat.mod_mod_of_dvd _ (by norm_num : 2 ∣ 8)] at h + rw [← Nat.mod_mod_of_dvd _ (by decide : 2 ∣ 8)] at h have h₁ := Nat.mod_lt (Fintype.card F) (by decide : 0 < 8) revert h₁ h generalize Fintype.card F % 8 = n diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean index 52d412108318c..36b611aecfb82 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean @@ -78,7 +78,7 @@ variable (hp : p ≠ 2) theorem exists_sq_eq_two_iff : IsSquare (2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 7 := by rw [FiniteField.isSquare_two_iff, card p] have h₁ := Prime.mod_two_eq_one_iff_ne_two.mpr hp - rw [← mod_mod_of_dvd p (by norm_num : 2 ∣ 8)] at h₁ + rw [← mod_mod_of_dvd p (by decide : 2 ∣ 8)] at h₁ have h₂ := mod_lt p (by norm_num : 0 < 8) revert h₂ h₁ generalize p % 8 = m; clear! p @@ -89,7 +89,7 @@ theorem exists_sq_eq_two_iff : IsSquare (2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 7 theorem exists_sq_eq_neg_two_iff : IsSquare (-2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 3 := by rw [FiniteField.isSquare_neg_two_iff, card p] have h₁ := Prime.mod_two_eq_one_iff_ne_two.mpr hp - rw [← mod_mod_of_dvd p (by norm_num : 2 ∣ 8)] at h₁ + rw [← mod_mod_of_dvd p (by decide : 2 ∣ 8)] at h₁ have h₂ := mod_lt p (by norm_num : 0 < 8) revert h₂ h₁ generalize p % 8 = m; clear! p diff --git a/Mathlib/NumberTheory/LegendreSymbol/ZModChar.lean b/Mathlib/NumberTheory/LegendreSymbol/ZModChar.lean index 7b7731b54bbf6..61dc2f08096cd 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/ZModChar.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/ZModChar.lean @@ -67,7 +67,7 @@ theorem χ₄_int_eq_if_mod_four (n : ℤ) : χ₄ n = if n % 2 = 0 then 0 else if n % 4 = 1 then 1 else -1 := by have help : ∀ m : ℤ, 0 ≤ m → m < 4 → χ₄ m = if m % 2 = 0 then 0 else if m = 1 then 1 else -1 := by decide - rw [← Int.emod_emod_of_dvd n (by norm_num : (2 : ℤ) ∣ 4), ← ZMod.int_cast_mod n 4] + rw [← Int.emod_emod_of_dvd n (by decide : (2 : ℤ) ∣ 4), ← ZMod.int_cast_mod n 4] exact help (n % 4) (Int.emod_nonneg n (by norm_num)) (Int.emod_lt n (by norm_num)) #align zmod.χ₄_int_eq_if_mod_four ZMod.χ₄_int_eq_if_mod_four @@ -88,7 +88,7 @@ theorem χ₄_eq_neg_one_pow {n : ℕ} (hn : n % 2 = 1) : χ₄ n = (-1) ^ (n / have help : ∀ m : ℕ, m < 4 → m % 2 = 1 → ite (m = 1) (1 : ℤ) (-1) = (-1) ^ (m / 2) := by decide exact help (n % 4) (Nat.mod_lt n (by norm_num)) - ((Nat.mod_mod_of_dvd n (by norm_num : 2 ∣ 4)).trans hn) + ((Nat.mod_mod_of_dvd n (by decide : 2 ∣ 4)).trans hn) #align zmod.χ₄_eq_neg_one_pow ZMod.χ₄_eq_neg_one_pow /-- If `n % 4 = 1`, then `χ₄ n = 1`. -/ @@ -162,7 +162,7 @@ theorem χ₈_int_eq_if_mod_eight (n : ℤ) : have help : ∀ m : ℤ, 0 ≤ m → m < 8 → χ₈ m = if m % 2 = 0 then 0 else if m = 1 ∨ m = 7 then 1 else -1 := by decide - rw [← Int.emod_emod_of_dvd n (by norm_num : (2 : ℤ) ∣ 8), ← ZMod.int_cast_mod n 8] + rw [← Int.emod_emod_of_dvd n (by decide : (2 : ℤ) ∣ 8), ← ZMod.int_cast_mod n 8] exact help (n % 8) (Int.emod_nonneg n (by norm_num)) (Int.emod_lt n (by norm_num)) #align zmod.χ₈_int_eq_if_mod_eight ZMod.χ₈_int_eq_if_mod_eight @@ -195,7 +195,7 @@ theorem χ₈'_int_eq_if_mod_eight (n : ℤ) : have help : ∀ m : ℤ, 0 ≤ m → m < 8 → χ₈' m = if m % 2 = 0 then 0 else if m = 1 ∨ m = 3 then 1 else -1 := by decide - rw [← Int.emod_emod_of_dvd n (by norm_num : (2 : ℤ) ∣ 8), ← ZMod.int_cast_mod n 8] + rw [← Int.emod_emod_of_dvd n (by decide : (2 : ℤ) ∣ 8), ← ZMod.int_cast_mod n 8] exact help (n % 8) (Int.emod_nonneg n (by norm_num)) (Int.emod_lt n (by norm_num)) #align zmod.χ₈'_int_eq_if_mod_eight ZMod.χ₈'_int_eq_if_mod_eight @@ -212,7 +212,7 @@ theorem χ₈'_eq_χ₄_mul_χ₈ (a : ZMod 8) : χ₈' a = χ₄ a * χ₈ a := #align zmod.χ₈'_eq_χ₄_mul_χ₈ ZMod.χ₈'_eq_χ₄_mul_χ₈ theorem χ₈'_int_eq_χ₄_mul_χ₈ (a : ℤ) : χ₈' a = χ₄ a * χ₈ a := by - rw [← @cast_int_cast 8 (ZMod 4) _ 4 _ (by norm_num) a] + rw [← @cast_int_cast 8 (ZMod 4) _ 4 _ (by decide) a] exact χ₈'_eq_χ₄_mul_χ₈ a #align zmod.χ₈'_int_eq_χ₄_mul_χ₈ ZMod.χ₈'_int_eq_χ₄_mul_χ₈ diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index e157f1c4438e8..00d9089a44505 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -114,7 +114,7 @@ theorem odd_sq_dvd_geom_sum₂_sub (hp : Odd p) : have : ∀ (x : ℕ), (hx : x ∈ range p) → a ^ (x + (p - 1 - x)) = a ^ (p - 1) := by intro x hx rw [← Nat.add_sub_assoc _ x, Nat.add_sub_cancel_left] - exact Nat.le_pred_of_lt (Finset.mem_range.mp hx) + exact Nat.le_sub_one_of_lt (Finset.mem_range.mp hx) rw [Finset.sum_congr rfl this] _ = mk (span {s}) @@ -129,7 +129,7 @@ theorem odd_sq_dvd_geom_sum₂_sub (hp : Odd p) : rintro (⟨⟩ | ⟨x⟩) hx · rw [Nat.cast_zero, mul_zero, mul_zero] · have : x.succ - 1 + (p - 1 - x.succ) = p - 2 := by - rw [← Nat.add_sub_assoc (Nat.le_pred_of_lt (Finset.mem_range.mp hx))] + rw [← Nat.add_sub_assoc (Nat.le_sub_one_of_lt (Finset.mem_range.mp hx))] exact congr_arg Nat.pred (Nat.add_sub_cancel_left _ _) rw [this] ring1 @@ -269,13 +269,13 @@ theorem Int.sq_mod_four_eq_one_of_odd {x : ℤ} : Odd x → x ^ 2 % 4 = 1 := by rcases hx with ⟨_, rfl⟩ ring_nf rw [add_assoc, ← add_mul, Int.add_mul_emod_self] - norm_num + decide #align int.sq_mod_four_eq_one_of_odd Int.sq_mod_four_eq_one_of_odd theorem Int.two_pow_two_pow_add_two_pow_two_pow {x y : ℤ} (hx : ¬2 ∣ x) (hxy : 4 ∣ x - y) (i : ℕ) : multiplicity 2 (x ^ 2 ^ i + y ^ 2 ^ i) = ↑(1 : ℕ) := by have hx_odd : Odd x := by rwa [Int.odd_iff_not_even, even_iff_two_dvd] - have hxy_even : Even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by norm_num) hxy) + have hxy_even : Even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by decide) hxy) have hy_odd : Odd y := by simpa using hx_odd.sub_even hxy_even refine' multiplicity.eq_coe_iff.mpr ⟨_, _⟩ · rw [pow_one, ← even_iff_two_dvd] @@ -292,7 +292,7 @@ theorem Int.two_pow_two_pow_add_two_pow_two_pow {x y : ℤ} (hx : ¬2 ∣ x) (hx suffices ∀ x : ℤ, Odd x → x ^ 2 ^ (i + 1) % 4 = 1 by rw [show (2 ^ (1 + 1) : ℤ) = 4 by norm_num, Int.dvd_iff_emod_eq_zero, Int.add_emod, this _ hx_odd, this _ hy_odd] - norm_num + decide intro x hx rw [pow_succ, mul_comm, pow_mul, Int.sq_mod_four_eq_one_of_odd hx.pow] #align int.two_pow_two_pow_add_two_pow_two_pow Int.two_pow_two_pow_add_two_pow_two_pow @@ -307,7 +307,7 @@ theorem Int.two_pow_two_pow_sub_pow_two_pow {x y : ℤ} (n : ℕ) (hxy : 4 ∣ x theorem Int.two_pow_sub_pow' {x y : ℤ} (n : ℕ) (hxy : 4 ∣ x - y) (hx : ¬2 ∣ x) : multiplicity 2 (x ^ n - y ^ n) = multiplicity 2 (x - y) + multiplicity (2 : ℤ) n := by have hx_odd : Odd x := by rwa [Int.odd_iff_not_even, even_iff_two_dvd] - have hxy_even : Even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by norm_num) hxy) + have hxy_even : Even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by decide) hxy) have hy_odd : Odd y := by simpa using hx_odd.sub_even hxy_even cases' n with n · simp only [pow_zero, sub_self, multiplicity.zero, Int.ofNat_zero, Nat.zero_eq, add_top] diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 333f2ef4a869c..9223c3f7ea701 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -3,11 +3,10 @@ Copyright (c) 2022 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ +import Mathlib.RingTheory.Discriminant import Mathlib.Algebra.Module.Zlattice import Mathlib.MeasureTheory.Group.GeometryOfNumbers -import Mathlib.MeasureTheory.Measure.Haar.NormedSpace import Mathlib.NumberTheory.NumberField.Embeddings -import Mathlib.RingTheory.Discriminant #align_import number_theory.number_field.canonical_embedding from "leanprover-community/mathlib"@"60da01b41bbe4206f05d34fd70c8dd7498717a30" @@ -21,28 +20,30 @@ into the type `(K →+* ℂ) → ℂ` of `ℂ`-vectors indexed by the complex em ## Main definitions and results -* `canonicalEmbedding`: the ring homorphism `K →+* ((K →+* ℂ) → ℂ)` defined by sending `x : K` to -the vector `(φ x)` indexed by `φ : K →+* ℂ`. +* `NumberField.canonicalEmbedding`: the ring homorphism `K →+* ((K →+* ℂ) → ℂ)` defined by +sending `x : K` to the vector `(φ x)` indexed by `φ : K →+* ℂ`. -* `canonicalEmbedding.integerLattice.inter_ball_finite`: the intersection of the +* `NumberField.canonicalEmbedding.integerLattice.inter_ball_finite`: the intersection of the image of the ring of integers by the canonical embedding and any ball centered at `0` of finite radius is finite. -* `mixedEmbedding`: the ring homomorphism from `K →+* ({ w // IsReal w } → ℝ) × +* `NumberField.mixedEmbedding`: the ring homomorphism from `K →+* ({ w // IsReal w } → ℝ) × ({ w // IsComplex w } → ℂ)` that sends `x ∈ K` to `(φ_w x)_w` where `φ_w` is the embedding associated to the infinite place `w`. In particular, if `w` is real then `φ_w : K →+* ℝ` and, if `w` is complex, `φ_w` is an arbitrary choice between the two complex embeddings defining the place `w`. -* `exists_ne_zero_mem_ringOfIntegers_lt`: let `f : InfinitePlace K → ℝ≥0`, if the product -`∏ w, f w` is large enough, then there exists a nonzero algebraic integer `a` in `K` such that -`w a < f w` for all infinite places `w`. +* `NumberField.mixedEmbedding.exists_ne_zero_mem_ringOfIntegers_lt`: let +`f : InfinitePlace K → ℝ≥0`, if the product `∏ w, f w` is large enough, then there exists a +nonzero algebraic integer `a` in `K` such that `w a < f w` for all infinite places `w`. ## Tags number field, infinite places -/ +local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 + variable (K : Type*) [Field K] namespace NumberField.canonicalEmbedding @@ -161,7 +162,7 @@ end NumberField.canonicalEmbedding namespace NumberField.mixedEmbedding -open NumberField NumberField.InfinitePlace NumberField.ComplexEmbedding FiniteDimensional +open NumberField NumberField.InfinitePlace FiniteDimensional /-- The space `ℝ^r₁ × ℂ^r₂` with `(r₁, r₂)` the signature of `K`. -/ local notation "E" K => @@ -183,9 +184,9 @@ instance [NumberField K] : Nontrivial (E K) := by protected theorem finrank [NumberField K] : finrank ℝ (E K) = finrank ℚ K := by classical rw [finrank_prod, finrank_pi, finrank_pi_fintype, Complex.finrank_real_complex, Finset.sum_const, - Finset.card_univ, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, - ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, - Nat.add_sub_of_le (Fintype.card_subtype_le _)] + Finset.card_univ, ← NrRealPlaces, ← NrComplexPlaces, ← card_real_embeddings, + Algebra.id.smul_eq_mul, mul_comm, ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, + Fintype.card_subtype_compl, Nat.add_sub_of_le (Fintype.card_subtype_le _)] theorem _root_.NumberField.mixedEmbedding_injective [NumberField K] : Function.Injective (NumberField.mixedEmbedding K) := by @@ -221,7 +222,7 @@ theorem commMap_canonical_eq_mixed (x : K) : /-- This is a technical result to ensure that the image of the `ℂ`-basis of `ℂ^n` defined in `canonicalEmbedding.latticeBasis` is a `ℝ`-basis of `ℝ^r₁ × ℂ^r₂`, see `mixedEmbedding.latticeBasis`. -/ -theorem disjoint_span_commMap_ker [NumberField K]: +theorem disjoint_span_commMap_ker [NumberField K] : Disjoint (Submodule.span ℝ (Set.range (canonicalEmbedding.latticeBasis K))) (LinearMap.ker (commMap K)) := by refine LinearMap.disjoint_ker.mpr (fun x h_mem h_zero => ?_) @@ -231,7 +232,7 @@ theorem disjoint_span_commMap_ker [NumberField K]: exact ⟨integralBasis K i, (canonicalEmbedding.latticeBasis_apply K i).symm⟩ ext1 φ rw [Pi.zero_apply] - by_cases hφ : IsReal φ + by_cases hφ : ComplexEmbedding.IsReal φ · rw [show x φ = (x φ).re by rw [eq_comm, ← Complex.conj_eq_iff_re, canonicalEmbedding.conj_apply _ h_mem, ComplexEmbedding.isReal_iff.mp hφ], ← Complex.ofReal_zero] @@ -248,6 +249,148 @@ theorem disjoint_span_commMap_ker [NumberField K]: end commMap +noncomputable section stdBasis + +open Classical Complex MeasureTheory MeasureTheory.Measure Zspan Matrix BigOperators + ComplexConjugate + +variable [NumberField K] + +/-- The type indexing the basis `stdBasis`. -/ +abbrev index := {w : InfinitePlace K // IsReal w} ⊕ ({w : InfinitePlace K // IsComplex w}) × (Fin 2) + +/-- The `ℝ`-basis of `({w // IsReal w} → ℝ) × ({ w // IsComplex w } → ℂ)` formed by the vector +equal to `1` at `w` and `0` elsewhere for `IsReal w` and by the couple of vectors equal to `1` +(resp. `I`) at `w` and `0` elsewhere for `IsComplex w`. -/ +def stdBasis : Basis (index K) ℝ (E K) := + Basis.prod (Pi.basisFun ℝ _) + (Basis.reindex (Pi.basis fun _ => basisOneI) (Equiv.sigmaEquivProd _ _)) + +variable {K} + +@[simp] +theorem stdBasis_apply_ofIsReal (x : E K) (w : {w : InfinitePlace K // IsReal w}) : + (stdBasis K).repr x (Sum.inl w) = x.1 w := rfl + +@[simp] +theorem stdBasis_apply_ofIsComplex_fst (x : E K) (w : {w : InfinitePlace K // IsComplex w}) : + (stdBasis K).repr x (Sum.inr ⟨w, 0⟩) = (x.2 w).re := rfl + +@[simp] +theorem stdBasis_apply_ofIsComplex_snd (x : E K) (w : {w : InfinitePlace K // IsComplex w}) : + (stdBasis K).repr x (Sum.inr ⟨w, 1⟩) = (x.2 w).im := rfl + +variable (K) + +theorem fundamentalDomain_stdBasis : + fundamentalDomain (stdBasis K) = + (Set.univ.pi fun _ => Set.Ico 0 1) ×ˢ + (Set.univ.pi fun _ => Complex.measurableEquivPi⁻¹' (Set.univ.pi fun _ => Set.Ico 0 1)) := by + ext + simp [stdBasis, mem_fundamentalDomain, Complex.measurableEquivPi] + +theorem volume_fundamentalDomain_stdBasis : + volume (fundamentalDomain (stdBasis K)) = 1 := by + rw [fundamentalDomain_stdBasis, volume_eq_prod, prod_prod, volume_pi, volume_pi, pi_pi, pi_pi, + Complex.volume_preserving_equiv_pi.measure_preimage ?_, volume_pi, pi_pi, Real.volume_Ico, + sub_zero, ENNReal.ofReal_one, Finset.prod_const_one, Finset.prod_const_one, + Finset.prod_const_one, one_mul] + exact MeasurableSet.pi Set.countable_univ (fun _ _ => measurableSet_Ico) + +/-- The `Equiv` between `index K` and `K →+* ℂ` defined by sending a real infinite place `w` to +the unique corresponding embedding `w.embedding`, and the pair `⟨w, 0⟩` (resp. `⟨w, 1⟩`) for a +complex infinite place `w` to `w.embedding` (resp. `conjugate w.embedding`). -/ +def indexEquiv : (index K) ≃ (K →+* ℂ) := by + refine Equiv.ofBijective (fun c => ?_) + ((Fintype.bijective_iff_surjective_and_card _).mpr ⟨?_, ?_⟩) + · cases c with + | inl w => exact w.val.embedding + | inr wj => rcases wj with ⟨w, j⟩ + exact if j = 0 then w.val.embedding else ComplexEmbedding.conjugate w.val.embedding + · intro φ + by_cases hφ : ComplexEmbedding.IsReal φ + · exact ⟨Sum.inl (InfinitePlace.mkReal ⟨φ, hφ⟩), by simp [embedding_mk_eq_of_isReal hφ]⟩ + · by_cases hw : (InfinitePlace.mk φ).embedding = φ + · exact ⟨Sum.inr ⟨InfinitePlace.mkComplex ⟨φ, hφ⟩, 0⟩, by simp [hw]⟩ + · exact ⟨Sum.inr ⟨InfinitePlace.mkComplex ⟨φ, hφ⟩, 1⟩, + by simp [(embedding_mk_eq φ).resolve_left hw]⟩ + · rw [Embeddings.card, ← mixedEmbedding.finrank K, + ← FiniteDimensional.finrank_eq_card_basis (stdBasis K)] + +variable {K} + +@[simp] +theorem indexEquiv_apply_ofIsReal (w : {w : InfinitePlace K // IsReal w}) : + (indexEquiv K) (Sum.inl w) = w.val.embedding := rfl + +@[simp] +theorem indexEquiv_apply_ofIsComplex_fst (w : {w : InfinitePlace K // IsComplex w}) : + (indexEquiv K) (Sum.inr ⟨w, 0⟩) = w.val.embedding := rfl + +@[simp] +theorem indexEquiv_apply_ofIsComplex_snd (w : {w : InfinitePlace K // IsComplex w}) : + (indexEquiv K) (Sum.inr ⟨w, 1⟩) = ComplexEmbedding.conjugate w.val.embedding := rfl + +variable (K) + +/-- The matrix that gives the representation on `stdBasis` of the image by `commMap` of an +element `x` of `(K →+* ℂ) → ℂ` fixed by the map `x_φ ↦ conj x_(conjugate φ)`, +see `stdBasis_repr_eq_matrixToStdBasis_mul`. -/ +def matrixToStdBasis : Matrix (index K) (index K) ℂ := + fromBlocks (diagonal fun _ => 1) 0 0 <| reindex (Equiv.prodComm _ _) (Equiv.prodComm _ _) + (blockDiagonal (fun _ => (2 : ℂ)⁻¹ • !![1, 1; - I, I])) + +theorem det_matrixToStdBasis : + (matrixToStdBasis K).det = (2⁻¹ * I) ^ NrComplexPlaces K := + calc + _ = ∏ k : { w : InfinitePlace K // IsComplex w }, det ((2 : ℂ)⁻¹ • !![1, 1; -I, I]) := by + rw [matrixToStdBasis, det_fromBlocks_zero₂₁, det_diagonal, Finset.prod_const_one, one_mul, + det_reindex_self, det_blockDiagonal] + _ = ∏ k : { w : InfinitePlace K // IsComplex w }, (2⁻¹ * Complex.I) := by + refine Finset.prod_congr (Eq.refl _) (fun _ _ => ?_) + field_simp; ring + _ = (2⁻¹ * Complex.I) ^ Fintype.card {w : InfinitePlace K // IsComplex w} := by + rw [Finset.prod_const, Fintype.card] + +/-- Let `x : (K →+* ℂ) → ℂ` such that `x_φ = conj x_(conj φ)` for all `φ : K →+* ℂ`, then the +representation of `commMap K x` on `stdBasis` is given (up to reindexing) by the product of +`matrixToStdBasis` by `x`. -/ +theorem stdBasis_repr_eq_matrixToStdBasis_mul (x : (K →+* ℂ) → ℂ) + (hx : ∀ φ, conj (x φ) = x (ComplexEmbedding.conjugate φ)) (c : index K) : + ((stdBasis K).repr (commMap K x) c : ℂ) = + (mulVec (matrixToStdBasis K) (x ∘ (indexEquiv K))) c := by + simp_rw [commMap, matrixToStdBasis, LinearMap.coe_mk, AddHom.coe_mk, + mulVec, dotProduct, Function.comp_apply, index, Fintype.sum_sum_type, + diagonal_one, reindex_apply, ← Finset.univ_product_univ, Finset.sum_product, + indexEquiv_apply_ofIsReal, Fin.sum_univ_two, indexEquiv_apply_ofIsComplex_fst, + indexEquiv_apply_ofIsComplex_snd, smul_of, smul_cons, smul_eq_mul, + mul_one, smul_empty, Equiv.prodComm_symm, Equiv.coe_prodComm] + cases c with + | inl w => + simp_rw [stdBasis_apply_ofIsReal, fromBlocks_apply₁₁, fromBlocks_apply₁₂, + one_apply, Matrix.zero_apply, ite_mul, one_mul, zero_mul, Finset.sum_ite_eq, + Finset.mem_univ, ite_true, add_zero, Finset.sum_const_zero, add_zero, + ← conj_eq_iff_re, hx (embedding w.val), conjugate_embedding_eq_of_isReal w.prop] + | inr c => + rcases c with ⟨w, j⟩ + fin_cases j + · simp_rw [Fin.mk_zero, stdBasis_apply_ofIsComplex_fst, fromBlocks_apply₂₁, + fromBlocks_apply₂₂, Matrix.zero_apply, submatrix_apply, + blockDiagonal_apply, Prod.swap_prod_mk, ite_mul, zero_mul, Finset.sum_const_zero, + zero_add, Finset.sum_add_distrib, Finset.sum_ite_eq, Finset.mem_univ, ite_true, + of_apply, cons_val', cons_val_zero, cons_val_one, + head_cons, ← hx (embedding w), re_eq_add_conj] + field_simp + · simp_rw [Fin.mk_one, stdBasis_apply_ofIsComplex_snd, fromBlocks_apply₂₁, + fromBlocks_apply₂₂, Matrix.zero_apply, submatrix_apply, + blockDiagonal_apply, Prod.swap_prod_mk, ite_mul, zero_mul, Finset.sum_const_zero, + zero_add, Finset.sum_add_distrib, Finset.sum_ite_eq, Finset.mem_univ, ite_true, + of_apply, cons_val', cons_val_zero, cons_val_one, + head_cons, ← hx (embedding w), im_eq_sub_conj] + ring_nf; field_simp + +end stdBasis + section integerLattice open Module FiniteDimensional @@ -266,8 +409,8 @@ noncomputable def latticeBasis [NumberField K] : refine basisOfLinearIndependentOfCardEqFinrank this ?_ rw [← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, finrank_prod, finrank_pi, finrank_pi_fintype, Complex.finrank_real_complex, Finset.sum_const, Finset.card_univ, - ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, ← card_complex_embeddings, - ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, + ← NrRealPlaces, ← NrComplexPlaces, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, + ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, Nat.add_sub_of_le (Fintype.card_subtype_le _)] @[simp] @@ -319,24 +462,22 @@ theorem convexBodyLt_convex : Convex ℝ (convexBodyLt K f) := open Classical Fintype MeasureTheory MeasureTheory.Measure BigOperators --- See: https://github.com/leanprover/lean4/issues/2220 -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) - variable [NumberField K] +instance : IsAddHaarMeasure (volume : Measure (E K)) := prod.instIsAddHaarMeasure volume volume + /-- The fudge factor that appears in the formula for the volume of `convexBodyLt`. -/ noncomputable abbrev convexBodyLtFactor : ℝ≥0∞ := - (2 : ℝ≥0∞) ^ card {w : InfinitePlace K // IsReal w} * - volume (ball (0 : ℂ) 1) ^ card {w : InfinitePlace K // IsComplex w} + (2 : ℝ≥0∞) ^ NrRealPlaces K * (NNReal.pi : ℝ≥0∞) ^ NrComplexPlaces K theorem convexBodyLtFactor_pos : 0 < (convexBodyLtFactor K) := by - refine mul_pos (NeZero.ne _) ?_ - exact ENNReal.pow_ne_zero (ne_of_gt (measure_ball_pos _ _ (by norm_num))) _ + refine mul_pos (NeZero.ne _) (ENNReal.pow_ne_zero ?_ _) + exact ne_of_gt (coe_pos.mpr Real.pi_pos) theorem convexBodyLtFactor_lt_top : (convexBodyLtFactor K) < ⊤ := by refine mul_lt_top ?_ ?_ · exact ne_of_lt (pow_lt_top (lt_top_iff_ne_top.mpr two_ne_top) _) - · exact ne_of_lt (pow_lt_top measure_ball_lt_top _) + · exact ne_of_lt (pow_lt_top coe_lt_top _) /-- The volume of `(ConvexBodyLt K f)` where `convexBodyLt K f` is the set of points `x` such that `‖x w‖ < f w` for all infinite places `w`. -/ @@ -344,18 +485,10 @@ theorem convexBodyLt_volume : volume (convexBodyLt K f) = (convexBodyLtFactor K) * ∏ w, (f w) ^ (mult w) := by calc _ = (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (2 * (f x.val))) * - ∏ x : {w // InfinitePlace.IsComplex w}, volume (ball (0 : ℂ) 1) * - ENNReal.ofReal (f x.val) ^ 2 := by - simp_rw [volume_eq_prod, prod_prod, volume_pi, pi_pi, Real.volume_ball] - conv_lhs => - congr; next => skip - congr; next => skip - ext i; rw [addHaar_ball _ _ (by exact (f i).prop), Complex.finrank_real_complex, mul_comm, - ENNReal.ofReal_pow (by exact (f i).prop)] - _ = (↑2 ^ card {w : InfinitePlace K // InfinitePlace.IsReal w} * - (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (f x.val))) * - (volume (ball (0 : ℂ) 1) ^ card {w : InfinitePlace K // IsComplex w} * - (∏ x : {w // IsComplex w}, ENNReal.ofReal (f x.val) ^ 2)) := by + ∏ x : {w // InfinitePlace.IsComplex w}, pi * ENNReal.ofReal (f x.val) ^ 2 := by + simp_rw [volume_eq_prod, prod_prod, volume_pi, pi_pi, Real.volume_ball, Complex.volume_ball] + _ = (↑2 ^ NrRealPlaces K * (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (f x.val))) * + (↑pi ^ NrComplexPlaces K * (∏ x : {w // IsComplex w}, ENNReal.ofReal (f x.val) ^ 2)) := by simp_rw [ofReal_mul (by norm_num : 0 ≤ (2 : ℝ)), Finset.prod_mul_distrib, Finset.prod_const, Finset.card_univ, ofReal_ofNat] _ = (convexBodyLtFactor K) * ((∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (f x.val)) * @@ -389,24 +522,81 @@ theorem adjust_f {w₁ : InfinitePlace K} (B : ℝ≥0) (hf : ∀ w, w ≠ w₁ end convexBodyLt +section convexBodySum + +open ENNReal BigOperators Classical MeasureTheory Fintype + +variable [NumberField K] (B : ℝ) + +/-- The convex body equal to the set of points `x : E` such that + `∑ w real, ‖x w‖ + 2 * ∑ w complex, ‖x w‖ ≤ B`. -/ +abbrev convexBodySum : Set (E K) := { x | ∑ w, ‖x.1 w‖ + 2 * ∑ w, ‖x.2 w‖ ≤ B } + +theorem convexBodySum_empty {B} (h : B < 0) : convexBodySum K B = ∅ := by + ext x + refine ⟨fun hx => ?_, fun h => h.elim⟩ + · rw [Set.mem_setOf] at hx + have : 0 ≤ ∑ w, ‖x.1 w‖ + 2 * ∑ w, ‖x.2 w‖ := by + refine add_nonneg ?_ ?_ + · exact Finset.sum_nonneg (fun _ _ => norm_nonneg _) + · exact mul_nonneg zero_le_two (Finset.sum_nonneg (fun _ _ => norm_nonneg _)) + linarith + +theorem convexBodySum_mem {x : K} : + mixedEmbedding K x ∈ (convexBodySum K B) ↔ + ∑ w : InfinitePlace K, (mult w) * w.val x ≤ B := by + simp_rw [Set.mem_setOf_eq, mixedEmbedding, RingHom.prod_apply, Pi.ringHom_apply, + ← Complex.norm_real, embedding_of_isReal_apply, norm_embedding_eq, mult, Nat.cast_ite, ite_mul, + Finset.sum_ite, Finset.filter_congr (fun _ _ => not_isReal_iff_isComplex), Finset.mul_sum, + ← Finset.sum_subtype_eq_sum_filter, Finset.subtype_univ, Nat.cast_one, one_mul, Nat.cast_ofNat] + rfl + +theorem convexBodySum_symmetric (x : E K) (hx : x ∈ (convexBodySum K B)) : + -x ∈ (convexBodySum K B) := by + simp_rw [Set.mem_setOf_eq, Prod.fst_neg, Prod.snd_neg, Pi.neg_apply, norm_neg] + exact hx + +theorem convexBodySum_convex : Convex ℝ (convexBodySum K B) := by + refine Convex_subadditive_le ?_ ?_ B + · intro x y + simp_rw [Prod.fst_add, Pi.add_apply, Prod.snd_add] + refine le_trans (add_le_add + (Finset.sum_le_sum (fun w _ => norm_add_le (x.1 w) (y.1 w))) + (mul_le_mul_of_nonneg_left + (Finset.sum_le_sum (fun w _ => norm_add_le (x.2 w) (y.2 w))) (by norm_num))) ?_ + simp_rw [Finset.sum_add_distrib, mul_add] + exact le_of_eq (by ring) + · intro _ _ h + simp_rw [Prod.smul_fst, Prod.smul_snd, Pi.smul_apply, smul_eq_mul, Complex.real_smul, norm_mul, + Complex.norm_real, Real.norm_of_nonneg h, ← Finset.mul_sum] + exact le_of_eq (by ring) + +end convexBodySum + section minkowski -open MeasureTheory MeasureTheory.Measure Classical NNReal ENNReal FiniteDimensional Zspan +open MeasureTheory MeasureTheory.Measure Classical FiniteDimensional Zspan Real + +open scoped ENNReal NNReal variable [NumberField K] /-- The bound that appears in Minkowski Convex Body theorem, see -`MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure`. -/ +`MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure`. See +`NumberField.mixedEmbedding.volume_fundamentalDomain_latticeBasis` for the computation of +`volume (fundamentalDomain (latticeBasis K))`. -/ noncomputable def minkowskiBound : ℝ≥0∞ := - volume (fundamentalDomain (latticeBasis K)) * 2 ^ (finrank ℝ (E K)) + volume (fundamentalDomain (latticeBasis K)) * (2:ℝ≥0∞) ^ (finrank ℝ (E K)) theorem minkowskiBound_lt_top : minkowskiBound K < ⊤ := by - refine mul_lt_top ?_ ?_ + refine ENNReal.mul_lt_top ?_ ?_ · exact ne_of_lt (fundamentalDomain_isBounded (latticeBasis K)).measure_lt_top - · exact ne_of_lt (pow_lt_top (lt_top_iff_ne_top.mpr two_ne_top) _) + · exact ne_of_lt (ENNReal.pow_lt_top (lt_top_iff_ne_top.mpr ENNReal.two_ne_top) _) variable {f : InfinitePlace K → ℝ≥0} +instance : IsAddHaarMeasure (volume : Measure (E K)) := prod.instIsAddHaarMeasure volume volume + /-- Assume that `f : InfinitePlace K → ℝ≥0` is such that `minkowskiBound K < volume (convexBodyLt K f)` where `convexBodyLt K f` is the set of points `x` such that `‖x w‖ < f w` for all infinite places `w` (see `convexBodyLt_volume` for @@ -414,7 +604,6 @@ the computation of this volume), then there exists a nonzero algebraic integer ` that `w a < f w` for all infinite places `w`. -/ theorem exists_ne_zero_mem_ringOfIntegers_lt (h : minkowskiBound K < volume (convexBodyLt K f)) : ∃ (a : 𝓞 K), a ≠ 0 ∧ ∀ w : InfinitePlace K, w a < f w := by - have : @IsAddHaarMeasure (E K) _ _ _ volume := prod.instIsAddHaarMeasure volume volume have h_fund := Zspan.isAddFundamentalDomain (latticeBasis K) volume have : Countable (Submodule.span ℤ (Set.range (latticeBasis K))).toAddSubgroup := by change Countable (Submodule.span ℤ (Set.range (latticeBasis K)): Set (E K)) @@ -427,6 +616,39 @@ theorem exists_ne_zero_mem_ringOfIntegers_lt (h : minkowskiBound K < volume (con rw [ne_eq, AddSubgroup.mk_eq_zero_iff, map_eq_zero, ← ne_eq] at h_nzr exact Subtype.ne_of_val_ne h_nzr +theorem exists_ne_zero_mem_ringOfIntegers_of_norm_le {B : ℝ} + (h : (minkowskiBound K) < volume (convexBodySum K B)) : + ∃ (a : 𝓞 K), a ≠ 0 ∧ |Algebra.norm ℚ (a:K)| ≤ (B / (finrank ℚ K)) ^ (finrank ℚ K) := by + have hB : 0 ≤ B := by + contrapose! h + rw [convexBodySum_empty K h, measure_empty] + exact zero_le (minkowskiBound K) + -- Some inequalities that will be useful later on + have h1 : 0 < (finrank ℚ K : ℝ)⁻¹ := inv_pos.mpr (Nat.cast_pos.mpr finrank_pos) + have h2 : 0 ≤ B / (finrank ℚ K) := div_nonneg hB (Nat.cast_nonneg _) + have h_fund := Zspan.isAddFundamentalDomain (latticeBasis K) volume + have : Countable (Submodule.span ℤ (Set.range (latticeBasis K))).toAddSubgroup := by + change Countable (Submodule.span ℤ (Set.range (latticeBasis K)): Set (E K)) + infer_instance + obtain ⟨⟨x, hx⟩, h_nzr, h_mem⟩ := exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure + h_fund h (convexBodySum_symmetric K B) (convexBodySum_convex K B) + rw [Submodule.mem_toAddSubgroup, mem_span_latticeBasis] at hx + obtain ⟨a, ha, rfl⟩ := hx + refine ⟨⟨a, ha⟩, ?_, ?_⟩ + · rw [ne_eq, AddSubgroup.mk_eq_zero_iff, map_eq_zero, ← ne_eq] at h_nzr + exact Subtype.ne_of_val_ne h_nzr + · rw [← rpow_nat_cast, ← rpow_le_rpow_iff (by simp only [Rat.cast_abs, abs_nonneg]) + (rpow_nonneg_of_nonneg h2 _) h1, ← rpow_mul h2, mul_inv_cancel (Nat.cast_ne_zero.mpr + (ne_of_gt finrank_pos)), rpow_one, le_div_iff' (Nat.cast_pos.mpr finrank_pos)] + refine le_trans ?_ ((convexBodySum_mem K B).mp h_mem) + rw [← le_div_iff' (Nat.cast_pos.mpr finrank_pos), ← sum_mult_eq, Nat.cast_sum] + refine le_trans ?_ (geom_mean_le_arith_mean Finset.univ _ _ (fun _ _ => Nat.cast_nonneg _) + ?_ (fun _ _ => AbsoluteValue.nonneg _ _)) + · simp_rw [← prod_eq_abs_norm, rpow_nat_cast] + exact le_of_eq rfl + · rw [← Nat.cast_sum, sum_mult_eq, Nat.cast_pos] + exact finrank_pos + end minkowski end NumberField.mixedEmbedding diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index c154063ab4fd5..1e7efa51bd67d 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ -import Mathlib.NumberTheory.NumberField.Basic +import Mathlib.NumberTheory.NumberField.CanonicalEmbedding import Mathlib.RingTheory.Localization.NormTrace /-! @@ -20,9 +20,11 @@ number field, discriminant -- TODO. Rewrite some of the FLT results on the disciminant using the definitions and results of -- this file +local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 + namespace NumberField -open NumberField Matrix +open Classical NumberField Matrix NumberField.InfinitePlace variable (K : Type*) [Field K] [NumberField K] @@ -41,6 +43,43 @@ theorem discr_eq_discr {ι : Type*} [Fintype ι] [DecidableEq ι] (b : Basis ι let b₀ := Basis.reindex (RingOfIntegers.basis K) (Basis.indexEquiv (RingOfIntegers.basis K) b) rw [Algebra.discr_eq_discr (𝓞 K) b b₀, Basis.coe_reindex, Algebra.discr_reindex] +open MeasureTheory MeasureTheory.Measure Zspan NumberField.mixedEmbedding + NumberField.InfinitePlace ENNReal NNReal Complex + +theorem _root_.NumberField.mixedEmbedding.volume_fundamentalDomain_latticeBasis : + volume (fundamentalDomain (latticeBasis K)) = + (2 : ℝ≥0∞)⁻¹ ^ (NrComplexPlaces K) * sqrt ‖discr K‖₊ := by + let f : Module.Free.ChooseBasisIndex ℤ (𝓞 K) ≃ (K →+* ℂ) := + (canonicalEmbedding.latticeBasis K).indexEquiv (Pi.basisFun ℂ _) + let e : (index K) ≃ Module.Free.ChooseBasisIndex ℤ (𝓞 K) := (indexEquiv K).trans f.symm + let M := (mixedEmbedding.stdBasis K).toMatrix ((latticeBasis K).reindex e.symm) + let N := Algebra.embeddingsMatrixReindex ℚ ℂ (integralBasis K ∘ f.symm) + RingHom.equivRatAlgHom + suffices M.map Complex.ofReal = (matrixToStdBasis K) * + (Matrix.reindex (indexEquiv K).symm (indexEquiv K).symm N).transpose by + calc volume (fundamentalDomain (latticeBasis K)) + _ = ‖((mixedEmbedding.stdBasis K).toMatrix ((latticeBasis K).reindex e.symm)).det‖₊ := by + rw [← fundamentalDomain_reindex _ e.symm, ← norm_toNNReal, measure_fundamentalDomain + ((latticeBasis K).reindex e.symm), volume_fundamentalDomain_stdBasis, mul_one] + rfl + _ = ‖(matrixToStdBasis K).det * N.det‖₊ := by + rw [← nnnorm_real, ← ofReal_eq_coe, RingHom.map_det, RingHom.mapMatrix_apply, this, + det_mul, det_transpose, det_reindex_self] + _ = (2 : ℝ≥0∞)⁻¹ ^ Fintype.card {w : InfinitePlace K // IsComplex w} * sqrt ‖N.det ^ 2‖₊ := by + have : ‖Complex.I‖₊ = 1 := by rw [← norm_toNNReal, norm_eq_abs, abs_I, Real.toNNReal_one] + rw [det_matrixToStdBasis, nnnorm_mul, nnnorm_pow, nnnorm_mul, this, mul_one, nnnorm_inv, + coe_mul, ENNReal.coe_pow, ← norm_toNNReal, IsROrC.norm_two, Real.toNNReal_ofNat, + coe_inv two_ne_zero, coe_ofNat, nnnorm_pow, NNReal.sqrt_sq] + _ = (2 : ℝ≥0∞)⁻¹ ^ Fintype.card { w // IsComplex w } * NNReal.sqrt ‖discr K‖₊ := by + rw [← Algebra.discr_eq_det_embeddingsMatrixReindex_pow_two, Algebra.discr_reindex, + ← coe_discr, map_intCast, ← Complex.nnnorm_int] + ext : 2 + dsimp only + rw [Matrix.map_apply, Basis.toMatrix_apply, Basis.coe_reindex, Function.comp, Equiv.symm_symm, + latticeBasis_apply, ← commMap_canonical_eq_mixed, Complex.ofReal_eq_coe, + stdBasis_repr_eq_matrixToStdBasis_mul K _ (fun _ => rfl)] + rfl + end NumberField namespace Rat diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 83380bf8ef22d..7c1fad2fe1be1 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -352,6 +352,11 @@ theorem isComplex_iff {w : InfinitePlace K} : | inr h => rwa [← ComplexEmbedding.isReal_conjugate_iff, h] at hφ #align number_field.infinite_place.is_complex_iff NumberField.InfinitePlace.isComplex_iff +@[simp] +theorem conjugate_embedding_eq_of_isReal {w : InfinitePlace K} (h : IsReal w) : + ComplexEmbedding.conjugate (embedding w) = embedding w := + ComplexEmbedding.isReal_iff.mpr (isReal_iff.mp h) + @[simp] theorem not_isReal_iff_isComplex {w : InfinitePlace K} : ¬IsReal w ↔ IsComplex w := by rw [isComplex_iff, isReal_iff] @@ -403,6 +408,19 @@ theorem card_filter_mk_eq [NumberField K] (w : InfinitePlace K) : · refine Finset.card_doubleton ?_ rwa [Ne.def, eq_comm, ← ComplexEmbedding.isReal_iff, ← isReal_iff] +open scoped BigOperators + +noncomputable instance NumberField.InfinitePlace.fintype [NumberField K] : + Fintype (InfinitePlace K) := Set.fintypeRange _ +#align number_field.infinite_place.number_field.infinite_place.fintype NumberField.InfinitePlace.NumberField.InfinitePlace.fintype + +theorem sum_mult_eq [NumberField K] : + ∑ w : InfinitePlace K, mult w = FiniteDimensional.finrank ℚ K := by + rw [← Embeddings.card K ℂ, Fintype.card, Finset.card_eq_sum_ones, ← Finset.univ.sum_fiberwise + (fun φ => InfinitePlace.mk φ)] + exact Finset.sum_congr rfl + (fun _ _ => by rw [Finset.sum_const, smul_eq_mul, mul_one, card_filter_mk_eq]) + /-- The map from real embeddings to real infinite places as an equiv -/ noncomputable def mkReal : { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } ≃ { w : InfinitePlace K // IsReal w } := by @@ -430,10 +448,6 @@ theorem mkComplex_coe (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } variable [NumberField K] -noncomputable instance NumberField.InfinitePlace.fintype : Fintype (InfinitePlace K) := - Set.fintypeRange _ -#align number_field.infinite_place.number_field.infinite_place.fintype NumberField.InfinitePlace.NumberField.InfinitePlace.fintype - open scoped BigOperators /-- The infinite part of the product formula : for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where @@ -455,14 +469,20 @@ theorem prod_eq_abs_norm (x : K) : open Fintype FiniteDimensional +variable (K) + +/-- The number of infinite real places of the number field `K`. -/ +noncomputable abbrev NrRealPlaces := card { w : InfinitePlace K // IsReal w } + +/-- The number of infinite complex places of the number field `K`. -/ +noncomputable abbrev NrComplexPlaces := card { w : InfinitePlace K // IsComplex w } + theorem card_real_embeddings : - card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } - = card { w : InfinitePlace K // IsReal w } := Fintype.card_congr mkReal + card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } = NrRealPlaces K := Fintype.card_congr mkReal #align number_field.infinite_place.card_real_embeddings NumberField.InfinitePlace.card_real_embeddings theorem card_complex_embeddings : - card { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } = - 2 * card { w : InfinitePlace K // IsComplex w } := by + card { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } = 2 * NrComplexPlaces K := by suffices ∀ w : { w : InfinitePlace K // IsComplex w }, (Finset.univ.filter fun φ : { φ // ¬ ComplexEmbedding.IsReal φ } => mkComplex φ = w).card = 2 by rw [Fintype.card, Finset.card_eq_sum_ones, ← Finset.sum_fiberwise _ (fun φ => mkComplex φ)] @@ -476,14 +496,13 @@ theorem card_complex_embeddings : · rwa [Subtype.mk_eq_mk, ← Subtype.ext_iff, ← Subtype.ext_iff] at h · refine ⟨⟨⟨φ, not_isReal_of_mk_isComplex (hφ.symm ▸ hw)⟩, ?_⟩, rfl⟩ rwa [Subtype.ext_iff, mkComplex_coe] - · simp_rw [mult, not_isReal_iff_isComplex.mpr hw] + · simp_rw [mult, not_isReal_iff_isComplex.mpr hw, ite_false] #align number_field.infinite_place.card_complex_embeddings NumberField.InfinitePlace.card_complex_embeddings theorem card_add_two_mul_card_eq_rank : - card { w : InfinitePlace K // IsReal w } + 2 * card { w : InfinitePlace K // IsComplex w } = - finrank ℚ K := by - rw [← card_real_embeddings, ← card_complex_embeddings] - rw [Fintype.card_subtype_compl, ← Embeddings.card K ℂ, Nat.add_sub_of_le] + NrRealPlaces K + 2 * NrComplexPlaces K = finrank ℚ K := by + rw [← card_real_embeddings, ← card_complex_embeddings, Fintype.card_subtype_compl, + ← Embeddings.card K ℂ, Nat.add_sub_of_le] exact Fintype.card_subtype_le _ end NumberField.InfinitePlace diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index ab40e8f28ec4b..6506641b724bd 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -41,6 +41,7 @@ fundamental system `fundSystem`. number field, units -/ +local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 open scoped NumberField @@ -226,7 +227,7 @@ theorem mult_log_place_eq_zero {x : (𝓞 K)ˣ} {w : InfinitePlace K} : mult w * Real.log (w x) = 0 ↔ w x = 1 := by rw [mul_eq_zero, or_iff_right, Real.log_eq_zero, or_iff_right, or_iff_left] · linarith [(map_nonneg _ _ : 0 ≤ w x)] - · simp only [ne_eq, map_eq_zero, coe_ne_zero x] + · simp only [ne_eq, map_eq_zero, coe_ne_zero x, not_false_eq_true] · refine (ne_of_gt ?_) rw [mult]; split_ifs <;> norm_num @@ -263,7 +264,7 @@ theorem log_le_of_logEmbedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : refine (le_trans ?_ hyp).trans ?_ · rw [← hw] exact tool _ (abs_nonneg _) - · refine (sum_le_card_nsmul univ _ _ + · refine (sum_le_card_nsmul univ _ _ (fun w _ => logEmbedding_component_le hr h w)).trans ?_ rw [nsmul_eq_mul] refine mul_le_mul ?_ le_rfl hr (Fintype.card (InfinitePlace K)).cast_nonneg @@ -321,9 +322,6 @@ sequence defining the same ideal and their quotient is the desired unit `u_w₁` open NumberField.mixedEmbedding NNReal --- See: https://github.com/leanprover/lean4/issues/2220 -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) - variable (w₁ : InfinitePlace K) {B : ℕ} (hB : minkowskiBound K < (convexBodyLtFactor K) * B) /-- This result shows that there always exists a next term in the sequence. -/ @@ -347,7 +345,7 @@ theorem seq_next {x : 𝓞 K} (hx : x ≠ 0) : simp_rw [← NNReal.coe_pow, ← NNReal.coe_prod] exact le_of_eq (congr_arg toReal h_gprod) · refine div_lt_self ?_ (by norm_num) - simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, hx] + simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, hx, not_false_eq_true] intro _ _ rw [ne_eq, Nonneg.mk_eq_zero, div_eq_zero_iff, map_eq_zero, not_or, ZeroMemClass.coe_eq_zero] exact ⟨hx, by norm_num⟩ @@ -400,20 +398,19 @@ theorem seq_norm_le (n : ℕ) : exact (seq_next K w₁ hB (seq K w₁ hB n).prop).choose_spec.2.2 /-- Construct a unit associated to the place `w₁`. The family, for `w₁ ≠ w₀`, formed by the -image by the `logEmbedding` of these units is `ℝ`-linearly independent, see +image by the `logEmbedding` of these units is `ℝ`-linearly independent, see `unitLattice_span_eq_top`. -/ -theorem exists_unit (w₁ : InfinitePlace K ) : +theorem exists_unit (w₁ : InfinitePlace K) : ∃ u : (𝓞 K)ˣ, ∀ w : InfinitePlace K, w ≠ w₁ → Real.log (w u) < 0 := by obtain ⟨B, hB⟩ : ∃ B : ℕ, minkowskiBound K < (convexBodyLtFactor K) * B := by - simp_rw [mul_comm] - refine ENNReal.exists_nat_mul_gt ?_ ?_ - exact ne_of_gt (convexBodyLtFactor_pos K) - exact ne_of_lt (minkowskiBound_lt_top K) + conv => congr; ext; rw [mul_comm] + exact ENNReal.exists_nat_mul_gt (ne_of_gt (convexBodyLtFactor_pos K)) + (ne_of_lt (minkowskiBound_lt_top K)) rsuffices ⟨n, m, hnm, h⟩ : ∃ n m, n < m ∧ (Ideal.span ({ (seq K w₁ hB n : 𝓞 K) }) = Ideal.span ({ (seq K w₁ hB m : 𝓞 K) })) · have hu := Ideal.span_singleton_eq_span_singleton.mp h refine ⟨hu.choose, fun w hw => Real.log_neg ?_ ?_⟩ - · simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, ne_zero] + · simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, ne_zero, not_false_eq_true] · calc _ = w ((seq K w₁ hB m : K) * (seq K w₁ hB n : K)⁻¹) := by rw [← congr_arg ((↑) : (𝓞 K) → K) hu.choose_spec, mul_comm, Submonoid.coe_mul, @@ -427,7 +424,7 @@ theorem exists_unit (w₁ : InfinitePlace K ) : (fun n => ?_) ?_ · rw [Set.mem_setOf_eq, Ideal.absNorm_span_singleton] refine ⟨?_, seq_norm_le K w₁ hB n⟩ - exact Nat.one_le_iff_ne_zero.mpr (Int.natAbs_ne_zero.mpr (seq_norm_ne_zero K w₁ hB n)) + exact Nat.one_le_iff_ne_zero.mpr (Int.natAbs_ne_zero.mpr (seq_norm_ne_zero K w₁ hB n)) · rw [show { I : Ideal (𝓞 K) | 1 ≤ Ideal.absNorm I ∧ Ideal.absNorm I ≤ B } = (⋃ n ∈ Set.Icc 1 B, { I : Ideal (𝓞 K) | Ideal.absNorm I = n }) by ext; simp] exact Set.Finite.biUnion (Set.finite_Icc _ _) (fun n hn => Ideal.finite_setOf_absNorm_eq hn.1) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index dec0489d9bc73..703c9a3dd3719 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -643,7 +643,7 @@ theorem range_pow_padicValNat_subset_divisors' {n : ℕ} [hp : Fact p.Prime] : #align range_pow_padic_val_nat_subset_divisors' range_pow_padicValNat_subset_divisors' /-- The `p`-adic valuation of `(p * n)!` is `n` more than that of `n!`. -/ -theorem padicValNat_factorial_mul (n : ℕ) [hp : Fact p.Prime]: +theorem padicValNat_factorial_mul (n : ℕ) [hp : Fact p.Prime] : padicValNat p (p * n) ! = padicValNat p n ! + n := by refine' PartENat.natCast_inj.mp _ rw [padicValNat_def' (Nat.Prime.ne_one hp.out) <| factorial_pos (p * n), Nat.cast_add, diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index b08fc29368fc3..c5980b3b24eea 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -508,6 +508,9 @@ theorem isCauSeq_nthHom (r : R) : IsCauSeq (padicNorm p) fun n => nthHom f r n : use k intro j hj refine' lt_of_le_of_lt _ hk + -- Need to do beta reduction first, as `norm_cast` doesn't. + -- Added to adapt to leanprover/lean4#2734. + beta_reduce norm_cast rw [← padicNorm.dvd_iff_norm_le] exact_mod_cast pow_dvd_nthHom_sub f_compat r k j hj diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index a68df8142d5d3..bba6c74e643fb 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -551,7 +551,7 @@ theorem y_strictMono {a : Solution₁ d} (h : IsFundamental a) : add_pos_of_pos_of_nonneg (mul_pos (x_zpow_pos h.x_pos _) h.2.1) (mul_nonneg _ (by rw [sub_nonneg]; exact h.1.le)) rcases hn.eq_or_lt with (rfl | hn) - · simp only [zpow_zero, y_one] + · simp only [zpow_zero, y_one, le_refl] · exact (y_zpow_pos h.x_pos h.2.1 hn).le refine' strictMono_int_of_lt_succ fun n => _ cases' le_or_lt 0 n with hn hn diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index 78a8edff6b0cd..f108fc103fc1c 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -551,8 +551,8 @@ theorem yn_modEq_a_sub_one : ∀ n, yn a1 n ≡ n [MOD a - 1] #align pell.yn_modeq_a_sub_one Pell.yn_modEq_a_sub_one theorem yn_modEq_two : ∀ n, yn a1 n ≡ n [MOD 2] - | 0 => by simp - | 1 => by simp + | 0 => by rfl + | 1 => by simp; rfl | n + 2 => (yn_modEq_two n).add_right_cancel <| by rw [yn_succ_succ, mul_assoc, (by ring : n + 2 + n = 2 * (n + 1))] diff --git a/Mathlib/NumberTheory/PythagoreanTriples.lean b/Mathlib/NumberTheory/PythagoreanTriples.lean index 5be1ba662c767..e909fe28d2c6a 100644 --- a/Mathlib/NumberTheory/PythagoreanTriples.lean +++ b/Mathlib/NumberTheory/PythagoreanTriples.lean @@ -29,7 +29,7 @@ the bulk of the proof below. theorem sq_ne_two_fin_zmod_four (z : ZMod 4) : z * z ≠ 2 := by change Fin 4 at z - fin_cases z <;> norm_num [Fin.ext_iff] + fin_cases z <;> decide #align sq_ne_two_fin_zmod_four sq_ne_two_fin_zmod_four theorem Int.sq_ne_two_mod_four (z : ℤ) : z * z % 4 ≠ 2 := by @@ -155,7 +155,8 @@ theorem even_odd_of_coprime (hc : Int.gcd x y = 1) : rw [show z * z = 4 * (x0 * x0 + x0 + y0 * y0 + y0) + 2 by rw [← h.eq] ring] - norm_num [Int.add_emod] + simp only [Int.add_emod, Int.mul_emod_right, zero_add] + decide #align pythagorean_triple.even_odd_of_coprime PythagoreanTriple.even_odd_of_coprime theorem gcd_dvd : (Int.gcd x y : ℤ) ∣ z := by @@ -243,10 +244,10 @@ theorem isPrimitiveClassified_of_coprime_of_zero_left (hc : Int.gcd x y = 1) (hx cases' Int.natAbs_eq y with hy hy · use 1, 0 rw [hy, hc, Int.gcd_zero_right] - norm_num + decide · use 0, 1 rw [hy, hc, Int.gcd_zero_left] - norm_num + decide #align pythagorean_triple.is_primitive_classified_of_coprime_of_zero_left PythagoreanTriple.isPrimitiveClassified_of_coprime_of_zero_left theorem coprime_of_coprime (hc : Int.gcd x y = 1) : Int.gcd y z = 1 := by @@ -339,7 +340,9 @@ private theorem coprime_sq_sub_sq_add_of_even_odd {m n : ℤ} (h : Int.gcd m n = have hnc : p = 2 ∨ p ∣ Int.natAbs n := prime_two_or_dvd_of_dvd_two_mul_pow_self_two hp h2n by_cases h2 : p = 2 -- Porting note: norm_num is not enough to close h3 - · have h3 : (m ^ 2 + n ^ 2) % 2 = 1 := by field_simp [sq, Int.add_emod, Int.mul_emod, hm, hn] + · have h3 : (m ^ 2 + n ^ 2) % 2 = 1 := by + simp only [sq, Int.add_emod, Int.mul_emod, hm, hn, dvd_refl, Int.emod_emod_of_dvd] + decide have h4 : (m ^ 2 + n ^ 2) % 2 = 0 := by apply Int.emod_eq_zero_of_dvd rwa [h2] at hp2 @@ -372,15 +375,17 @@ private theorem coprime_sq_sub_mul_of_even_odd {m n : ℤ} (h : Int.gcd m n = 1) rw [hp2'] apply mt Int.emod_eq_zero_of_dvd -- Porting note: norm_num is not enough to close this - field_simp [sq, Int.sub_emod, Int.mul_emod, hm, hn] + simp only [sq, Nat.cast_ofNat, Int.sub_emod, Int.mul_emod, hm, hn, + mul_zero, EuclideanDomain.zero_mod, mul_one, zero_sub] + decide apply mt (Int.dvd_gcd (Int.coe_nat_dvd_left.mpr hpm)) hnp - apply (or_self_iff _).mp + apply or_self_iff.mp apply Int.Prime.dvd_mul' hp rw [(by ring : n * n = -(m ^ 2 - n ^ 2) + m * m)] exact hp1.neg_right.add ((Int.coe_nat_dvd_left.2 hpm).mul_right _) rw [Int.gcd_comm] at hnp apply mt (Int.dvd_gcd (Int.coe_nat_dvd_left.mpr hpn)) hnp - apply (or_self_iff _).mp + apply or_self_iff.mp apply Int.Prime.dvd_mul' hp rw [(by ring : m * m = m ^ 2 - n ^ 2 + n * n)] apply dvd_add hp1 @@ -485,6 +490,11 @@ theorem isPrimitiveClassified_of_coprime_of_odd_of_pos (hc : Int.gcd x y = 1) (h let m := (q.den : ℤ) let n := q.num have hm0 : m ≠ 0 := by + -- Added to adapt to leanprover/lean4#2734. + -- Without `unfold_let`, `norm_cast` can't see the coercion. + -- One might try `zeta := true` in `Tactic.NormCast.derive`, + -- but that seems to break many other things. + unfold_let m norm_cast apply Rat.den_nz q have hq2 : q = n / m := (Rat.num_div_den q).symm @@ -523,7 +533,7 @@ theorem isPrimitiveClassified_of_coprime_of_odd_of_pos (hc : Int.gcd x y = 1) (h Int.dvd_gcd (Int.dvd_of_emod_eq_zero hn2) (Int.dvd_of_emod_eq_zero hm2) rw [hnmcp] at h1 revert h1 - norm_num + decide · -- m even, n odd apply h.isPrimitiveClassified_aux hc hzpos hm2n2 hv2 hw2 _ hmncp · apply Or.intro_left diff --git a/Mathlib/NumberTheory/Rayleigh.lean b/Mathlib/NumberTheory/Rayleigh.lean index 9b511a02cc709..6730de57895ac 100644 --- a/Mathlib/NumberTheory/Rayleigh.lean +++ b/Mathlib/NumberTheory/Rayleigh.lean @@ -117,8 +117,8 @@ theorem compl_beattySeq {r s : ℝ} (hrs : r.IsConjugateExponent s) : ext j by_cases h₁ : j ∈ {beattySeq r k | k} <;> by_cases h₂ : j ∈ {beattySeq' s k | k} · exact (Set.not_disjoint_iff.2 ⟨j, h₁, h₂⟩ (Beatty.no_collision hrs)).elim - · simp only [Set.mem_compl_iff, h₁, h₂] - · simp only [Set.mem_compl_iff, h₁, h₂] + · simp only [Set.mem_compl_iff, h₁, h₂, not_true_eq_false] + · simp only [Set.mem_compl_iff, h₁, h₂, not_false_eq_true] · have ⟨k, h₁₁, h₁₂⟩ := (Beatty.hit_or_miss hrs.pos).resolve_left h₁ have ⟨m, h₂₁, h₂₂⟩ := (Beatty.hit_or_miss' hrs.symm.pos).resolve_left h₂ exact (Beatty.no_anticollision hrs ⟨j, k, m, h₁₁, h₁₂, h₂₁, h₂₂⟩).elim diff --git a/Mathlib/NumberTheory/SumFourSquares.lean b/Mathlib/NumberTheory/SumFourSquares.lean index 05e44b1be541c..dae91282da785 100644 --- a/Mathlib/NumberTheory/SumFourSquares.lean +++ b/Mathlib/NumberTheory/SumFourSquares.lean @@ -187,7 +187,7 @@ protected theorem Prime.sum_four_squares {p : ℕ} (hp : p.Prime) : rcases (zero_le r).eq_or_gt with rfl | hr₀ · replace hr : f a = 0 ∧ f b = 0 ∧ f c = 0 ∧ f d = 0; · simpa [and_assoc] using hr obtain ⟨⟨a, rfl⟩, ⟨b, rfl⟩, ⟨c, rfl⟩, ⟨d, rfl⟩⟩ : m ∣ a ∧ m ∣ b ∧ m ∣ c ∧ m ∣ d - · simp only [← ZMod.nat_cast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero] + · simp only [← ZMod.nat_cast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero, and_self] have : m * m ∣ m * p := habcd ▸ ⟨a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2, by ring⟩ rw [mul_dvd_mul_iff_left hm₀.ne'] at this exact (hp.eq_one_or_self_of_dvd _ this).elim hm₁.ne' hmp.ne diff --git a/Mathlib/NumberTheory/SumTwoSquares.lean b/Mathlib/NumberTheory/SumTwoSquares.lean index 4936e6c381629..4829dcc69b667 100644 --- a/Mathlib/NumberTheory/SumTwoSquares.lean +++ b/Mathlib/NumberTheory/SumTwoSquares.lean @@ -109,7 +109,7 @@ theorem ZMod.isSquare_neg_one_iff {n : ℕ} (hn : Squarefree n) : refine' ⟨fun H q hqp hqd => hqp.mod_four_ne_three_of_dvd_isSquare_neg_one hqd H, fun H => _⟩ induction' n using induction_on_primes with p n hpp ih · exact False.elim (hn.ne_zero rfl) - · exact ⟨0, by simp only [Fin.zero_mul, neg_eq_zero, Fin.one_eq_zero_iff]⟩ + · exact ⟨0, by simp only [mul_zero, eq_iff_true_of_subsingleton]⟩ · haveI : Fact p.Prime := ⟨hpp⟩ have hcp : p.Coprime n := by by_contra hc diff --git a/Mathlib/NumberTheory/ZetaFunction.lean b/Mathlib/NumberTheory/ZetaFunction.lean index 840977beb5e3e..18a650a539655 100644 --- a/Mathlib/NumberTheory/ZetaFunction.lean +++ b/Mathlib/NumberTheory/ZetaFunction.lean @@ -30,7 +30,7 @@ I haven't checked exactly what they are). * `differentiableAt_completed_zeta` : the function `Λ(s)` is differentiable away from `s = 0` and `s = 1`. * `differentiableAt_riemannZeta` : the function `ζ(s)` is differentiable away from `s = 1`. -* `zeta_eq_tsum_of_one_lt_re` : for `1 < re s`, we have +* `zeta_eq_tsum_one_div_nat_add_one_cpow` : for `1 < re s`, we have `ζ(s) = ∑' (n : ℕ), 1 / (n + 1) ^ s`. * `riemannCompletedZeta₀_one_sub`, `riemannCompletedZeta_one_sub`, and `riemannZeta_one_sub` : functional equation relating values at `s` and `1 - s` diff --git a/Mathlib/Order/Atoms.lean b/Mathlib/Order/Atoms.lean index 1a12752c55db4..0d860dfeedfc2 100644 --- a/Mathlib/Order/Atoms.lean +++ b/Mathlib/Order/Atoms.lean @@ -1032,7 +1032,7 @@ theorem isAtom_iff_eq_single [DecidableEq ι] [∀ i, PartialOrder (π i)] rw [Function.update_noteq hij, hbot _ hij, bot_apply] case mpr => rintro ⟨i, a, h, rfl⟩ - refine isAtom_single h + exact isAtom_single h instance isAtomic [∀ i, PartialOrder (π i)] [∀ i, OrderBot (π i)] [∀ i, IsAtomic (π i)] : IsAtomic (∀ i, π i) where @@ -1067,7 +1067,7 @@ instance isAtomistic [∀ i, CompleteLattice (π i)] [∀ i, IsAtomistic (π i)] rintro _ ⟨⟨_, ⟨j, a, ha, rfl⟩, hle⟩, rfl⟩ by_cases hij : i = j; case neg => simp [Function.update_noteq hij] subst hij; simp only [Function.update_same] - refine le_sSup ⟨ha, by simpa using hle i⟩ + exact le_sSup ⟨ha, by simpa using hle i⟩ instance isCoatomistic [∀ i, CompleteLattice (π i)] [∀ i, IsCoatomistic (π i)] : IsCoatomistic (∀ i, π i) := diff --git a/Mathlib/Order/Birkhoff.lean b/Mathlib/Order/Birkhoff.lean new file mode 100644 index 0000000000000..919bf57422d7a --- /dev/null +++ b/Mathlib/Order/Birkhoff.lean @@ -0,0 +1,233 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Fintype.Order +import Mathlib.Order.Irreducible +import Mathlib.Order.UpperLower.Basic + +/-! +# Birkhoff's representation theorem + +This file proves the Birkhoff representation theorem: Any finite distributive lattice can be +represented as a sublattice of some powerset algebra. + +Precisely, any nonempty finite distributive lattice is isomorphic to the lattice of lower sets of +its irreducible elements. And conversely it is isomorphic to the order of its irreducible lower +sets. + +## Main declarations + +For a nonempty finite distributive lattice `α`: +* `OrderIso.lowerSetSupIrred`: `α` is isomorphic to the lattice of lower sets of its irreducible + elements. +* `OrderIso.supIrredLowerSet`: `α` is isomorphic to the order of its irreducible lower sets. +* `OrderEmbedding.birkhoffSet`, `OrderEmbedding.birkhoffFinset`: Order embedding of `α` into the + powerset lattice of its irreducible elements. +* `LatticeHom.birkhoffSet`, `LatticeHom.birkhoffFinet`: Same as the previous two, but bundled as + an injective lattice homomorphism. +* `exists_birkhoff_representation`: `α` embeds into some powerset algebra. You should prefer using + this over the explicit Birkhoff embedding because the Birkhoff embedding is littered with + decidability footguns that this existential-packaged version can afford to avoid. + +## See also + +This correspondence between finite distributive lattices and finite boolean algebras is made +functorial in... TODO: Actually do it. + +## Tags + +birkhoff, representation, stone duality, lattice embedding +-/ + +open Finset Function OrderDual + +variable {α : Type*} + +namespace UpperSet +variable [SemilatticeInf α] {s : UpperSet α} {a : α} + +@[simp] lemma infIrred_Ici (a : α) : InfIrred (Ici a) := by + refine ⟨fun h ↦ Ici_ne_top h.eq_top, fun s t hst ↦ ?_⟩ + have := mem_Ici_iff.2 (le_refl a) + rw [←hst] at this + exact this.imp (fun ha ↦ le_antisymm (le_Ici.2 ha) $ hst.ge.trans inf_le_left) fun ha ↦ + le_antisymm (le_Ici.2 ha) $ hst.ge.trans inf_le_right + +variable [Finite α] + +@[simp] lemma infIrred_iff_of_finite : InfIrred s ↔ ∃ a, Ici a = s := by + refine' ⟨fun hs ↦ _, _⟩ + · obtain ⟨a, ha, has⟩ := (s : Set α).toFinite.exists_minimal_wrt id _ (coe_nonempty.2 hs.ne_top) + exact ⟨a, (hs.2 $ erase_inf_Ici ha $ by simpa [eq_comm] using has).resolve_left + (lt_erase.2 ha).ne'⟩ + · rintro ⟨a, rfl⟩ + exact infIrred_Ici _ + +end UpperSet + +namespace LowerSet +variable [SemilatticeSup α] {s : LowerSet α} {a : α} + +@[simp] lemma supIrred_Iic (a : α) : SupIrred (Iic a) := by + refine' ⟨fun h ↦ Iic_ne_bot h.eq_bot, fun s t hst ↦ _⟩ + have := mem_Iic_iff.2 (le_refl a) + rw [←hst] at this + exact this.imp (fun ha ↦ (le_sup_left.trans_eq hst).antisymm $ Iic_le.2 ha) fun ha ↦ + (le_sup_right.trans_eq hst).antisymm $ Iic_le.2 ha + +variable [Finite α] + +@[simp] lemma supIrred_iff_of_finite : SupIrred s ↔ ∃ a, Iic a = s := by + refine' ⟨fun hs ↦ _, _⟩ + · obtain ⟨a, ha, has⟩ := (s : Set α).toFinite.exists_maximal_wrt id _ (coe_nonempty.2 hs.ne_bot) + exact ⟨a, (hs.2 $ erase_sup_Iic ha $ by simpa [eq_comm] using has).resolve_left + (erase_lt.2 ha).ne⟩ + · rintro ⟨a, rfl⟩ + exact supIrred_Iic _ + +end LowerSet + +section DistribLattice +variable [DistribLattice α] + +section Fintype +variable [Fintype α] [OrderBot α] + +open scoped Classical + +/-- **Birkhoff's Representation Theorem**. Any nonempty finite distributive lattice is isomorphic to +the lattice of lower sets of its sup-irreducible elements. -/ +noncomputable def OrderIso.lowerSetSupIrred : α ≃o LowerSet {a : α // SupIrred a} := + Equiv.toOrderIso + { toFun := fun a ↦ ⟨{b | ↑b ≤ a}, fun b c hcb hba ↦ hba.trans' hcb⟩ + invFun := fun s ↦ (s : Set {a : α // SupIrred a}).toFinset.sup (↑) + left_inv := fun a ↦ by + refine' le_antisymm (Finset.sup_le fun b ↦ Set.mem_toFinset.1) _ + obtain ⟨s, rfl, hs⟩ := exists_supIrred_decomposition a + exact Finset.sup_le fun i hi ↦ + le_sup_of_le (b := ⟨i, hs hi⟩) (Set.mem_toFinset.2 $ le_sup (f := id) hi) le_rfl + right_inv := fun s ↦ by + ext a + dsimp + refine' ⟨fun ha ↦ _, fun ha ↦ _⟩ + · obtain ⟨i, hi, ha⟩ := a.2.supPrime.le_finset_sup.1 ha + exact s.lower ha (Set.mem_toFinset.1 hi) + · dsimp + exact le_sup (Set.mem_toFinset.2 ha) } + (fun b c hbc d ↦ le_trans' hbc) fun s t hst ↦ Finset.sup_mono $ Set.toFinset_mono hst + +-- We remove this instance locally to let `Set.toFinset_Iic` fire. When the instance is present, +-- `simp` refuses to use that lemma because TC inference synthesizes `Set.fintypeIic`, which is not +-- defeq to the instance it finds in the term. +attribute [-instance] Set.fintypeIic + +/-- Any nonempty finite distributive lattice is isomorphic to its lattice of sup-irreducible lower +sets. This is the other unitor of Birkhoff's representation theorem. -/ +noncomputable def OrderIso.supIrredLowerSet : α ≃o {s : LowerSet α // SupIrred s} := + Equiv.toOrderIso + { toFun := fun a ↦ ⟨LowerSet.Iic a, LowerSet.supIrred_Iic _⟩ + invFun := fun s ↦ ((s : LowerSet α) : Set α).toFinset.sup id + left_inv := fun a ↦ by + have : LocallyFiniteOrder α := Fintype.toLocallyFiniteOrder + simp + right_inv := by + classical + have : LocallyFiniteOrder α := Fintype.toLocallyFiniteOrder + rintro ⟨s, hs⟩ + obtain ⟨a, rfl⟩ := LowerSet.supIrred_iff_of_finite.1 hs + simp } + (fun b c hbc d ↦ le_trans' hbc) fun s t hst ↦ Finset.sup_mono $ Set.toFinset_mono hst + +end Fintype + +variable (α) + +namespace OrderEmbedding +variable [Fintype α] [@DecidablePred α SupIrred] + +/-- **Birkhoff's Representation Theorem**. Any finite distributive lattice can be embedded in a +powerset lattice. -/ +noncomputable def birkhoffSet : α ↪o Set {a : α // SupIrred a} := by + by_cases IsEmpty α + · exact OrderEmbedding.ofIsEmpty + rw [not_isEmpty_iff] at h + have := Fintype.toOrderBot α + exact OrderIso.lowerSetSupIrred.toOrderEmbedding.trans ⟨⟨_, SetLike.coe_injective⟩, Iff.rfl⟩ + +/-- **Birkhoff's Representation Theorem**. Any finite distributive lattice can be embedded in a +powerset lattice. -/ +noncomputable def birkhoffFinset : α ↪o Finset {a : α // SupIrred a} := by + exact (birkhoffSet _).trans Fintype.finsetOrderIsoSet.symm.toOrderEmbedding + +variable {α} + +@[simp] lemma coe_birkhoffFinset (a : α) : birkhoffFinset α a = birkhoffSet α a := by + classical + -- TODO: This should be a single `simp` call but `simp` refuses to use + -- `OrderIso.coe_toOrderEmbedding` and `Fintype.coe_finsetOrderIsoSet_symm` + simp [birkhoffFinset] + rw [OrderIso.coe_toOrderEmbedding, Fintype.coe_finsetOrderIsoSet_symm] + simp + +@[simp] lemma birkhoffSet_sup (a b : α) : + birkhoffSet α (a ⊔ b) = birkhoffSet α a ∪ birkhoffSet α b := by + unfold OrderEmbedding.birkhoffSet; split <;> simp + +@[simp] lemma birkhoffSet_inf (a b : α) : + birkhoffSet α (a ⊓ b) = birkhoffSet α a ∩ birkhoffSet α b := by + unfold OrderEmbedding.birkhoffSet; split <;> simp + +variable [DecidableEq α] + +@[simp] lemma birkhoffFinset_sup (a b : α) : + birkhoffFinset α (a ⊔ b) = birkhoffFinset α a ∪ birkhoffFinset α b := by + dsimp [OrderEmbedding.birkhoffFinset] + rw [birkhoffSet_sup, OrderIso.coe_toOrderEmbedding] + simpa using OrderIso.map_sup _ _ _ + +@[simp] lemma birkhoffFinset_inf (a b : α) : + birkhoffFinset α (a ⊓ b) = birkhoffFinset α a ∩ birkhoffFinset α b := by + dsimp [OrderEmbedding.birkhoffFinset] + rw [birkhoffSet_inf, OrderIso.coe_toOrderEmbedding] + simpa using OrderIso.map_inf _ _ _ + +variable [OrderBot α] + +@[simp] lemma birkhoffSet_apply (a : α) : birkhoffSet α a = OrderIso.lowerSetSupIrred a := by + simp [birkhoffSet]; convert rfl + +end OrderEmbedding + +namespace LatticeHom +variable [Fintype α] [DecidableEq α] [@DecidablePred α SupIrred] + +/-- **Birkhoff's Representation Theorem**. Any finite distributive lattice can be embedded in a +powerset lattice. -/ +noncomputable def birkhoffSet : LatticeHom α (Set {a : α // SupIrred a}) where + toFun := OrderEmbedding.birkhoffSet α + map_sup' := OrderEmbedding.birkhoffSet_sup + map_inf' := OrderEmbedding.birkhoffSet_inf + +/-- **Birkhoff's Representation Theorem**. Any finite distributive lattice can be embedded in a +powerset lattice. -/ +noncomputable def birkhoffFinset : LatticeHom α (Finset {a : α // SupIrred a}) where + toFun := OrderEmbedding.birkhoffFinset α + map_sup' := OrderEmbedding.birkhoffFinset_sup + map_inf' := OrderEmbedding.birkhoffFinset_inf + +lemma birkhoffFinset_injective : Injective (birkhoffFinset α) := + (OrderEmbedding.birkhoffFinset α).injective + +end LatticeHom + +lemma exists_birkhoff_representation.{u} (α : Type u) [Finite α] [DistribLattice α] : + ∃ (β : Type u) (_ : DecidableEq β) (_ : Fintype β) (f : LatticeHom α (Finset β)), + Injective f := by + classical + cases nonempty_fintype α + exact ⟨{a : α // SupIrred a}, _ , by infer_instance, _, LatticeHom.birkhoffFinset_injective _⟩ + +end DistribLattice diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index 78f8b54bbacdf..36303f1944a2f 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -785,20 +785,13 @@ instance Pi.booleanAlgebra {ι : Type u} {α : ι → Type v} [∀ i, BooleanAlg top_le_sup_compl := fun _ _ => BooleanAlgebra.top_le_sup_compl _ } #align pi.boolean_algebra Pi.booleanAlgebra -instance : BooleanAlgebra Bool := - { Bool.linearOrder, Bool.boundedOrder with - sup := or, - le_sup_left := Bool.left_le_or, - le_sup_right := Bool.right_le_or, - sup_le := fun _ _ _ => Bool.or_le, - inf := and, - inf_le_left := Bool.and_le_left, - inf_le_right := Bool.and_le_right, - le_inf := fun _ _ _ => Bool.le_and, - le_sup_inf := by decide, - compl := not, - inf_compl_le_bot := fun a => a.and_not_self.le, - top_le_sup_compl := fun a => a.or_not_self.ge } +instance Bool.instBooleanAlgebra : BooleanAlgebra Bool where + __ := Bool.linearOrder + __ := Bool.boundedOrder + __ := Bool.instDistribLattice + compl := not + inf_compl_le_bot a := a.and_not_self.le + top_le_sup_compl a := a.or_not_self.ge @[simp] theorem Bool.sup_eq_bor : (· ⊔ ·) = or := diff --git a/Mathlib/Order/BoundedOrder.lean b/Mathlib/Order/BoundedOrder.lean index 9a806f27bea35..75a44a714c647 100644 --- a/Mathlib/Order/BoundedOrder.lean +++ b/Mathlib/Order/BoundedOrder.lean @@ -942,9 +942,9 @@ open Bool instance Bool.boundedOrder : BoundedOrder Bool where top := true - le_top _ := le_true + le_top := Bool.le_true bot := false - bot_le _ := false_le + bot_le := Bool.false_le @[simp] theorem top_eq_true : ⊤ = true := diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index 7cce552a3c98e..9b0166eed10dc 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -67,7 +67,7 @@ def IsLeast (s : Set α) (a : α) : Prop := a ∈ s ∧ a ∈ lowerBounds s #align is_least IsLeast -/-- `a` is a greatest element of a set `s`; for a partial order, it is unique if exists -/ +/-- `a` is a greatest element of a set `s`; for a partial order, it is unique if exists. -/ def IsGreatest (s : Set α) (a : α) : Prop := a ∈ s ∧ a ∈ upperBounds s #align is_greatest IsGreatest diff --git a/Mathlib/Order/Circular.lean b/Mathlib/Order/Circular.lean index 476cc42122201..8599dc8812675 100644 --- a/Mathlib/Order/Circular.lean +++ b/Mathlib/Order/Circular.lean @@ -289,7 +289,7 @@ section CircularOrder variable {α : Type*} [CircularOrder α] theorem btw_refl_left_right (a b : α) : btw a b a := - (or_self_iff _).1 (btw_total a b a) + or_self_iff.1 (btw_total a b a) #align btw_refl_left_right btw_refl_left_right theorem btw_rfl_left_right {a b : α} : btw a b a := diff --git a/Mathlib/Order/Closure.lean b/Mathlib/Order/Closure.lean index 8f587431dd0ba..281b32c20f8c6 100644 --- a/Mathlib/Order/Closure.lean +++ b/Mathlib/Order/Closure.lean @@ -47,8 +47,7 @@ place when using concrete closure operators such as `ConvexHull`. * https://en.wikipedia.org/wiki/Closure_operator#Closure_operators_on_partially_ordered_sets -/ - -universe u +open Set /-! ### Closure operator -/ @@ -272,7 +271,25 @@ end SemilatticeSup section CompleteLattice -variable [CompleteLattice α] (c : ClosureOperator α) +variable [CompleteLattice α] (c : ClosureOperator α) {p : α → Prop} + +/-- Define a closure operator from a predicate that's preserved under infima. -/ +def ofPred (p : α → Prop) (hsinf : ∀ s, (∀ a ∈ s, p a) → p (sInf s)) : ClosureOperator α := + ClosureOperator.mk₃ (fun a ↦ ⨅ b : {b // p b ∧ a ≤ b}, b) p + (fun a ↦ by simp [forall_swap]) + (fun a ↦ hsinf _ $ forall_range_iff.2 $ fun b ↦ b.2.1) + (fun a b hab hb ↦ iInf_le_of_le ⟨b, hb, hab⟩ le_rfl) + +/-- This lemma shows that the image of `x` of a closure operator built from the `ofPred` constructor +respects `p`, the property that was fed into it. -/ +lemma ofPred_spec {sinf} (x : α) : p (ofPred p sinf x) := closure_mem_mk₃ _ + +/-- The property `p` fed into the `ofPred` constructor exactly corresponds to being closed. -/ +@[simp] lemma closed_ofPred {hsinf} : (ofPred p hsinf).closed = {x | p x} := by + ext; exact mem_mk₃_closed + +/-- The property `p` fed into the `ofPred` constructor exactly corresponds to being closed. -/ +lemma mem_closed_ofPred {hsinf} {x : α} : x ∈ (ofPred p hsinf).closed ↔ p x := mem_mk₃_closed @[simp] theorem closure_iSup_closure (f : ι → α) : c (⨆ i, c (f i)) = c (⨆ i, f i) := diff --git a/Mathlib/Order/CompactlyGenerated.lean b/Mathlib/Order/CompactlyGenerated.lean index 8a1a1f1c6e01e..87579a53fedc2 100644 --- a/Mathlib/Order/CompactlyGenerated.lean +++ b/Mathlib/Order/CompactlyGenerated.lean @@ -53,6 +53,8 @@ This is demonstrated by means of the following four lemmas: complete lattice, well-founded, compact -/ +open Set + variable {ι : Sort*} {α : Type*} [CompleteLattice α] {f : ι → α} namespace CompleteLattice @@ -438,6 +440,28 @@ theorem CompleteLattice.setIndependent_iff_finite {s : Set α} : exact ⟨ha, Set.Subset.trans ht (Set.diff_subset _ _)⟩⟩ #align complete_lattice.set_independent_iff_finite CompleteLattice.setIndependent_iff_finite +lemma CompleteLattice.independent_iff_supIndep_of_injOn {ι : Type*} {f : ι → α} + (hf : InjOn f {i | f i ≠ ⊥}) : + CompleteLattice.Independent f ↔ ∀ (s : Finset ι), s.SupIndep f := by + refine ⟨fun h ↦ h.supIndep', fun h ↦ CompleteLattice.independent_def'.mpr fun i ↦ ?_⟩ + simp_rw [disjoint_iff, inf_sSup_eq_iSup_inf_sup_finset, iSup_eq_bot, ← disjoint_iff] + intro s hs + classical + rw [← Finset.sup_erase_bot] + set t := s.erase ⊥ + replace hf : InjOn f (f ⁻¹' t) := fun i hi j _ hij ↦ by refine hf ?_ ?_ hij <;> aesop + have : (Finset.erase (insert i (t.preimage _ hf)) i).image f = t := by + ext a + simp only [Finset.mem_preimage, Finset.mem_erase, ne_eq, Finset.mem_insert, true_or, not_true, + Finset.erase_insert_eq_erase, not_and, Finset.mem_image] + refine ⟨by aesop, fun ⟨ha, has⟩ ↦ ?_⟩ + obtain ⟨j, hj, rfl⟩ := hs has + exact ⟨j, ⟨hj, ha, has⟩, rfl⟩ + rw [← this, Finset.sup_image] + specialize h (insert i (t.preimage _ hf)) + rw [Finset.supIndep_iff_disjoint_erase] at h + exact h i (Finset.mem_insert_self i _) + theorem CompleteLattice.setIndependent_iUnion_of_directed {η : Type*} {s : η → Set α} (hs : Directed (· ⊆ ·) s) (h : ∀ i, CompleteLattice.SetIndependent (s i)) : CompleteLattice.SetIndependent (⋃ i, s i) := by diff --git a/Mathlib/Order/CompleteBooleanAlgebra.lean b/Mathlib/Order/CompleteBooleanAlgebra.lean index 59267a5e4668d..d07886f54c749 100644 --- a/Mathlib/Order/CompleteBooleanAlgebra.lean +++ b/Mathlib/Order/CompleteBooleanAlgebra.lean @@ -131,7 +131,7 @@ theorem iSup_iInf_eq [CompletelyDistribLattice α] {f : ∀ a, κ a → α} : refine le_trans ?_ (le_iSup _ a) refine le_iInf fun b => ?_ obtain ⟨h, rfl, rfl⟩ := ha b - refine iInf_le _ _ + exact iInf_le _ _ instance (priority := 100) CompletelyDistribLattice.toCompleteDistribLattice [CompletelyDistribLattice α] : CompleteDistribLattice α where @@ -182,7 +182,7 @@ instance (priority := 100) CompleteLinearOrder.toCompletelyDistribLattice [Compl have : ∀ a, lhs ≤ g a (f a) := fun a => (h (g a (f a))).resolve_left (by simpa using hf a) refine le_trans ?_ (le_iSup _ f) - refine le_iInf fun a => this _ + exact le_iInf fun a => this _ section Frame diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index 352b2f5381ba7..e1ab0d32f75a5 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -693,6 +693,10 @@ theorem iSup_congr (h : ∀ i, f i = g i) : ⨆ i, f i = ⨆ i, g i := congr_arg _ <| funext h #align supr_congr iSup_congr +theorem biSup_congr {p : ι → Prop} (h : ∀ i, p i → f i = g i) : + ⨆ (i) (_ : p i), f i = ⨆ (i) (_ : p i), g i := + iSup_congr <| fun i ↦ iSup_congr (h i) + theorem Function.Surjective.iSup_comp {f : ι → ι'} (hf : Surjective f) (g : ι' → α) : ⨆ x, g (f x) = ⨆ y, g y := by simp only [iSup._eq_1] @@ -758,6 +762,10 @@ theorem iInf_congr (h : ∀ i, f i = g i) : ⨅ i, f i = ⨅ i, g i := congr_arg _ <| funext h #align infi_congr iInf_congr +theorem biInf_congr {p : ι → Prop} (h : ∀ i, p i → f i = g i) : + ⨅ (i) (_ : p i), f i = ⨅ (i) (_ : p i), g i := + biSup_congr (α := αᵒᵈ) h + theorem Function.Surjective.iInf_comp {f : ι → ι'} (hf : Surjective f) (g : ι' → α) : ⨅ x, g (f x) = ⨅ y, g y := @Function.Surjective.iSup_comp αᵒᵈ _ _ _ f hf g diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index b50356751f203..5af75177e59f0 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -1135,8 +1135,7 @@ open Function variable [IsWellOrder α (· < ·)] -theorem sInf_eq_argmin_on (hs : s.Nonempty) : - sInf s = argminOn id (@IsWellFounded.wf α (· < ·) _) s hs := +theorem sInf_eq_argmin_on (hs : s.Nonempty) : sInf s = argminOn id wellFounded_lt s hs := IsLeast.csInf_eq ⟨argminOn_mem _ _ _ _, fun _ ha => argminOn_le id _ _ ha⟩ #align Inf_eq_argmin_on sInf_eq_argmin_on @@ -1364,7 +1363,7 @@ theorem isGLB_sInf' {β : Type*} [ConditionallyCompleteLattice β] {s : Set (Wit contrapose! h rintro (⟨⟩ | a) ha · exact mem_singleton ⊤ - · exact (h ⟨a, ha⟩).elim + · exact (not_nonempty_iff_eq_empty.2 h ⟨a, ha⟩).elim · intro b hb rw [← some_le_some] exact ha hb diff --git a/Mathlib/Order/Disjoint.lean b/Mathlib/Order/Disjoint.lean index ab8ef9a9853f5..812ba28c3688c 100644 --- a/Mathlib/Order/Disjoint.lean +++ b/Mathlib/Order/Disjoint.lean @@ -42,6 +42,10 @@ def Disjoint (a b : α) : Prop := ∀ ⦃x⦄, x ≤ a → x ≤ b → x ≤ ⊥ #align disjoint Disjoint +@[simp] +theorem disjoint_of_subsingleton [Subsingleton α] : Disjoint a b := + fun x _ _ ↦ le_of_eq (Subsingleton.elim x ⊥) + theorem disjoint_comm : Disjoint a b ↔ Disjoint b a := forall_congr' fun _ ↦ forall_swap #align disjoint.comm disjoint_comm diff --git a/Mathlib/Order/Estimator.lean b/Mathlib/Order/Estimator.lean index b8b597a7c56a4..763f956bfc634 100644 --- a/Mathlib/Order/Estimator.lean +++ b/Mathlib/Order/Estimator.lean @@ -7,6 +7,7 @@ import Mathlib.Order.RelClasses import Mathlib.Data.Set.Image import Mathlib.Order.Hom.Basic import Mathlib.Lean.Thunk +import Mathlib.Data.Prod.Lex /-! # Improvable lower bounds. @@ -32,6 +33,7 @@ the exact value. -/ set_option autoImplicit true +set_option relaxedAutoImplicit true /-- Given `[EstimatorData a ε]` @@ -63,18 +65,38 @@ class Estimator [Preorder α] (a : Thunk α) (ε : Type*) extends EstimatorData open EstimatorData Set -/-- -A trivial estimator. --/ -instance [Preorder α] (a : α) : Estimator (Thunk.pure a) { x // x = a } where - bound e := e +section trivial + +variable [Preorder α] + +/-- A trivial estimator, containing the actual value. -/ +abbrev Estimator.trivial (a : α) : Type* := { b : α // b = a } + +instance : Bot (Estimator.trivial a) := ⟨⟨a, rfl⟩⟩ + +instance : WellFoundedGT Unit where + wf := ⟨fun .unit => ⟨.unit, fun.⟩⟩ + +instance (a : α) : WellFoundedGT (Estimator.trivial a) := + let f : Estimator.trivial a ≃o Unit := RelIso.relIsoOfUniqueOfRefl _ _ + let f' : Estimator.trivial a ↪o Unit := f.toOrderEmbedding + f'.wellFoundedGT + +instance {a : α} : Estimator (Thunk.pure a) (Estimator.trivial a) where + bound b := b.val improve _ := none - bound_le e := e.property.le - improve_spec e := e.property + bound_le b := b.prop.le + improve_spec b := b.prop + +end trivial + +section improveUntil + +variable [Preorder α] attribute [local instance] WellFoundedGT.toWellFoundedRelation in /-- Implementation of `Estimator.improveUntil`. -/ -def Estimator.improveUntilAux [PartialOrder α] +def Estimator.improveUntilAux (a : Thunk α) (p : α → Bool) [Estimator a ε] [WellFoundedGT (range (bound a : ε → α))] (e : ε) (r : Bool) : Except (Option ε) ε := @@ -91,12 +113,11 @@ termination_by Estimator.improveUntilAux p I e r => (⟨_, mem_range_self e⟩ : Improve an estimate until it satisfies a predicate, or else return the best available estimate, if any improvement was made. -/ -def Estimator.improveUntil [PartialOrder α] (a : Thunk α) (p : α → Bool) +def Estimator.improveUntil (a : Thunk α) (p : α → Bool) [Estimator a ε] [WellFoundedGT (range (bound a : ε → α))] (e : ε) : Except (Option ε) ε := Estimator.improveUntilAux a p e false -variable [PartialOrder α] attribute [local instance] WellFoundedGT.toWellFoundedRelation in /-- @@ -132,25 +153,63 @@ theorem Estimator.improveUntil_spec | .ok e' => p (bound a e') := Estimator.improveUntilAux_spec a p e false -variable [∀ a : α, WellFoundedGT { x // x ≤ a }] +end improveUntil -instance [Estimator a ε] : WellFoundedGT (range (bound a : ε → α)) := - let f : range (bound a : ε → α) ↪o { x // x ≤ a.get } := - Subtype.orderEmbedding (by rintro _ ⟨e, rfl⟩; exact Estimator.bound_le e) - f.wellFoundedGT +/-! Estimators for sums. -/ +section add + +variable [Preorder α] + +@[simps] +instance [Add α] {a b : Thunk α} (εa εb : Type*) [EstimatorData a εa] [EstimatorData b εb] : + EstimatorData (a + b) (εa × εb) where + bound e := bound a e.1 + bound b e.2 + improve e := match improve a e.1 with + | some e' => some { e with fst := e' } + | none => match improve b e.2 with + | some e' => some { e with snd := e' } + | none => none + +instance (a b : Thunk ℕ) {εa εb : Type*} [Estimator a εa] [Estimator b εb] : + Estimator (a + b) (εa × εb) where + bound_le e := + Nat.add_le_add (Estimator.bound_le e.1) (Estimator.bound_le e.2) + improve_spec e := by + dsimp + have s₁ := Estimator.improve_spec (a := a) e.1 + have s₂ := Estimator.improve_spec (a := b) e.2 + revert s₁ s₂ + cases h₁ : improve a e.fst <;> cases h₂ : improve b e.snd <;> intro s₁ s₂ <;> simp_all only + · apply Nat.add_lt_add_left s₂ + · apply Nat.add_lt_add_right s₁ + · apply Nat.add_lt_add_right s₁ + +end add + +/-! Estimator for the first component of a pair. -/ +section fst + +variable [PartialOrder α] [PartialOrder β] /-- An estimator for `(a, b)` can be turned into an estimator for `a`, simply by repeatedly running `improve` until the first factor "improves". The hypothesis that `>` is well-founded on `{ q // q ≤ (a, b) }` ensures this terminates. -/ -structure Estimator.fst [Preorder α] [Preorder β] +structure Estimator.fst (p : Thunk (α × β)) (ε : Type*) [Estimator p ε] where /-- The wrapped bound for a value in `α × β`, which we will use as a bound for the first component. -/ inner : ε -instance [PartialOrder α] [DecidableRel ((· : α) < ·)] [PartialOrder β] {a : Thunk α} {b : Thunk β} +variable [∀ a : α, WellFoundedGT { x // x ≤ a }] + +instance [Estimator a ε] : WellFoundedGT (range (bound a : ε → α)) := + let f : range (bound a : ε → α) ↪o { x // x ≤ a.get } := + Subtype.orderEmbedding (by rintro _ ⟨e, rfl⟩; exact Estimator.bound_le e) + f.wellFoundedGT + +instance [DecidableRel ((· : α) < ·)] {a : Thunk α} {b : Thunk β} (ε : Type*) [Estimator (a.prod b) ε] [∀ (p : α × β), WellFoundedGT { q // q ≤ p }] : EstimatorData a (Estimator.fst (a.prod b) ε) where bound e := (bound (a.prod b) e.inner).1 @@ -162,8 +221,7 @@ instance [PartialOrder α] [DecidableRel ((· : α) < ·)] [PartialOrder β] {a /-- Given an estimator for a pair, we can extract an estimator for the first factor. -/ -- This isn't an instance as at the sole use case we need to provide -- the instance arguments by hand anyway. -def Estimator.fstInst [PartialOrder α] [DecidableRel ((· : α) < ·)] [PartialOrder β] - [∀ (p : α × β), WellFoundedGT { q // q ≤ p }] +def Estimator.fstInst [DecidableRel ((· : α) < ·)] [∀ (p : α × β), WellFoundedGT { q // q ≤ p }] (a : Thunk α) (b : Thunk β) {ε : Type*} (i : Estimator (a.prod b) ε) : Estimator a (Estimator.fst (a.prod b) ε) where bound_le e := (Estimator.bound_le e.inner : bound (a.prod b) e.inner ≤ (a.get, b.get)).1 @@ -179,3 +237,5 @@ def Estimator.fstInst [PartialOrder α] [DecidableRel ((· : α) < ·)] [Partial eq_of_le_of_not_lt (Estimator.bound_le e.inner : bound (a.prod b) e.inner ≤ (a.get, b.get)).1 w | .ok e' => exact fun w => w + +end fst diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index 68510ec097ce0..f7d105a86cd15 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -2053,7 +2053,8 @@ theorem Monotone.piecewise_eventually_eq_iUnion {β : α → Type*} [Preorder ι · refine (eventually_ge_atTop i).mono fun j hij ↦ ?_ simp only [Set.piecewise_eq_of_mem, hs hij hi, subset_iUnion _ _ hi] · refine eventually_of_forall fun i ↦ ?_ - simp only [Set.piecewise_eq_of_not_mem, not_exists.1 ha i, mt mem_iUnion.1 ha] + simp only [Set.piecewise_eq_of_not_mem, not_exists.1 ha i, mt mem_iUnion.1 ha, + not_false_eq_true, exists_false] theorem Antitone.piecewise_eventually_eq_iInter {β : α → Type*} [Preorder ι] {s : ι → Set α} [∀ i, DecidablePred (· ∈ s i)] [DecidablePred (· ∈ ⋂ i, s i)] diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index fc9c7730569ba..cd80ca1c4b3b2 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -944,6 +944,10 @@ theorem mem_prod_self_iff {s} : s ∈ la ×ˢ la ↔ ∃ t ∈ la, t ×ˢ t ⊆ la.basis_sets.prod_self.mem_iff #align filter.mem_prod_self_iff Filter.mem_prod_self_iff +lemma eventually_prod_self_iff {r : α → α → Prop} : + (∀ᶠ x in la ×ˢ la, r x.1 x.2) ↔ ∃ t ∈ la, ∀ x ∈ t, ∀ y ∈ t, r x y := + mem_prod_self_iff.trans <| by simp only [prod_subset_iff, mem_setOf_eq] + theorem HasAntitoneBasis.prod {ι : Type*} [LinearOrder ι] {f : Filter α} {g : Filter β} {s : ι → Set α} {t : ι → Set β} (hf : HasAntitoneBasis f s) (hg : HasAntitoneBasis g t) : HasAntitoneBasis (f ×ˢ g) fun n => s n ×ˢ t n := diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index a7beb9f1bb35b..2e2316a4a63e1 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -1310,6 +1310,12 @@ theorem Eventually.exists {p : α → Prop} {f : Filter α} [NeBot f] (hp : ∀ hp.frequently.exists #align filter.eventually.exists Filter.Eventually.exists +lemma frequently_iff_neBot {p : α → Prop} : (∃ᶠ x in l, p x) ↔ NeBot (l ⊓ 𝓟 {x | p x}) := by + rw [neBot_iff, Ne.def, inf_principal_eq_bot]; rfl + +lemma frequently_mem_iff_neBot {s : Set α} : (∃ᶠ x in l, x ∈ s) ↔ NeBot (l ⊓ 𝓟 s) := + frequently_iff_neBot + theorem frequently_iff_forall_eventually_exists_and {p : α → Prop} {f : Filter α} : (∃ᶠ x in f, p x) ↔ ∀ {q : α → Prop}, (∀ᶠ x in f, q x) → ∃ x, p x ∧ q x := ⟨fun hp q hq => (hp.and_eventually hq).exists, fun H hp => by @@ -1318,7 +1324,7 @@ theorem frequently_iff_forall_eventually_exists_and {p : α → Prop} {f : Filte theorem frequently_iff {f : Filter α} {P : α → Prop} : (∃ᶠ x in f, P x) ↔ ∀ {U}, U ∈ f → ∃ x ∈ U, P x := by - simp only [frequently_iff_forall_eventually_exists_and, exists_prop, @and_comm (P _)] + simp only [frequently_iff_forall_eventually_exists_and, @and_comm (P _)] rfl #align filter.frequently_iff Filter.frequently_iff @@ -1334,7 +1340,7 @@ theorem not_frequently {p : α → Prop} {f : Filter α} : (¬∃ᶠ x in f, p x @[simp] theorem frequently_true_iff_neBot (f : Filter α) : (∃ᶠ _ in f, True) ↔ NeBot f := by - simp [Filter.Frequently, -not_eventually, eventually_false_iff_eq_bot, neBot_iff] + simp [frequently_iff_neBot] #align filter.frequently_true_iff_ne_bot Filter.frequently_true_iff_neBot @[simp] @@ -2600,6 +2606,10 @@ theorem map_inf' {f g : Filter α} {m : α → β} {t : Set α} (htf : t ∈ f) simp only [map_map, ← map_inf Subtype.coe_injective, map_inf h] #align filter.map_inf' Filter.map_inf' +lemma disjoint_of_map {α β : Type*} {F G : Filter α} {f : α → β} + (h : Disjoint (map f F) (map f G)) : Disjoint F G := + disjoint_iff.mpr <| map_eq_bot_iff.mp <| le_bot_iff.mp <| trans map_inf_le (disjoint_iff.mp h) + theorem disjoint_map {m : α → β} (hm : Injective m) {f₁ f₂ : Filter α} : Disjoint (map m f₁) (map m f₂) ↔ Disjoint f₁ f₂ := by simp only [disjoint_iff, ← map_inf hm, map_eq_bot_iff] diff --git a/Mathlib/Order/Filter/ENNReal.lean b/Mathlib/Order/Filter/ENNReal.lean index 82cf5da48a02f..b43fbd981c410 100644 --- a/Mathlib/Order/Filter/ENNReal.lean +++ b/Mathlib/Order/Filter/ENNReal.lean @@ -65,7 +65,7 @@ theorem limsup_const_mul [CountableInterFilter f] {u : α → ℝ≥0∞} {a : have h_top_le : (f.limsup fun x : α => ite (u x = 0) (0 : ℝ≥0∞) ⊤) = ⊤ := eq_top_iff.mpr (le_limsup_of_frequently_le hu_mul) have hfu : f.limsup u ≠ 0 := mt limsup_eq_zero_iff.1 hu - simp only [ha_top, top_mul', hfu, h_top_le] + simp only [ha_top, top_mul', h_top_le, hfu, ite_false] #align ennreal.limsup_const_mul ENNReal.limsup_const_mul theorem limsup_mul_le [CountableInterFilter f] (u v : α → ℝ≥0∞) : diff --git a/Mathlib/Order/Filter/Germ.lean b/Mathlib/Order/Filter/Germ.lean index b0f462beede9b..a027428eda1f7 100644 --- a/Mathlib/Order/Filter/Germ.lean +++ b/Mathlib/Order/Filter/Germ.lean @@ -829,7 +829,7 @@ instance orderedCommMonoid [OrderedCommMonoid β] : OrderedCommMonoid (Germ l β inductionOn h fun _h => H.mono fun _x H => mul_le_mul_left' H _ } @[to_additive] -instance orderedCancelCommMonoid [OrderedCancelCommMonoid β] : +instance orderedCancelCommMonoid [OrderedCancelCommMonoid β] : OrderedCancelCommMonoid (Germ l β) := { Germ.orderedCommMonoid with le_of_mul_le_mul_left := fun f g h => diff --git a/Mathlib/Order/Filter/IndicatorFunction.lean b/Mathlib/Order/Filter/IndicatorFunction.lean index 5e405dfff62a7..e3b129df67696 100644 --- a/Mathlib/Order/Filter/IndicatorFunction.lean +++ b/Mathlib/Order/Filter/IndicatorFunction.lean @@ -31,7 +31,7 @@ theorem mulIndicator_eventuallyEq (hf : f =ᶠ[l ⊓ 𝓟 s] g) (hs : s =ᶠ[l] (eventually_inf_principal.1 hf).mp <| hs.mem_iff.mono fun x hst hfg => by_cases (fun hxs : x ∈ s => by simp only [*, hst.1 hxs, mulIndicator_of_mem]) - (fun hxs => by simp only [mulIndicator_of_not_mem, hxs, mt hst.2 hxs]) + (fun hxs => by simp only [mulIndicator_of_not_mem, hxs, mt hst.2 hxs, not_false_eq_true]) #align indicator_eventually_eq indicator_eventuallyEq end One diff --git a/Mathlib/Order/Filter/Pi.lean b/Mathlib/Order/Filter/Pi.lean index 9b85bbddf27c9..04e2f708641ff 100644 --- a/Mathlib/Order/Filter/Pi.lean +++ b/Mathlib/Order/Filter/Pi.lean @@ -5,7 +5,7 @@ Authors: Yury G. Kudryashov, Alex Kontorovich -/ import Mathlib.Order.Filter.Bases -#align_import order.filter.pi from "leanprover-community/mathlib"@"1f0096e6caa61e9c849ec2adbd227e960e9dff58" +#align_import order.filter.pi from "leanprover-community/mathlib"@"ce64cd319bb6b3e82f31c2d38e79080d377be451" /-! # (Co)product of a family of filters @@ -29,6 +29,7 @@ open Classical Filter namespace Filter variable {ι : Type*} {α : ι → Type*} {f f₁ f₂ : (i : ι) → Filter (α i)} {s : (i : ι) → Set (α i)} + {p : ∀ i, α i → Prop} section Pi @@ -103,6 +104,14 @@ theorem pi_mem_pi_iff [∀ i, NeBot (f i)] {I : Set ι} (hI : I.Finite) : ⟨fun h _i hi => mem_of_pi_mem_pi h hi, pi_mem_pi hI⟩ #align filter.pi_mem_pi_iff Filter.pi_mem_pi_iff +theorem Eventually.eval_pi {i : ι} (hf : ∀ᶠ x : α i in f i, p i x) : + ∀ᶠ x : ∀ i : ι, α i in pi f, p i (x i) := (tendsto_eval_pi _ _).eventually hf +#align filter.eventually.eval_pi Filter.Eventually.eval_pi + +theorem eventually_pi [Finite ι] (hf : ∀ i, ∀ᶠ x in f i, p i x) : + ∀ᶠ x : ∀ i, α i in pi f, ∀ i, p i (x i) := eventually_all.2 fun _i => (hf _).eval_pi +#align filter.eventually_pi Filter.eventually_pi + theorem hasBasis_pi {ι' : ι → Type} {s : ∀ i, ι' i → Set (α i)} {p : ∀ i, ι' i → Prop} (h : ∀ i, (f i).HasBasis (p i) (s i)) : (pi f).HasBasis (fun If : Set ι × ∀ i, ι' i => If.1.Finite ∧ ∀ i ∈ If.1, p i (If.2 i)) diff --git a/Mathlib/Order/Filter/Prod.lean b/Mathlib/Order/Filter/Prod.lean index f5c0ed27063cb..37ab53ad4fefa 100644 --- a/Mathlib/Order/Filter/Prod.lean +++ b/Mathlib/Order/Filter/Prod.lean @@ -57,8 +57,7 @@ protected def prod (f : Filter α) (g : Filter β) : Filter (α × β) := instance instSProd : SProd (Filter α) (Filter β) (Filter (α × β)) where sprod := Filter.prod -theorem prod_mem_prod {s : Set α} {t : Set β} {f : Filter α} {g : Filter β} (hs : s ∈ f) - (ht : t ∈ g) : s ×ˢ t ∈ f ×ˢ g := +theorem prod_mem_prod (hs : s ∈ f) (ht : t ∈ g) : s ×ˢ t ∈ f ×ˢ g := inter_mem_inf (preimage_mem_comap hs) (preimage_mem_comap ht) #align filter.prod_mem_prod Filter.prod_mem_prod @@ -73,8 +72,7 @@ theorem mem_prod_iff {s : Set (α × β)} {f : Filter α} {g : Filter β} : #align filter.mem_prod_iff Filter.mem_prod_iff @[simp] -theorem prod_mem_prod_iff {s : Set α} {t : Set β} {f : Filter α} {g : Filter β} [f.NeBot] - [g.NeBot] : s ×ˢ t ∈ f ×ˢ g ↔ s ∈ f ∧ t ∈ g := +theorem prod_mem_prod_iff [f.NeBot] [g.NeBot] : s ×ˢ t ∈ f ×ˢ g ↔ s ∈ f ∧ t ∈ g := ⟨fun h => let ⟨_s', hs', _t', ht', H⟩ := mem_prod_iff.1 h (prod_subset_prod_iff.1 H).elim @@ -84,7 +82,7 @@ theorem prod_mem_prod_iff {s : Set α} {t : Set β} {f : Filter α} {g : Filter fun h => prod_mem_prod h.1 h.2⟩ #align filter.prod_mem_prod_iff Filter.prod_mem_prod_iff -theorem mem_prod_principal {f : Filter α} {s : Set (α × β)} {t : Set β} : +theorem mem_prod_principal {s : Set (α × β)} : s ∈ f ×ˢ 𝓟 t ↔ { a | ∀ b ∈ t, (a, b) ∈ s } ∈ f := by rw [← @exists_mem_subset_iff _ f, mem_prod_iff] refine' exists_congr fun u => Iff.rfl.and ⟨_, fun h => ⟨t, mem_principal_self t, _⟩⟩ @@ -94,7 +92,7 @@ theorem mem_prod_principal {f : Filter α} {s : Set (α × β)} {t : Set β} : exact h hx y hy #align filter.mem_prod_principal Filter.mem_prod_principal -theorem mem_prod_top {f : Filter α} {s : Set (α × β)} : +theorem mem_prod_top {s : Set (α × β)} : s ∈ f ×ˢ (⊤ : Filter β) ↔ { a | ∀ b, (a, b) ∈ s } ∈ f := by rw [← principal_univ, mem_prod_principal] simp only [mem_univ, forall_true_left] @@ -111,7 +109,7 @@ theorem comap_prod (f : α → β × γ) (b : Filter β) (c : Filter γ) : erw [comap_inf, Filter.comap_comap, Filter.comap_comap] #align filter.comap_prod Filter.comap_prod -theorem prod_top {f : Filter α} : f ×ˢ (⊤ : Filter β) = f.comap Prod.fst := by +theorem prod_top : f ×ˢ (⊤ : Filter β) = f.comap Prod.fst := by dsimp only [SProd.sprod] rw [Filter.prod, comap_top, inf_top_eq] #align filter.prod_top Filter.prod_top @@ -126,28 +124,27 @@ theorem prod_sup (f : Filter α) (g₁ g₂ : Filter β) : f ×ˢ (g₁ ⊔ g₂ rw [Filter.prod, comap_sup, inf_sup_left, ← Filter.prod, ← Filter.prod] #align filter.prod_sup Filter.prod_sup -theorem eventually_prod_iff {p : α × β → Prop} {f : Filter α} {g : Filter β} : +theorem eventually_prod_iff {p : α × β → Prop} : (∀ᶠ x in f ×ˢ g, p x) ↔ ∃ pa : α → Prop, (∀ᶠ x in f, pa x) ∧ ∃ pb : β → Prop, (∀ᶠ y in g, pb y) ∧ ∀ {x}, pa x → ∀ {y}, pb y → p (x, y) := by simpa only [Set.prod_subset_iff] using @mem_prod_iff α β p f g #align filter.eventually_prod_iff Filter.eventually_prod_iff -theorem tendsto_fst {f : Filter α} {g : Filter β} : Tendsto Prod.fst (f ×ˢ g) f := +theorem tendsto_fst : Tendsto Prod.fst (f ×ˢ g) f := tendsto_inf_left tendsto_comap #align filter.tendsto_fst Filter.tendsto_fst -theorem tendsto_snd {f : Filter α} {g : Filter β} : Tendsto Prod.snd (f ×ˢ g) g := +theorem tendsto_snd : Tendsto Prod.snd (f ×ˢ g) g := tendsto_inf_right tendsto_comap #align filter.tendsto_snd Filter.tendsto_snd -theorem Tendsto.prod_mk {f : Filter α} {g : Filter β} {h : Filter γ} {m₁ : α → β} {m₂ : α → γ} +theorem Tendsto.prod_mk {h : Filter γ} {m₁ : α → β} {m₂ : α → γ} (h₁ : Tendsto m₁ f g) (h₂ : Tendsto m₂ f h) : Tendsto (fun x => (m₁ x, m₂ x)) f (g ×ˢ h) := tendsto_inf.2 ⟨tendsto_comap_iff.2 h₁, tendsto_comap_iff.2 h₂⟩ #align filter.tendsto.prod_mk Filter.Tendsto.prod_mk -theorem tendsto_prod_swap {α1 α2 : Type*} {a1 : Filter α1} {a2 : Filter α2} : - Tendsto (Prod.swap : α1 × α2 → α2 × α1) (a1 ×ˢ a2) (a2 ×ˢ a1) := +theorem tendsto_prod_swap : Tendsto (Prod.swap : α × β → β × α) (f ×ˢ g) (g ×ˢ f) := tendsto_snd.prod_mk tendsto_fst #align filter.tendsto_prod_swap Filter.tendsto_prod_swap @@ -186,7 +183,7 @@ theorem Eventually.curry {la : Filter α} {lb : Filter β} {p : α × β → Pro /-- A fact that is eventually true about all pairs `l ×ˢ l` is eventually true about all diagonal pairs `(i, i)` -/ -theorem Eventually.diag_of_prod {f : Filter α} {p : α × α → Prop} (h : ∀ᶠ i in f ×ˢ f, p i) : +theorem Eventually.diag_of_prod {p : α × α → Prop} (h : ∀ᶠ i in f ×ˢ f, p i) : ∀ᶠ i in f, p (i, i) := by obtain ⟨t, ht, s, hs, hst⟩ := eventually_prod_iff.1 h apply (ht.and hs).mono fun x hx => hst hx.1 hx.2 @@ -253,13 +250,13 @@ theorem prod_comm : f ×ˢ g = map (fun p : β × α => (p.2, p.1)) (g ×ˢ f) : rfl #align filter.prod_comm Filter.prod_comm -theorem mem_prod_iff_left {s : Set (α × β)} {f : Filter α} {g : Filter β} : +theorem mem_prod_iff_left {s : Set (α × β)} : s ∈ f ×ˢ g ↔ ∃ t ∈ f, ∀ᶠ y in g, ∀ x ∈ t, (x, y) ∈ s := by simp only [mem_prod_iff, prod_subset_iff] refine exists_congr fun _ => Iff.rfl.and <| Iff.trans ?_ exists_mem_subset_iff exact exists_congr fun _ => Iff.rfl.and forall₂_swap -theorem mem_prod_iff_right {s : Set (α × β)} {f : Filter α} {g : Filter β} : +theorem mem_prod_iff_right {s : Set (α × β)} : s ∈ f ×ˢ g ↔ ∃ t ∈ g, ∀ᶠ x in f, ∀ y ∈ t, (x, y) ∈ s := by rw [prod_comm, mem_map, mem_prod_iff_left]; rfl @@ -309,24 +306,24 @@ theorem prod_assoc_symm (f : Filter α) (g : Filter β) (h : Filter γ) : Function.comp, Equiv.prodAssoc_apply] #align filter.prod_assoc_symm Filter.prod_assoc_symm -theorem tendsto_prodAssoc {f : Filter α} {g : Filter β} {h : Filter γ} : +theorem tendsto_prodAssoc {h : Filter γ} : Tendsto (Equiv.prodAssoc α β γ) ((f ×ˢ g) ×ˢ h) (f ×ˢ (g ×ˢ h)) := (prod_assoc f g h).le #align filter.tendsto_prod_assoc Filter.tendsto_prodAssoc -theorem tendsto_prodAssoc_symm {f : Filter α} {g : Filter β} {h : Filter γ} : +theorem tendsto_prodAssoc_symm {h : Filter γ} : Tendsto (Equiv.prodAssoc α β γ).symm (f ×ˢ (g ×ˢ h)) ((f ×ˢ g) ×ˢ h) := (prod_assoc_symm f g h).le #align filter.tendsto_prod_assoc_symm Filter.tendsto_prodAssoc_symm /-- A useful lemma when dealing with uniformities. -/ -theorem map_swap4_prod {f : Filter α} {g : Filter β} {h : Filter γ} {k : Filter δ} : +theorem map_swap4_prod {h : Filter γ} {k : Filter δ} : map (fun p : (α × β) × γ × δ => ((p.1.1, p.2.1), (p.1.2, p.2.2))) ((f ×ˢ g) ×ˢ (h ×ˢ k)) = (f ×ˢ h) ×ˢ (g ×ˢ k) := by simp_rw [map_swap4_eq_comap, SProd.sprod, Filter.prod, comap_inf, comap_comap]; ac_rfl #align filter.map_swap4_prod Filter.map_swap4_prod -theorem tendsto_swap4_prod {f : Filter α} {g : Filter β} {h : Filter γ} {k : Filter δ} : +theorem tendsto_swap4_prod {h : Filter γ} {k : Filter δ} : Tendsto (fun p : (α × β) × γ × δ => ((p.1.1, p.2.1), (p.1.2, p.2.2))) ((f ×ˢ g) ×ˢ (h ×ˢ k)) ((f ×ˢ h) ×ˢ (g ×ˢ k)) := map_swap4_prod.le @@ -377,9 +374,7 @@ protected theorem map_prod (m : α × β → γ) (f : Filter α) (g : Filter β) exact fun ⟨s, hs, t, ht, h⟩ => ⟨t, ht, s, hs, fun ⟨x, y⟩ ⟨hx, hy⟩ => h x hx y hy⟩ #align filter.map_prod Filter.map_prod -theorem prod_eq {f : Filter α} {g : Filter β} : f ×ˢ g = (f.map Prod.mk).seq g := by - have h := f.map_prod id g - rwa [map_id] at h +theorem prod_eq : f ×ˢ g = (f.map Prod.mk).seq g := f.map_prod id g #align filter.prod_eq Filter.prod_eq theorem prod_inf_prod {f₁ f₂ : Filter α} {g₁ g₂ : Filter β} : @@ -387,10 +382,10 @@ theorem prod_inf_prod {f₁ f₂ : Filter α} {g₁ g₂ : Filter β} : simp only [SProd.sprod, Filter.prod, comap_inf, inf_comm, inf_assoc, inf_left_comm] #align filter.prod_inf_prod Filter.prod_inf_prod -theorem inf_prod {f₁ f₂ : Filter α} {g : Filter β} : (f₁ ⊓ f₂) ×ˢ g = (f₁ ×ˢ g) ⊓ (f₂ ×ˢ g) := by +theorem inf_prod {f₁ f₂ : Filter α} : (f₁ ⊓ f₂) ×ˢ g = (f₁ ×ˢ g) ⊓ (f₂ ×ˢ g) := by rw [prod_inf_prod, inf_idem] -theorem prod_inf {f : Filter α} {g₁ g₂ : Filter β} : f ×ˢ (g₁ ⊓ g₂) = (f ×ˢ g₁) ⊓ (f ×ˢ g₂) := by +theorem prod_inf {g₁ g₂ : Filter β} : f ×ˢ (g₁ ⊓ g₂) = (f ×ˢ g₁) ⊓ (f ×ˢ g₂) := by rw [prod_inf_prod, inf_idem] @[simp] @@ -410,7 +405,7 @@ theorem map_pure_prod (f : α → β → γ) (a : α) (B : Filter β) : #align filter.map_pure_prod Filter.map_pure_prod @[simp] -theorem prod_pure {f : Filter α} {b : β} : f ×ˢ pure b = map (fun a => (a, b)) f := by +theorem prod_pure {b : β} : f ×ˢ pure b = map (fun a => (a, b)) f := by rw [prod_eq, seq_pure, map_map]; rfl #align filter.prod_pure Filter.prod_pure @@ -418,53 +413,59 @@ theorem prod_pure_pure {a : α} {b : β} : (pure a : Filter α) ×ˢ (pure b : Filter β) = pure (a, b) := by simp #align filter.prod_pure_pure Filter.prod_pure_pure -theorem prod_eq_bot {f : Filter α} {g : Filter β} : f ×ˢ g = ⊥ ↔ f = ⊥ ∨ g = ⊥ := by +theorem prod_eq_bot : f ×ˢ g = ⊥ ↔ f = ⊥ ∨ g = ⊥ := by simp_rw [← empty_mem_iff_bot, mem_prod_iff, subset_empty_iff, prod_eq_empty_iff, ← exists_prop, Subtype.exists', exists_or, exists_const, Subtype.exists, exists_prop, exists_eq_right] #align filter.prod_eq_bot Filter.prod_eq_bot -@[simp] theorem prod_bot {f : Filter α} : f ×ˢ (⊥ : Filter β) = ⊥ := prod_eq_bot.2 <| Or.inr rfl +@[simp] theorem prod_bot : f ×ˢ (⊥ : Filter β) = ⊥ := prod_eq_bot.2 <| Or.inr rfl #align filter.prod_bot Filter.prod_bot -@[simp] theorem bot_prod {g : Filter β} : (⊥ : Filter α) ×ˢ g = ⊥ := prod_eq_bot.2 <| Or.inl rfl +@[simp] theorem bot_prod : (⊥ : Filter α) ×ˢ g = ⊥ := prod_eq_bot.2 <| Or.inl rfl #align filter.bot_prod Filter.bot_prod -theorem prod_neBot {f : Filter α} {g : Filter β} : NeBot (f ×ˢ g) ↔ NeBot f ∧ NeBot g := by +theorem prod_neBot : NeBot (f ×ˢ g) ↔ NeBot f ∧ NeBot g := by simp only [neBot_iff, Ne, prod_eq_bot, not_or] #align filter.prod_ne_bot Filter.prod_neBot -theorem NeBot.prod {f : Filter α} {g : Filter β} (hf : NeBot f) (hg : NeBot g) : NeBot (f ×ˢ g) := - prod_neBot.2 ⟨hf, hg⟩ +theorem NeBot.prod (hf : NeBot f) (hg : NeBot g) : NeBot (f ×ˢ g) := prod_neBot.2 ⟨hf, hg⟩ #align filter.ne_bot.prod Filter.NeBot.prod -instance prod_neBot' {f : Filter α} {g : Filter β} [hf : NeBot f] [hg : NeBot g] : - NeBot (f ×ˢ g) := - hf.prod hg +instance prod_neBot' [hf : NeBot f] [hg : NeBot g] : NeBot (f ×ˢ g) := hf.prod hg #align filter.prod_ne_bot' Filter.prod_neBot' +@[simp] +lemma disjoint_prod {f' : Filter α} {g' : Filter β} : + Disjoint (f ×ˢ g) (f' ×ˢ g') ↔ Disjoint f f' ∨ Disjoint g g' := by + simp only [disjoint_iff, prod_inf_prod, prod_eq_bot] + +/-- `p ∧ q` occurs frequently along the product of two filters +iff both `p` and `q` occur frequently along the corresponding filters. -/ +theorem frequently_prod_and {p : α → Prop} {q : β → Prop} : + (∃ᶠ x in f ×ˢ g, p x.1 ∧ q x.2) ↔ (∃ᶠ a in f, p a) ∧ ∃ᶠ b in g, q b := by + simp only [frequently_iff_neBot, ← prod_neBot, ← prod_inf_prod, prod_principal_principal] + rfl + theorem tendsto_prod_iff {f : α × β → γ} {x : Filter α} {y : Filter β} {z : Filter γ} : Tendsto f (x ×ˢ y) z ↔ ∀ W ∈ z, ∃ U ∈ x, ∃ V ∈ y, ∀ x y, x ∈ U → y ∈ V → f (x, y) ∈ W := by simp only [tendsto_def, mem_prod_iff, prod_sub_preimage_iff, exists_prop, iff_self_iff] #align filter.tendsto_prod_iff Filter.tendsto_prod_iff -theorem le_prod {f : Filter (α × β)} {g : Filter α} {g' : Filter β} : - (f ≤ g ×ˢ g') ↔ Tendsto Prod.fst f g ∧ Tendsto Prod.snd f g' := by - dsimp only [SProd.sprod] - unfold Filter.prod - simp only [le_inf_iff, ← map_le_iff_le_comap, Tendsto] - -theorem tendsto_prod_iff' {f : Filter α} {g : Filter β} {g' : Filter γ} {s : α → β × γ} : +theorem tendsto_prod_iff' {g' : Filter γ} {s : α → β × γ} : Tendsto s f (g ×ˢ g') ↔ Tendsto (fun n => (s n).1) f g ∧ Tendsto (fun n => (s n).2) f g' := by dsimp only [SProd.sprod] unfold Filter.prod simp only [tendsto_inf, tendsto_comap_iff, (· ∘ ·)] #align filter.tendsto_prod_iff' Filter.tendsto_prod_iff' +theorem le_prod {f : Filter (α × β)} {g : Filter α} {g' : Filter β} : + (f ≤ g ×ˢ g') ↔ Tendsto Prod.fst f g ∧ Tendsto Prod.snd f g' := + tendsto_prod_iff' + end Prod /-! ### Coproducts of filters -/ - section Coprod variable {f : Filter α} {g : Filter β} diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index 986d57f9f76a4..b1972fd76d791 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -740,6 +740,15 @@ def toOrderHom {X Y : Type*} [Preorder X] [Preorder Y] (f : X ↪o Y) : X →o Y #align order_embedding.to_order_hom OrderEmbedding.toOrderHom #align order_embedding.to_order_hom_coe OrderEmbedding.toOrderHom_coe +/-- The trivial embedding from an empty preorder to another preorder -/ +@[simps] def ofIsEmpty [IsEmpty α] : α ↪o β where + toFun := isEmptyElim + inj' := isEmptyElim + map_rel_iff' {a} := isEmptyElim a + +@[simp, norm_cast] +lemma coe_ofIsEmpty [IsEmpty α] : (ofIsEmpty : α ↪o β) = (isEmptyElim : α → β) := rfl + end OrderEmbedding section RelHom diff --git a/Mathlib/Order/Hom/CompleteLattice.lean b/Mathlib/Order/Hom/CompleteLattice.lean index 84c84dd517c0c..612327af887a9 100644 --- a/Mathlib/Order/Hom/CompleteLattice.lean +++ b/Mathlib/Order/Hom/CompleteLattice.lean @@ -125,7 +125,7 @@ export sInfHomClass (map_sInf) attribute [simp] map_sSup map_sInf -theorem map_iSup [SupSet α] [SupSet β] [sSupHomClass F α β] (f : F) (g : ι → α) : +@[simp] theorem map_iSup [SupSet α] [SupSet β] [sSupHomClass F α β] (f : F) (g : ι → α) : f (⨆ i, g i) = ⨆ i, f (g i) := by simp [iSup, ← Set.range_comp, Function.comp] #align map_supr map_iSup @@ -133,7 +133,7 @@ theorem map_iSup₂ [SupSet α] [SupSet β] [sSupHomClass F α β] (f : F) (g : f (⨆ (i) (j), g i j) = ⨆ (i) (j), f (g i j) := by simp_rw [map_iSup] #align map_supr₂ map_iSup₂ -theorem map_iInf [InfSet α] [InfSet β] [sInfHomClass F α β] (f : F) (g : ι → α) : +@[simp] theorem map_iInf [InfSet α] [InfSet β] [sInfHomClass F α β] (f : F) (g : ι → α) : f (⨅ i, g i) = ⨅ i, f (g i) := by simp [iInf, ← Set.range_comp, Function.comp] #align map_infi map_iInf diff --git a/Mathlib/Order/Lattice.lean b/Mathlib/Order/Lattice.lean index 4ff7c915a913e..223bb8e80ab7c 100644 --- a/Mathlib/Order/Lattice.lean +++ b/Mathlib/Order/Lattice.lean @@ -1550,5 +1550,5 @@ instance [LinearOrder α] : LinearOrder (ULift.{v} α) := end ULift --To avoid noncomputability poisoning from `Bool.completeBooleanAlgebra` -instance : DistribLattice Bool := +instance Bool.instDistribLattice : DistribLattice Bool := inferInstance diff --git a/Mathlib/Order/LocallyFinite.lean b/Mathlib/Order/LocallyFinite.lean index b5c6794182c66..8156f0d79f8f5 100644 --- a/Mathlib/Order/LocallyFinite.lean +++ b/Mathlib/Order/LocallyFinite.lean @@ -176,7 +176,7 @@ def LocallyFiniteOrder.ofIcc' (α : Type*) [Preorder α] [DecidableRel ((· ≤ #align locally_finite_order.of_Icc' LocallyFiniteOrder.ofIcc' /-- A constructor from a definition of `Finset.Icc` alone, the other ones being derived by removing -the ends. As opposed to `LocallyFiniteOrder.ofIcc`, this one requires `PartialOrder` but only +the ends. As opposed to `LocallyFiniteOrder.ofIcc'`, this one requires `PartialOrder` but only `DecidableEq`. -/ def LocallyFiniteOrder.ofIcc (α : Type*) [PartialOrder α] [DecidableEq α] (finsetIcc : α → α → Finset α) (mem_Icc : ∀ a b x, x ∈ finsetIcc a b ↔ a ≤ x ∧ x ≤ b) : @@ -1081,8 +1081,8 @@ instance locallyFiniteOrder : LocallyFiniteOrder (WithTop α) where rw [some_mem_insertNone] simp | (a : α), (b : α), ⊤ => by - simp only [some, le_eq_subset, mem_map, mem_Icc, le_top, top_le_iff, and_false, iff_false, - not_exists, not_and, and_imp, Embedding.some, forall_const] + simp only [Embedding.some, mem_map, mem_Icc, and_false, exists_const, some, le_top, + top_le_iff] | (a : α), (b : α), (x : α) => by simp only [some, le_eq_subset, Embedding.some, mem_map, mem_Icc, Embedding.coeFn_mk, some_le_some] @@ -1396,7 +1396,7 @@ section Finite variable {α : Type*} {s : Set α} -theorem Set.finite_iff_bddAbove [SemilatticeSup α] [LocallyFiniteOrder α] [OrderBot α]: +theorem Set.finite_iff_bddAbove [SemilatticeSup α] [LocallyFiniteOrder α] [OrderBot α] : s.Finite ↔ BddAbove s := ⟨fun h ↦ ⟨h.toFinset.sup id, fun x hx ↦ Finset.le_sup (f := id) (by simpa)⟩, fun ⟨m, hm⟩ ↦ (Set.finite_Icc ⊥ m).subset (fun x hx ↦ ⟨bot_le, hm hx⟩)⟩ diff --git a/Mathlib/Order/PartialSups.lean b/Mathlib/Order/PartialSups.lean index 3a537e7546be8..e04b1ff22b5c5 100644 --- a/Mathlib/Order/PartialSups.lean +++ b/Mathlib/Order/PartialSups.lean @@ -59,34 +59,38 @@ theorem partialSups_succ (f : ℕ → α) (n : ℕ) : rfl #align partial_sups_succ partialSups_succ -theorem le_partialSups_of_le (f : ℕ → α) {m n : ℕ} (h : m ≤ n) : f m ≤ partialSups f n := by - induction' n with n ih - · rw [nonpos_iff_eq_zero.mp h, partialSups_zero] - · cases' h with h h - · exact le_sup_right - · exact (ih h).trans le_sup_left +lemma partialSups_iff_forall {f : ℕ → α} (p : α → Prop) + (hp : ∀ {a b}, p (a ⊔ b) ↔ p a ∧ p b) : ∀ {n : ℕ}, p (partialSups f n) ↔ ∀ k ≤ n, p (f k) + | 0 => by simp + | (n + 1) => by simp [hp, partialSups_iff_forall, ← Nat.lt_succ_iff, ← Nat.forall_lt_succ] + +@[simp] +lemma partialSups_le_iff {f : ℕ → α} {n : ℕ} {a : α} : partialSups f n ≤ a ↔ ∀ k ≤ n, f k ≤ a := + partialSups_iff_forall (· ≤ a) sup_le_iff + +theorem le_partialSups_of_le (f : ℕ → α) {m n : ℕ} (h : m ≤ n) : f m ≤ partialSups f n := + partialSups_le_iff.1 le_rfl m h #align le_partial_sups_of_le le_partialSups_of_le theorem le_partialSups (f : ℕ → α) : f ≤ partialSups f := fun _n => le_partialSups_of_le f le_rfl #align le_partial_sups le_partialSups theorem partialSups_le (f : ℕ → α) (n : ℕ) (a : α) (w : ∀ m, m ≤ n → f m ≤ a) : - partialSups f n ≤ a := by - induction' n with n ih - · apply w 0 le_rfl - · exact sup_le (ih fun m p => w m (Nat.le_succ_of_le p)) (w (n + 1) le_rfl) + partialSups f n ≤ a := + partialSups_le_iff.2 w #align partial_sups_le partialSups_le +@[simp] +lemma upperBounds_range_partialSups (f : ℕ → α) : + upperBounds (Set.range (partialSups f)) = upperBounds (Set.range f) := by + ext a + simp only [mem_upperBounds, Set.forall_range_iff, partialSups_le_iff] + exact ⟨fun h _ ↦ h _ _ le_rfl, fun h _ _ _ ↦ h _⟩ + @[simp] theorem bddAbove_range_partialSups {f : ℕ → α} : - BddAbove (Set.range (partialSups f)) ↔ BddAbove (Set.range f) := by - apply exists_congr fun a => _ - intro a - constructor - · rintro h b ⟨i, rfl⟩ - exact (le_partialSups _ _).trans (h (Set.mem_range_self i)) - · rintro h b ⟨i, rfl⟩ - exact partialSups_le _ _ _ fun _ _ => h (Set.mem_range_self _) + BddAbove (Set.range (partialSups f)) ↔ BddAbove (Set.range f) := + .of_eq <| congr_arg Set.Nonempty <| upperBounds_range_partialSups f #align bdd_above_range_partial_sups bddAbove_range_partialSups theorem Monotone.partialSups_eq {f : ℕ → α} (hf : Monotone f) : (partialSups f : ℕ → α) = f := by @@ -96,16 +100,10 @@ theorem Monotone.partialSups_eq {f : ℕ → α} (hf : Monotone f) : (partialSup · rw [partialSups_succ, ih, sup_eq_right.2 (hf (Nat.le_succ _))] #align monotone.partial_sups_eq Monotone.partialSups_eq -theorem partialSups_mono : Monotone (partialSups : (ℕ → α) → ℕ →o α) := by - rintro f g h n - induction' n with n ih - · exact h 0 - · exact sup_le_sup ih (h _) +theorem partialSups_mono : Monotone (partialSups : (ℕ → α) → ℕ →o α) := fun _f _g h _n ↦ + partialSups_le_iff.2 fun k hk ↦ (h k).trans (le_partialSups_of_le _ hk) #align partial_sups_mono partialSups_mono -theorem partialSups_apply_mono (f : ℕ → α) : Monotone (partialSups f) := - fun n _ hnm => partialSups_le f n _ (fun _ hm'n => le_partialSups_of_le _ (hm'n.trans hnm)) - /-- `partialSups` forms a Galois insertion with the coercion from monotone functions to functions. -/ def partialSups.gi : GaloisInsertion (partialSups : (ℕ → α) → ℕ →o α) (↑) where @@ -120,42 +118,42 @@ def partialSups.gi : GaloisInsertion (partialSups : (ℕ → α) → ℕ →o α #align partial_sups.gi partialSups.gi theorem partialSups_eq_sup'_range (f : ℕ → α) (n : ℕ) : - partialSups f n = (Finset.range (n + 1)).sup' ⟨n, Finset.self_mem_range_succ n⟩ f := by - induction' n with n ih - · simp - · dsimp [partialSups] at ih ⊢ - simp_rw [@Finset.range_succ n.succ] - rw [ih, Finset.sup'_insert, sup_comm] + partialSups f n = (Finset.range (n + 1)).sup' ⟨n, Finset.self_mem_range_succ n⟩ f := + eq_of_forall_ge_iff fun _ ↦ by simp [Nat.lt_succ_iff] #align partial_sups_eq_sup'_range partialSups_eq_sup'_range end SemilatticeSup theorem partialSups_eq_sup_range [SemilatticeSup α] [OrderBot α] (f : ℕ → α) (n : ℕ) : - partialSups f n = (Finset.range (n + 1)).sup f := by - induction' n with n ih - · simp - · dsimp [partialSups] at ih ⊢ - rw [Finset.range_succ, Finset.sup_insert, sup_comm, ih] + partialSups f n = (Finset.range (n + 1)).sup f := + eq_of_forall_ge_iff fun _ ↦ by simp [Nat.lt_succ_iff] #align partial_sups_eq_sup_range partialSups_eq_sup_range +@[simp] +lemma disjoint_partialSups_left [DistribLattice α] [OrderBot α] {f : ℕ → α} {n : ℕ} {x : α} : + Disjoint (partialSups f n) x ↔ ∀ k ≤ n, Disjoint (f k) x := + partialSups_iff_forall (Disjoint · x) disjoint_sup_left + +@[simp] +lemma disjoint_partialSups_right [DistribLattice α] [OrderBot α] {f : ℕ → α} {n : ℕ} {x : α} : + Disjoint x (partialSups f n) ↔ ∀ k ≤ n, Disjoint x (f k) := + partialSups_iff_forall (Disjoint x) disjoint_sup_right + /- Note this lemma requires a distributive lattice, so is not useful (or true) in situations such as submodules. -/ theorem partialSups_disjoint_of_disjoint [DistribLattice α] [OrderBot α] (f : ℕ → α) - (h : Pairwise (Disjoint on f)) {m n : ℕ} (hmn : m < n) : Disjoint (partialSups f m) (f n) := by - induction' m with m ih - · exact h hmn.ne - · rw [partialSups_succ, disjoint_sup_left] - exact ⟨ih (Nat.lt_of_succ_lt hmn), h hmn.ne⟩ + (h : Pairwise (Disjoint on f)) {m n : ℕ} (hmn : m < n) : Disjoint (partialSups f m) (f n) := + disjoint_partialSups_left.2 fun _k hk ↦ h <| (hk.trans_lt hmn).ne #align partial_sups_disjoint_of_disjoint partialSups_disjoint_of_disjoint section ConditionallyCompleteLattice variable [ConditionallyCompleteLattice α] -theorem partialSups_eq_ciSup_Iic (f : ℕ → α) (n : ℕ) : partialSups f n = ⨆ i : Set.Iic n, f i := by - have : Set.Iio (n + 1) = Set.Iic n := Set.ext fun _ => Nat.lt_succ_iff - rw [partialSups_eq_sup'_range, Finset.sup'_eq_csSup_image, Finset.coe_range, iSup, this] - simp only [Set.range, Subtype.exists, Set.mem_Iic, exists_prop, (· '' ·)] +theorem partialSups_eq_ciSup_Iic (f : ℕ → α) (n : ℕ) : partialSups f n = ⨆ i : Set.Iic n, f i := + eq_of_forall_ge_iff fun _ ↦ by + rw [ciSup_set_le_iff Set.nonempty_Iic ((Set.finite_le_nat _).image _).bddAbove, + partialSups_le_iff]; rfl #align partial_sups_eq_csupr_Iic partialSups_eq_ciSup_Iic @[simp] diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 71aadb05040af..200aa382c423a 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -128,7 +128,7 @@ protected theorem IsTotal.isTrichotomous (r) [IsTotal α r] : IsTrichotomous α -- see Note [lower instance priority] instance (priority := 100) IsTotal.to_isRefl (r) [IsTotal α r] : IsRefl α r := - ⟨fun a => (or_self_iff _).1 <| total_of r a a⟩ + ⟨fun a => or_self_iff.1 <| total_of r a a⟩ theorem ne_of_irrefl {r} [IsIrrefl α r] : ∀ {x y : α}, r x y → x ≠ y | _, _, h, rfl => irrefl _ h @@ -356,6 +356,9 @@ def WellFoundedGT (α : Type*) [LT α] : Prop := IsWellFounded α (· > ·) #align well_founded_gt WellFoundedGT +lemma wellFounded_lt [LT α] [WellFoundedLT α] : @WellFounded α (· < ·) := IsWellFounded.wf +lemma wellFounded_gt [LT α] [WellFoundedGT α] : @WellFounded α (· > ·) := IsWellFounded.wf + -- See note [lower instance priority] instance (priority := 100) (α : Type*) [LT α] [h : WellFoundedLT α] : WellFoundedGT αᵒᵈ := h diff --git a/Mathlib/Order/SemiconjSup.lean b/Mathlib/Order/SemiconjSup.lean index a6e7f451f98d6..de9614b794d1d 100644 --- a/Mathlib/Order/SemiconjSup.lean +++ b/Mathlib/Order/SemiconjSup.lean @@ -3,12 +3,12 @@ Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ +import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Logic.Function.Conjugate import Mathlib.Order.Bounds.OrderIso import Mathlib.Order.ConditionallyCompleteLattice.Basic -import Mathlib.Order.RelIso.Group import Mathlib.Order.OrdContinuous -import Mathlib.Algebra.Hom.Equiv.Units.Basic +import Mathlib.Order.RelIso.Group #align_import order.semiconj_Sup from "leanprover-community/mathlib"@"422e70f7ce183d2900c586a8cda8381e788a0c62" diff --git a/Mathlib/Order/Sublattice.lean b/Mathlib/Order/Sublattice.lean new file mode 100644 index 0000000000000..7e5524b0b3b0e --- /dev/null +++ b/Mathlib/Order/Sublattice.lean @@ -0,0 +1,291 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Order.SupClosed + +/-! +# Sublattices + +This file defines sublattices. + +## TODO + +Subsemilattices, if people care about them. + +## Tags + +sublattice +-/ + +open Function Set + +variable {ι : Sort*} (α β γ : Type*) [Lattice α] [Lattice β] [Lattice γ] + +/-- A sublattice of a lattice is a set containing the suprema and infima of any of its elements. -/ +structure Sublattice where + /-- The underlying set of a sublattice. **Do not use directly**. Instead, use the coercion + `Sublattice α → Set α`, which Lean should automatically insert for you in most cases. -/ + carrier : Set α + supClosed' : SupClosed carrier + infClosed' : InfClosed carrier + +variable {α β γ} + +namespace Sublattice +variable {L M : Sublattice α} {f : LatticeHom α β} {s t : Set α} {a : α} + +initialize_simps_projections Sublattice (carrier → coe) + +instance instSetLike : SetLike (Sublattice α) α where + coe L := L.carrier + coe_injective' L M h := by cases L; congr + +/-- Turn a set closed under supremum and infimum into a sublattice. -/ +abbrev ofIsSublattice (s : Set α) (hs : IsSublattice s) : Sublattice α := ⟨s, hs.1, hs.2⟩ + +lemma coe_inj : (L : Set α) = M ↔ L = M := SetLike.coe_set_eq + +@[simp] lemma supClosed (L : Sublattice α) : SupClosed (L : Set α) := L.supClosed' +@[simp] lemma infClosed (L : Sublattice α) : InfClosed (L : Set α) := L.infClosed' +@[simp] lemma isSublattice (L : Sublattice α) : IsSublattice (L : Set α) := + ⟨L.supClosed, L.infClosed⟩ + +@[simp] lemma mem_carrier : a ∈ L.carrier ↔ a ∈ L := Iff.rfl +@[simp] lemma mem_mk (h_sup h_inf) : a ∈ mk s h_sup h_inf ↔ a ∈ s := Iff.rfl +@[simp, norm_cast] lemma coe_mk (h_sup h_inf) : mk s h_sup h_inf = s := rfl +@[simp] lemma mk_le_mk (hs_sup hs_inf ht_sup ht_inf) : + mk s hs_sup hs_inf ≤ mk t ht_sup ht_inf ↔ s ⊆ t := Iff.rfl +@[simp] lemma mk_lt_mk (hs_sup hs_inf ht_sup ht_inf) : + mk s hs_sup hs_inf < mk t ht_sup ht_inf ↔ s ⊂ t := Iff.rfl + +/-- Copy of a sublattice with a new `carrier` equal to the old one. Useful to fix definitional +equalities.-/ +protected def copy (L : Sublattice α) (s : Set α) (hs : s = L) : Sublattice α where + carrier := s + supClosed' := hs.symm ▸ L.supClosed' + infClosed' := hs.symm ▸ L.infClosed' + +@[simp, norm_cast] lemma coe_copy (L : Sublattice α) (s : Set α) (hs) : L.copy s hs = s := rfl + +lemma copy_eq (L : Sublattice α) (s : Set α) (hs) : L.copy s hs = L := SetLike.coe_injective hs + +/-- Two sublattices are equal if they have the same elements. -/ +lemma ext : (∀ a, a ∈ L ↔ a ∈ M) → L = M := SetLike.ext + +/-- A sublattice of a lattice inherits a supremum. -/ +instance instSupCoe : Sup L where + sup a b := ⟨a ⊔ b, L.supClosed a.2 b.2⟩ + +/-- A sublattice of a lattice inherits an infimum. -/ +instance instInfCoe : Inf L where + inf a b := ⟨a ⊓ b, L.infClosed a.2 b.2⟩ + +@[simp, norm_cast] lemma coe_sup (a b : L) : a ⊔ b = (a : α) ⊔ b := rfl +@[simp, norm_cast] lemma coe_inf (a b : L) : a ⊓ b = (a : α) ⊓ b := rfl +@[simp] lemma mk_sup_mk (a b : α) (ha hb) : (⟨a, ha⟩ ⊔ ⟨b, hb⟩ : L) = ⟨a ⊔ b, L.supClosed ha hb⟩ := + rfl +@[simp] lemma mk_inf_mk (a b : α) (ha hb) : (⟨a, ha⟩ ⊓ ⟨b, hb⟩ : L) = ⟨a ⊓ b, L.infClosed ha hb⟩ := + rfl + +/-- A sublattice of a lattice inherits a lattice structure. -/ +instance instLatticeCoe (L : Sublattice α) : Lattice L := + Subtype.coe_injective.lattice _ (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) + +/-- A sublattice of a distributive lattice inherits a distributive lattice structure. -/ +instance instDistribLatticeCoe {α : Type*} [DistribLattice α] (L : Sublattice α) : + DistribLattice L := + Subtype.coe_injective.distribLattice _ (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) + +/-- The natural lattice hom from a sublattice to the original lattice. -/ +def subtype (L : Sublattice α) : LatticeHom L α where + toFun := ((↑) : L → α) + map_sup' _ _ := rfl + map_inf' _ _ := rfl + +@[simp, norm_cast] lemma coe_subtype (L : Sublattice α) : L.subtype = ((↑) : L → α) := rfl +lemma subtype_apply (L : Sublattice α) (a : L) : L.subtype a = a := rfl + +lemma subtype_injective (L : Sublattice α) : Injective $ subtype L := Subtype.coe_injective + +/-- The inclusion homomorphism from a sublattice `L` to a bigger sublattice `M`. -/ +def inclusion (h : L ≤ M) : LatticeHom L M where + toFun := Set.inclusion h + map_sup' _ _ := rfl + map_inf' _ _ := rfl + +@[simp] lemma coe_inclusion (h : L ≤ M) : inclusion h = Set.inclusion h := rfl +lemma inclusion_apply (h : L ≤ M) (a : L) : inclusion h a = Set.inclusion h a := rfl + +lemma inclusion_injective (h : L ≤ M) : Injective $ inclusion h := Set.inclusion_injective h + +@[simp] lemma inclusion_rfl (L : Sublattice α) : inclusion le_rfl = LatticeHom.id L := rfl +@[simp] lemma subtype_comp_inclusion (h : L ≤ M) : M.subtype.comp (inclusion h) = L.subtype := rfl + +/-- The maximum sublattice of a lattice. -/ +instance instTop : Top (Sublattice α) where + top.carrier := univ + top.supClosed' := supClosed_univ + top.infClosed' := infClosed_univ + +/-- The empty sublattice of a lattice. -/ +instance instBot : Bot (Sublattice α) where + bot.carrier := ∅ + bot.supClosed' := supClosed_empty + bot.infClosed' := infClosed_empty + +/-- The inf of two sublattices is their intersection. -/ +instance instInf : Inf (Sublattice α) where + inf L M := { carrier := L ∩ M + supClosed' := L.supClosed.inter M.supClosed + infClosed' := L.infClosed.inter M.infClosed } + +/-- The inf of sublattices is their intersection. -/ +instance instInfSet : InfSet (Sublattice α) where + sInf S := { carrier := ⨅ L ∈ S, L + supClosed' := supClosed_sInter $ forall_range_iff.2 $ fun L ↦ supClosed_sInter $ + forall_range_iff.2 fun _ ↦ L.supClosed + infClosed' := infClosed_sInter $ forall_range_iff.2 $ fun L ↦ infClosed_sInter $ + forall_range_iff.2 fun _ ↦ L.infClosed } + +instance instInhabited : Inhabited (Sublattice α) := ⟨⊥⟩ + +/-- The top sublattice is isomorphic to the lattice. + +This is the sublattice version of `Equiv.Set.univ α`. -/ +def topEquiv : (⊤ : Sublattice α) ≃o α where + toEquiv := Equiv.Set.univ _ + map_rel_iff' := Iff.rfl + +@[simp, norm_cast] lemma coe_top : (⊤ : Sublattice α) = (univ : Set α) := rfl +@[simp, norm_cast] lemma coe_bot : (⊥ : Sublattice α) = (∅ : Set α) := rfl +@[simp, norm_cast] lemma coe_inf' (L M : Sublattice α) : L ⊓ M = (L : Set α) ∩ M := rfl +@[simp, norm_cast] lemma coe_sInf (S : Set (Sublattice α)) : sInf S = ⋂ L ∈ S, (L : Set α) := rfl +@[simp, norm_cast] lemma coe_iInf (f : ι → Sublattice α) : ⨅ i, f i = ⋂ i, (f i : Set α) := by + simp [iInf] + +@[simp, norm_cast] lemma coe_eq_univ : L = (univ : Set α) ↔ L = ⊤ := by rw [←coe_top, coe_inj] +@[simp, norm_cast] lemma coe_eq_empty : L = (∅ : Set α) ↔ L = ⊥ := by rw [←coe_bot, coe_inj] + +@[simp] lemma not_mem_bot (a : α) : a ∉ (⊥ : Sublattice α) := id +@[simp] lemma mem_top (a : α) : a ∈ (⊤ : Sublattice α) := mem_univ _ +@[simp] lemma mem_inf : a ∈ L ⊓ M ↔ a ∈ L ∧ a ∈ M := Iff.rfl +@[simp] lemma mem_sInf {S : Set (Sublattice α)} : a ∈ sInf S ↔ ∀ L ∈ S, a ∈ L := by + rw [←SetLike.mem_coe]; simp +@[simp] lemma mem_iInf {f : ι → Sublattice α} : a ∈ ⨅ i, f i ↔ ∀ i, a ∈ f i := by + rw [←SetLike.mem_coe]; simp + +/-- Sublattices of a lattice form a complete lattice. -/ +instance instCompleteLattice : CompleteLattice (Sublattice α) where + bot := ⊥ + bot_le := fun _S _a ↦ False.elim + top := ⊤ + le_top := fun _S a _ha ↦ mem_top a + inf := (· ⊓ ·) + le_inf := fun _L _M _N hM hN _a ha ↦ ⟨hM ha, hN ha⟩ + inf_le_left := fun _L _M _a ↦ And.left + inf_le_right := fun _L _M _a ↦ And.right + __ := completeLatticeOfInf (Sublattice α) + fun _s ↦ IsGLB.of_image SetLike.coe_subset_coe isGLB_biInf + +lemma subsingleton_iff : Subsingleton (Sublattice α) ↔ IsEmpty α := + ⟨fun _ ↦ univ_eq_empty_iff.1 $ coe_inj.2 $ Subsingleton.elim ⊤ ⊥, + fun _ ↦ SetLike.coe_injective.subsingleton⟩ + +instance [IsEmpty α] : Unique (Sublattice α) where + uniq _ := @Subsingleton.elim _ (subsingleton_iff.2 ‹_›) _ _ + +/-- The preimage of a sublattice along a lattice homomorphism. -/ +def comap (f : LatticeHom α β) (L : Sublattice β) : Sublattice α where + carrier := f ⁻¹' L + supClosed' := L.supClosed.preimage _ + infClosed' := L.infClosed.preimage _ + +@[simp, norm_cast] lemma coe_comap (L : Sublattice β) (f : LatticeHom α β) : L.comap f = f ⁻¹' L := + rfl + +@[simp] lemma mem_comap {L : Sublattice β} : a ∈ L.comap f ↔ f a ∈ L := Iff.rfl + +lemma comap_mono : Monotone (comap f) := fun _ _ ↦ preimage_mono + +@[simp] lemma comap_id (L : Sublattice α) : L.comap (LatticeHom.id _) = L := rfl + +@[simp] lemma comap_comap (L : Sublattice γ) (g : LatticeHom β γ) (f : LatticeHom α β) : + (L.comap g).comap f = L.comap (g.comp f) := rfl + +/-- The image of a sublattice along a monoid homomorphism is a sublattice. -/ +def map (f : LatticeHom α β) (L : Sublattice α) : Sublattice β where + carrier := f '' L + supClosed' := L.supClosed.image f + infClosed' := L.infClosed.image f + +@[simp] lemma coe_map (f : LatticeHom α β) (L : Sublattice α) : (L.map f : Set β) = f '' L := rfl +@[simp] lemma mem_map {b : β} : b ∈ L.map f ↔ ∃ a ∈ L, f a = b := Iff.rfl + +lemma mem_map_of_mem (f : LatticeHom α β) {a : α} : a ∈ L → f a ∈ L.map f := mem_image_of_mem f +lemma apply_coe_mem_map (f : LatticeHom α β) (a : L) : f a ∈ L.map f := mem_map_of_mem f a.prop + +lemma map_mono : Monotone (map f) := fun _ _ ↦ image_subset _ + +@[simp] lemma map_id : L.map (LatticeHom.id α) = L := SetLike.coe_injective $ image_id _ + +@[simp] lemma map_map (g : LatticeHom β γ) (f : LatticeHom α β) : + (L.map f).map g = L.map (g.comp f) := SetLike.coe_injective $ image_image _ _ _ + +lemma mem_map_equiv {f : α ≃o β} {a : β} : a ∈ L.map f ↔ f.symm a ∈ L := Set.mem_image_equiv + +lemma apply_mem_map_iff (hf : Injective f) : f a ∈ L.map f ↔ a ∈ L := hf.mem_set_image + +lemma map_equiv_eq_comap_symm (f : α ≃o β) (L : Sublattice α) : + L.map f = L.comap (f.symm : LatticeHom β α) := + SetLike.coe_injective $ f.toEquiv.image_eq_preimage L + +lemma comap_equiv_eq_map_symm (f : β ≃o α) (L : Sublattice α) : + L.comap f = L.map (f.symm : LatticeHom α β) := (map_equiv_eq_comap_symm f.symm L).symm + +lemma map_symm_eq_iff_eq_map {M : Sublattice β} {e : β ≃o α} : + L.map ↑e.symm = M ↔ L = M.map ↑e := by + simp_rw [←coe_inj]; exact (Equiv.eq_image_iff_symm_image_eq _ _ _).symm + +lemma map_le_iff_le_comap {f : LatticeHom α β} {M : Sublattice β} : L.map f ≤ M ↔ L ≤ M.comap f := + image_subset_iff + +lemma gc_map_comap (f : LatticeHom α β) : GaloisConnection (map f) (comap f) := + fun _ _ ↦ map_le_iff_le_comap + +@[simp] lemma map_bot (f : LatticeHom α β) : (⊥ : Sublattice α).map f = ⊥ := (gc_map_comap f).l_bot + +lemma map_sup (f : LatticeHom α β) (L M : Sublattice α) : (L ⊔ M).map f = L.map f ⊔ M.map f := + (gc_map_comap f).l_sup + +lemma map_iSup (f : LatticeHom α β) (L : ι → Sublattice α) : (⨆ i, L i).map f = ⨆ i, (L i).map f := + (gc_map_comap f).l_iSup + +@[simp] lemma comap_top (f : LatticeHom α β) : (⊤ : Sublattice β).comap f = ⊤ := + (gc_map_comap f).u_top + +lemma comap_inf (L M : Sublattice β) (f : LatticeHom α β) : + (L ⊓ M).comap f = L.comap f ⊓ M.comap f := (gc_map_comap f).u_inf + +lemma comap_iInf (f : LatticeHom α β) (s : ι → Sublattice β) : + (iInf s).comap f = ⨅ i, (s i).comap f := (gc_map_comap f).u_iInf + +lemma map_inf_le (L M : Sublattice α) (f : LatticeHom α β) : map f (L ⊓ M) ≤ map f L ⊓ map f M := + map_mono.map_inf_le _ _ + +lemma le_comap_sup (L M : Sublattice β) (f : LatticeHom α β) : + comap f L ⊔ comap f M ≤ comap f (L ⊔ M) := comap_mono.le_map_sup _ _ + +lemma le_comap_iSup (f : LatticeHom α β) (L : ι → Sublattice β) : + ⨆ i, (L i).comap f ≤ (⨆ i, L i).comap f := comap_mono.le_map_iSup + +lemma map_inf (L M : Sublattice α) (f : LatticeHom α β) (hf : Injective f) : + map f (L ⊓ M) = map f L ⊓ map f M := by + rw [← SetLike.coe_set_eq] + simp [Set.image_inter hf] + +lemma map_top (f : LatticeHom α β) (h : Surjective f) : Sublattice.map f ⊤ = ⊤ := + SetLike.coe_injective $ by simp [h.range_eq] + +end Sublattice diff --git a/Mathlib/Order/SupClosed.lean b/Mathlib/Order/SupClosed.lean index 0981e5d291bfe..4a32fc9cf4b34 100644 --- a/Mathlib/Order/SupClosed.lean +++ b/Mathlib/Order/SupClosed.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Christopher Hoskin -/ import Mathlib.Data.Finset.Lattice +import Mathlib.Data.Set.Finite import Mathlib.Order.Closure import Mathlib.Order.UpperLower.Basic @@ -18,18 +19,20 @@ is automatically complete. All dually for `⊓`. * `SupClosed`: Predicate for a set to be closed under join (`a ∈ s` and `b ∈ s` imply `a ⊔ b ∈ s`). * `InfClosed`: Predicate for a set to be closed under meet (`a ∈ s` and `b ∈ s` imply `a ⊓ b ∈ s`). +* `IsSublattice`: Predicate for a set to be closed under meet and join. * `supClosure`: Sup-closure. Smallest sup-closed set containing a given set. * `infClosure`: Inf-closure. Smallest inf-closed set containing a given set. +* `latticeClosure`: Smallest sublattice containing a given set. * `SemilatticeSup.toCompleteSemilatticeSup`: A join-semilattice where every sup-closed set has a least upper bound is automatically complete. * `SemilatticeInf.toCompleteSemilatticeInf`: A meet-semilattice where every inf-closed set has a greatest lower bound is automatically complete. -/ -variable {α : Type*} +variable {F α β : Type*} section SemilatticeSup -variable [SemilatticeSup α] +variable [SemilatticeSup α] [SemilatticeSup β] section Set variable {ι : Sort*} {S : Set (Set α)} {f : ι → Set α} {s t : Set α} {a : α} @@ -56,6 +59,24 @@ lemma SupClosed.directedOn (hs : SupClosed s) : DirectedOn (· ≤ ·) s := lemma IsUpperSet.supClosed (hs : IsUpperSet s) : SupClosed s := fun _a _ _b ↦ hs le_sup_right +lemma SupClosed.preimage [SupHomClass F β α] (hs : SupClosed s) (f : F) : SupClosed (f ⁻¹' s) := + fun a ha b hb ↦ by simpa [map_sup] using hs ha hb + +lemma SupClosed.image [SupHomClass F α β] (hs : SupClosed s) (f : F) : SupClosed (f '' s) := by + rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ + rw [←map_sup] + exact Set.mem_image_of_mem _ $ hs ha hb + +lemma supClosed_range [SupHomClass F α β] (f : F) : SupClosed (Set.range f) := by + simpa using supClosed_univ.image f + +lemma SupClosed.prod {t : Set β} (hs : SupClosed s) (ht : SupClosed t) : SupClosed (s ×ˢ t) := + fun _a ha _b hb ↦ ⟨hs ha.1 hb.1, ht ha.2 hb.2⟩ + +lemma supClosed_pi {ι : Type*} {α : ι → Type*} [∀ i, SemilatticeSup (α i)] {s : Set ι} + {t : ∀ i, Set (α i)} (ht : ∀ i ∈ s, SupClosed (t i)) : SupClosed (s.pi t) := + fun _a ha _b hb _i hi ↦ ht _ hi (ha _ hi) (hb _ hi) + end Set section Finset @@ -75,7 +96,7 @@ end Finset end SemilatticeSup section SemilatticeInf -variable [SemilatticeInf α] +variable [SemilatticeInf α] [SemilatticeInf β] section Set variable {ι : Sort*} {S : Set (Set α)} {f : ι → Set α} {s t : Set α} {a : α} @@ -102,6 +123,24 @@ lemma InfClosed.codirectedOn (hs : InfClosed s) : DirectedOn (· ≥ ·) s := lemma IsLowerSet.infClosed (hs : IsLowerSet s) : InfClosed s := λ _a _ _b ↦ hs inf_le_right +lemma InfClosed.preimage [InfHomClass F β α] (hs : InfClosed s) (f : F) : InfClosed (f ⁻¹' s) := + fun a ha b hb ↦ by simpa [map_inf] using hs ha hb + +lemma InfClosed.image [InfHomClass F α β] (hs : InfClosed s) (f : F) : InfClosed (f '' s) := by + rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ + rw [←map_inf] + exact Set.mem_image_of_mem _ $ hs ha hb + +lemma infClosed_range [InfHomClass F α β] (f : F) : InfClosed (Set.range f) := by + simpa using infClosed_univ.image f + +lemma InfClosed.prod {t : Set β} (hs : InfClosed s) (ht : InfClosed t) : InfClosed (s ×ˢ t) := + fun _a ha _b hb ↦ ⟨hs ha.1 hb.1, ht ha.2 hb.2⟩ + +lemma infClosed_pi {ι : Type*} {α : ι → Type*} [∀ i, SemilatticeInf (α i)] {s : Set ι} + {t : ∀ i, Set (α i)} (ht : ∀ i ∈ s, InfClosed (t i)) : InfClosed (s.pi t) := + fun _a ha _b hb _i hi ↦ ht _ hi (ha _ hi) (hb _ hi) + end Set section Finset @@ -123,22 +162,71 @@ end SemilatticeInf open Finset OrderDual section Lattice -variable [Lattice α] +variable {ι : Sort*} [Lattice α] [Lattice β] {S : Set (Set α)} {f : ι → Set α} {s t : Set α} {a : α} + +open Set + +/-- A set `s` is a *sublattice* if `a ⊔ b ∈ s` and `a ⊓ b ∈ s` for all `a ∈ s`, `b ∈ s`. +Note: This is not the preferred way to declare a sublattice. One should instead use `Sublattice`. +TODO: Define `Sublattice`. -/ +structure IsSublattice (s : Set α) : Prop where + supClosed : SupClosed s + infClosed : InfClosed s + +@[simp] lemma isSublattice_empty : IsSublattice (∅ : Set α) := ⟨supClosed_empty, infClosed_empty⟩ +@[simp] lemma isSublattice_singleton : IsSublattice ({a} : Set α) := + ⟨supClosed_singleton, infClosed_singleton⟩ + +@[simp] lemma isSublattice_univ : IsSublattice (Set.univ : Set α) := + ⟨supClosed_univ, infClosed_univ⟩ + +lemma IsSublattice.inter (hs : IsSublattice s) (ht : IsSublattice t) : IsSublattice (s ∩ t) := + ⟨hs.1.inter ht.1, hs.2.inter ht.2⟩ -@[simp] lemma supClosed_preimage_toDual {s : Set αᵒᵈ} : SupClosed (toDual ⁻¹' s) ↔ InfClosed s := -Iff.rfl +lemma isSublattice_sInter (hS : ∀ s ∈ S, IsSublattice s) : IsSublattice (⋂₀ S) := + ⟨supClosed_sInter fun _s hs ↦ (hS _ hs).1, infClosed_sInter fun _s hs ↦ (hS _ hs).2⟩ -@[simp] lemma infClosed_preimage_toDual {s : Set αᵒᵈ} : InfClosed (toDual ⁻¹' s) ↔ SupClosed s := -Iff.rfl +lemma isSublattice_iInter (hf : ∀ i, IsSublattice (f i)) : IsSublattice (⋂ i, f i) := + ⟨supClosed_iInter fun _i ↦ (hf _).1, infClosed_iInter fun _i ↦ (hf _).2⟩ -@[simp] lemma supClosed_preimage_ofDual {s : Set α} : SupClosed (ofDual ⁻¹' s) ↔ InfClosed s := -Iff.rfl +lemma IsSublattice.preimage [LatticeHomClass F β α] (hs : IsSublattice s) (f : F) : + IsSublattice (f ⁻¹' s) := ⟨hs.1.preimage _, hs.2.preimage _⟩ -@[simp] lemma infClosed_preimage_ofDual {s : Set α} : InfClosed (ofDual ⁻¹' s) ↔ SupClosed s := -Iff.rfl +lemma IsSublattice.image [LatticeHomClass F α β] (hs : IsSublattice s) (f : F) : + IsSublattice (f '' s) := ⟨hs.1.image _, hs.2.image _⟩ + +lemma IsSublattice_range [LatticeHomClass F α β] (f : F) : IsSublattice (Set.range f) := + ⟨supClosed_range _, infClosed_range _⟩ + +lemma IsSublattice.prod {t : Set β} (hs : IsSublattice s) (ht : IsSublattice t) : + IsSublattice (s ×ˢ t) := ⟨hs.1.prod ht.1, hs.2.prod ht.2⟩ + +lemma isSublattice_pi {ι : Type*} {α : ι → Type*} [∀ i, Lattice (α i)] {s : Set ι} + {t : ∀ i, Set (α i)} (ht : ∀ i ∈ s, IsSublattice (t i)) : IsSublattice (s.pi t) := + ⟨supClosed_pi λ _i hi ↦ (ht _ hi).1, infClosed_pi λ _i hi ↦ (ht _ hi).2⟩ + +@[simp] lemma supClosed_preimage_toDual {s : Set αᵒᵈ} : + SupClosed (toDual ⁻¹' s) ↔ InfClosed s := Iff.rfl + +@[simp] lemma infClosed_preimage_toDual {s : Set αᵒᵈ} : + InfClosed (toDual ⁻¹' s) ↔ SupClosed s := Iff.rfl + +@[simp] lemma supClosed_preimage_ofDual {s : Set α} : + SupClosed (ofDual ⁻¹' s) ↔ InfClosed s := Iff.rfl + +@[simp] lemma infClosed_preimage_ofDual {s : Set α} : + InfClosed (ofDual ⁻¹' s) ↔ SupClosed s := Iff.rfl + +@[simp] lemma isSublattice_preimage_toDual {s : Set αᵒᵈ} : + IsSublattice (toDual ⁻¹' s) ↔ IsSublattice s := ⟨fun h ↦ ⟨h.2, h.1⟩, fun h ↦ ⟨h.2, h.1⟩⟩ + +@[simp] lemma isSublattice_preimage_ofDual : + IsSublattice (ofDual ⁻¹' s) ↔ IsSublattice s := ⟨fun h ↦ ⟨h.2, h.1⟩, fun h ↦ ⟨h.2, h.1⟩⟩ alias ⟨_, InfClosed.dual⟩ := supClosed_preimage_ofDual alias ⟨_, SupClosed.dual⟩ := infClosed_preimage_ofDual +alias ⟨_, IsSublattice.dual⟩ := isSublattice_preimage_ofDual +alias ⟨_, IsSublattice.of_dual⟩ := isSublattice_preimage_toDual end Lattice @@ -151,12 +239,17 @@ variable [LinearOrder α] @[simp] protected lemma LinearOrder.infClosed (s : Set α) : InfClosed s := λ a ha b hb ↦ by cases le_total a b <;> simp [*] +@[simp] protected lemma LinearOrder.isSublattice (s : Set α) : IsSublattice s := + ⟨LinearOrder.supClosed _, LinearOrder.infClosed _⟩ + end LinearOrder /-! ## Closure -/ +open Finset + section SemilatticeSup -variable [SemilatticeSup α] {s : Set α} {a : α} +variable [SemilatticeSup α] [SemilatticeSup β] {s t : Set α} {a b : α} /-- Every set in a join-semilattice generates a set closed under join. -/ def supClosure : ClosureOperator (Set α) := ClosureOperator.mk₃ @@ -196,10 +289,34 @@ supClosure.idempotent _ @[simp] lemma isLUB_supClosure : IsLUB (supClosure s) a ↔ IsLUB s a := by simp [IsLUB] +lemma sup_mem_supClosure (ha : a ∈ s) (hb : b ∈ s) : a ⊔ b ∈ supClosure s := + supClosed_supClosure (subset_supClosure ha) (subset_supClosure hb) + +lemma finsetSup'_mem_supClosure {ι : Type*} {t : Finset ι} (ht : t.Nonempty) {f : ι → α} + (hf : ∀ i ∈ t, f i ∈ s) : t.sup' ht f ∈ supClosure s := + supClosed_supClosure.finsetSup'_mem _ $ fun _i hi ↦ subset_supClosure $ hf _ hi + +@[simp] lemma supClosure_closed : supClosure.closed = {s : Set α | SupClosed s} := by + ext; exact ClosureOperator.mem_mk₃_closed + +lemma supClosure_min (hst : s ⊆ t) (ht : SupClosed t) : supClosure s ⊆ t := + ClosureOperator.closure_le_mk₃_iff hst ht + +/-- The semilatice generated by a finite set is finite. -/ +protected lemma Set.Finite.supClosure (hs : s.Finite) : (supClosure s).Finite := by + lift s to Finset α using hs + classical + refine' ((s.powerset.filter Finset.Nonempty).attach.image + fun t ↦ t.1.sup' (mem_filter.1 t.2).2 id).finite_toSet.subset _ + rintro _ ⟨t, ht, hts, rfl⟩ + simp only [id_eq, coe_image, mem_image, mem_coe, mem_attach, true_and, Subtype.exists, + Finset.mem_powerset, Finset.not_nonempty_iff_eq_empty, mem_filter] + exact ⟨t, ⟨hts, ht⟩, rfl⟩ + end SemilatticeSup section SemilatticeInf -variable [SemilatticeInf α] {s : Set α} {a : α} +variable [SemilatticeInf α] [SemilatticeInf β] {s t : Set α} {a b : α} /-- Every set in a join-semilattice generates a set closed under join. -/ def infClosure : ClosureOperator (Set α) := ClosureOperator.mk₃ @@ -238,9 +355,101 @@ infClosure.idempotent _ exact le_inf' _ _ λ b hb ↦ ha $ hts hb @[simp] lemma isGLB_infClosure : IsGLB (infClosure s) a ↔ IsGLB s a := by simp [IsGLB] +open Finset + +lemma inf_mem_infClosure (ha : a ∈ s) (hb : b ∈ s) : a ⊓ b ∈ infClosure s := + infClosed_infClosure (subset_infClosure ha) (subset_infClosure hb) + +lemma finsetInf'_mem_infClosure {ι : Type*} {t : Finset ι} (ht : t.Nonempty) {f : ι → α} + (hf : ∀ i ∈ t, f i ∈ s) : t.inf' ht f ∈ infClosure s := + infClosed_infClosure.finsetInf'_mem _ $ fun _i hi ↦ subset_infClosure $ hf _ hi + +@[simp] lemma infClosure_closed : infClosure.closed = {s : Set α | InfClosed s} := by + ext; exact ClosureOperator.mem_mk₃_closed + +lemma infClosure_min (hst : s ⊆ t) (ht : InfClosed t) : infClosure s ⊆ t := + ClosureOperator.closure_le_mk₃_iff hst ht + +/-- The semilatice generated by a finite set is finite. -/ +protected lemma Set.Finite.infClosure (hs : s.Finite) : (infClosure s).Finite := by + lift s to Finset α using hs + classical + refine' ((s.powerset.filter Finset.Nonempty).attach.image + fun t ↦ t.1.inf' (mem_filter.1 t.2).2 id).finite_toSet.subset _ + rintro _ ⟨t, ht, hts, rfl⟩ + simp only [id_eq, coe_image, mem_image, mem_coe, mem_attach, true_and, Subtype.exists, + Finset.mem_powerset, Finset.not_nonempty_iff_eq_empty, mem_filter] + exact ⟨t, ⟨hts, ht⟩, rfl⟩ end SemilatticeInf +section Lattice +variable [Lattice α] {s t : Set α} + +/-- Every set in a join-semilattice generates a set closed under join. -/ +def latticeClosure : ClosureOperator (Set α) := + ClosureOperator.ofPred IsSublattice $ fun _ ↦ isSublattice_sInter + +@[simp] lemma subset_latticeClosure : s ⊆ latticeClosure s := latticeClosure.le_closure _ + +@[simp] lemma isSublattice_latticeClosure : IsSublattice (latticeClosure s) := + ClosureOperator.ofPred_spec _ + +lemma latticeClosure_min (hst : s ⊆ t) (ht : IsSublattice t) : latticeClosure s ⊆ t := by + rw [latticeClosure, ClosureOperator.ofPred] + exact ClosureOperator.closure_le_mk₃_iff hst ht + +lemma latticeClosure_mono : Monotone (latticeClosure : Set α → Set α) := latticeClosure.monotone + +@[simp] lemma latticeClosure_eq_self : latticeClosure s = s ↔ IsSublattice s := + ClosureOperator.mem_closed_ofPred + +@[simp] lemma latticeClosure_closed : latticeClosure.closed = {s : Set α | IsSublattice s} := + ClosureOperator.closed_ofPred + +alias ⟨_, IsSublattice.latticeClosure_eq⟩ := latticeClosure_eq_self + +lemma latticeClosure_idem (s : Set α) : latticeClosure (latticeClosure s) = latticeClosure s := + latticeClosure.idempotent _ + +@[simp] lemma latticeClosure_empty : latticeClosure (∅ : Set α) = ∅ := by simp +@[simp] lemma latticeClosure_singleton (a : α) : latticeClosure {a} = {a} := by simp +@[simp] lemma latticeClosure_univ : latticeClosure (Set.univ : Set α) = Set.univ := by simp + +end Lattice + +section DistribLattice +variable [DistribLattice α] [DistribLattice β] {s : Set α} + +open Finset + +protected lemma SupClosed.infClosure (hs : SupClosed s) : SupClosed (infClosure s) := by + rintro _ ⟨t, ht, hts, rfl⟩ _ ⟨u, hu, hus, rfl⟩ + rw [inf'_sup_inf'] + exact finsetInf'_mem_infClosure _ + fun i hi ↦ hs (hts (mem_product.1 hi).1) (hus (mem_product.1 hi).2) + +protected lemma InfClosed.supClosure (hs : InfClosed s) : InfClosed (supClosure s) := by + rintro _ ⟨t, ht, hts, rfl⟩ _ ⟨u, hu, hus, rfl⟩ + rw [sup'_inf_sup'] + exact finsetSup'_mem_supClosure _ + fun i hi ↦ hs (hts (mem_product.1 hi).1) (hus (mem_product.1 hi).2) + +@[simp] lemma supClosure_infClosure (s : Set α) : supClosure (infClosure s) = latticeClosure s := + le_antisymm (supClosure_min (infClosure_min subset_latticeClosure isSublattice_latticeClosure.2) + isSublattice_latticeClosure.1) $ latticeClosure_min (subset_infClosure.trans subset_supClosure) + ⟨supClosed_supClosure, infClosed_infClosure.supClosure⟩ + +@[simp] lemma infClosure_supClosure (s : Set α) : infClosure (supClosure s) = latticeClosure s := + le_antisymm (infClosure_min (supClosure_min subset_latticeClosure isSublattice_latticeClosure.1) + isSublattice_latticeClosure.2) $ latticeClosure_min (subset_supClosure.trans subset_infClosure) + ⟨supClosed_supClosure.infClosure, infClosed_infClosure⟩ + +lemma Set.Finite.latticeClosure (hs : s.Finite) : (latticeClosure s).Finite := by + rw [←supClosure_infClosure]; exact hs.infClosure.supClosure + +end DistribLattice + /-- A join-semilattice where every sup-closed set has a least upper bound is automatically complete. -/ def SemilatticeSup.toCompleteSemilatticeSup [SemilatticeSup α] (sSup : Set α → α) diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index 47c89babee919..6eddec5929187 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -72,6 +72,7 @@ theorem SupIndep.subset (ht : t.SupIndep f) (h : s ⊆ t) : s.SupIndep f := fun ht (hu.trans h) (h hi) #align finset.sup_indep.subset Finset.SupIndep.subset +@[simp] theorem supIndep_empty (f : ι → α) : (∅ : Finset ι).SupIndep f := fun _ _ a ha => (not_mem_empty a ha).elim #align finset.sup_indep_empty Finset.supIndep_empty @@ -406,10 +407,10 @@ theorem Independent.setIndependent_range (ht : Independent t) : SetIndependent < exact ht.comp' surjective_onto_range #align complete_lattice.independent.set_independent_range CompleteLattice.Independent.setIndependent_range -theorem Independent.injective (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Injective t := by - intro i j h +theorem Independent.injOn (ht : Independent t) : InjOn t {i | t i ≠ ⊥} := by + rintro i _ j (hj : t j ≠ ⊥) h by_contra' contra - apply h_ne_bot j + apply hj suffices t j ≤ ⨆ (k) (_ : k ≠ i), t k by replace ht := (ht i).mono_right this rwa [h, disjoint_self] at ht @@ -417,6 +418,10 @@ theorem Independent.injective (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ · exact Ne.symm contra -- Porting note: needs explicit `f` exact @le_iSup₂ _ _ _ _ (fun x _ => t x) j contra + +theorem Independent.injective (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Injective t := by + suffices univ = {i | t i ≠ ⊥} by rw [injective_iff_injOn_univ, this]; exact ht.injOn + aesop #align complete_lattice.independent.injective CompleteLattice.Independent.injective theorem independent_pair {i j : ι} (hij : i ≠ j) (huniv : ∀ k, k = i ∨ k = j) : @@ -473,6 +478,10 @@ alias ⟨CompleteLattice.Independent.supIndep, Finset.SupIndep.independent⟩ := #align complete_lattice.independent.sup_indep CompleteLattice.Independent.supIndep #align finset.sup_indep.independent Finset.SupIndep.independent +theorem CompleteLattice.Independent.supIndep' [CompleteLattice α] {f : ι → α} (s : Finset ι) + (h : CompleteLattice.Independent f) : s.SupIndep f := + CompleteLattice.Independent.supIndep (h.comp Subtype.coe_injective) + /-- A variant of `CompleteLattice.independent_iff_supIndep` for `Fintype`s. -/ theorem CompleteLattice.independent_iff_supIndep_univ [CompleteLattice α] [Fintype ι] {f : ι → α} : CompleteLattice.Independent f ↔ Finset.univ.SupIndep f := by diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index b8c03b9a7b8f7..7dc5ec72c1f30 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -507,17 +507,17 @@ theorem coe_max [LinearOrder α] (x y : α) : ((max x y : α) : WithBot α) = ma rfl #align with_bot.coe_max WithBot.coe_max -theorem wellFounded_lt [LT α] (h : @WellFounded α (· < ·)) : - @WellFounded (WithBot α) (· < ·) := +instance instWellFoundedLT [LT α] [WellFoundedLT α] : WellFoundedLT (WithBot α) where + wf := have not_lt_bot : ∀ a : WithBot α, ¬ a < ⊥ := (fun.) have acc_bot := ⟨_, by simp [not_lt_bot]⟩ .intro fun | ⊥ => acc_bot - | (a : α) => (h.1 a).rec fun a _ ih => + | (a : α) => (wellFounded_lt.1 a).rec fun a _ ih => .intro _ fun | ⊥, _ => acc_bot | (b : α), hlt => ih _ (some_lt_some.1 hlt) -#align with_bot.well_founded_lt WithBot.wellFounded_lt +#align with_bot.well_founded_lt WithBot.instWellFoundedLT instance denselyOrdered [LT α] [DenselyOrdered α] [NoMinOrder α] : DenselyOrdered (WithBot α) := ⟨fun a b => @@ -1294,11 +1294,11 @@ theorem coe_max [LinearOrder α] (x y : α) : (↑(max x y) : WithTop α) = max rfl #align with_top.coe_max WithTop.coe_max -theorem wellFounded_lt [LT α] (h : @WellFounded α (· < ·)) : - @WellFounded (WithTop α) (· < ·) := +instance instWellFoundedLT [LT α] [WellFoundedLT α] : WellFoundedLT (WithTop α) where + wf := have not_top_lt : ∀ a : WithTop α, ¬ ⊤ < a := (fun.) have acc_some (a : α) : Acc ((· < ·) : WithTop α → WithTop α → Prop) a := - (h.1 a).rec fun _ _ ih => + (wellFounded_lt.1 a).rec fun _ _ ih => .intro _ fun | (b : α), hlt => ih _ (some_lt_some.1 hlt) | ⊤, hlt => nomatch not_top_lt _ hlt @@ -1307,16 +1307,14 @@ theorem wellFounded_lt [LT α] (h : @WellFounded α (· < ·)) : | ⊤ => .intro _ fun | (b : α), _ => acc_some b | ⊤, hlt => nomatch not_top_lt _ hlt -#align with_top.well_founded_lt WithTop.wellFounded_lt +#align with_top.well_founded_lt WithTop.instWellFoundedLT open OrderDual -theorem wellFounded_gt [LT α] (h : @WellFounded α (· > ·)) : - @WellFounded (WithTop α) (· > ·) := - ⟨fun a => by - -- ideally, use rel_hom_class.acc, but that is defined later - have : Acc (· < ·) (WithTop.toDual a) := WellFounded.apply (WithBot.wellFounded_lt - (by convert h using 1)) _ +instance instWellFoundedGT [LT α] [WellFoundedGT α] : WellFoundedGT (WithTop α) where + wf := ⟨fun a => by + -- ideally, use RelHomClass.acc, but that is defined later + have : Acc (· < ·) (WithTop.toDual a) := wellFounded_lt.apply _ revert this generalize ha : WithBot.toDual a = b intro ac @@ -1324,14 +1322,12 @@ theorem wellFounded_gt [LT α] (h : @WellFounded α (· > ·)) : induction' ac with _ H IH generalizing a subst ha exact ⟨_, fun a' h => IH (WithTop.toDual a') (toDual_lt_toDual.mpr h) _ rfl⟩⟩ -#align with_top.well_founded_gt WithTop.wellFounded_gt - -theorem _root_.WithBot.wellFounded_gt [LT α] (h : @WellFounded α (· > ·)) : - @WellFounded (WithBot α) (· > ·) := - ⟨fun a => by - -- ideally, use rel_hom_class.acc, but that is defined later - have : Acc (· < ·) (WithBot.toDual a) := - WellFounded.apply (WithTop.wellFounded_lt (by convert h using 1)) _ +#align with_top.well_founded_gt WithTop.instWellFoundedGT + +instance _root_.WithBot.instWellFoundedGT [LT α] [WellFoundedGT α] : WellFoundedGT (WithBot α) where + wf := ⟨fun a => by + -- ideally, use RelHomClass.acc, but that is defined later + have : Acc (· < ·) (WithBot.toDual a) := wellFounded_lt.apply _ revert this generalize ha : WithBot.toDual a = b intro ac @@ -1339,7 +1335,7 @@ theorem _root_.WithBot.wellFounded_gt [LT α] (h : @WellFounded α (· > ·)) : induction' ac with _ H IH generalizing a subst ha exact ⟨_, fun a' h => IH (WithBot.toDual a') (toDual_lt_toDual.mpr h) _ rfl⟩⟩ -#align with_bot.well_founded_gt WithBot.wellFounded_gt +#align with_bot.well_founded_gt WithBot.instWellFoundedGT instance trichotomous.lt [Preorder α] [IsTrichotomous α (· < ·)] : IsTrichotomous (WithTop α) (· < ·) := @@ -1351,8 +1347,7 @@ instance trichotomous.lt [Preorder α] [IsTrichotomous α (· < ·)] : · simpa [some_eq_coe, IsTrichotomous, coe_eq_coe] using @trichotomous α (· < ·) _ a b⟩ #align with_top.trichotomous.lt WithTop.trichotomous.lt -instance IsWellOrder.lt [Preorder α] [h : IsWellOrder α (· < ·)] : - IsWellOrder (WithTop α) (· < ·) where wf := wellFounded_lt h.wf +instance IsWellOrder.lt [Preorder α] [IsWellOrder α (· < ·)] : IsWellOrder (WithTop α) (· < ·) where #align with_top.is_well_order.lt WithTop.IsWellOrder.lt instance trichotomous.gt [Preorder α] [IsTrichotomous α (· > ·)] : @@ -1365,8 +1360,7 @@ instance trichotomous.gt [Preorder α] [IsTrichotomous α (· > ·)] : · simpa [some_eq_coe, IsTrichotomous, coe_eq_coe] using @trichotomous α (· > ·) _ a b⟩ #align with_top.trichotomous.gt WithTop.trichotomous.gt -instance IsWellOrder.gt [Preorder α] [h : IsWellOrder α (· > ·)] : - IsWellOrder (WithTop α) (· > ·) where wf := wellFounded_gt h.wf +instance IsWellOrder.gt [Preorder α] [IsWellOrder α (· > ·)] : IsWellOrder (WithTop α) (· > ·) where #align with_top.is_well_order.gt WithTop.IsWellOrder.gt instance _root_.WithBot.trichotomous.lt [Preorder α] [h : IsTrichotomous α (· < ·)] : @@ -1374,9 +1368,8 @@ instance _root_.WithBot.trichotomous.lt [Preorder α] [h : IsTrichotomous α (· @WithTop.trichotomous.gt αᵒᵈ _ h #align with_bot.trichotomous.lt WithBot.trichotomous.lt -instance _root_.WithBot.isWellOrder.lt [Preorder α] [h : IsWellOrder α (· < ·)] : - IsWellOrder (WithBot α) (· < ·) := - @WithTop.IsWellOrder.gt αᵒᵈ _ h +instance _root_.WithBot.isWellOrder.lt [Preorder α] [IsWellOrder α (· < ·)] : + IsWellOrder (WithBot α) (· < ·) where #align with_bot.is_well_order.lt WithBot.isWellOrder.lt instance _root_.WithBot.trichotomous.gt [Preorder α] [h : IsTrichotomous α (· > ·)] : diff --git a/Mathlib/Probability/Density.lean b/Mathlib/Probability/Density.lean index 75d8094b946fd..e565539d42f77 100644 --- a/Mathlib/Probability/Density.lean +++ b/Mathlib/Probability/Density.lean @@ -5,6 +5,7 @@ Authors: Kexing Ying -/ import Mathlib.MeasureTheory.Decomposition.RadonNikodym import Mathlib.MeasureTheory.Measure.Haar.OfBasis +import Mathlib.Probability.Independence.Basic #align_import probability.density from "leanprover-community/mathlib"@"c14c8fcde993801fca8946b0d80131a1a81d1520" @@ -136,13 +137,25 @@ theorem lintegral_eq_measure_univ {X : Ω → E} [HasPDF X ℙ μ] : Measure.map_apply (HasPDF.measurable X ℙ μ) MeasurableSet.univ, Set.preimage_univ] #align measure_theory.pdf.lintegral_eq_measure_univ MeasureTheory.pdf.lintegral_eq_measure_univ +theorem unique [IsFiniteMeasure ℙ] {X : Ω → E} [HasPDF X ℙ μ] (f : E → ℝ≥0∞) + (hmf : AEMeasurable f μ) : ℙ.map X = μ.withDensity f ↔ pdf X ℙ μ =ᵐ[μ] f := by + rw [map_eq_withDensity_pdf X ℙ μ] + apply withDensity_eq_iff (measurable_pdf X ℙ μ).aemeasurable hmf + rw [lintegral_eq_measure_univ] + exact measure_ne_top _ _ + +theorem unique' [SigmaFinite μ] {X : Ω → E} [HasPDF X ℙ μ] (f : E → ℝ≥0∞) (hmf : AEMeasurable f μ) : + ℙ.map X = μ.withDensity f ↔ pdf X ℙ μ =ᵐ[μ] f := + map_eq_withDensity_pdf X ℙ μ ▸ + withDensity_eq_iff_of_sigmaFinite (measurable_pdf X ℙ μ).aemeasurable hmf + nonrec theorem ae_lt_top [IsFiniteMeasure ℙ] {μ : Measure E} {X : Ω → E} : ∀ᵐ x ∂μ, pdf X ℙ μ x < ∞ := by by_cases hpdf : HasPDF X ℙ μ · haveI := hpdf refine' ae_lt_top (measurable_pdf X ℙ μ) _ rw [lintegral_eq_measure_univ] - exact (measure_lt_top _ _).ne + exact measure_ne_top _ _ · simp [pdf, hpdf] #align measure_theory.pdf.ae_lt_top MeasureTheory.pdf.ae_lt_top @@ -411,6 +424,33 @@ end IsUniform end +section TwoVariables + +open ProbabilityTheory + +variable {F : Type*} [MeasurableSpace F] {ν : Measure F} {X : Ω → E} {Y : Ω → F} + +/-- Random variables are independent iff their joint density is a product of marginal densities. -/ +theorem indepFun_iff_pdf_prod_eq_pdf_mul_pdf + [IsFiniteMeasure ℙ] [SigmaFinite μ] [SigmaFinite ν] [HasPDF (fun ω ↦ (X ω, Y ω)) ℙ (μ.prod ν)] : + IndepFun X Y ℙ ↔ + pdf (fun ω ↦ (X ω, Y ω)) ℙ (μ.prod ν) =ᵐ[μ.prod ν] fun z ↦ pdf X ℙ μ z.1 * pdf Y ℙ ν z.2 := by + have : HasPDF X ℙ μ := quasiMeasurePreserving_hasPDF' (μ := μ.prod ν) (X := fun ω ↦ (X ω, Y ω)) + quasiMeasurePreserving_fst + have : HasPDF Y ℙ ν := quasiMeasurePreserving_hasPDF' (μ := μ.prod ν) (X := fun ω ↦ (X ω, Y ω)) + quasiMeasurePreserving_snd + have h₀ : (ℙ.map X).prod (ℙ.map Y) = + (μ.prod ν).withDensity fun z ↦ pdf X ℙ μ z.1 * pdf Y ℙ ν z.2 := + prod_eq fun s t hs ht ↦ by rw [withDensity_apply _ (hs.prod ht), ← prod_restrict, + lintegral_prod_mul (measurable_pdf X ℙ μ).aemeasurable (measurable_pdf Y ℙ ν).aemeasurable, + map_eq_set_lintegral_pdf X ℙ μ hs, map_eq_set_lintegral_pdf Y ℙ ν ht] + rw [indepFun_iff_map_prod_eq_prod_map_map (HasPDF.measurable X ℙ μ) (HasPDF.measurable Y ℙ ν), + ← unique, h₀] + exact (((measurable_pdf X ℙ μ).comp measurable_fst).mul + ((measurable_pdf Y ℙ ν).comp measurable_snd)).aemeasurable + +end TwoVariables + end pdf end MeasureTheory diff --git a/Mathlib/Probability/Distributions/Gaussian.lean b/Mathlib/Probability/Distributions/Gaussian.lean index 34970b4575577..ff463656637fe 100644 --- a/Mathlib/Probability/Distributions/Gaussian.lean +++ b/Mathlib/Probability/Distributions/Gaussian.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Lorenzo Luccioli, Rémy Degenne +Authors: Lorenzo Luccioli, Rémy Degenne, Alexander Bentkamp -/ import Mathlib.Analysis.SpecialFunctions.Gaussian import Mathlib.Probability.Notation @@ -21,6 +21,13 @@ We define a Gaussian measure over the reals. If `v = 0`, this is `dirac μ`, otherwise it is defined as the measure with density `gaussianPdf μ v` with respect to the Lebesgue measure. +## Main results + +* `gaussianReal_add_const`: if `X` is a random variable with Gaussian distribution with mean `μ` and + variance `v`, then `X + y` is Gaussian with mean `μ + y` and variance `v`. +* `gaussianReal_const_mul`: if `X` is a random variable with Gaussian distribution with mean `μ` and + variance `v`, then `c * X` is Gaussian with mean `c * μ` and variance `c^2 * v`. + -/ open scoped ENNReal NNReal Real @@ -114,6 +121,40 @@ lemma integral_gaussianPdfReal_eq_one (μ : ℝ) {v : ℝ≥0} (hv : v ≠ 0) : (ae_of_all _ (gaussianPdfReal_nonneg _ _)), ← ENNReal.ofReal_one] at h rwa [← ENNReal.ofReal_eq_ofReal_iff (integral_nonneg (gaussianPdfReal_nonneg _ _)) zero_le_one] +lemma gaussianPdfReal_sub {μ : ℝ} {v : ℝ≥0} (x y : ℝ) : + gaussianPdfReal μ v (x - y) = gaussianPdfReal (μ + y) v x := by + simp only [gaussianPdfReal] + rw [sub_add_eq_sub_sub_swap] + +lemma gaussianPdfReal_add {μ : ℝ} {v : ℝ≥0} (x y : ℝ) : + gaussianPdfReal μ v (x + y) = gaussianPdfReal (μ - y) v x := by + rw [sub_eq_add_neg, ← gaussianPdfReal_sub, sub_eq_add_neg, neg_neg] + +lemma gaussianPdfReal_inv_mul {μ : ℝ} {v : ℝ≥0} {c : ℝ} (hc : c ≠ 0) (x : ℝ) : + gaussianPdfReal μ v (c⁻¹ * x) = |c| * gaussianPdfReal (c * μ) (⟨c^2, sq_nonneg _⟩ * v) x := by + simp only [gaussianPdfReal._eq_1, gt_iff_lt, zero_lt_two, zero_le_mul_left, NNReal.zero_le_coe, + Real.sqrt_mul', one_div, mul_inv_rev, NNReal.coe_mul, NNReal.coe_mk, NNReal.coe_pos] + rw [← mul_assoc] + refine congr_arg₂ _ ?_ ?_ + · field_simp + rw [Real.sqrt_sq_eq_abs] + ring_nf + calc (Real.sqrt ↑v)⁻¹ * (Real.sqrt 2)⁻¹ * (Real.sqrt π)⁻¹ + = (Real.sqrt ↑v)⁻¹ * (Real.sqrt 2)⁻¹ * (Real.sqrt π)⁻¹ * (|c| * |c|⁻¹) := by + rw [mul_inv_cancel, mul_one] + simp only [ne_eq, abs_eq_zero, hc, not_false_eq_true] + _ = (Real.sqrt ↑v)⁻¹ * (Real.sqrt 2)⁻¹ * (Real.sqrt π)⁻¹ * |c| * |c|⁻¹ := by ring + · congr 1 + field_simp + congr 1 + ring + +lemma gaussianPdfReal_mul {μ : ℝ} {v : ℝ≥0} {c : ℝ} (hc : c ≠ 0) (x : ℝ) : + gaussianPdfReal μ v (c * x) + = |c⁻¹| * gaussianPdfReal (c⁻¹ * μ) (⟨(c^2)⁻¹, inv_nonneg.mpr (sq_nonneg _)⟩ * v) x := by + conv_lhs => rw [← inv_inv c, gaussianPdfReal_inv_mul (inv_ne_zero hc)] + simp + /-- The pdf of a Gaussian distribution on ℝ with mean `μ` and variance `v`. -/ noncomputable def gaussianPdf (μ : ℝ) (v : ℝ≥0) (x : ℝ) : ℝ≥0∞ := ENNReal.ofReal (gaussianPdfReal μ v x) @@ -192,6 +233,124 @@ lemma rnDeriv_gaussianReal (μ : ℝ) (v : ℝ≥0) : · rw [gaussianReal_of_var_ne_zero _ hv] exact Measure.rnDeriv_withDensity _ (measurable_gaussianPdf μ v) +section Transformations + +variable {μ : ℝ} {v : ℝ≥0} + +lemma _root_.MeasurableEmbedding.gaussianReal_comap_apply (hv : v ≠ 0) + {f : ℝ → ℝ} (hf : MeasurableEmbedding f) + {f' : ℝ → ℝ} (h_deriv : ∀ x, HasDerivAt f (f' x) x) {s : Set ℝ} (hs : MeasurableSet s) : + (gaussianReal μ v).comap f s + = ENNReal.ofReal (∫ x in s, |f' x| * gaussianPdfReal μ v (f x)) := by + rw [gaussianReal_of_var_ne_zero _ hv, gaussianPdf_def] + exact hf.withDensity_ofReal_comap_apply_eq_integral_abs_deriv_mul' hs h_deriv + (ae_of_all _ (gaussianPdfReal_nonneg _ _)) (integrable_gaussianPdfReal _ _) + +lemma _root_.MeasurableEquiv.gaussianReal_map_symm_apply (hv : v ≠ 0) (f : ℝ ≃ᵐ ℝ) {f' : ℝ → ℝ} + (h_deriv : ∀ x, HasDerivAt f (f' x) x) {s : Set ℝ} (hs : MeasurableSet s) : + (gaussianReal μ v).map f.symm s + = ENNReal.ofReal (∫ x in s, |f' x| * gaussianPdfReal μ v (f x)) := by + rw [gaussianReal_of_var_ne_zero _ hv, gaussianPdf_def] + exact f.withDensity_ofReal_map_symm_apply_eq_integral_abs_deriv_mul' hs h_deriv + (ae_of_all _ (gaussianPdfReal_nonneg _ _)) (integrable_gaussianPdfReal _ _) + +/-- The map of a Gaussian distribution by addition of a constant is a Gaussian. -/ +lemma gaussianReal_map_add_const (y : ℝ) : + (gaussianReal μ v).map (· + y) = gaussianReal (μ + y) v := by + by_cases hv : v = 0 + · simp only [hv, ne_eq, not_true, gaussianReal_zero_var] + exact Measure.map_dirac (measurable_id'.add_const _) _ + let e : ℝ ≃ᵐ ℝ := (Homeomorph.addRight y).symm.toMeasurableEquiv + have he' : ∀ x, HasDerivAt e ((fun _ ↦ 1) x) x := fun _ ↦ (hasDerivAt_id _).sub_const y + change (gaussianReal μ v).map e.symm = gaussianReal (μ + y) v + ext s' hs' + rw [MeasurableEquiv.gaussianReal_map_symm_apply hv e he' hs'] + simp only [abs_neg, abs_one, MeasurableEquiv.coe_mk, Equiv.coe_fn_mk, one_mul, ne_eq] + rw [gaussianReal_apply_eq_integral _ hv hs'] + simp [gaussianPdfReal_sub _ y, Homeomorph.addRight, ← sub_eq_add_neg] + +/-- The map of a Gaussian distribution by addition of a constant is a Gaussian. -/ +lemma gaussianReal_map_const_add (y : ℝ) : + (gaussianReal μ v).map (y + ·) = gaussianReal (μ + y) v := by + simp_rw [add_comm y] + exact gaussianReal_map_add_const y + +/-- The map of a Gaussian distribution by multiplication by a constant is a Gaussian. -/ +lemma gaussianReal_map_const_mul (c : ℝ) : + (gaussianReal μ v).map (c * ·) = gaussianReal (c * μ) (⟨c^2, sq_nonneg _⟩ * v) := by + by_cases hv : v = 0 + · simp only [hv, mul_zero, ne_eq, not_true, gaussianReal_zero_var] + exact Measure.map_dirac (measurable_id'.const_mul c) μ + by_cases hc : c = 0 + · simp only [hc, zero_mul, ne_eq, abs_zero, mul_eq_zero] + rw [Measure.map_const] + simp only [ne_eq, measure_univ, one_smul, mul_eq_zero] + convert (gaussianReal_zero_var 0).symm + simp only [ne_eq, zero_pow', mul_eq_zero, hv, or_false, not_false_eq_true] + rfl + let e : ℝ ≃ᵐ ℝ := (Homeomorph.mulLeft₀ c hc).symm.toMeasurableEquiv + have he' : ∀ x, HasDerivAt e ((fun _ ↦ c⁻¹) x) x := by + suffices ∀ x, HasDerivAt (fun x => c⁻¹ * x) (c⁻¹ * 1) x by rwa [mul_one] at this + exact fun _ ↦ HasDerivAt.const_mul _ (hasDerivAt_id _) + change (gaussianReal μ v).map e.symm = gaussianReal (c * μ) (⟨c^2, sq_nonneg _⟩ * v) + ext s' hs' + rw [MeasurableEquiv.gaussianReal_map_symm_apply hv e he' hs'] + simp only [MeasurableEquiv.coe_mk, Equiv.coe_fn_mk, ne_eq, mul_eq_zero] + rw [gaussianReal_apply_eq_integral _ _ hs'] + swap + · simp only [ne_eq, mul_eq_zero, hv, or_false] + rw [← NNReal.coe_eq] + simp [hc] + simp only [Homeomorph.mulLeft₀, Equiv.toFun_as_coe, Equiv.mulLeft₀_apply, Equiv.invFun_as_coe, + Equiv.mulLeft₀_symm_apply, Homeomorph.toMeasurableEquiv_coe, Homeomorph.homeomorph_mk_coe_symm, + Equiv.coe_fn_symm_mk, gaussianPdfReal_inv_mul hc] + congr with x + suffices |c⁻¹| * |c| = 1 by rw [← mul_assoc, this, one_mul] + rw [abs_inv, inv_mul_cancel] + rwa [ne_eq, abs_eq_zero] + +/-- The map of a Gaussian distribution by multiplication by a constant is a Gaussian. -/ +lemma gaussianReal_map_mul_const (c : ℝ) : + (gaussianReal μ v).map (· * c) = gaussianReal (c * μ) (⟨c^2, sq_nonneg _⟩ * v) := by + simp_rw [mul_comm _ c] + exact gaussianReal_map_const_mul c + +variable {Ω : Type} [MeasureSpace Ω] + +/-- If `X` is a real random variable with Gaussian law with mean `μ` and variance `v`, then `X + y` +has Gaussian law with mean `μ + y` and variance `v`. -/ +lemma gaussianReal_add_const {X : Ω → ℝ} (hX : Measure.map X ℙ = gaussianReal μ v) (y : ℝ) : + Measure.map (fun ω ↦ X ω + y) ℙ = gaussianReal (μ + y) v := by + have hXm : AEMeasurable X := aemeasurable_of_map_neZero (by rw [hX]; infer_instance) + change Measure.map ((fun ω ↦ ω + y) ∘ X) ℙ = gaussianReal (μ + y) v + rw [← AEMeasurable.map_map_of_aemeasurable (measurable_id'.add_const _).aemeasurable hXm, hX, + gaussianReal_map_add_const y] + +/-- If `X` is a real random variable with Gaussian law with mean `μ` and variance `v`, then `y + X` +has Gaussian law with mean `μ + y` and variance `v`. -/ +lemma gaussianReal_const_add {X : Ω → ℝ} (hX : Measure.map X ℙ = gaussianReal μ v) (y : ℝ) : + Measure.map (fun ω ↦ y + X ω) ℙ = gaussianReal (μ + y) v := by + simp_rw [add_comm y] + exact gaussianReal_add_const hX y + +/-- If `X` is a real random variable with Gaussian law with mean `μ` and variance `v`, then `c * X` +has Gaussian law with mean `c * μ` and variance `c^2 * v`. -/ +lemma gaussianReal_const_mul {X : Ω → ℝ} (hX : Measure.map X ℙ = gaussianReal μ v) (c : ℝ) : + Measure.map (fun ω ↦ c * X ω) ℙ = gaussianReal (c * μ) (⟨c^2, sq_nonneg _⟩ * v) := by + have hXm : AEMeasurable X := aemeasurable_of_map_neZero (by rw [hX]; infer_instance) + change Measure.map ((fun ω ↦ c * ω) ∘ X) ℙ = gaussianReal (c * μ) (⟨c^2, sq_nonneg _⟩ * v) + rw [← AEMeasurable.map_map_of_aemeasurable (measurable_id'.const_mul c).aemeasurable hXm, hX] + exact gaussianReal_map_const_mul c + +/-- If `X` is a real random variable with Gaussian law with mean `μ` and variance `v`, then `X * c` +has Gaussian law with mean `c * μ` and variance `c^2 * v`. -/ +lemma gaussianReal_mul_const {X : Ω → ℝ} (hX : Measure.map X ℙ = gaussianReal μ v) (c : ℝ) : + Measure.map (fun ω ↦ X ω * c) ℙ = gaussianReal (c * μ) (⟨c^2, sq_nonneg _⟩ * v) := by + simp_rw [mul_comm _ c] + exact gaussianReal_const_mul hX c + +end Transformations + end GaussianReal end ProbabilityTheory diff --git a/Mathlib/Probability/Independence/Basic.lean b/Mathlib/Probability/Independence/Basic.lean index 976ae8945b3b6..57d3b388ee2ff 100644 --- a/Mathlib/Probability/Independence/Basic.lean +++ b/Mathlib/Probability/Independence/Basic.lean @@ -5,7 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Independence.Kernel -#align_import probability.independence.basic from "leanprover-community/mathlib"@"2f8347015b12b0864dfaf366ec4909eb70c78740" +#align_import probability.independence.basic from "leanprover-community/mathlib"@"001ffdc42920050657fd45bd2b8bfbec8eaaeb29" /-! # Independence of sets of sets and measure spaces (σ-algebras) @@ -66,15 +66,13 @@ when defining `μ` in the example above, the measurable space used is the last o Part A, Chapter 4. -/ -set_option autoImplicit true - -open MeasureTheory MeasurableSpace +open MeasureTheory MeasurableSpace Set open scoped BigOperators MeasureTheory ENNReal namespace ProbabilityTheory -variable {Ω ι : Type*} +variable {Ω ι β γ : Type*} {κ : ι → Type*} section Definitions @@ -145,12 +143,23 @@ def IndepFun {β γ} [MeasurableSpace Ω] [MeasurableSpace β] [MeasurableSpace end Definitions section Definition_lemmas +variable {π : ι → Set (Set Ω)} {m : ι → MeasurableSpace Ω} {_ : MeasurableSpace Ω} {μ : Measure Ω} + {S : Finset ι} {s : ι → Set Ω} lemma iIndepSets_iff [MeasurableSpace Ω] (π : ι → Set (Set Ω)) (μ : Measure Ω) : iIndepSets π μ ↔ ∀ (s : Finset ι) {f : ι → Set Ω} (_H : ∀ i, i ∈ s → f i ∈ π i), μ (⋂ i ∈ s, f i) = ∏ i in s, μ (f i) := by simp only [iIndepSets, kernel.iIndepSets, ae_dirac_eq, Filter.eventually_pure, kernel.const_apply] +lemma iIndepSets.meas_biInter (h : iIndepSets π μ) (s : Finset ι) {f : ι → Set Ω} + (hf : ∀ i, i ∈ s → f i ∈ π i) : μ (⋂ i ∈ s, f i) = ∏ i in s, μ (f i) := + (iIndepSets_iff _ _).1 h s hf + +lemma iIndepSets.meas_iInter [Fintype ι] (h : iIndepSets π μ) (hs : ∀ i, s i ∈ π i) : + μ (⋂ i, s i) = ∏ i, μ (s i) := by simp [←h.meas_biInter _ fun _i _ ↦ hs _] +set_option linter.uppercaseLean3 false in +#align probability_theory.Indep_sets.meas_Inter ProbabilityTheory.iIndepSets.meas_iInter + lemma IndepSets_iff [MeasurableSpace Ω] (s1 s2 : Set (Set Ω)) (μ : Measure Ω) : IndepSets s1 s2 μ ↔ ∀ t1 t2 : Set Ω, t1 ∈ s1 → t2 ∈ s2 → (μ (t1 ∩ t2) = μ t1 * μ t2) := by simp only [IndepSets, kernel.IndepSets, ae_dirac_eq, Filter.eventually_pure, kernel.const_apply] @@ -159,16 +168,25 @@ lemma iIndep_iff_iIndepSets (m : ι → MeasurableSpace Ω) [MeasurableSpace Ω] iIndep m μ ↔ iIndepSets (fun x ↦ {s | MeasurableSet[m x] s}) μ := by simp only [iIndep, iIndepSets, kernel.iIndep] +lemma iIndep.iIndepSets' {m : ι → MeasurableSpace Ω} (hμ : iIndep m μ) : + iIndepSets (fun x ↦ {s | MeasurableSet[m x] s}) μ := (iIndep_iff_iIndepSets _ _).1 hμ + lemma iIndep_iff (m : ι → MeasurableSpace Ω) [MeasurableSpace Ω] (μ : Measure Ω) : iIndep m μ ↔ ∀ (s : Finset ι) {f : ι → Set Ω} (_H : ∀ i, i ∈ s → MeasurableSet[m i] (f i)), μ (⋂ i ∈ s, f i) = ∏ i in s, μ (f i) := by simp only [iIndep_iff_iIndepSets, iIndepSets_iff]; rfl -lemma Indep_iff_IndepSets (m₁ m₂ : MeasurableSpace Ω) [MeasurableSpace Ω] (μ : Measure Ω ) : +lemma iIndep.meas_biInter (hμ : iIndep m μ) (hs : ∀ i, i ∈ S → MeasurableSet[m i] (s i)) : + μ (⋂ i ∈ S, s i) = ∏ i in S, μ (s i) := (iIndep_iff _ _).1 hμ _ hs + +lemma iIndep.meas_iInter [Fintype ι] (hμ : iIndep m μ) (hs : ∀ i, MeasurableSet[m i] (s i)) : + μ (⋂ i, s i) = ∏ i, μ (s i) := by simp [←hμ.meas_biInter fun _ _ ↦ hs _] + +lemma Indep_iff_IndepSets (m₁ m₂ : MeasurableSpace Ω) [MeasurableSpace Ω] (μ : Measure Ω) : Indep m₁ m₂ μ ↔ IndepSets {s | MeasurableSet[m₁] s} {s | MeasurableSet[m₂] s} μ := by simp only [Indep, IndepSets, kernel.Indep] -lemma Indep_iff (m₁ m₂ : MeasurableSpace Ω) [MeasurableSpace Ω] (μ : Measure Ω ) : +lemma Indep_iff (m₁ m₂ : MeasurableSpace Ω) [MeasurableSpace Ω] (μ : Measure Ω) : Indep m₁ m₂ μ ↔ ∀ t1 t2, MeasurableSet[m₁] t1 → MeasurableSet[m₂] t2 → μ (t1 ∩ t2) = μ t1 * μ t2 := by rw [Indep_iff_IndepSets, IndepSets_iff]; rfl @@ -197,6 +215,10 @@ lemma iIndepFun_iff_iIndep [MeasurableSpace Ω] {β : ι → Type*} iIndepFun m f μ ↔ iIndep (fun x ↦ (m x).comap (f x)) μ := by simp only [iIndepFun, iIndep, kernel.iIndepFun] +protected lemma iIndepFun.iIndep {m : ∀ i, MeasurableSpace (κ i)} {f : ∀ x : ι, Ω → κ x} + (hf : iIndepFun m f μ) : + iIndep (fun x ↦ (m x).comap (f x)) μ := hf + lemma iIndepFun_iff [MeasurableSpace Ω] {β : ι → Type*} (m : ∀ x : ι, MeasurableSpace (β x)) (f : ∀ x : ι, Ω → β x) (μ : Measure Ω) : iIndepFun m f μ ↔ ∀ (s : Finset ι) {f' : ι → Set Ω} @@ -204,6 +226,14 @@ lemma iIndepFun_iff [MeasurableSpace Ω] {β : ι → Type*} μ (⋂ i ∈ s, f' i) = ∏ i in s, μ (f' i) := by simp only [iIndepFun_iff_iIndep, iIndep_iff] +lemma iIndepFun.meas_biInter {m : ∀ i, MeasurableSpace (κ i)} {f : ∀ x : ι, Ω → κ x} + (hf : iIndepFun m f μ) (hs : ∀ i, i ∈ S → MeasurableSet[(m i).comap (f i)] (s i)) : + μ (⋂ i ∈ S, s i) = ∏ i in S, μ (s i) := hf.iIndep.meas_biInter hs + +lemma iIndepFun.meas_iInter [Fintype ι] {m : ∀ i, MeasurableSpace (κ i)} {f : ∀ x : ι, Ω → κ x} + (hf : iIndepFun m f μ) (hs : ∀ i, MeasurableSet[(m i).comap (f i)] (s i)) : + μ (⋂ i, s i) = ∏ i, μ (s i) := hf.iIndep.meas_iInter hs + lemma IndepFun_iff_Indep [MeasurableSpace Ω] [mβ : MeasurableSpace β] [mγ : MeasurableSpace γ] (f : Ω → β) (g : Ω → γ) (μ : Measure Ω) : IndepFun f g μ ↔ Indep (MeasurableSpace.comap f mβ) (MeasurableSpace.comap g mγ) μ := by @@ -215,6 +245,12 @@ lemma IndepFun_iff {β γ} [MeasurableSpace Ω] [mβ : MeasurableSpace β] [mγ → MeasurableSet[MeasurableSpace.comap g mγ] t2 → μ (t1 ∩ t2) = μ t1 * μ t2 := by rw [IndepFun_iff_Indep, Indep_iff] +lemma IndepFun.meas_inter [mβ : MeasurableSpace β] [mγ : MeasurableSpace γ] {f : Ω → β} {g : Ω → γ} + (hfg : IndepFun f g μ) {s t : Set Ω} (hs : MeasurableSet[mβ.comap f] s) + (ht : MeasurableSet[mγ.comap g] t) : + μ (s ∩ t) = μ s * μ t := + (IndepFun_iff _ _ _).1 hfg _ _ hs ht + end Definition_lemmas section Indep @@ -472,7 +508,7 @@ theorem iIndepSets.iIndep [IsProbabilityMeasure μ] (m : ι → MeasurableSpace iIndep m μ := kernel.iIndepSets.iIndep m h_le π h_pi h_generate h_ind set_option linter.uppercaseLean3 false in -#align probability_theory.Indep_sets.Indep ProbabilityTheory.IndepSets.indep +#align probability_theory.Indep_sets.Indep ProbabilityTheory.iIndepSets.iIndep end FromPiSystemsToMeasurableSpaces @@ -486,7 +522,7 @@ We prove the following equivalences on `IndepSet`, for measurable sets `s, t`. -/ -variable {s t : Set Ω} (S T : Set (Set Ω)) +variable {_ : MeasurableSpace Ω} {μ : Measure Ω} {s t : Set Ω} (S T : Set (Set Ω)) theorem indepSet_iff_indepSets_singleton {_m0 : MeasurableSpace Ω} (hs_meas : MeasurableSet s) (ht_meas : MeasurableSet t) (μ : Measure Ω := by volume_tac) @@ -521,6 +557,45 @@ theorem indep_iff_forall_indepSet (m₁ m₂ : MeasurableSpace Ω) {_m0 : Measur kernel.indep_iff_forall_indepSet m₁ m₂ _ _ #align probability_theory.indep_iff_forall_indep_set ProbabilityTheory.indep_iff_forall_indepSet +theorem iIndep_comap_mem_iff {f : ι → Set Ω} : + iIndep (fun i => MeasurableSpace.comap (· ∈ f i) ⊤) μ ↔ iIndepSet f μ := + kernel.iIndep_comap_mem_iff +set_option linter.uppercaseLean3 false in +#align probability_theory.Indep_comap_mem_iff ProbabilityTheory.iIndep_comap_mem_iff + +alias ⟨_, iIndepSet.iIndep_comap_mem⟩ := iIndep_comap_mem_iff +set_option linter.uppercaseLean3 false in +#align probability_theory.Indep_set.Indep_comap_mem ProbabilityTheory.iIndepSet.iIndep_comap_mem + +theorem iIndepSets_singleton_iff {s : ι → Set Ω} : + iIndepSets (fun i ↦ {s i}) μ ↔ ∀ t, μ (⋂ i ∈ t, s i) = ∏ i in t, μ (s i) := by + simp_rw [iIndepSets, kernel.iIndepSets_singleton_iff, ae_dirac_eq, Filter.eventually_pure, + kernel.const_apply] +set_option linter.uppercaseLean3 false in +#align probability_theory.Indep_sets_singleton_iff ProbabilityTheory.iIndepSets_singleton_iff + +variable [IsProbabilityMeasure μ] + +theorem iIndepSet_iff_iIndepSets_singleton {f : ι → Set Ω} (hf : ∀ i, MeasurableSet (f i)) : + iIndepSet f μ ↔ iIndepSets (fun i ↦ {f i}) μ := + kernel.iIndepSet_iff_iIndepSets_singleton hf +set_option linter.uppercaseLean3 false in +#align probability_theory.Indep_set_iff_Indep_sets_singleton ProbabilityTheory.iIndepSet_iff_iIndepSets_singleton + +theorem iIndepSet_iff_meas_biInter {f : ι → Set Ω} (hf : ∀ i, MeasurableSet (f i)) : + iIndepSet f μ ↔ ∀ s, μ (⋂ i ∈ s, f i) = ∏ i in s, μ (f i) := by + simp_rw [iIndepSet, kernel.iIndepSet_iff_meas_biInter hf, ae_dirac_eq, Filter.eventually_pure, + kernel.const_apply] +set_option linter.uppercaseLean3 false in +#align probability_theory.Indep_set_iff_measure_Inter_eq_prod ProbabilityTheory.iIndepSet_iff_meas_biInter + +theorem iIndepSets.iIndepSet_of_mem {π : ι → Set (Set Ω)} {f : ι → Set Ω} + (hfπ : ∀ i, f i ∈ π i) (hf : ∀ i, MeasurableSet (f i)) + (hπ : iIndepSets π μ) : iIndepSet f μ := + kernel.iIndepSets.iIndepSet_of_mem hfπ hf hπ +set_option linter.uppercaseLean3 false in +#align probability_theory.Indep_sets.Indep_set_of_mem ProbabilityTheory.iIndepSets.iIndepSet_of_mem + end IndepSet section IndepFun @@ -559,6 +634,21 @@ theorem indepFun_iff_indepSet_preimage {mβ : MeasurableSpace β} {mβ' : Measur Filter.eventually_pure, kernel.const_apply] #align probability_theory.indep_fun_iff_indep_set_preimage ProbabilityTheory.indepFun_iff_indepSet_preimage +theorem indepFun_iff_map_prod_eq_prod_map_map {mβ : MeasurableSpace β} {mβ' : MeasurableSpace β'} + [IsFiniteMeasure μ] (hf : Measurable f) (hg : Measurable g) : + IndepFun f g μ ↔ μ.map (fun ω ↦ (f ω, g ω)) = (μ.map f).prod (μ.map g) := by + rw [indepFun_iff_measure_inter_preimage_eq_mul] + have h₀ {s : Set β} {t : Set β'} (hs : MeasurableSet s) (ht : MeasurableSet t) : + μ (f ⁻¹' s) * μ (g ⁻¹' t) = μ.map f s * μ.map g t ∧ + μ (f ⁻¹' s ∩ g ⁻¹' t) = μ.map (fun ω ↦ (f ω, g ω)) (s ×ˢ t) := + ⟨by rw [Measure.map_apply hf hs, Measure.map_apply hg ht], + (Measure.map_apply (hf.prod_mk hg) (hs.prod ht)).symm⟩ + constructor + · refine fun h ↦ (Measure.prod_eq fun s t hs ht ↦ ?_).symm + rw [← (h₀ hs ht).1, ← (h₀ hs ht).2, h s t hs ht] + · intro h s t hs ht + rw [(h₀ hs ht).1, (h₀ hs ht).2, h, Measure.prod_prod] + @[symm] nonrec theorem IndepFun.symm {_ : MeasurableSpace β} {f g : Ω → β} (hfg : IndepFun f g μ) : IndepFun g f μ := hfg.symm diff --git a/Mathlib/Probability/Independence/Kernel.lean b/Mathlib/Probability/Independence/Kernel.lean index d3131e24bd68a..9ce35e757abe2 100644 --- a/Mathlib/Probability/Independence/Kernel.lean +++ b/Mathlib/Probability/Independence/Kernel.lean @@ -114,6 +114,50 @@ def IndepFun {β γ} {_mΩ : MeasurableSpace Ω} [mβ : MeasurableSpace β] [mγ end Definitions +section ByDefinition + +variable {β : ι → Type*} {mβ : ∀ i, MeasurableSpace (β i)} + {_mα : MeasurableSpace α} {m : ι → MeasurableSpace Ω} {_mΩ : MeasurableSpace Ω} + {κ : kernel α Ω} {μ : Measure α} + {π : ι → Set (Set Ω)} {s : ι → Set Ω} {S : Finset ι} {f : ∀ x : ι, Ω → β x} + +lemma iIndepSets.meas_biInter (h : iIndepSets π κ μ) (s : Finset ι) + {f : ι → Set Ω} (hf : ∀ i, i ∈ s → f i ∈ π i) : + ∀ᵐ a ∂μ, κ a (⋂ i ∈ s, f i) = ∏ i in s, κ a (f i) := h s hf + +lemma iIndepSets.meas_iInter [Fintype ι] (h : iIndepSets π κ μ) (hs : ∀ i, s i ∈ π i) : + ∀ᵐ a ∂μ, κ a (⋂ i, s i) = ∏ i, κ a (s i) := by + filter_upwards [h.meas_biInter Finset.univ (fun _i _ ↦ hs _)] with a ha using by simp [← ha] + +lemma iIndep.iIndepSets' (hμ : iIndep m κ μ) : + iIndepSets (fun x ↦ {s | MeasurableSet[m x] s}) κ μ := hμ + +lemma iIndep.meas_biInter (hμ : iIndep m κ μ) (hs : ∀ i, i ∈ S → MeasurableSet[m i] (s i)) : + ∀ᵐ a ∂μ, κ a (⋂ i ∈ S, s i) = ∏ i in S, κ a (s i) := hμ _ hs + +lemma iIndep.meas_iInter [Fintype ι] (h : iIndep m κ μ) (hs : ∀ i, MeasurableSet[m i] (s i)) : + ∀ᵐ a ∂μ, κ a (⋂ i, s i) = ∏ i, κ a (s i) := by + filter_upwards [h.meas_biInter (fun i (_ : i ∈ Finset.univ) ↦ hs _)] with a ha + simp [← ha] + +protected lemma iIndepFun.iIndep (hf : iIndepFun mβ f κ μ) : + iIndep (fun x ↦ (mβ x).comap (f x)) κ μ := hf + +lemma iIndepFun.meas_biInter (hf : iIndepFun mβ f κ μ) + (hs : ∀ i, i ∈ S → MeasurableSet[(mβ i).comap (f i)] (s i)) : + ∀ᵐ a ∂μ, κ a (⋂ i ∈ S, s i) = ∏ i in S, κ a (s i) := hf.iIndep.meas_biInter hs + +lemma iIndepFun.meas_iInter [Fintype ι] (hf : iIndepFun mβ f κ μ) + (hs : ∀ i, MeasurableSet[(mβ i).comap (f i)] (s i)) : + ∀ᵐ a ∂μ, κ a (⋂ i, s i) = ∏ i, κ a (s i) := hf.iIndep.meas_iInter hs + +lemma IndepFun.meas_inter {β γ : Type*} [mβ : MeasurableSpace β] [mγ : MeasurableSpace γ] + {f : Ω → β} {g : Ω → γ} (hfg : IndepFun f g κ μ) + {s t : Set Ω} (hs : MeasurableSet[mβ.comap f] s) (ht : MeasurableSet[mγ.comap g] t) : + ∀ᵐ a ∂μ, κ a (s ∩ t) = κ a s * κ a t := hfg _ _ hs ht + +end ByDefinition + section Indep variable {_mα : MeasurableSpace α} @@ -228,6 +272,20 @@ theorem IndepSets.bInter {s : ι → Set (Set Ω)} {s' : Set (Set Ω)} {_mΩ : M rcases h with ⟨n, hn, h⟩ exact h t1 t2 (Set.biInter_subset_of_mem hn ht1) ht2 +theorem iIndep_comap_mem_iff {f : ι → Set Ω} {_mΩ : MeasurableSpace Ω} + {κ : kernel α Ω} {μ : Measure α} : + iIndep (fun i => MeasurableSpace.comap (· ∈ f i) ⊤) κ μ ↔ iIndepSet f κ μ := by + simp_rw [← generateFrom_singleton]; rfl + +theorem iIndepSets_singleton_iff {s : ι → Set Ω} {_mΩ : MeasurableSpace Ω} + {κ : kernel α Ω} {μ : Measure α} : + iIndepSets (fun i ↦ {s i}) κ μ ↔ + ∀ S : Finset ι, ∀ᵐ a ∂μ, κ a (⋂ i ∈ S, s i) = ∏ i in S, κ a (s i) := by + refine ⟨fun h S ↦ h S (fun i _ ↦ rfl), fun h S f hf ↦ ?_⟩ + filter_upwards [h S] with a ha + have : ∀ i ∈ S, κ a (f i) = κ a (s i) := fun i hi ↦ by rw [hf i hi] + rwa [Finset.prod_congr rfl this, Set.iInter₂_congr hf] + theorem indepSets_singleton_iff {s t : Set Ω} {_mΩ : MeasurableSpace Ω} {κ : kernel α Ω} {μ : Measure α} : IndepSets {s} {t} κ μ ↔ ∀ᵐ a ∂μ, κ a (s ∩ t) = κ a s * κ a t := @@ -597,8 +655,29 @@ We prove the following equivalences on `IndepSet`, for measurable sets `s, t`. * `IndepSet s t κ μ ↔ IndepSets {s} {t} κ μ`. -/ +variable {_mα : MeasurableSpace α} -variable {s t : Set Ω} (S T : Set (Set Ω)) {_mα : MeasurableSpace α} +theorem iIndepSet_iff_iIndepSets_singleton {_mΩ : MeasurableSpace Ω} {κ : kernel α Ω} + [IsMarkovKernel κ] {μ : Measure α} {f : ι → Set Ω} + (hf : ∀ i, MeasurableSet (f i)) : + iIndepSet f κ μ ↔ iIndepSets (fun i ↦ {f i}) κ μ := + ⟨iIndep.iIndepSets fun _ ↦ rfl, + iIndepSets.iIndep _ (fun i ↦ generateFrom_le <| by rintro t (rfl : t = _); exact hf _) _ + (fun _ ↦ IsPiSystem.singleton _) fun _ ↦ rfl⟩ + +theorem iIndepSet_iff_meas_biInter {_mΩ : MeasurableSpace Ω} {κ : kernel α Ω} + [IsMarkovKernel κ] {μ : Measure α} {f : ι → Set Ω} (hf : ∀ i, MeasurableSet (f i)) : + iIndepSet f κ μ ↔ ∀ s, ∀ᵐ a ∂μ, κ a (⋂ i ∈ s, f i) = ∏ i in s, κ a (f i) := + (iIndepSet_iff_iIndepSets_singleton hf).trans iIndepSets_singleton_iff + +theorem iIndepSets.iIndepSet_of_mem {_mΩ : MeasurableSpace Ω} {κ : kernel α Ω} + [IsMarkovKernel κ] {μ : Measure α} {π : ι → Set (Set Ω)} {f : ι → Set Ω} + (hfπ : ∀ i, f i ∈ π i) (hf : ∀ i, MeasurableSet (f i)) + (hπ : iIndepSets π κ μ) : + iIndepSet f κ μ := + (iIndepSet_iff_meas_biInter hf).2 fun _t ↦ hπ.meas_biInter _ fun _i _ ↦ hfπ _ + +variable {s t : Set Ω} (S T : Set (Set Ω)) theorem indepSet_iff_indepSets_singleton {m0 : MeasurableSpace Ω} (hs_meas : MeasurableSet s) (ht_meas : MeasurableSet t) (κ : kernel α Ω) (μ : Measure α) diff --git a/Mathlib/Probability/Integration.lean b/Mathlib/Probability/Integration.lean index 10faa875ee53e..164def57b5de1 100644 --- a/Mathlib/Probability/Integration.lean +++ b/Mathlib/Probability/Integration.lean @@ -172,7 +172,7 @@ theorem IndepFun.integrable_left_of_integrable_mul {β : Type*} [MeasurableSpace have A : (∫⁻ ω, ‖X ω * Y ω‖₊ ∂μ) < ∞ := h'XY.2 simp only [nnnorm_mul, ENNReal.coe_mul] at A rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun'' hX.ennnorm hY.ennnorm J, H] at A - simp only [ENNReal.top_mul I] at A + simp only [ENNReal.top_mul I, lt_self_iff_false] at A #align probability_theory.indep_fun.integrable_left_of_integrable_mul ProbabilityTheory.IndepFun.integrable_left_of_integrable_mul /-- If the product of two independent real-valued random variables is integrable and the @@ -194,7 +194,7 @@ theorem IndepFun.integrable_right_of_integrable_mul {β : Type*} [MeasurableSpac have A : (∫⁻ ω, ‖X ω * Y ω‖₊ ∂μ) < ∞ := h'XY.2 simp only [nnnorm_mul, ENNReal.coe_mul] at A rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun'' hX.ennnorm hY.ennnorm J, H] at A - simp only [ENNReal.mul_top I] at A + simp only [ENNReal.mul_top I, lt_self_iff_false] at A #align probability_theory.indep_fun.integrable_right_of_integrable_mul ProbabilityTheory.IndepFun.integrable_right_of_integrable_mul /-- The (Bochner) integral of the product of two independent, nonnegative random diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index ae67dc630686d..27ef08a67dd46 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -697,15 +697,15 @@ theorem lintegral_prodMkLeft (κ : kernel α β) (ca : γ × α) (g : β → ℝ #align probability_theory.kernel.lintegral_prod_mk_left ProbabilityTheory.kernel.lintegral_prodMkLeft instance IsMarkovKernel.prodMkLeft (κ : kernel α β) [IsMarkovKernel κ] : - IsMarkovKernel (prodMkLeft γ κ) := by rw [prodMkLeft]; infer_instance + IsMarkovKernel (prodMkLeft γ κ) := by rw [kernel.prodMkLeft]; infer_instance #align probability_theory.kernel.is_markov_kernel.prod_mk_left ProbabilityTheory.kernel.IsMarkovKernel.prodMkLeft instance IsFiniteKernel.prodMkLeft (κ : kernel α β) [IsFiniteKernel κ] : - IsFiniteKernel (prodMkLeft γ κ) := by rw [prodMkLeft]; infer_instance + IsFiniteKernel (prodMkLeft γ κ) := by rw [kernel.prodMkLeft]; infer_instance #align probability_theory.kernel.is_finite_kernel.prod_mk_left ProbabilityTheory.kernel.IsFiniteKernel.prodMkLeft instance IsSFiniteKernel.prodMkLeft (κ : kernel α β) [IsSFiniteKernel κ] : - IsSFiniteKernel (prodMkLeft γ κ) := by rw [prodMkLeft]; infer_instance + IsSFiniteKernel (prodMkLeft γ κ) := by rw [kernel.prodMkLeft]; infer_instance #align probability_theory.kernel.is_s_finite_kernel.prod_mk_left ProbabilityTheory.kernel.IsSFiniteKernel.prodMkLeft /-- Define a `kernel (β × α) γ` from a `kernel (α × β) γ` by taking the comap of `Prod.swap`. -/ @@ -728,15 +728,15 @@ theorem lintegral_swapLeft (κ : kernel (α × β) γ) (a : β × α) (g : γ #align probability_theory.kernel.lintegral_swap_left ProbabilityTheory.kernel.lintegral_swapLeft instance IsMarkovKernel.swapLeft (κ : kernel (α × β) γ) [IsMarkovKernel κ] : - IsMarkovKernel (swapLeft κ) := by rw [swapLeft]; infer_instance + IsMarkovKernel (swapLeft κ) := by rw [kernel.swapLeft]; infer_instance #align probability_theory.kernel.is_markov_kernel.swap_left ProbabilityTheory.kernel.IsMarkovKernel.swapLeft instance IsFiniteKernel.swapLeft (κ : kernel (α × β) γ) [IsFiniteKernel κ] : - IsFiniteKernel (swapLeft κ) := by rw [swapLeft]; infer_instance + IsFiniteKernel (swapLeft κ) := by rw [kernel.swapLeft]; infer_instance #align probability_theory.kernel.is_finite_kernel.swap_left ProbabilityTheory.kernel.IsFiniteKernel.swapLeft instance IsSFiniteKernel.swapLeft (κ : kernel (α × β) γ) [IsSFiniteKernel κ] : - IsSFiniteKernel (swapLeft κ) := by rw [swapLeft]; infer_instance + IsSFiniteKernel (swapLeft κ) := by rw [kernel.swapLeft]; infer_instance #align probability_theory.kernel.is_s_finite_kernel.swap_left ProbabilityTheory.kernel.IsSFiniteKernel.swapLeft /-- Define a `kernel α (γ × β)` from a `kernel α (β × γ)` by taking the map of `Prod.swap`. -/ @@ -759,15 +759,15 @@ theorem lintegral_swapRight (κ : kernel α (β × γ)) (a : α) {g : γ × β #align probability_theory.kernel.lintegral_swap_right ProbabilityTheory.kernel.lintegral_swapRight instance IsMarkovKernel.swapRight (κ : kernel α (β × γ)) [IsMarkovKernel κ] : - IsMarkovKernel (swapRight κ) := by rw [swapRight]; infer_instance + IsMarkovKernel (swapRight κ) := by rw [kernel.swapRight]; infer_instance #align probability_theory.kernel.is_markov_kernel.swap_right ProbabilityTheory.kernel.IsMarkovKernel.swapRight instance IsFiniteKernel.swapRight (κ : kernel α (β × γ)) [IsFiniteKernel κ] : - IsFiniteKernel (swapRight κ) := by rw [swapRight]; infer_instance + IsFiniteKernel (swapRight κ) := by rw [kernel.swapRight]; infer_instance #align probability_theory.kernel.is_finite_kernel.swap_right ProbabilityTheory.kernel.IsFiniteKernel.swapRight instance IsSFiniteKernel.swapRight (κ : kernel α (β × γ)) [IsSFiniteKernel κ] : - IsSFiniteKernel (swapRight κ) := by rw [swapRight]; infer_instance + IsSFiniteKernel (swapRight κ) := by rw [kernel.swapRight]; infer_instance #align probability_theory.kernel.is_s_finite_kernel.swap_right ProbabilityTheory.kernel.IsSFiniteKernel.swapRight /-- Define a `kernel α β` from a `kernel α (β × γ)` by taking the map of the first projection. -/ @@ -789,15 +789,15 @@ theorem lintegral_fst (κ : kernel α (β × γ)) (a : α) {g : β → ℝ≥0 #align probability_theory.kernel.lintegral_fst ProbabilityTheory.kernel.lintegral_fst instance IsMarkovKernel.fst (κ : kernel α (β × γ)) [IsMarkovKernel κ] : IsMarkovKernel (fst κ) := by - rw [fst]; infer_instance + rw [kernel.fst]; infer_instance #align probability_theory.kernel.is_markov_kernel.fst ProbabilityTheory.kernel.IsMarkovKernel.fst instance IsFiniteKernel.fst (κ : kernel α (β × γ)) [IsFiniteKernel κ] : IsFiniteKernel (fst κ) := by - rw [fst]; infer_instance + rw [kernel.fst]; infer_instance #align probability_theory.kernel.is_finite_kernel.fst ProbabilityTheory.kernel.IsFiniteKernel.fst instance IsSFiniteKernel.fst (κ : kernel α (β × γ)) [IsSFiniteKernel κ] : IsSFiniteKernel (fst κ) := - by rw [fst]; infer_instance + by rw [kernel.fst]; infer_instance #align probability_theory.kernel.is_s_finite_kernel.fst ProbabilityTheory.kernel.IsSFiniteKernel.fst /-- Define a `kernel α γ` from a `kernel α (β × γ)` by taking the map of the second projection. -/ @@ -819,15 +819,15 @@ theorem lintegral_snd (κ : kernel α (β × γ)) (a : α) {g : γ → ℝ≥0 #align probability_theory.kernel.lintegral_snd ProbabilityTheory.kernel.lintegral_snd instance IsMarkovKernel.snd (κ : kernel α (β × γ)) [IsMarkovKernel κ] : IsMarkovKernel (snd κ) := by - rw [snd]; infer_instance + rw [kernel.snd]; infer_instance #align probability_theory.kernel.is_markov_kernel.snd ProbabilityTheory.kernel.IsMarkovKernel.snd instance IsFiniteKernel.snd (κ : kernel α (β × γ)) [IsFiniteKernel κ] : IsFiniteKernel (snd κ) := by - rw [snd]; infer_instance + rw [kernel.snd]; infer_instance #align probability_theory.kernel.is_finite_kernel.snd ProbabilityTheory.kernel.IsFiniteKernel.snd instance IsSFiniteKernel.snd (κ : kernel α (β × γ)) [IsSFiniteKernel κ] : IsSFiniteKernel (snd κ) := - by rw [snd]; infer_instance + by rw [kernel.snd]; infer_instance #align probability_theory.kernel.is_s_finite_kernel.snd ProbabilityTheory.kernel.IsSFiniteKernel.snd end FstSnd @@ -932,15 +932,15 @@ theorem lintegral_prod (κ : kernel α β) [IsSFiniteKernel κ] (η : kernel α #align probability_theory.kernel.lintegral_prod ProbabilityTheory.kernel.lintegral_prod instance IsMarkovKernel.prod (κ : kernel α β) [IsMarkovKernel κ] (η : kernel α γ) - [IsMarkovKernel η] : IsMarkovKernel (κ ×ₖ η) := by rw [prod]; infer_instance + [IsMarkovKernel η] : IsMarkovKernel (κ ×ₖ η) := by rw [kernel.prod]; infer_instance #align probability_theory.kernel.is_markov_kernel.prod ProbabilityTheory.kernel.IsMarkovKernel.prod instance IsFiniteKernel.prod (κ : kernel α β) [IsFiniteKernel κ] (η : kernel α γ) - [IsFiniteKernel η] : IsFiniteKernel (κ ×ₖ η) := by rw [prod]; infer_instance + [IsFiniteKernel η] : IsFiniteKernel (κ ×ₖ η) := by rw [kernel.prod]; infer_instance #align probability_theory.kernel.is_finite_kernel.prod ProbabilityTheory.kernel.IsFiniteKernel.prod instance IsSFiniteKernel.prod (κ : kernel α β) (η : kernel α γ) : - IsSFiniteKernel (κ ×ₖ η) := by rw [prod]; infer_instance + IsSFiniteKernel (κ ×ₖ η) := by rw [kernel.prod]; infer_instance #align probability_theory.kernel.is_s_finite_kernel.prod ProbabilityTheory.kernel.IsSFiniteKernel.prod end Prod diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean index e53c30c8b9f42..5290b82b20420 100644 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ b/Mathlib/Probability/Kernel/Disintegration.lean @@ -128,7 +128,7 @@ theorem set_lintegral_condKernelReal_prod {s : Set α} (hs : MeasurableSet s) {t prod_iUnion, measure_iUnion] · simp_rw [hf_eq] · intro i j hij - rw [Function.onFun, disjoint_prod] + rw [Function.onFun, Set.disjoint_prod] exact Or.inr (hf_disj hij) · exact fun i => MeasurableSet.prod hs (hf_meas i) #align probability_theory.set_lintegral_cond_kernel_real_prod ProbabilityTheory.set_lintegral_condKernelReal_prod diff --git a/Mathlib/Probability/Kernel/WithDensity.lean b/Mathlib/Probability/Kernel/WithDensity.lean index 392b43ca1ae83..9d86218cbec7a 100644 --- a/Mathlib/Probability/Kernel/WithDensity.lean +++ b/Mathlib/Probability/Kernel/WithDensity.lean @@ -203,6 +203,8 @@ theorem isSFiniteKernel_withDensity_of_isFiniteKernel (κ : kernel α β) [IsFin refine' isSFiniteKernel_sum fun n => _ suffices IsFiniteKernel (withDensity κ (fs n)) by haveI := this; infer_instance refine' isFiniteKernel_withDensity_of_bounded _ (ENNReal.coe_ne_top : ↑n + 1 ≠ ∞) fun a b => _ + -- After leanprover/lean4#2734, we need to do beta reduction before `norm_cast` + beta_reduce norm_cast calc fs n a b ≤ min (f a b) (n + 1) := tsub_le_self diff --git a/Mathlib/Probability/Martingale/Convergence.lean b/Mathlib/Probability/Martingale/Convergence.lean index 63ea1272dbcee..d4383382f162c 100644 --- a/Mathlib/Probability/Martingale/Convergence.lean +++ b/Mathlib/Probability/Martingale/Convergence.lean @@ -399,7 +399,7 @@ theorem Integrable.tendsto_ae_condexp (hg : Integrable g μ) by_cases hnm : n ≤ m · exact ⟨m, (ℱ.mono hnm _ hs).inter ht⟩ · exact ⟨n, hs.inter (ℱ.mono (not_le.1 hnm).le _ ht)⟩ - · simp only [measure_empty, WithTop.zero_lt_top, Measure.restrict_empty, integral_zero_measure, + · simp only [measure_empty, ENNReal.zero_lt_top, Measure.restrict_empty, integral_zero_measure, forall_true_left] · rintro t ⟨n, ht⟩ - exact this n _ ht diff --git a/Mathlib/Probability/Martingale/Upcrossing.lean b/Mathlib/Probability/Martingale/Upcrossing.lean index 9b3b3b33a3e46..525b5c80d4e9a 100644 --- a/Mathlib/Probability/Martingale/Upcrossing.lean +++ b/Mathlib/Probability/Martingale/Upcrossing.lean @@ -321,7 +321,7 @@ theorem upperCrossingTime_bound_eq (f : ℕ → Ω → ℝ) (N : ℕ) (ω : Ω) refine' strictMonoOn_Iic_of_lt_succ fun m hm => upperCrossingTime_lt_succ hab _ rw [Nat.lt_pred_iff] at hm convert Nat.find_min _ hm - convert StrictMonoOn.Iic_id_le hmono N (Nat.le_pred_of_lt hN') + convert StrictMonoOn.Iic_id_le hmono N (Nat.le_sub_one_of_lt hN') · rw [not_lt] at hN' exact upperCrossingTime_stabilize hN' (Nat.find_spec (exists_upperCrossingTime_eq f N ω hab)) #align measure_theory.upper_crossing_time_bound_eq MeasureTheory.upperCrossingTime_bound_eq @@ -699,7 +699,7 @@ theorem crossing_pos_eq (hab : a < b) : split_ifs with h₁ h₂ h₂ · simp_rw [← sub_le_iff_le_add, hf ω] · refine' False.elim (h₂ _) - simp_all only [Set.mem_Ici] + simp_all only [Set.mem_Ici, not_true_eq_false] · refine' False.elim (h₁ _) simp_all only [Set.mem_Ici] · rfl @@ -709,7 +709,7 @@ theorem crossing_pos_eq (hab : a < b) : split_ifs with h₁ h₂ h₂ · simp_rw [hf' ω] · refine' False.elim (h₂ _) - simp_all only [Set.mem_Iic] + simp_all only [Set.mem_Iic, not_true_eq_false] · refine' False.elim (h₁ _) simp_all only [Set.mem_Iic] · rfl diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 5e8ae6fddf8c6..57580ccf3cf08 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -465,7 +465,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : · rintro ⟨i, j⟩ hij simp only [mem_sigma, mem_range] at hij simp only [hij.1, hij.2, mem_sigma, mem_range, mem_filter, and_true_iff] - exact hij.2.trans_le (u_mono (Nat.le_pred_of_lt hij.1)) + exact hij.2.trans_le (u_mono (Nat.le_sub_one_of_lt hij.1)) · rintro ⟨i, j⟩ hij simp only [mem_sigma, mem_range, mem_filter] at hij simp only [hij.2.1, hij.2.2, mem_sigma, mem_range, and_self_iff] diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean index 09342ffafe3d9..2f0cda5a33e8f 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston -/ import Mathlib.Algebra.Homology.Opposite +import Mathlib.Algebra.Homology.ShortComplex.HomologicalComplex import Mathlib.RepresentationTheory.GroupCohomology.Resolution #align_import representation_theory.group_cohomology.basic from "leanprover-community/mathlib"@"cc5dd6244981976cc9da7afc4eee5682b037a013" @@ -30,6 +31,10 @@ This gives us for free a proof that our $d^n$ squares to zero. It also gives us $\mathrm{H}^n(G, A) \cong \mathrm{Ext}^n(k, A),$ where $\mathrm{Ext}$ is taken in the category `Rep k G`. +To talk about cohomology in low degree, please see the file +`RepresentationTheory.GroupCohomology.LowDegree`, which gives simpler expressions for `H⁰, H¹, H²` +than the definition `groupCohomology` in this file. + ## Main definitions * `groupCohomology.linearYonedaObjResolution A`: a complex whose objects are the representation @@ -38,7 +43,7 @@ $\mathrm{H}^n(G, A)$. * `groupCohomology.inhomogeneousCochains A`: a complex whose objects are $\mathrm{Fun}(G^n, A)$ and whose cohomology is the group cohomology $\mathrm{H}^n(G, A).$ * `groupCohomology.inhomogeneousCochainsIso A`: an isomorphism between the above two complexes. -* `group_cohomology A n`: this is $\mathrm{H}^n(G, A),$ defined as the $n$th cohomology of the +* `groupCohomology A n`: this is $\mathrm{H}^n(G, A),$ defined as the $n$th cohomology of the second complex, `inhomogeneousCochains A`. * `groupCohomologyIsoExt A n`: an isomorphism $\mathrm{H}^n(G, A) \cong \mathrm{Ext}^n(k, A)$ (where $\mathrm{Ext}$ is taken in the category `Rep k G`) induced by `inhomogeneousCochainsIso A`. @@ -199,6 +204,11 @@ noncomputable abbrev inhomogeneousCochains : CochainComplex (ModuleCat k) ℕ := exact map_zero _ #align group_cohomology.inhomogeneous_cochains groupCohomology.inhomogeneousCochains +@[simp] +theorem inhomogeneousCochains.d_def (n : ℕ) : + (inhomogeneousCochains A).d n (n + 1) = inhomogeneousCochains.d n A := + CochainComplex.of_d _ _ _ _ + /-- Given a `k`-linear `G`-representation `A`, the complex of inhomogeneous cochains is isomorphic to `Hom(P, A)`, where `P` is the standard resolution of `k` as a trivial `G`-representation. -/ def inhomogeneousCochainsIso : inhomogeneousCochains A ≅ linearYonedaObjResolution A := by @@ -225,7 +235,8 @@ def groupCohomology [Group G] (A : Rep k G) (n : ℕ) : ModuleCat k := `Extⁿ(k, A)` (taken in `Rep k G`), where `k` is a trivial `k`-linear `G`-representation. -/ def groupCohomologyIsoExt [Group G] (A : Rep k G) (n : ℕ) : groupCohomology A n ≅ ((Ext k (Rep k G) n).obj (Opposite.op <| Rep.trivial k G k)).obj A := - homologyObjIsoOfHomotopyEquiv (HomotopyEquiv.ofIso (inhomogeneousCochainsIso _)) _ ≪≫ - HomologicalComplex.homologyUnop _ _ ≪≫ (extIso k G A n).symm + ((inhomogeneousCochains A).homology'IsoHomology n).symm ≪≫ + homologyObjIsoOfHomotopyEquiv (HomotopyEquiv.ofIso (inhomogeneousCochainsIso _)) _ ≪≫ + HomologicalComplex.homology'Unop _ _ ≪≫ (extIso k G A n).symm set_option linter.uppercaseLean3 false in #align group_cohomology_iso_Ext groupCohomologyIsoExt diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean new file mode 100644 index 0000000000000..1e01758844926 --- /dev/null +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -0,0 +1,325 @@ +/- +Copyright (c) 2023 Amelia Livingston. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Amelia Livingston + +-/ +import Mathlib.RepresentationTheory.GroupCohomology.Basic +import Mathlib.RepresentationTheory.Invariants + +/-! +# The low-degree cohomology of a `k`-linear `G`-representation + +Let `k` be a commutative ring and `G` a group. This file gives simple expressions for +the group cohomology of a `k`-linear `G`-representation `A` in degrees 0, 1 and 2. + +In `RepresentationTheory.GroupCohomology.Basic`, we define the `n`th group cohomology of `A` to be +the cohomology of a complex `inhomogeneousCochains A`, whose objects are `(Fin n → G) → A`; this is +unnecessarily unwieldy in low degree. Moreover, cohomology of a complex is defined as an abstract +cokernel, whereas the definitions here are explicit quotients of cocycles by coboundaries. + +Later this file will contain an identification between the definition in +`RepresentationTheory.GroupCohomology.Basic`, `groupCohomology A n`, and the `Hn A` in this file, +for `n = 0, 1, 2`. + +## Main definitions + +* `groupCohomology.H0 A`: the invariants `Aᴳ` of the `G`-representation on `A`. +* `groupCohomology.H1 A`: 1-cocycles (i.e. `Z¹(G, A) := Ker(d¹ : Fun(G, A) → Fun(G², A)`) modulo +1-coboundaries (i.e. `B¹(G, A) := Im(d⁰: A → Fun(G, A))`). +* `groupCohomology.H2 A`: 2-cocycles (i.e. `Z²(G, A) := Ker(d² : Fun(G², A) → Fun(G³, A)`) modulo +2-coboundaries (i.e. `B²(G, A) := Im(d¹: Fun(G, A) → Fun(G², A))`). + +## TODO + +* Identify `Hn A` as defined in this file with `groupCohomology A n` for `n = 0, 1, 2`. +* Properties of `H1`, like the isomorphism `H1(G, A) ≃ Hom(G, A)` when the representation +is trivial +* The relationship between `H2` and group extensions +* The inflation-restriction exact sequence +* Nonabelian group cohomology + +-/ + +universe v u + +noncomputable section + +open CategoryTheory Limits Representation + +variable {k G : Type u} [CommRing k] [Group G] (A : Rep k G) + +namespace groupCohomology + +section Cochains + +/-- The 0th object in the complex of inhomogeneous cochains of `A : Rep k G` is isomorphic +to `A` as a `k`-module. -/ +def zeroCochainsLinearEquiv : (inhomogeneousCochains A).X 0 ≃ₗ[k] A := + LinearEquiv.funUnique (Fin 0 → G) k A + +/-- The 1st object in the complex of inhomogeneous cochains of `A : Rep k G` is isomorphic +to `Fun(G, A)` as a `k`-module. -/ +def oneCochainsLinearEquiv : (inhomogeneousCochains A).X 1 ≃ₗ[k] G → A := + LinearEquiv.funCongrLeft k A (Equiv.funUnique (Fin 1) G).symm + +/-- The 2nd object in the complex of inhomogeneous cochains of `A : Rep k G` is isomorphic +to `Fun(G², A)` as a `k`-module. -/ +def twoCochainsLinearEquiv : (inhomogeneousCochains A).X 2 ≃ₗ[k] G × G → A := + LinearEquiv.funCongrLeft k A <| (piFinTwoEquiv fun _ => G).symm + +/-- The 3rd object in the complex of inhomogeneous cochains of `A : Rep k G` is isomorphic +to `Fun(G³, A)` as a `k`-module. -/ +def threeCochainsLinearEquiv : (inhomogeneousCochains A).X 3 ≃ₗ[k] G × G × G → A := + LinearEquiv.funCongrLeft k A <| ((Equiv.piFinSucc 2 G).trans + ((Equiv.refl G).prodCongr (piFinTwoEquiv fun _ => G))).symm + +end Cochains +section Differentials + +/-- The 0th differential in the complex of inhomogeneous cochains of `A : Rep k G`, as a +`k`-linear map `A → Fun(G, A)`. It sends `(a, g) ↦ ρ_A(g)(a) - a.` -/ +@[simps] +def dZero : A →ₗ[k] G → A where + toFun m g := A.ρ g m - m + map_add' x y := funext fun g => by simp only [map_add, add_sub_add_comm]; rfl + map_smul' r x := funext fun g => by dsimp; rw [map_smul, smul_sub] + +theorem dZero_ker_eq_invariants : LinearMap.ker (dZero A) = invariants A.ρ := by + ext x + simp only [LinearMap.mem_ker, mem_invariants, ← @sub_eq_zero _ _ _ x, Function.funext_iff] + rfl + +/-- The 1st differential in the complex of inhomogeneous cochains of `A : Rep k G`, as a +`k`-linear map `Fun(G, A) → Fun(G × G, A)`. It sends +`(f, (g₁, g₂)) ↦ ρ_A(g₁)(f(g₂)) - f(g₁g₂) + f(g₁).` -/ +@[simps] +def dOne : (G → A) →ₗ[k] G × G → A where + toFun f g := A.ρ g.1 (f g.2) - f (g.1 * g.2) + f g.1 + map_add' x y := funext fun g => by dsimp; rw [map_add, add_add_add_comm, add_sub_add_comm] + map_smul' r x := funext fun g => by dsimp; rw [map_smul, smul_add, smul_sub] + +/-- The 2nd differential in the complex of inhomogeneous cochains of `A : Rep k G`, as a +`k`-linear map `Fun(G × G, A) → Fun(G × G × G, A)`. It sends +`(f, (g₁, g₂, g₃)) ↦ ρ_A(g₁)(f(g₂, g₃)) - f(g₁g₂, g₃) + f(g₁, g₂g₃) - f(g₁, g₂).` -/ +@[simps] +def dTwo : (G × G → A) →ₗ[k] G × G × G → A where + toFun f g := + A.ρ g.1 (f (g.2.1, g.2.2)) - f (g.1 * g.2.1, g.2.2) + f (g.1, g.2.1 * g.2.2) - f (g.1, g.2.1) + map_add' x y := + funext fun g => by + dsimp + rw [map_add, add_sub_add_comm (A.ρ _ _), add_sub_assoc, add_sub_add_comm, add_add_add_comm, + add_sub_assoc, add_sub_assoc] + map_smul' r x := funext fun g => by dsimp; simp only [map_smul, smul_add, smul_sub] + +/-- Let `C(G, A)` denote the complex of inhomogeneous cochains of `A : Rep k G`. This lemma +says `dZero` gives a simpler expression for the 0th differential: that is, the following +square commutes: +``` + C⁰(G, A) ---d⁰---> C¹(G, A) + | | + | | + | | + v v + A ---- dZero ---> Fun(G, A) +``` +where the vertical arrows are `zeroCochainsLinearEquiv` and `oneCochainsLinearEquiv` respectively. +-/ +theorem dZero_comp_eq : dZero A ∘ₗ (zeroCochainsLinearEquiv A) = + oneCochainsLinearEquiv A ∘ₗ (inhomogeneousCochains A).d 0 1 := by + ext x y + show A.ρ y (x default) - x default = _ + ({0} : Finset _).sum _ + simp_rw [Fin.coe_fin_one, zero_add, pow_one, neg_smul, one_smul, + Finset.sum_singleton, sub_eq_add_neg] + rcongr i <;> exact Fin.elim0 i + +/-- Let `C(G, A)` denote the complex of inhomogeneous cochains of `A : Rep k G`. This lemma +says `dOne` gives a simpler expression for the 1st differential: that is, the following +square commutes: +``` + C¹(G, A) ---d¹-----> C²(G, A) + | | + | | + | | + v v + Fun(G, A) -dOne-> Fun(G × G, A) +``` +where the vertical arrows are `oneCochainsLinearEquiv` and `twoCochainsLinearEquiv` respectively. +-/ +theorem dOne_comp_eq : dOne A ∘ₗ oneCochainsLinearEquiv A = + twoCochainsLinearEquiv A ∘ₗ (inhomogeneousCochains A).d 1 2 := by + ext x y + show A.ρ y.1 (x _) - x _ + x _ = _ + _ + rw [Fin.sum_univ_two] + simp only [Fin.val_zero, zero_add, pow_one, neg_smul, one_smul, Fin.val_one, + Nat.one_add, neg_one_sq, sub_eq_add_neg, add_assoc] + rcongr i <;> rw [Subsingleton.elim i 0] <;> rfl + +/-- Let `C(G, A)` denote the complex of inhomogeneous cochains of `A : Rep k G`. This lemma +says `dTwo` gives a simpler expression for the 2nd differential: that is, the following +square commutes: +``` + C²(G, A) -------d²-----> C³(G, A) + | | + | | + | | + v v + Fun(G × G, A) --dTwo--> Fun(G × G × G, A) +``` +where the vertical arrows are `twoCochainsLinearEquiv` and `threeCochainsLinearEquiv` respectively. +-/ +theorem dTwo_comp_eq : + dTwo A ∘ₗ twoCochainsLinearEquiv A = + threeCochainsLinearEquiv A ∘ₗ (inhomogeneousCochains A).d 2 3 := by + ext x y + show A.ρ y.1 (x _) - x _ + x _ - x _ = _ + _ + dsimp + rw [Fin.sum_univ_three] + simp only [sub_eq_add_neg, add_assoc, Fin.val_zero, zero_add, pow_one, neg_smul, + one_smul, Fin.val_one, Fin.val_two, pow_succ (-1 : k) 2, neg_sq, Nat.one_add, one_pow, mul_one] + rcongr i <;> fin_cases i <;> rfl + +theorem dOne_comp_dZero : dOne A ∘ₗ dZero A = 0 := by + ext x g + simp only [LinearMap.coe_comp, Function.comp_apply, dOne_apply A, dZero_apply A, map_sub, + map_mul, LinearMap.mul_apply, sub_sub_sub_cancel_left, sub_add_sub_cancel, sub_self] + rfl + +theorem dTwo_comp_dOne : dTwo A ∘ₗ dOne A = 0 := by + show ModuleCat.ofHom (dOne A) ≫ ModuleCat.ofHom (dTwo A) = _ + have h1 : _ ≫ ModuleCat.ofHom (dOne A) = _ ≫ _ := congr_arg ModuleCat.ofHom (dOne_comp_eq A) + have h2 : _ ≫ ModuleCat.ofHom (dTwo A) = _ ≫ _ := congr_arg ModuleCat.ofHom (dTwo_comp_eq A) + simp only [←LinearEquiv.toModuleIso_hom] at h1 h2 + simp only [(Iso.eq_inv_comp _).2 h2, (Iso.eq_inv_comp _).2 h1, + Category.assoc, Iso.hom_inv_id_assoc, HomologicalComplex.d_comp_d_assoc, zero_comp, comp_zero] + +end Differentials + +section Cocycles + +/-- The 1-cocycles `Z¹(G, A)` of `A : Rep k G`, defined as the kernel of the map +`Fun(G, A) → Fun(G × G, A)` sending `(f, (g₁, g₂)) ↦ ρ_A(g₁)(f(g₂)) - f(g₁g₂) + f(g₁).` -/ +def oneCocycles : Submodule k (G → A) := LinearMap.ker (dOne A) + +/-- The 2-cocycles `Z²(G, A)` of `A : Rep k G`, defined as the kernel of the map +`Fun(G × G, A) → Fun(G × G × G, A)` sending +`(f, (g₁, g₂, g₃)) ↦ ρ_A(g₁)(f(g₂, g₃)) - f(g₁g₂, g₃) + f(g₁, g₂g₃) - f(g₁, g₂).` -/ +def twoCocycles : Submodule k (G × G → A) := LinearMap.ker (dTwo A) + +variable {A} + +theorem mem_oneCocycles_def (f : G → A) : + f ∈ oneCocycles A ↔ ∀ g : G × G, A.ρ g.1 (f g.2) - f (g.1 * g.2) + f g.1 = 0 := + LinearMap.mem_ker.trans Function.funext_iff + +theorem mem_oneCocycles_iff (f : G → A) : + f ∈ oneCocycles A ↔ ∀ g : G × G, f (g.1 * g.2) = A.ρ g.1 (f g.2) + f g.1 := by + simp_rw [mem_oneCocycles_def, sub_add_eq_add_sub, sub_eq_zero, eq_comm] + +theorem oneCocycles_map_one (f : oneCocycles A) : f.1 1 = 0 := by + have := (mem_oneCocycles_def f.1).1 f.2 (1, 1) + simpa only [map_one, LinearMap.one_apply, mul_one, sub_self, zero_add] using this + +theorem oneCocycles_map_inv (f : oneCocycles A) (g : G) : + A.ρ g (f.1 g⁻¹) = - f.1 g := by + rw [← add_eq_zero_iff_eq_neg, ← oneCocycles_map_one f, ← mul_inv_self g, + (mem_oneCocycles_iff f.1).1 f.2 (g, g⁻¹)] + +theorem mem_twoCocycles_def (f : G × G → A) : + f ∈ twoCocycles A ↔ ∀ g : G × G × G, + A.ρ g.1 (f (g.2.1, g.2.2)) - f (g.1 * g.2.1, g.2.2) + + f (g.1, g.2.1 * g.2.2) - f (g.1, g.2.1) = 0 := + LinearMap.mem_ker.trans Function.funext_iff + +theorem mem_twoCocycles_iff (f : G × G → A) : + f ∈ twoCocycles A ↔ ∀ g : G × G × G, + f (g.1 * g.2.1, g.2.2) + f (g.1, g.2.1) = + A.ρ g.1 (f (g.2.1, g.2.2)) + f (g.1, g.2.1 * g.2.2) := by + simp_rw [mem_twoCocycles_def, sub_eq_zero, sub_add_eq_add_sub, sub_eq_iff_eq_add, eq_comm, + add_comm (f (Prod.fst _, _))] + +theorem twoCocycles_map_one_fst (f : twoCocycles A) (g : G) : + f.1 (1, g) = f.1 (1, 1) := by + have := ((mem_twoCocycles_iff f.1).1 f.2 (1, (1, g))).symm + simpa only [map_one, LinearMap.one_apply, one_mul, add_right_inj, this] + +theorem twoCocycles_map_one_snd (f : twoCocycles A) (g : G) : + f.1 (g, 1) = A.ρ g (f.1 (1, 1)) := by + have := (mem_twoCocycles_iff f.1).1 f.2 (g, (1, 1)) + simpa only [mul_one, add_left_inj, this] + +lemma twoCocycles_ρ_map_inv_sub_map_inv (f : twoCocycles A) (g : G) : + A.ρ g (f.1 (g⁻¹, g)) - f.1 (g, g⁻¹) + = f.1 (1, 1) - f.1 (g, 1) := by + have := (mem_twoCocycles_iff f.1).1 f.2 (g, g⁻¹, g) + simp only [mul_right_inv, mul_left_inv, twoCocycles_map_one_fst _ g] + at this + exact sub_eq_sub_iff_add_eq_add.2 this.symm + +end Cocycles +section Coboundaries + +/-- The 1-coboundaries `B¹(G, A)` of `A : Rep k G`, defined as the image of the map +`A → Fun(G, A)` sending `(a, g) ↦ ρ_A(g)(a) - a.` -/ +def oneCoboundaries : Submodule k (oneCocycles A) := + LinearMap.range ((dZero A).codRestrict (oneCocycles A) fun c => + LinearMap.ext_iff.1 (dOne_comp_dZero A) c) + +/-- The 2-coboundaries `B²(G, A)` of `A : Rep k G`, defined as the image of the map +`Fun(G, A) → Fun(G × G, A)` sending `(f, (g₁, g₂)) ↦ ρ_A(g₁)(f(g₂)) - f(g₁g₂) + f(g₁).` -/ +def twoCoboundaries : Submodule k (twoCocycles A) := + LinearMap.range ((dOne A).codRestrict (twoCocycles A) fun c => + LinearMap.ext_iff.1 (dTwo_comp_dOne.{u} A) c) + +variable {A} + +theorem mem_oneCoboundaries_of_dZero_apply (x : A) : + ⟨dZero A x, LinearMap.ext_iff.1 (dOne_comp_dZero A) x⟩ ∈ oneCoboundaries A := +LinearMap.mem_range_self _ _ + +theorem mem_oneCoboundaries_of_mem_range (f : G → A) (h : f ∈ LinearMap.range (dZero A)) : + ⟨f, LinearMap.range_le_ker_iff.2 (dOne_comp_dZero A) h⟩ ∈ oneCoboundaries A := by + rcases h with ⟨x, rfl⟩; exact ⟨x, rfl⟩ + +theorem mem_range_of_mem_oneCoboundaries (f : oneCocycles A) (h : f ∈ oneCoboundaries A) : + f.1 ∈ LinearMap.range (dZero A) := by + rcases h with ⟨x, rfl⟩; exact ⟨x, rfl⟩ + +theorem mem_twoCoboundaries_of_dOne_apply (x : G → A) : + ⟨dOne A x, LinearMap.ext_iff.1 (dTwo_comp_dOne A) x⟩ ∈ twoCoboundaries A := +LinearMap.mem_range_self _ _ + +theorem mem_twoCoboundaries_of_mem_range (f : G × G → A) (h : f ∈ LinearMap.range (dOne A)) : + ⟨f, LinearMap.range_le_ker_iff.2 (dTwo_comp_dOne A) h⟩ ∈ twoCoboundaries A := by + rcases h with ⟨x, rfl⟩; exact ⟨x, rfl⟩ + +theorem mem_range_of_mem_twoCoboundaries (f : twoCocycles A) (h : f ∈ twoCoboundaries A) : + (twoCocycles A).subtype f ∈ LinearMap.range (dOne A) := by + rcases h with ⟨x, rfl⟩; exact ⟨x, rfl⟩ + +end Coboundaries +section Cohomology + +/-- We define the 0th group cohomology of a `k`-linear `G`-representation `A`, `H⁰(G, A)`, to be +the invariants of the representation, `Aᴳ`. -/ +abbrev H0 := A.ρ.invariants + +/-- We define the 1st group cohomology of a `k`-linear `G`-representation `A`, `H¹(G, A)`, to be +1-cocycles (i.e. `Z¹(G, A) := Ker(d¹ : Fun(G, A) → Fun(G², A)`) modulo 1-coboundaries +(i.e. `B¹(G, A) := Im(d⁰: A → Fun(G, A))`). -/ +abbrev H1 := oneCocycles A ⧸ oneCoboundaries A + +/-- The quotient map `Z¹(G, A) → H¹(G, A).` -/ +def H1_π : oneCocycles A →ₗ[k] H1 A := (oneCoboundaries A).mkQ + +/-- We define the 2nd group cohomology of a `k`-linear `G`-representation `A`, `H²(G, A)`, to be +2-cocycles (i.e. `Z²(G, A) := Ker(d² : Fun(G², A) → Fun(G³, A)`) modulo 2-coboundaries +(i.e. `B²(G, A) := Im(d¹: Fun(G, A) → Fun(G², A))`). -/ +abbrev H2 := twoCocycles A ⧸ twoCoboundaries A + +/-- The quotient map `Z²(G, A) → H²(G, A).` -/ +def H2_π : twoCocycles A →ₗ[k] H2 A := (twoCoboundaries A).mkQ + +end Cohomology +end groupCohomology diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index 11cdca5ecc256..ea5a1f83710ee 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -675,16 +675,16 @@ theorem εToSingle₀_comp_eq : exact (forget₂ToModuleCatHomotopyEquiv_f_0_eq k G).symm #align group_cohomology.resolution.ε_to_single₀_comp_eq groupCohomology.resolution.εToSingle₀_comp_eq -theorem quasiIsoOfForget₂εToSingle₀ : - QuasiIso (((forget₂ _ (ModuleCat.{u} k)).mapHomologicalComplex _).map (εToSingle₀ k G)) := by - have h : QuasiIso (forget₂ToModuleCatHomotopyEquiv k G).hom := HomotopyEquiv.toQuasiIso _ +theorem quasiIso'OfForget₂εToSingle₀ : + QuasiIso' (((forget₂ _ (ModuleCat.{u} k)).mapHomologicalComplex _).map (εToSingle₀ k G)) := by + have h : QuasiIso' (forget₂ToModuleCatHomotopyEquiv k G).hom := HomotopyEquiv.toQuasiIso' _ rw [← εToSingle₀_comp_eq k G] at h haveI := h - exact quasiIso_of_comp_right _ ((ChainComplex.single₀MapHomologicalComplex _).hom.app _) -#align group_cohomology.resolution.quasi_iso_of_forget₂_ε_to_single₀ groupCohomology.resolution.quasiIsoOfForget₂εToSingle₀ + exact quasiIso'_of_comp_right _ ((ChainComplex.single₀MapHomologicalComplex _).hom.app _) +#align group_cohomology.resolution.quasi_iso_of_forget₂_ε_to_single₀ groupCohomology.resolution.quasiIso'OfForget₂εToSingle₀ -instance : QuasiIso (εToSingle₀ k G) := - (forget₂ _ (ModuleCat.{u} k)).quasiIso_of_map_quasiIso _ (quasiIsoOfForget₂εToSingle₀ k G) +instance : QuasiIso' (εToSingle₀ k G) := + (forget₂ _ (ModuleCat.{u} k)).quasiIso'_of_map_quasiIso' _ (quasiIso'OfForget₂εToSingle₀ k G) end Exactness @@ -710,7 +710,7 @@ standard resolution of `k` called `groupCohomology.resolution k G`. -/ def groupCohomology.extIso (V : Rep k G) (n : ℕ) : ((Ext k (Rep k G) n).obj (Opposite.op <| Rep.trivial k G k)).obj V ≅ (((((linearYoneda k (Rep k G)).obj V).rightOp.mapHomologicalComplex _).obj - (groupCohomology.resolution k G)).homology + (groupCohomology.resolution k G)).homology' n).unop := (((linearYoneda k (Rep k G)).obj V).rightOp.leftDerivedObjIso n (groupCohomology.projectiveResolution k G)).unop.symm set_option linter.uppercaseLean3 false in diff --git a/Mathlib/RingTheory/Adjoin/Field.lean b/Mathlib/RingTheory/Adjoin/Field.lean index f69f10c6e67d1..1f829408485a2 100644 --- a/Mathlib/RingTheory/Adjoin/Field.lean +++ b/Mathlib/RingTheory/Adjoin/Field.lean @@ -41,16 +41,23 @@ def AlgEquiv.adjoinSingletonEquivAdjoinRootMinpoly {R : Type*} [CommRing R] [Alg rwa [Minpoly.toAdjoin_apply', liftHom_mk, ← Subalgebra.coe_eq_zero, aeval_subalgebra_coe] at hP₁ #align alg_equiv.adjoin_singleton_equiv_adjoin_root_minpoly AlgEquiv.adjoinSingletonEquivAdjoinRootMinpoly +/-- Produce an algebra homomorphism `Adjoin R {x} →ₐ[R] T` sending `x` to +a root of `x`'s minimal polynomial in `T`. -/ +noncomputable def Algebra.adjoin.liftSingleton {S T : Type*} + [CommRing S] [CommRing T] [Algebra F S] [Algebra F T] + (x : S) (y : T) (h : aeval y (minpoly F x) = 0) : + Algebra.adjoin F {x} →ₐ[F] T := + (AdjoinRoot.liftHom _ y h).comp (AlgEquiv.adjoinSingletonEquivAdjoinRootMinpoly F x).toAlgHom + open Finset -set_option synthInstance.maxHeartbeats 40000 in /-- If `K` and `L` are field extensions of `F` and we have `s : Finset K` such that the minimal polynomial of each `x ∈ s` splits in `L` then `Algebra.adjoin F s` embeds in `L`. -/ -theorem lift_of_splits {F K L : Type*} [Field F] [Field K] [Field L] [Algebra F K] [Algebra F L] - (s : Finset K) : (∀ x ∈ s, IsIntegral F x ∧ Polynomial.Splits (algebraMap F L) (minpoly F x)) → - Nonempty (Algebra.adjoin F (↑s : Set K) →ₐ[F] L) := by +theorem Polynomial.lift_of_splits {F K L : Type*} [Field F] [Field K] [Field L] [Algebra F K] + [Algebra F L] (s : Finset K) : (∀ x ∈ s, IsIntegral F x ∧ + Splits (algebraMap F L) (minpoly F x)) → Nonempty (Algebra.adjoin F (s : Set K) →ₐ[F] L) := by classical - refine' Finset.induction_on s (fun _ => _) fun a s _ ih H => _ + refine Finset.induction_on s (fun _ ↦ ?_) fun a s _ ih H ↦ ?_ · rw [coe_empty, Algebra.adjoin_empty] exact ⟨(Algebra.ofId F L).comp (Algebra.botEquiv F K)⟩ rw [forall_mem_insert] at H @@ -58,26 +65,20 @@ theorem lift_of_splits {F K L : Type*} [Field F] [Field K] [Field L] [Algebra F cases' ih H3 with f choose H3 _ using H3 rw [coe_insert, Set.insert_eq, Set.union_comm, Algebra.adjoin_union_eq_adjoin_adjoin] - letI := (f : Algebra.adjoin F (↑s : Set K) →+* L).toAlgebra - haveI : FiniteDimensional F (Algebra.adjoin F (↑s : Set K)) := - ((Submodule.fg_iff_finiteDimensional _).1 - (FG_adjoin_of_finite s.finite_toSet H3)).of_subalgebra_toSubmodule - letI := fieldOfFiniteDimensional F (Algebra.adjoin F (↑s : Set K)) - have H5 : IsIntegral (Algebra.adjoin F (s : Set K)) a := isIntegral_of_isScalarTower H1 - have H6 : (minpoly (Algebra.adjoin F (s : Set K)) a).Splits - (algebraMap (Algebra.adjoin F (s : Set K)) L) := by - have : Polynomial.map (algebraMap F (Algebra.adjoin F (s : Set K))) (minpoly F a) ≠ 0 := - Polynomial.map_ne_zero <| minpoly.ne_zero H1 - refine' Polynomial.splits_of_splits_of_dvd _ this - ((Polynomial.splits_map_iff _ _).2 _) (minpoly.dvd _ _ _) + set Ks := Algebra.adjoin F (s : Set K) + haveI : FiniteDimensional F Ks := ((Submodule.fg_iff_finiteDimensional _).1 + (FG_adjoin_of_finite s.finite_toSet H3)).of_subalgebra_toSubmodule + letI := fieldOfFiniteDimensional F Ks + letI := (f : Ks →+* L).toAlgebra + have H5 : IsIntegral Ks a := isIntegral_of_isScalarTower H1 + have H6 : (minpoly Ks a).Splits (algebraMap Ks L) := by + refine splits_of_splits_of_dvd _ ((minpoly.monic H1).map (algebraMap F Ks)).ne_zero + ((splits_map_iff _ _).2 ?_) (minpoly.dvd _ _ ?_) · rw [← IsScalarTower.algebraMap_eq] exact H2 · rw [Polynomial.aeval_map_algebraMap, minpoly.aeval] - obtain ⟨y, hy⟩ := Polynomial.exists_root_of_splits _ H6 (ne_of_lt (minpoly.degree_pos H5)).symm - refine' ⟨Subalgebra.ofRestrictScalars F _ _⟩ - refine' (AdjoinRoot.liftHom (minpoly (Algebra.adjoin F (↑s : Set K)) a) y hy).comp _ - exact (AlgEquiv.adjoinSingletonEquivAdjoinRootMinpoly - (Algebra.adjoin F (↑s : Set K)) a).toAlgHom -#align lift_of_splits lift_of_splits + obtain ⟨y, hy⟩ := Polynomial.exists_root_of_splits _ H6 (minpoly.degree_pos H5).ne' + exact ⟨Subalgebra.ofRestrictScalars F _ <| Algebra.adjoin.liftSingleton Ks a y hy⟩ +#align lift_of_splits Polynomial.lift_of_splits end Embeddings diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index 018a122326783..327729e48459d 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -289,16 +289,30 @@ variable [Algebra K L] theorem Algebra.IsAlgebraic.algHom_bijective (ha : Algebra.IsAlgebraic K L) (f : L →ₐ[K] L) : Function.Bijective f := by - refine' ⟨f.toRingHom.injective, fun b => _⟩ + refine' ⟨f.injective, fun b ↦ _⟩ obtain ⟨p, hp, he⟩ := ha b - let f' : p.rootSet L → p.rootSet L := (rootSet_maps_to' (fun x => x) f).restrict f _ _ - have : Function.Surjective f' := - Finite.injective_iff_surjective.1 fun _ _ h => - Subtype.eq <| f.toRingHom.injective <| Subtype.ext_iff.1 h + let f' : p.rootSet L → p.rootSet L := (rootSet_maps_to' (fun x ↦ x) f).restrict f _ _ + have : f'.Surjective := Finite.injective_iff_surjective.1 + fun _ _ h ↦ Subtype.eq <| f.injective <| Subtype.ext_iff.1 h obtain ⟨a, ha⟩ := this ⟨b, mem_rootSet.2 ⟨hp, he⟩⟩ exact ⟨a, Subtype.ext_iff.1 ha⟩ #align algebra.is_algebraic.alg_hom_bijective Algebra.IsAlgebraic.algHom_bijective +theorem Algebra.IsAlgebraic.algHom_bijective₂ [Field R] [Algebra K R] + (ha : Algebra.IsAlgebraic K L) (f : L →ₐ[K] R) (g : R →ₐ[K] L) : + Function.Bijective f ∧ Function.Bijective g := + (g.injective.bijective₂_of_surjective f.injective (ha.algHom_bijective <| g.comp f).2).symm + +theorem Algebra.IsAlgebraic.bijective_of_isScalarTower (ha : Algebra.IsAlgebraic K L) + [Field R] [Algebra K R] [Algebra L R] [IsScalarTower K L R] (f : R →ₐ[K] L) : + Function.Bijective f := + (ha.algHom_bijective₂ (IsScalarTower.toAlgHom K L R) f).2 + +theorem Algebra.IsAlgebraic.bijective_of_isScalarTower' [Field R] [Algebra K R] + (ha : Algebra.IsAlgebraic K R) [Algebra L R] [IsScalarTower K L R] (f : R →ₐ[K] L) : + Function.Bijective f := + (ha.algHom_bijective₂ f (IsScalarTower.toAlgHom K L R)).1 + theorem AlgHom.bijective [FiniteDimensional K L] (ϕ : L →ₐ[K] L) : Function.Bijective ϕ := (Algebra.isAlgebraic_of_finite K L).algHom_bijective ϕ #align alg_hom.bijective AlgHom.bijective diff --git a/Mathlib/RingTheory/Congruence.lean b/Mathlib/RingTheory/Congruence.lean index 031849cef08f4..75742f4f91c78 100644 --- a/Mathlib/RingTheory/Congruence.lean +++ b/Mathlib/RingTheory/Congruence.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Algebra.GroupRingAction.Basic -import Mathlib.Algebra.Hom.Ring.Defs +import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Algebra.Ring.InjSurj import Mathlib.GroupTheory.Congruence diff --git a/Mathlib/RingTheory/Coprime/Basic.lean b/Mathlib/RingTheory/Coprime/Basic.lean index 2965bce57df2e..08bb6c238c26c 100644 --- a/Mathlib/RingTheory/Coprime/Basic.lean +++ b/Mathlib/RingTheory/Coprime/Basic.lean @@ -3,11 +3,11 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Ken Lee, Chris Hughes -/ -import Mathlib.Tactic.Ring -import Mathlib.GroupTheory.GroupAction.Units -import Mathlib.Algebra.Ring.Divisibility.Basic -import Mathlib.Algebra.Hom.Ring.Defs import Mathlib.Algebra.GroupPower.Ring +import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Ring.Hom.Defs +import Mathlib.GroupTheory.GroupAction.Units +import Mathlib.Tactic.Ring #align_import ring_theory.coprime.basic from "leanprover-community/mathlib"@"a95b16cbade0f938fc24abd05412bde1e84bab9b" diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 4026ecc90a541..6c7d0f1cf3af3 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -297,8 +297,8 @@ theorem integrallyClosed : IsIntegrallyClosed A := by -- It suffices to show that for integral `x`, -- `A[x]` (which is a fractional ideal) is in fact equal to `A`. refine ⟨fun {x hx} => ?_⟩ - rw [← Set.mem_range, ← Algebra.mem_bot, ← Subalgebra.mem_toSubmodule, Algebra.toSubmodule_bot, ← - coe_spanSingleton A⁰ (1 : FractionRing A), spanSingleton_one, ← + rw [← Set.mem_range, ← Algebra.mem_bot, ← Subalgebra.mem_toSubmodule, Algebra.toSubmodule_bot, + Submodule.one_eq_span, ←coe_spanSingleton A⁰ (1 : FractionRing A), spanSingleton_one, ← FractionalIdeal.adjoinIntegral_eq_one_of_isUnit x hx (h.isUnit _)] · exact mem_adjoinIntegral_self A⁰ x hx · exact fun h => one_ne_zero (eq_zero_iff.mp h 1 (Algebra.adjoin A {x}).one_mem) @@ -366,9 +366,9 @@ theorem exists_multiset_prod_cons_le_and_prod_not_le [IsDedekindDomain A] (hNF : -- Let `Z` be a minimal set of prime ideals such that their product is contained in `J`. obtain ⟨Z₀, hZ₀⟩ := PrimeSpectrum.exists_primeSpectrum_prod_le_and_ne_bot_of_domain hNF hI0 obtain ⟨Z, ⟨hZI, hprodZ⟩, h_eraseZ⟩ := - Multiset.wellFounded_lt.has_min - (fun Z => (Z.map PrimeSpectrum.asIdeal).prod ≤ I ∧ (Z.map PrimeSpectrum.asIdeal).prod ≠ ⊥) - ⟨Z₀, hZ₀⟩ + wellFounded_lt.has_min + {Z | (Z.map PrimeSpectrum.asIdeal).prod ≤ I ∧ (Z.map PrimeSpectrum.asIdeal).prod ≠ ⊥} + ⟨Z₀, hZ₀.1, hZ₀.2⟩ have hZM : Multiset.prod (Z.map PrimeSpectrum.asIdeal) ≤ M := le_trans hZI hIM have hZ0 : Z ≠ 0 := by rintro rfl; simp [hM.ne_top] at hZM obtain ⟨_, hPZ', hPM⟩ := (hM.isPrime.multiset_prod_le (mt Multiset.map_eq_zero.mp hZ0)).mp hZM diff --git a/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean b/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean index f66d5a11d1b29..30605185ddced 100644 --- a/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean +++ b/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ -import Mathlib.Algebra.Hom.Equiv.TypeTags +import Mathlib.Algebra.Group.Equiv.TypeTags import Mathlib.Data.ZMod.Quotient import Mathlib.RingTheory.DedekindDomain.AdicValuation import Mathlib.RingTheory.Norm diff --git a/Mathlib/RingTheory/Derivation/Basic.lean b/Mathlib/RingTheory/Derivation/Basic.lean index 22c58a4f5f579..e6873ac4a86f2 100644 --- a/Mathlib/RingTheory/Derivation/Basic.lean +++ b/Mathlib/RingTheory/Derivation/Basic.lean @@ -40,8 +40,9 @@ equality. We also require that `D 1 = 0`. See `Derivation.mk'` for a constructor assumption from the Leibniz rule when `M` is cancellative. TODO: update this when bimodules are defined. -/ -structure Derivation (R : Type*) (A : Type*) [CommSemiring R] [CommSemiring A] [Algebra R A] - (M : Type*) [AddCommMonoid M] [Module A M] [Module R M] extends A →ₗ[R] M where +structure Derivation (R : Type*) (A : Type*) (M : Type*) + [CommSemiring R] [CommSemiring A] [AddCommMonoid M] [Algebra R A] [Module A M] [Module R M] + extends A →ₗ[R] M where protected map_one_eq_zero' : toLinearMap 1 = 0 protected leibniz' (a b : A) : toLinearMap (a * b) = a • toLinearMap b + b • toLinearMap a #align derivation Derivation @@ -53,11 +54,11 @@ namespace Derivation section -variable {R : Type*} [CommSemiring R] +variable {R : Type*} {A : Type*} {B : Type*} {M : Type*} +variable [CommSemiring R] [CommSemiring A] [CommSemiring B] [AddCommMonoid M] +variable [Algebra R A] [Algebra R B] +variable [Module A M] [Module B M] [Module R M] -variable {A : Type*} [CommSemiring A] [Algebra R A] - -variable {M : Type*} [AddCommMonoid M] [Module A M] [Module R M] variable (D : Derivation R A M) {D1 D2 : Derivation R A M} (r : R) (a b : A) @@ -77,6 +78,11 @@ theorem toFun_eq_coe : D.toFun = ⇑D := rfl #align derivation.to_fun_eq_coe Derivation.toFun_eq_coe +/-- See Note [custom simps projection] -/ +def Simps.apply (D : Derivation R A M) : A → M := D + +initialize_simps_projections Derivation (toFun → apply) + attribute [coe] toLinearMap instance hasCoeToLinearMap : Coe (Derivation R A M) (A →ₗ[R] M) := @@ -325,6 +331,17 @@ def _root_.LinearEquiv.compDer : Derivation R A M ≃ₗ[R] Derivation R A N := end PushForward +variable (A) in +/-- For a tower `R → A → B` and an `R`-derivation `B → M`, we may compose with `A → B` to obtain an +`R`-derivation `A → M`. -/ +@[simps!] +def compAlgebraMap [Algebra A B] [IsScalarTower R A B] [IsScalarTower A B M] + (d : Derivation R B M) : Derivation R A M where + map_one_eq_zero' := by simp + leibniz' a b := by simp + toLinearMap := d.toLinearMap.comp (IsScalarTower.toAlgHom R A B).toLinearMap +#align derivation.comp_algebra_map Derivation.compAlgebraMap + section RestrictScalars variable {S : Type*} [CommSemiring S] @@ -461,4 +478,5 @@ end end + end Derivation diff --git a/Mathlib/RingTheory/Filtration.lean b/Mathlib/RingTheory/Filtration.lean index ac2e16ead92b9..eae7939d7c5d7 100644 --- a/Mathlib/RingTheory/Filtration.lean +++ b/Mathlib/RingTheory/Filtration.lean @@ -269,7 +269,7 @@ protected def submodule : Submodule (reesAlgebra I) (PolynomialModule R M) where rw [Subalgebra.smul_def, PolynomialModule.smul_apply] apply Submodule.sum_mem rintro ⟨j, k⟩ e - rw [Finset.Nat.mem_antidiagonal] at e + rw [Finset.mem_antidiagonal] at e subst e exact F.pow_smul_le j k (Submodule.smul_mem_smul (r.2 j) (hf k)) #align ideal.filtration.submodule Ideal.Filtration.submodule diff --git a/Mathlib/RingTheory/FiniteType.lean b/Mathlib/RingTheory/FiniteType.lean index ef7ccbaac2026..3389f8d56c282 100644 --- a/Mathlib/RingTheory/FiniteType.lean +++ b/Mathlib/RingTheory/FiniteType.lean @@ -8,6 +8,7 @@ import Mathlib.GroupTheory.Finiteness import Mathlib.RingTheory.Adjoin.Tower import Mathlib.RingTheory.Finiteness import Mathlib.RingTheory.Noetherian +import Mathlib.Data.Polynomial.Module #align_import ring_theory.finite_type from "leanprover-community/mathlib"@"bb168510ef455e9280a152e7f31673cabd3d7496" @@ -709,70 +710,33 @@ end MonoidAlgebra section Vasconcelos -variable {R : Type*} [CommRing R] {M : Type*} [AddCommGroup M] [Module R M] (f : M →ₗ[R] M) - -noncomputable section - -/-- The structure of a module `M` over a ring `R` as a module over `R[X]` when given a -choice of how `X` acts by choosing a linear map `f : M →ₗ[R] M` -/ -def modulePolynomialOfEndo : Module R[X] M := - Module.compHom M (Polynomial.aeval f).toRingHom -#align module_polynomial_of_endo modulePolynomialOfEndo - -theorem modulePolynomialOfEndo_smul_def (n : R[X]) (a : M) : - @HSMul.hSMul _ _ _ (by letI := modulePolynomialOfEndo f; infer_instance) n a = - Polynomial.aeval f n a := - rfl -#align module_polynomial_of_endo_smul_def modulePolynomialOfEndo_smul_def - -attribute [local simp] modulePolynomialOfEndo_smul_def - -theorem modulePolynomialOfEndo.isScalarTower : - @IsScalarTower R R[X] M _ - (by - letI := modulePolynomialOfEndo f - infer_instance) - _ := by - let _ := modulePolynomialOfEndo f - constructor - intro x y z - simp -#align module_polynomial_of_endo.is_scalar_tower modulePolynomialOfEndo.isScalarTower - -open Polynomial Module - /-- A theorem/proof by Vasconcelos, given a finite module `M` over a commutative ring, any surjective endomorphism of `M` is also injective. Based on, https://math.stackexchange.com/a/239419/31917, https://www.ams.org/journals/tran/1969-138-00/S0002-9947-1969-0238839-5/. This is similar to `IsNoetherian.injective_of_surjective_endomorphism` but only applies in the commutative case, but does not use a Noetherian hypothesis. -/ -theorem Module.Finite.injective_of_surjective_endomorphism [hfg : Finite R M] +theorem Module.Finite.injective_of_surjective_endomorphism {R : Type*} [CommRing R] {M : Type*} + [AddCommGroup M] [Module R M] [Finite R M] (f : M →ₗ[R] M) (f_surj : Function.Surjective f) : Function.Injective f := by - let _ := modulePolynomialOfEndo f - haveI : IsScalarTower R R[X] M := modulePolynomialOfEndo.isScalarTower f - have hfgpoly : Finite R[X] M := Finite.of_restrictScalars_finite R _ _ - have X_mul : ∀ o, (X : R[X]) • o = f o := by - intro - rw [modulePolynomialOfEndo_smul_def, aeval_X] - have : (⊤ : Submodule R[X] M) ≤ Ideal.span {(X : R[X])} • ⊤ := by - intro a ha - obtain ⟨y, rfl⟩ := f_surj a - rw [← X_mul y] + have : (⊤ : Submodule R[X] (AEval' f)) ≤ Ideal.span {(X : R[X])} • ⊤ + · intro a _ + obtain ⟨y, rfl⟩ := f_surj.comp (AEval'.of f).symm.surjective a + rw [Function.comp_apply, ←AEval'.of_symm_X_smul] exact Submodule.smul_mem_smul (Ideal.mem_span_singleton.mpr (dvd_refl _)) trivial obtain ⟨F, hFa, hFb⟩ := - Submodule.exists_sub_one_mem_and_smul_eq_zero_of_fg_of_le_smul _ (⊤ : Submodule R[X] M) - (finite_def.mp hfgpoly) this + Submodule.exists_sub_one_mem_and_smul_eq_zero_of_fg_of_le_smul _ (⊤ : Submodule R[X] (AEval' f)) + (finite_def.mp inferInstance) this rw [← LinearMap.ker_eq_bot, LinearMap.ker_eq_bot'] intro m hm + rw [← map_eq_zero_iff (AEval'.of f) (AEval'.of f).injective] + set m' := Module.AEval'.of f m rw [Ideal.mem_span_singleton'] at hFa obtain ⟨G, hG⟩ := hFa - suffices (F - 1) • m = 0 by - have Fmzero := hFb m (by simp) + suffices (F - 1) • m' = 0 by + have Fmzero := hFb m' (by simp) rwa [← sub_add_cancel F 1, add_smul, one_smul, this, zero_add] at Fmzero - rw [← hG, mul_smul, X_mul m, hm, smul_zero] + rw [← hG, mul_smul, AEval'.X_smul_of, hm, map_zero, smul_zero] #align module.finite.injective_of_surjective_endomorphism Module.Finite.injective_of_surjective_endomorphism -end - end Vasconcelos diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index a5060a8a63a04..f4943b805f947 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -147,7 +147,7 @@ theorem fg_bot : (⊥ : Submodule R M).FG := theorem _root_.Subalgebra.fg_bot_toSubmodule {R A : Type*} [CommSemiring R] [Semiring A] [Algebra R A] : (⊥ : Subalgebra R A).toSubmodule.FG := - ⟨{1}, by simp [Algebra.toSubmodule_bot]⟩ + ⟨{1}, by simp [Algebra.toSubmodule_bot, one_eq_span]⟩ #align subalgebra.fg_bot_to_submodule Subalgebra.fg_bot_toSubmodule theorem fg_unit {R A : Type*} [CommSemiring R] [Semiring A] [Algebra R A] (I : (Submodule R A)ˣ) : diff --git a/Mathlib/RingTheory/Flat.lean b/Mathlib/RingTheory/Flat.lean index 1e902b721bade..97033513813eb 100644 --- a/Mathlib/RingTheory/Flat.lean +++ b/Mathlib/RingTheory/Flat.lean @@ -4,6 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.RingTheory.Noetherian +import Mathlib.Algebra.DirectSum.Module +import Mathlib.Algebra.DirectSum.Finsupp +import Mathlib.LinearAlgebra.DirectSum.TensorProduct +import Mathlib.LinearAlgebra.FreeModule.Basic +import Mathlib.Algebra.Module.Projective #align_import ring_theory.flat from "leanprover-community/mathlib"@"62c0a4ef1441edb463095ea02a06e87f3dfe135c" @@ -23,6 +28,14 @@ This result is not yet formalised. * `Module.Flat`: the predicate asserting that an `R`-module `M` is flat. +## Main theorems + +* `Module.Flat.of_retract`: retracts of flat modules are flat +* `Module.Flat.of_linearEquiv`: modules linearly equivalent to a flat modules are flat +* `Module.Flat.directSum`: arbitrary direct sums of flat modules are flat +* `Module.Flat.of_free`: free modules are flat +* `Module.Flat.of_projective`: projective modules are flat + ## TODO * Show that tensoring with a flat module preserves injective morphisms. @@ -46,13 +59,13 @@ This result is not yet formalised. -/ -universe u v +universe u v w namespace Module -open Function (Injective) +open Function (Injective Surjective) -open LinearMap (lsmul) +open LinearMap (lsmul rTensor lTensor) open TensorProduct @@ -64,7 +77,9 @@ class Flat (R : Type u) (M : Type v) [CommRing R] [AddCommGroup M] [Module R M] namespace Flat -open TensorProduct LinearMap Submodule +variable (R : Type u) [CommRing R] + +open LinearMap Submodule instance self (R : Type u) [CommRing R] : Flat R R := ⟨by @@ -76,6 +91,120 @@ instance self (R : Type u) [CommRing R] : Flat R R := lift.tmul, Submodule.subtype_apply, Algebra.id.smul_eq_mul, lsmul_apply]⟩ #align module.flat.self Module.Flat.self +variable (M : Type v) [AddCommGroup M] [Module R M] + +lemma iff_rTensor_injective : + Flat R M ↔ (∀ ⦃I : Ideal R⦄ (_ : I.FG), Injective (rTensor M I.subtype)) := by + have aux : ∀ (I : Ideal R), ((TensorProduct.lid R M).comp (rTensor M I.subtype)) = + (TensorProduct.lift ((lsmul R M).comp I.subtype)) + · intro I; apply TensorProduct.ext'; intro x y; simp + constructor + · intro F I hI + erw [← Equiv.comp_injective _ (TensorProduct.lid R M).toEquiv] + have h₁ := F.out hI + rw [← aux] at h₁ + refine h₁ + · intro h₁ + constructor + intro I hI + rw [← aux] + simp [h₁ hI] + +variable (N : Type w) [AddCommGroup N] [Module R N] + +/-- A retract of a flat `R`-module is flat. -/ +lemma of_retract [f : Flat R M] (i : N →ₗ[R] M) (r : M →ₗ[R] N) (h : r.comp i = LinearMap.id) : + Flat R N := by + rw [iff_rTensor_injective] at * + intro I hI + have h₁ : Function.Injective (lTensor R i) + · apply Function.RightInverse.injective (g := (lTensor R r)) + intro x + rw [← LinearMap.comp_apply, ← lTensor_comp, h] + simp + rw [← Function.Injective.of_comp_iff h₁ (rTensor N I.subtype), ← LinearMap.coe_comp] + rw [LinearMap.lTensor_comp_rTensor, ←LinearMap.rTensor_comp_lTensor] + rw [LinearMap.coe_comp, Function.Injective.of_comp_iff (f hI)] + apply Function.RightInverse.injective (g := lTensor _ r) + intro x + rw [← LinearMap.comp_apply, ← lTensor_comp, h] + simp + +/-- A `R`-module linearly equivalent to a flat `R`-module is flat. -/ +lemma of_linearEquiv [f : Flat R M] (e : N ≃ₗ[R] M) : Flat R N := by + have h : e.symm.toLinearMap.comp e.toLinearMap = LinearMap.id := by simp + exact of_retract _ _ _ e.toLinearMap e.symm.toLinearMap h + +end Flat + +namespace Flat + +open DirectSum LinearMap Submodule + +variable (R : Type u) [CommRing R] + +/-- A direct sum of flat `R`-modules is flat. -/ +instance directSum (ι : Type v) (M : ι → Type w) [(i : ι) → AddCommGroup (M i)] + [(i : ι) → Module R (M i)] [F : (i : ι) → (Flat R (M i))] : Flat R (⨁ i, M i) := by + classical + rw [iff_rTensor_injective] + intro I hI + rw [← Equiv.comp_injective _ (TensorProduct.lid R (⨁ i, M i)).toEquiv] + set η₁ := TensorProduct.lid R (⨁ i, M i) + set η := (fun i ↦ TensorProduct.lid R (M i)) + set φ := (fun i ↦ rTensor (M i) I.subtype) + set π := (fun i ↦ component R ι (fun j ↦ M j) i) + set ψ := (TensorProduct.directSumRight R {x // x ∈ I} (fun i ↦ M i)).symm.toLinearMap with psi_def + set ρ := rTensor (⨁ i, M i) I.subtype + set τ := (fun i ↦ component R ι (fun j ↦ ({x // x ∈ I} ⊗[R] (M j))) i) + rw [← Equiv.injective_comp (TensorProduct.directSumRight _ _ _).symm.toEquiv] + rw [LinearEquiv.coe_toEquiv, ← LinearEquiv.coe_coe, ← LinearMap.coe_comp] + rw [LinearEquiv.coe_toEquiv, ← LinearEquiv.coe_coe, ← LinearMap.coe_comp] + rw [← psi_def, injective_iff_map_eq_zero ((η₁.comp ρ).comp ψ)] + have h₁ : ∀ (i : ι), (π i).comp ((η₁.comp ρ).comp ψ) = (η i).comp ((φ i).comp (τ i)) + · intro i + apply DirectSum.linearMap_ext + intro j + apply TensorProduct.ext' + intro a m + simp only [coe_comp, LinearEquiv.coe_coe, Function.comp_apply, directSumRight_symm_lof_tmul, + rTensor_tmul, coeSubtype, lid_tmul, map_smul] + rw [DirectSum.component.of, DirectSum.component.of] + by_cases h₂ : j = i + · subst j; simp + · simp [h₂] + intro a ha; rw [DirectSum.ext_iff R]; intro i + have f := LinearMap.congr_arg (f:= (π i)) ha + erw [LinearMap.congr_fun (h₁ i) a] at f + rw [(map_zero (π i) : (π i) 0 = (0 : M i))] at f + have h₂ := (F i) + rw [iff_rTensor_injective] at h₂ + have h₃ := h₂ hI + simp only [coe_comp, LinearEquiv.coe_coe, Function.comp_apply, AddEquivClass.map_eq_zero_iff, + h₃, LinearMap.map_eq_zero_iff] at f + simp [f] + +/-- Free `R`-modules over discrete types are flat. -/ +instance finsupp (ι : Type v) : Flat R (ι →₀ R) := + let _ := Classical.decEq ι + of_linearEquiv R _ _ (finsuppLEquivDirectSum R R ι) + +variable (M : Type v) [AddCommGroup M] [Module R M] + +instance of_free [Free R M] : Flat R M := of_linearEquiv R _ _ (Free.repr R M) + +/-- A projective module with a discrete type of generator is flat -/ +lemma of_projective_surjective (ι : Type w) [Projective R M] (p : (ι →₀ R) →ₗ[R] M) + (hp : Surjective p) : Flat R M := by + have h := Module.projective_lifting_property p (LinearMap.id) hp + cases h with + | _ e he => exact of_retract R _ _ _ _ he + +instance of_projective [h : Projective R M] : Flat R M := by + rw [Module.projective_def'] at h + cases h with + | _ e he => exact of_retract R _ _ _ _ he + end Flat end Module diff --git a/Mathlib/RingTheory/HahnSeries.lean b/Mathlib/RingTheory/HahnSeries.lean index 0236bf7c3d1e6..0d1246a682392 100644 --- a/Mathlib/RingTheory/HahnSeries.lean +++ b/Mathlib/RingTheory/HahnSeries.lean @@ -1126,7 +1126,7 @@ def toPowerSeries : HahnSeries ℕ R ≃+* PowerSeries R where classical refine' sum_filter_ne_zero.symm.trans ((sum_congr _ fun _ _ => rfl).trans sum_filter_ne_zero) ext m - simp only [Nat.mem_antidiagonal, mem_addAntidiagonal, and_congr_left_iff, mem_filter, + simp only [mem_antidiagonal, mem_addAntidiagonal, and_congr_left_iff, mem_filter, mem_support] rintro h rw [and_iff_right (left_ne_zero_of_mul h), and_iff_right (right_ne_zero_of_mul h)] @@ -1241,7 +1241,7 @@ def toMvPowerSeries {σ : Type*} [Fintype σ] : HahnSeries (σ →₀ ℕ) R ≃ refine' sum_filter_ne_zero.symm.trans ((sum_congr _ fun _ _ => rfl).trans sum_filter_ne_zero) ext m simp only [and_congr_left_iff, mem_addAntidiagonal, mem_filter, mem_support, - Finsupp.mem_antidiagonal] + Finset.mem_antidiagonal] rintro h rw [and_iff_right (left_ne_zero_of_mul h), and_iff_right (right_ne_zero_of_mul h)] #align hahn_series.to_mv_power_series HahnSeries.toMvPowerSeries @@ -1270,12 +1270,7 @@ def toPowerSeriesAlg : HahnSeries ℕ A ≃ₐ[R] PowerSeries A := { toPowerSeries with commutes' := fun r => by ext n - simp only [algebraMap_apply, PowerSeries.algebraMap_apply, C_apply, - coeff_toPowerSeries] - cases' n with n - · simp [PowerSeries.coeff_zero_eq_constantCoeff, single_coeff_same] - · simp [n.succ_ne_zero, Ne.def, not_false_iff, single_coeff_of_ne] - rw [PowerSeries.coeff_C, if_neg n.succ_ne_zero] } + cases n <;> simp [algebraMap_apply, PowerSeries.algebraMap_apply] } #align hahn_series.to_power_series_alg HahnSeries.toPowerSeriesAlg variable (Γ) [StrictOrderedSemiring Γ] diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 3d35a1aefb38b..f9c4a54300481 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -2035,6 +2035,21 @@ end Basis end Ideal +section span_range +variable {α R : Type*} [Semiring R] + +theorem Finsupp.mem_ideal_span_range_iff_exists_finsupp {x : R} {v : α → R} : + x ∈ Ideal.span (Set.range v) ↔ ∃ c : α →₀ R, (c.sum fun i a => a * v i) = x := + Finsupp.mem_span_range_iff_exists_finsupp + +/-- An element `x` lies in the span of `v` iff it can be written as sum `∑ cᵢ • vᵢ = x`. +-/ +theorem mem_ideal_span_range_iff_exists_fun [Fintype α] {x : R} {v : α → R} : + x ∈ Ideal.span (Set.range v) ↔ ∃ c : α → R, ∑ i, c i * v i = x := + mem_span_range_iff_exists_fun _ + +end span_range + theorem Associates.mk_ne_zero' {R : Type*} [CommSemiring R] {r : R} : Associates.mk (Ideal.span {r} : Ideal R) ≠ 0 ↔ r ≠ 0 := by rw [Associates.mk_ne_zero, Ideal.zero_eq_bot, Ne.def, Ideal.span_singleton_eq_bot] diff --git a/Mathlib/RingTheory/Int/Basic.lean b/Mathlib/RingTheory/Int/Basic.lean index 1fad6d1dc062d..ba62b5f36f031 100644 --- a/Mathlib/RingTheory/Int/Basic.lean +++ b/Mathlib/RingTheory/Int/Basic.lean @@ -37,8 +37,7 @@ instance : WfDvdMonoid ℕ := ⟨by refine' RelHomClass.wellFounded - (⟨fun x : ℕ => if x = 0 then (⊤ : ℕ∞) else x, _⟩ : DvdNotUnit →r (· < ·)) - (WithTop.wellFounded_lt Nat.lt_wfRel.wf) + (⟨fun x : ℕ => if x = 0 then (⊤ : ℕ∞) else x, _⟩ : DvdNotUnit →r (· < ·)) wellFounded_lt intro a b h cases' a with a · exfalso @@ -294,7 +293,7 @@ theorem prime_two_or_dvd_of_dvd_two_mul_pow_self_two {m : ℤ} {p : ℕ} (hp : N exact le_antisymm (Nat.le_of_dvd zero_lt_two hp2) (Nat.Prime.two_le hp) · apply Or.intro_right rw [sq, Int.natAbs_mul] at hpp - exact (or_self_iff _).mp ((Nat.Prime.dvd_mul hp).mp hpp) + exact or_self_iff.mp ((Nat.Prime.dvd_mul hp).mp hpp) #align prime_two_or_dvd_of_dvd_two_mul_pow_self_two prime_two_or_dvd_of_dvd_two_mul_pow_self_two theorem Int.exists_prime_and_dvd {n : ℤ} (hn : n.natAbs ≠ 1) : ∃ p, Prime p ∧ p ∣ n := by diff --git a/Mathlib/RingTheory/IntegralClosure.lean b/Mathlib/RingTheory/IntegralClosure.lean index 46d2a6b194385..042cabacd61d8 100644 --- a/Mathlib/RingTheory/IntegralClosure.lean +++ b/Mathlib/RingTheory/IntegralClosure.lean @@ -250,9 +250,7 @@ theorem FG_adjoin_of_finite {s : Set A} (hfs : s.Finite) (his : ∀ x ∈ s, IsI (fun _ => ⟨{1}, Submodule.ext fun x => by - erw [Algebra.adjoin_empty, Finset.coe_singleton, ← one_eq_span, one_eq_range, - LinearMap.mem_range, Algebra.mem_bot] - rfl⟩) + rw [Algebra.adjoin_empty, Finset.coe_singleton, ← one_eq_span, Algebra.toSubmodule_bot]⟩) (fun {a s} _ _ ih his => by rw [← Set.union_singleton, Algebra.adjoin_union_coe_submodule] exact @@ -460,12 +458,8 @@ theorem Algebra.IsIntegral.finite (h : Algebra.IsIntegral R A) [h' : Algebra.Fin exact Algebra.smul_def _ _ #align algebra.is_integral.finite Algebra.IsIntegral.finite -theorem Algebra.IsIntegral.of_finite [h : Module.Finite R A] : Algebra.IsIntegral R A := by - apply RingHom.Finite.to_isIntegral - rw [RingHom.Finite] - convert h - ext - exact (Algebra.smul_def _ _).symm +theorem Algebra.IsIntegral.of_finite [h : Module.Finite R A] : Algebra.IsIntegral R A := + fun _ ↦ isIntegral_of_mem_of_FG ⊤ h.1 _ trivial #align algebra.is_integral.of_finite Algebra.IsIntegral.of_finite /-- finite = integral + finite type -/ @@ -734,7 +728,7 @@ theorem normalizeScaleRoots_coeff_mul_leadingCoeff_pow (i : ℕ) (hp : 1 ≤ nat · simp [h₁] · rw [h₂, leadingCoeff, ← pow_succ, tsub_add_cancel_of_le hp] · rw [mul_assoc, ← pow_add, tsub_add_cancel_of_le] - apply Nat.le_pred_of_lt + apply Nat.le_sub_one_of_lt rw [lt_iff_le_and_ne] exact ⟨le_natDegree_of_ne_zero h₁, h₂⟩ #align normalize_scale_roots_coeff_mul_leading_coeff_pow normalizeScaleRoots_coeff_mul_leadingCoeff_pow diff --git a/Mathlib/RingTheory/Kaehler.lean b/Mathlib/RingTheory/Kaehler.lean index 1c811437c9aca..1e34ef46ad200 100644 --- a/Mathlib/RingTheory/Kaehler.lean +++ b/Mathlib/RingTheory/Kaehler.lean @@ -482,9 +482,8 @@ noncomputable def KaehlerDifferential.kerTotal : Submodule S (S →₀ S) := unsuppress_compilation in -- Porting note: was `local notation x "𝖣" y => (KaehlerDifferential.kerTotal R S).mkQ (single y x)` --- but `notation3` wants an explicit expansion to be able to generate a pretty printer. -local notation3 x "𝖣" y => - FunLike.coe (Submodule.mkQ (KaehlerDifferential.kerTotal R S)) (single y x) +-- but not having `FunLike.coe` leads to `kerTotal_mkQ_single_smul` failing. +local notation3 x "𝖣" y => FunLike.coe (KaehlerDifferential.kerTotal R S).mkQ (single y x) theorem KaehlerDifferential.kerTotal_mkQ_single_add (x y z) : (z𝖣x + y) = (z𝖣x) + z𝖣y := by rw [← map_add, eq_comm, ← sub_eq_zero, ← map_sub (Submodule.mkQ (kerTotal R S)), @@ -636,18 +635,6 @@ variable (A B : Type*) [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] variable [Algebra A B] [Algebra S B] [IsScalarTower R A B] [IsScalarTower R S B] -variable {R B} - -/-- For a tower `R → A → B` and an `R`-derivation `B → M`, we may compose with `A → B` to obtain an -`R`-derivation `A → M`. -/ -def Derivation.compAlgebraMap [Module A M] [Module B M] [IsScalarTower A B M] - (d : Derivation R B M) : Derivation R A M where - map_one_eq_zero' := by simp - leibniz' a b := by simp - toLinearMap := d.toLinearMap.comp (IsScalarTower.toAlgHom R A B).toLinearMap -#align derivation.comp_algebra_map Derivation.compAlgebraMap - -variable (R B) variable [SMulCommClass S A B] /-- The map `Ω[A⁄R] →ₗ[A] Ω[B⁄R]` given a square diff --git a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean index 75de3d7aef968..ed0570655eda9 100644 --- a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean @@ -98,7 +98,7 @@ theorem homogeneousSubmodule_mul [CommSemiring R] (m n : ℕ) : simp_all only [Ne.def, not_false_iff, zero_mul, mul_zero] specialize hφ aux.1 specialize hψ aux.2 - rw [Finsupp.mem_antidiagonal] at hde + rw [Finset.mem_antidiagonal] at hde classical have hd' : d.support ⊆ d.support ∪ e.support := Finset.subset_union_left _ _ have he' : e.support ⊆ d.support ∪ e.support := Finset.subset_union_right _ _ @@ -204,7 +204,7 @@ theorem prod {ι : Type*} (s : Finset ι) (φ : ι → MvPolynomial σ R) (n : #align mv_polynomial.is_homogeneous.prod MvPolynomial.IsHomogeneous.prod theorem totalDegree (hφ : IsHomogeneous φ n) (h : φ ≠ 0) : totalDegree φ = n := by - rw [totalDegree] + rw [MvPolynomial.totalDegree] apply le_antisymm · apply Finset.sup_le intro d hd diff --git a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean index 6b1840e725a4c..f93639b52657e 100644 --- a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean @@ -93,6 +93,7 @@ private theorem pairMap_mem_pairs {k : ℕ} (t : Finset σ × σ) (h : t ∈ pai simp only [card_erase_of_mem h1, tsub_le_iff_right, mem_erase, ne_eq, h1] refine ⟨le_step h, ?_⟩ by_contra h2 + simp only [not_true_eq_false, and_true, not_forall, not_false_eq_true, exists_prop] at h2 rw [← h2] at h exact not_le_of_lt (sub_lt (card_pos.mpr ⟨t.snd, h1⟩) zero_lt_one) h · rw [pairMap_of_snd_nmem_fst σ h1] @@ -184,8 +185,9 @@ private theorem sum_filter_pairs_eq_sum_filter_antidiagonal_powersetCard_sum (k have hnk' : n.fst ≤ k := by apply le_of_lt; aesop aesop · simp_all only [mem_antidiagonal, mem_filter, mem_pairs, disjiUnion_eq_biUnion, - add_tsub_cancel_of_le] - · simp_all only [mem_antidiagonal, mem_filter, mem_pairs, disjiUnion_eq_biUnion, implies_true] + add_tsub_cancel_of_le, and_true] + · simp_all only [mem_antidiagonal, mem_filter, mem_pairs, disjiUnion_eq_biUnion, + implies_true, and_true] simp only [← hdisj, disj_equiv] private theorem disjoint_filter_pairs_lt_filter_pairs_eq (k : ℕ) : @@ -280,7 +282,7 @@ theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : psum σ R k = rw [mem_filter, mem_antidiagonal, mem_singleton] refine' ⟨_, fun ha ↦ by aesop⟩ rintro ⟨ha, ⟨_, ha0⟩⟩ - rw [← ha, Nat.eq_zero_of_nonpos a.fst ha0, zero_add, ← Nat.eq_zero_of_nonpos a.fst ha0] + rw [← ha, Nat.eq_zero_of_not_pos ha0, zero_add, ← Nat.eq_zero_of_not_pos ha0] rw [this, sum_singleton] at sub_both_sides simp only [_root_.pow_zero, esymm_zero, mul_one, one_mul, filter_filter] at sub_both_sides exact sub_both_sides.symm diff --git a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean index adcef6c6d5d0e..e8f70f29aafde 100644 --- a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean @@ -185,7 +185,7 @@ theorem weightedHomogeneousSubmodule_mul (w : σ → M) (m n : M) : contrapose! H by_cases h : coeff d φ = 0 <;> simp_all only [Ne.def, not_false_iff, zero_mul, mul_zero] - rw [← Finsupp.mem_antidiagonal.mp hde, ← hφ aux.1, ← hψ aux.2, map_add] + rw [← mem_antidiagonal.mp hde, ← hφ aux.1, ← hψ aux.2, map_add] #align mv_polynomial.weighted_homogeneous_submodule_mul MvPolynomial.weightedHomogeneousSubmodule_mul /-- Monomials are weighted homogeneous. -/ diff --git a/Mathlib/RingTheory/Nilpotent.lean b/Mathlib/RingTheory/Nilpotent.lean index a9d30343e14c6..727a2537c7cfe 100644 --- a/Mathlib/RingTheory/Nilpotent.lean +++ b/Mathlib/RingTheory/Nilpotent.lean @@ -78,6 +78,11 @@ theorem IsNilpotent.map [MonoidWithZero R] [MonoidWithZero S] {r : R} {F : Type* rw [← map_pow, hr.choose_spec, map_zero] #align is_nilpotent.map IsNilpotent.map +lemma IsNilpotent.map_iff [MonoidWithZero R] [MonoidWithZero S] {r : R} {F : Type*} + [MonoidWithZeroHomClass F R S] {f : F} (hf : Function.Injective f) : + IsNilpotent (f r) ↔ IsNilpotent r := + ⟨fun ⟨k, hk⟩ ↦ ⟨k, (map_eq_zero_iff f hf).mp <| by rwa [map_pow]⟩, fun h ↦ h.map f⟩ + theorem IsNilpotent.sub_one_isUnit [Ring R] {r : R} (hnil : IsNilpotent r) : IsUnit (r - 1) := by obtain ⟨n, hn⟩ := hnil refine' ⟨⟨r - 1, -∑ i in Finset.range n, r ^ i, _, _⟩, rfl⟩ @@ -181,7 +186,7 @@ theorem isNilpotent_add (hx : IsNilpotent x) (hy : IsNilpotent y) : IsNilpotent apply Finset.sum_eq_zero rintro ⟨i, j⟩ hij suffices x ^ i * y ^ j = 0 by simp only [this, nsmul_eq_mul, mul_zero] - cases' Nat.le_or_le_of_add_eq_add_pred (Finset.Nat.mem_antidiagonal.mp hij) with hi hj + cases' Nat.le_or_le_of_add_eq_add_pred (Finset.mem_antidiagonal.mp hij) with hi hj · rw [pow_eq_zero_of_le hi hn, zero_mul] · rw [pow_eq_zero_of_le hj hm, mul_zero] #align commute.is_nilpotent_add Commute.isNilpotent_add @@ -322,3 +327,16 @@ theorem IsNilpotent.mapQ (hnp : IsNilpotent f) : IsNilpotent (p.mapQ p f hp) := #align module.End.is_nilpotent.mapq Module.End.IsNilpotent.mapQ end Module.End + +lemma NoZeroSMulDivisors.isReduced (R M : Type*) + [MonoidWithZero R] [Zero M] [MulActionWithZero R M] [Nontrivial M] [NoZeroSMulDivisors R M] : + IsReduced R := by + refine ⟨fun x ⟨k, hk⟩ ↦ ?_⟩ + induction' k with k ih + · rw [Nat.zero_eq, pow_zero] at hk + exact eq_zero_of_zero_eq_one hk.symm x + · obtain ⟨m : M, hm : m ≠ 0⟩ := exists_ne (0 : M) + have : x ^ (k + 1) • m = 0 := by simp only [hk, zero_smul] + rw [pow_succ, mul_smul] at this + rcases eq_zero_or_eq_zero_of_smul_eq_zero this with rfl | hx; rfl + exact ih <| (eq_zero_or_eq_zero_of_smul_eq_zero hx).resolve_right hm diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 1cd38c230aa0f..e466850a53d3c 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -631,7 +631,7 @@ theorem _root_.Polynomial.coeff_prod_mem_ideal_pow_tsub {ι : Type*} (s : Finset · rw [sum_insert ha, prod_insert ha, coeff_mul] apply sum_mem rintro ⟨i, j⟩ e - obtain rfl : i + j = k := Nat.mem_antidiagonal.mp e + obtain rfl : i + j = k := mem_antidiagonal.mp e apply Ideal.pow_le_pow add_tsub_add_le_tsub_add_tsub rw [pow_add] exact @@ -703,14 +703,14 @@ theorem isPrime_map_C_iff_isPrime (P : Ideal R) : let m := Nat.find hf let n := Nat.find hg refine' ⟨m + n, _⟩ - rw [coeff_mul, ← Finset.insert_erase ((@Finset.Nat.mem_antidiagonal _ (m, n)).mpr rfl), + rw [coeff_mul, ← Finset.insert_erase ((Finset.mem_antidiagonal (a := (m,n))).mpr rfl), Finset.sum_insert (Finset.not_mem_erase _ _), (P.add_mem_iff_left _).not] · apply mt h.2 rw [not_or] exact ⟨Nat.find_spec hf, Nat.find_spec hg⟩ apply P.sum_mem rintro ⟨i, j⟩ hij - rw [Finset.mem_erase, Finset.Nat.mem_antidiagonal] at hij + rw [Finset.mem_erase, Finset.mem_antidiagonal] at hij simp only [Ne.def, Prod.mk.inj_iff, not_and_or] at hij obtain hi | hj : i < m ∨ j < n := by rw [or_iff_not_imp_left, not_lt, le_iff_lt_or_eq] @@ -860,8 +860,7 @@ instance (priority := 100) {R : Type*} [CommRing R] [IsDomain R] [WfDvdMonoid R] (⟨fun p : R[X] => ((if p = 0 then ⊤ else ↑p.degree : WithTop (WithBot ℕ)), p.leadingCoeff), _⟩ : DvdNotUnit →r Prod.Lex (· < ·) DvdNotUnit) - (WellFounded.prod_lex (WithTop.wellFounded_lt <| WithBot.wellFounded_lt Nat.lt_wfRel.wf) - ‹WfDvdMonoid R›.wellFounded_dvdNotUnit) + (wellFounded_lt.prod_lex ‹WfDvdMonoid R›.wellFounded_dvdNotUnit) rintro a b ⟨ane0, ⟨c, ⟨not_unit_c, rfl⟩⟩⟩ dsimp rw [Polynomial.degree_mul, if_neg ane0] diff --git a/Mathlib/RingTheory/Polynomial/Bernstein.lean b/Mathlib/RingTheory/Polynomial/Bernstein.lean index 826dbb78cce20..43ad2835716ae 100644 --- a/Mathlib/RingTheory/Polynomial/Bernstein.lean +++ b/Mathlib/RingTheory/Polynomial/Bernstein.lean @@ -189,7 +189,7 @@ theorem iterate_derivative_at_0 (n ν : ℕ) : rw [this, ascPochhammer_eval_succ] rw_mod_cast [tsub_add_cancel_of_le (h'.trans n.pred_le)] · simp only [not_le] at h - rw [tsub_eq_zero_iff_le.mpr (Nat.le_pred_of_lt h), eq_zero_of_lt R h] + rw [tsub_eq_zero_iff_le.mpr (Nat.le_sub_one_of_lt h), eq_zero_of_lt R h] simp [pos_iff_ne_zero.mp (pos_of_gt h)] #align bernstein_polynomial.iterate_derivative_at_0 bernsteinPolynomial.iterate_derivative_at_0 diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean index c2528bfbaa470..53eaf0060c3c2 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean @@ -339,7 +339,7 @@ theorem degree_cyclotomic (n : ℕ) (R : Type*) [Ring R] [Nontrivial R] : rw [← map_cyclotomic_int] rw [degree_map_eq_of_leadingCoeff_ne_zero (Int.castRingHom R) _] · cases' n with k - · simp only [cyclotomic, degree_one, dif_pos, Nat.totient_zero, WithTop.coe_zero] + · simp only [cyclotomic, degree_one, dif_pos, Nat.totient_zero, CharP.cast_eq_zero] rw [← degree_cyclotomic' (Complex.isPrimitiveRoot_exp k.succ (Nat.succ_ne_zero k))] exact (int_cyclotomic_spec k.succ).2.1 simp only [(int_cyclotomic_spec n).right.right, eq_intCast, Monic.leadingCoeff, Int.cast_one, diff --git a/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean b/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean index 06dd5a2f76c6c..5b7fbd652fba8 100644 --- a/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean +++ b/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean @@ -133,7 +133,8 @@ theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S} tsub_add_cancel_of_le one_le_deg] exact degree_eq_natDegree hp · have : i.1 ≤ p.natDegree - 1 := - Nat.le_pred_of_lt (lt_of_le_of_ne (le_natDegree_of_ne_zero (mem_support_iff.mp i.2)) hi) + Nat.le_sub_one_of_lt + (lt_of_le_of_ne (le_natDegree_of_ne_zero (mem_support_iff.mp i.2)) hi) rw [integralNormalization_coeff_ne_natDegree hi, mul_assoc, ← pow_add, tsub_add_cancel_of_le this] _ = f p.leadingCoeff ^ (natDegree p - 1) * eval₂ f z p := by diff --git a/Mathlib/RingTheory/Polynomial/Pochhammer.lean b/Mathlib/RingTheory/Polynomial/Pochhammer.lean index f81166c85ba97..916dc9f3c71e9 100644 --- a/Mathlib/RingTheory/Polynomial/Pochhammer.lean +++ b/Mathlib/RingTheory/Polynomial/Pochhammer.lean @@ -310,7 +310,7 @@ theorem descPochhammer_succ_right (n : ℕ) : rw [Nat.succ_eq_one_add, Nat.cast_add, Nat.cast_one, sub_add_eq_sub_sub] @[simp] -theorem descPochhammer_natDegree (n : ℕ) [NoZeroDivisors R] [Nontrivial R]: +theorem descPochhammer_natDegree (n : ℕ) [NoZeroDivisors R] [Nontrivial R] : (descPochhammer R n).natDegree = n := by induction' n with n hn · simp diff --git a/Mathlib/RingTheory/PolynomialAlgebra.lean b/Mathlib/RingTheory/PolynomialAlgebra.lean index 102b7915a3a6c..ab16bc0d5aa33 100644 --- a/Mathlib/RingTheory/PolynomialAlgebra.lean +++ b/Mathlib/RingTheory/PolynomialAlgebra.lean @@ -86,7 +86,7 @@ theorem toFunLinear_mul_tmul_mul_aux_1 (p : R[X]) (k : ℕ) (h : Decidable ¬p.c theorem toFunLinear_mul_tmul_mul_aux_2 (k : ℕ) (a₁ a₂ : A) (p₁ p₂ : R[X]) : a₁ * a₂ * (algebraMap R A) ((p₁ * p₂).coeff k) = - (Finset.Nat.antidiagonal k).sum fun x => + (Finset.antidiagonal k).sum fun x => a₁ * (algebraMap R A) (coeff p₁ x.1) * (a₂ * (algebraMap R A) (coeff p₂ x.2)) := by simp_rw [mul_assoc, Algebra.commutes, ← Finset.mul_sum, mul_assoc, ← Finset.mul_sum] congr diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index 1c42b834f9cf8..a5b551b1371f2 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -72,6 +72,8 @@ noncomputable section open BigOperators Polynomial +open Finset (antidiagonal mem_antidiagonal) + /-- Multivariate formal power series, where `σ` is the index set of the variables and `R` is the coefficient ring.-/ def MvPowerSeries (σ : Type*) (R : Type*) := @@ -212,10 +214,10 @@ instance : AddMonoidWithOne (MvPowerSeries σ R) := instance : Mul (MvPowerSeries σ R) := letI := Classical.decEq σ - ⟨fun φ ψ n => ∑ p in Finsupp.antidiagonal n, coeff R p.1 φ * coeff R p.2 ψ⟩ + ⟨fun φ ψ n => ∑ p in antidiagonal n, coeff R p.1 φ * coeff R p.2 ψ⟩ theorem coeff_mul [DecidableEq σ] : - coeff R n (φ * ψ) = ∑ p in Finsupp.antidiagonal n, coeff R p.1 φ * coeff R p.2 ψ := by + coeff R n (φ * ψ) = ∑ p in antidiagonal n, coeff R p.1 φ * coeff R p.2 ψ := by refine Finset.sum_congr ?_ fun _ _ => rfl rw [Subsingleton.elim (Classical.decEq σ) ‹DecidableEq σ›] #align mv_power_series.coeff_mul MvPowerSeries.coeff_mul @@ -235,7 +237,8 @@ theorem coeff_monomial_mul (a : R) : ∀ p ∈ antidiagonal m, coeff R (p : (σ →₀ ℕ) × (σ →₀ ℕ)).1 (monomial R n a) * coeff R p.2 φ ≠ 0 → p.1 = n := fun p _ hp => eq_of_coeff_monomial_ne_zero (left_ne_zero_of_mul hp) - rw [coeff_mul, ← Finset.sum_filter_of_ne this, antidiagonal_filter_fst_eq, Finset.sum_ite_index] + rw [coeff_mul, ← Finset.sum_filter_of_ne this, Finset.filter_fst_eq_antidiagonal _ n, + Finset.sum_ite_index] simp only [Finset.sum_singleton, coeff_monomial_same, Finset.sum_empty] #align mv_power_series.coeff_monomial_mul MvPowerSeries.coeff_monomial_mul @@ -246,7 +249,8 @@ theorem coeff_mul_monomial (a : R) : ∀ p ∈ antidiagonal m, coeff R (p : (σ →₀ ℕ) × (σ →₀ ℕ)).1 φ * coeff R p.2 (monomial R n a) ≠ 0 → p.2 = n := fun p _ hp => eq_of_coeff_monomial_ne_zero (right_ne_zero_of_mul hp) - rw [coeff_mul, ← Finset.sum_filter_of_ne this, antidiagonal_filter_snd_eq, Finset.sum_ite_index] + rw [coeff_mul, ← Finset.sum_filter_of_ne this, Finset.filter_snd_eq_antidiagonal _ n, + Finset.sum_ite_index] simp only [Finset.sum_singleton, coeff_monomial_same, Finset.sum_empty] #align mv_power_series.coeff_mul_monomial MvPowerSeries.coeff_mul_monomial @@ -769,8 +773,9 @@ theorem X_pow_dvd_iff {s : σ} {n : ℕ} {φ : MvPowerSeries σ R} : rintro ⟨i, j⟩ hij rw [coeff_X_pow, if_neg, zero_mul] contrapose! h + dsimp at h subst i - rw [Finsupp.mem_antidiagonal] at hij + rw [mem_antidiagonal] at hij rw [← hij, Finsupp.add_apply, Finsupp.single_eq_same] exact Nat.le_add_right n _ · intro h @@ -781,7 +786,7 @@ theorem X_pow_dvd_iff {s : σ} {n : ℕ} {φ : MvPowerSeries σ R} : · rw [coeff_X_pow, if_pos rfl, one_mul] simpa using congr_arg (fun m : σ →₀ ℕ => coeff R m φ) H.symm · rintro ⟨i, j⟩ hij hne - rw [Finsupp.mem_antidiagonal] at hij + rw [mem_antidiagonal] at hij rw [coeff_X_pow] split_ifs with hi · exfalso @@ -794,10 +799,10 @@ theorem X_pow_dvd_iff {s : σ} {n : ℕ} {φ : MvPowerSeries σ R} : · intro hni exfalso apply hni - rwa [Finsupp.mem_antidiagonal, add_comm] + rwa [mem_antidiagonal, add_comm] · rw [h, coeff_mul, Finset.sum_eq_zero] · rintro ⟨i, j⟩ hij - rw [Finsupp.mem_antidiagonal] at hij + rw [mem_antidiagonal] at hij rw [coeff_X_pow] split_ifs with hi · exfalso @@ -844,7 +849,7 @@ protected noncomputable def inv.aux (a : R) (φ : MvPowerSeries σ R) : MvPowerS if n = 0 then a else -a * - ∑ x in n.antidiagonal, if _ : x.2 < n then coeff R x.1 φ * inv.aux a φ x.2 else 0 + ∑ x in antidiagonal n, if _ : x.2 < n then coeff R x.1 φ * inv.aux a φ x.2 else 0 termination_by _ n => n #align mv_power_series.inv.aux MvPowerSeries.inv.aux @@ -853,11 +858,11 @@ theorem coeff_inv_aux [DecidableEq σ] (n : σ →₀ ℕ) (a : R) (φ : MvPower if n = 0 then a else -a * - ∑ x in n.antidiagonal, if x.2 < n then coeff R x.1 φ * coeff R x.2 (inv.aux a φ) else 0 := + ∑ x in antidiagonal n, if x.2 < n then coeff R x.1 φ * coeff R x.2 (inv.aux a φ) else 0 := show inv.aux a φ n = _ by + cases Subsingleton.elim ‹DecidableEq σ› (Classical.decEq σ) rw [inv.aux] - -- unify `Decidable` instances - convert rfl + rfl #align mv_power_series.coeff_inv_aux MvPowerSeries.coeff_inv_aux /-- A multivariate formal power series is invertible if the constant coefficient is invertible.-/ @@ -870,7 +875,7 @@ theorem coeff_invOfUnit [DecidableEq σ] (n : σ →₀ ℕ) (φ : MvPowerSeries if n = 0 then ↑u⁻¹ else -↑u⁻¹ * - ∑ x in n.antidiagonal, + ∑ x in antidiagonal n, if x.2 < n then coeff R x.1 φ * coeff R x.2 (invOfUnit φ u) else 0 := by convert coeff_inv_aux n (↑u⁻¹) φ #align mv_power_series.coeff_inv_of_unit MvPowerSeries.coeff_invOfUnit @@ -891,7 +896,7 @@ theorem mul_invOfUnit (φ : MvPowerSeries σ R) (u : Rˣ) (h : constantCoeff σ simp [coeff_mul, support_single_ne_zero, h] else by classical - have : ((0 : σ →₀ ℕ), n) ∈ n.antidiagonal := by rw [Finsupp.mem_antidiagonal, zero_add] + have : ((0 : σ →₀ ℕ), n) ∈ antidiagonal n := by rw [mem_antidiagonal, zero_add] rw [coeff_one, if_neg H, coeff_mul, ← Finset.insert_erase this, Finset.sum_insert (Finset.not_mem_erase _ _), coeff_zero_eq_constantCoeff_apply, h, coeff_invOfUnit, if_neg H, neg_mul, mul_neg, Units.mul_inv_cancel_left, ← @@ -899,7 +904,7 @@ theorem mul_invOfUnit (φ : MvPowerSeries σ R) (u : Rˣ) (h : constantCoeff σ Finset.insert_erase this, if_neg (not_lt_of_ge <| le_rfl), zero_add, add_comm, ← sub_eq_add_neg, sub_eq_zero, Finset.sum_congr rfl] rintro ⟨i, j⟩ hij - rw [Finset.mem_erase, Finsupp.mem_antidiagonal] at hij + rw [Finset.mem_erase, mem_antidiagonal] at hij cases' hij with h₁ h₂ subst n rw [if_pos] @@ -970,7 +975,7 @@ theorem coeff_inv [DecidableEq σ] (n : σ →₀ ℕ) (φ : MvPowerSeries σ k) if n = 0 then (constantCoeff σ k φ)⁻¹ else -(constantCoeff σ k φ)⁻¹ * - ∑ x in n.antidiagonal, if x.2 < n then coeff k x.1 φ * coeff k x.2 φ⁻¹ else 0 := + ∑ x in antidiagonal n, if x.2 < n then coeff k x.1 φ * coeff k x.2 φ⁻¹ else 0 := coeff_inv_aux n _ φ #align mv_power_series.coeff_inv MvPowerSeries.coeff_inv @@ -1458,10 +1463,17 @@ set_option linter.uppercaseLean3 false in @[simp] theorem coeff_zero_C (a : R) : coeff R 0 (C R a) = a := by - rw [← monomial_zero_eq_C_apply, coeff_monomial_same 0 a] + rw [coeff_C, if_pos rfl] set_option linter.uppercaseLean3 false in #align power_series.coeff_zero_C PowerSeries.coeff_zero_C +theorem coeff_ne_zero_C {a : R} {n : ℕ} (h : n ≠ 0) : coeff R n (C R a) = 0 := by + rw [coeff_C, if_neg h] + +@[simp] +theorem coeff_succ_C {a : R} {n : ℕ} : coeff R (n + 1) (C R a) = 0 := + coeff_ne_zero_C n.succ_ne_zero + theorem X_eq : (X : R⟦X⟧) = monomial R 1 1 := rfl set_option linter.uppercaseLean3 false in @@ -1516,7 +1528,7 @@ theorem coeff_zero_one : coeff R 0 (1 : R⟦X⟧) = 1 := #align power_series.coeff_zero_one PowerSeries.coeff_zero_one theorem coeff_mul (n : ℕ) (φ ψ : R⟦X⟧) : - coeff R n (φ * ψ) = ∑ p in Finset.Nat.antidiagonal n, coeff R p.1 φ * coeff R p.2 ψ := by + coeff R n (φ * ψ) = ∑ p in antidiagonal n, coeff R p.1 φ * coeff R p.2 ψ := by -- `rw` can't see that `PowerSeries = MvPowerSeries Unit`, so use `.trans` refine (MvPowerSeries.coeff_mul _ φ ψ).trans ?_ rw [Finsupp.antidiagonal_single, Finset.sum_map] @@ -1619,10 +1631,10 @@ theorem coeff_mul_X_pow (p : R⟦X⟧) (n d : ℕ) : rw [coeff_X_pow, if_neg, mul_zero] rintro rfl apply h2 - rw [Finset.Nat.mem_antidiagonal, add_right_cancel_iff] at h1 + rw [mem_antidiagonal, add_right_cancel_iff] at h1 subst h1 rfl - · exact fun h1 => (h1 (Finset.Nat.mem_antidiagonal.2 rfl)).elim + · exact fun h1 => (h1 (mem_antidiagonal.2 rfl)).elim set_option linter.uppercaseLean3 false in #align power_series.coeff_mul_X_pow PowerSeries.coeff_mul_X_pow @@ -1634,11 +1646,11 @@ theorem coeff_X_pow_mul (p : R⟦X⟧) (n d : ℕ) : rw [coeff_X_pow, if_neg, zero_mul] rintro rfl apply h2 - rw [Finset.Nat.mem_antidiagonal, add_comm, add_right_cancel_iff] at h1 + rw [mem_antidiagonal, add_comm, add_right_cancel_iff] at h1 subst h1 rfl · rw [add_comm] - exact fun h1 => (h1 (Finset.Nat.mem_antidiagonal.2 rfl)).elim + exact fun h1 => (h1 (mem_antidiagonal.2 rfl)).elim set_option linter.uppercaseLean3 false in #align power_series.coeff_X_pow_mul PowerSeries.coeff_X_pow_mul @@ -1648,7 +1660,7 @@ theorem coeff_mul_X_pow' (p : R⟦X⟧) (n d : ℕ) : · rw [← tsub_add_cancel_of_le h, coeff_mul_X_pow, add_tsub_cancel_right] · refine' (coeff_mul _ _ _).trans (Finset.sum_eq_zero fun x hx => _) rw [coeff_X_pow, if_neg, mul_zero] - exact ((le_of_add_le_right (Finset.Nat.mem_antidiagonal.mp hx).le).trans_lt <| not_le.mp h).ne + exact ((le_of_add_le_right (mem_antidiagonal.mp hx).le).trans_lt <| not_le.mp h).ne set_option linter.uppercaseLean3 false in #align power_series.coeff_mul_X_pow' PowerSeries.coeff_mul_X_pow' @@ -1659,7 +1671,7 @@ theorem coeff_X_pow_mul' (p : R⟦X⟧) (n d : ℕ) : simp · refine' (coeff_mul _ _ _).trans (Finset.sum_eq_zero fun x hx => _) rw [coeff_X_pow, if_neg, zero_mul] - have := Finset.Nat.mem_antidiagonal.mp hx + have := mem_antidiagonal.mp hx rw [add_comm] at this exact ((le_of_add_le_right this.le).trans_lt <| not_le.mp h).ne set_option linter.uppercaseLean3 false in @@ -1784,7 +1796,7 @@ noncomputable def rescale (a : R) : R⟦X⟧ →+* R⟦X⟧ where ext rw [PowerSeries.coeff_mul, PowerSeries.coeff_mk, PowerSeries.coeff_mul, Finset.mul_sum] apply sum_congr rfl - simp only [coeff_mk, Prod.forall, Nat.mem_antidiagonal] + simp only [coeff_mk, Prod.forall, mem_antidiagonal] intro b c H rw [← H, pow_add, mul_mul_mul_comm] #align power_series.rescale PowerSeries.rescale @@ -1955,7 +1967,7 @@ theorem coeff_inv_aux (n : ℕ) (a : R) (φ : R⟦X⟧) : if n = 0 then a else -a * - ∑ x in Finset.Nat.antidiagonal n, + ∑ x in antidiagonal n, if x.2 < n then coeff R x.1 φ * coeff R x.2 (inv.aux a φ) else 0 := by -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [coeff, inv.aux, MvPowerSeries.coeff_inv_aux] @@ -1965,8 +1977,8 @@ theorem coeff_inv_aux (n : ℕ) (a : R) (φ : R⟦X⟧) : symm apply Finset.sum_bij fun (p : ℕ × ℕ) _h => (single () p.1, single () p.2) · rintro ⟨i, j⟩ hij - rw [Finset.Nat.mem_antidiagonal] at hij - rw [Finsupp.mem_antidiagonal, ← Finsupp.single_add, hij] + rw [mem_antidiagonal] at hij + rw [mem_antidiagonal, ← Finsupp.single_add, hij] · rintro ⟨i, j⟩ _hij by_cases H : j < n · rw [if_pos H, if_pos] @@ -1987,8 +1999,8 @@ theorem coeff_inv_aux (n : ℕ) (a : R) (φ : R⟦X⟧) : simpa only [Prod.mk.inj_iff, Finsupp.unique_single_eq_iff] using id · rintro ⟨f, g⟩ hfg refine' ⟨(f (), g ()), _, _⟩ - · rw [Finsupp.mem_antidiagonal] at hfg - rw [Finset.Nat.mem_antidiagonal, ← Finsupp.add_apply, hfg, Finsupp.single_eq_same] + · rw [mem_antidiagonal] at hfg + rw [mem_antidiagonal, ← Finsupp.add_apply, hfg, Finsupp.single_eq_same] · rw [Prod.mk.inj_iff] dsimp exact ⟨Finsupp.unique_single f, Finsupp.unique_single g⟩ @@ -2004,7 +2016,7 @@ theorem coeff_invOfUnit (n : ℕ) (φ : R⟦X⟧) (u : Rˣ) : if n = 0 then ↑u⁻¹ else -↑u⁻¹ * - ∑ x in Finset.Nat.antidiagonal n, + ∑ x in antidiagonal n, if x.2 < n then coeff R x.1 φ * coeff R x.2 (invOfUnit φ u) else 0 := coeff_inv_aux n (↑u⁻¹ : R) φ #align power_series.coeff_inv_of_unit PowerSeries.coeff_invOfUnit @@ -2095,7 +2107,7 @@ theorem eq_zero_or_eq_zero_of_mul_eq_zero [NoZeroDivisors R] (φ ψ : R⟦X⟧) · specialize hm₂ _ hi push_neg at hm₂ rw [hm₂, zero_mul] - rw [Finset.Nat.mem_antidiagonal] at hij + rw [mem_antidiagonal] at hij push_neg at hi hj suffices m < i by have : m + n < i + j := add_lt_add_of_lt_of_le this hj @@ -2106,7 +2118,7 @@ theorem eq_zero_or_eq_zero_of_mul_eq_zero [NoZeroDivisors R] (φ ψ : R⟦X⟧) simpa [Ne.def, Prod.mk.inj_iff] using (add_right_inj m).mp hij · contrapose! intro - rw [Finset.Nat.mem_antidiagonal] + rw [mem_antidiagonal] #align power_series.eq_zero_or_eq_zero_of_mul_eq_zero PowerSeries.eq_zero_or_eq_zero_of_mul_eq_zero instance [NoZeroDivisors R] : NoZeroDivisors R⟦X⟧ where @@ -2210,7 +2222,7 @@ theorem coeff_inv (n) (φ : PowerSeries k) : if n = 0 then (constantCoeff k φ)⁻¹ else -(constantCoeff k φ)⁻¹ * - ∑ x in Finset.Nat.antidiagonal n, + ∑ x in antidiagonal n, if x.2 < n then coeff k x.1 φ * coeff k x.2 φ⁻¹ else 0 := by rw [inv_eq_inv_aux, coeff_inv_aux n (constantCoeff k φ)⁻¹ φ] #align power_series.coeff_inv PowerSeries.coeff_inv @@ -2472,7 +2484,7 @@ theorem order_mul_ge (φ ψ : R⟦X⟧) : order φ + order ψ ≤ order (φ * ψ · rw [coeff_of_lt_order i hi, zero_mul] by_cases hj : ↑j < order ψ · rw [coeff_of_lt_order j hj, mul_zero] - rw [not_lt] at hi hj; rw [Finset.Nat.mem_antidiagonal] at hij + rw [not_lt] at hi hj; rw [mem_antidiagonal] at hij exfalso apply ne_of_lt (lt_of_lt_of_le hn <| add_le_add hi hj) rw [← Nat.cast_add, hij] @@ -2502,13 +2514,13 @@ theorem order_monomial_of_ne_zero (n : ℕ) (a : R) (h : a ≠ 0) : order (monom with any other power series is `0`. -/ theorem coeff_mul_of_lt_order {φ ψ : R⟦X⟧} {n : ℕ} (h : ↑n < ψ.order) : coeff R n (φ * ψ) = 0 := by - suffices : coeff R n (φ * ψ) = ∑ p in Finset.Nat.antidiagonal n, 0 + suffices : coeff R n (φ * ψ) = ∑ p in antidiagonal n, 0 rw [this, Finset.sum_const_zero] rw [coeff_mul] apply Finset.sum_congr rfl intro x hx refine' mul_eq_zero_of_right (coeff R x.fst φ) (coeff_of_lt_order x.snd (lt_of_le_of_lt _ h)) - rw [Finset.Nat.mem_antidiagonal] at hx + rw [mem_antidiagonal] at hx norm_cast linarith #align power_series.coeff_mul_of_lt_order PowerSeries.coeff_mul_of_lt_order @@ -2535,7 +2547,8 @@ theorem X_pow_order_dvd (h : (order φ).Dom) : X ^ (order φ).get h ∣ φ := by refine' ⟨PowerSeries.mk fun n => coeff R (n + (order φ).get h) φ, _⟩ ext n simp only [coeff_mul, coeff_X_pow, coeff_mk, boole_mul, Finset.sum_ite, - Finset.Nat.filter_fst_eq_antidiagonal, Finset.sum_const_zero, add_zero] + Finset.sum_const_zero, add_zero] + rw [Finset.filter_fst_eq_antidiagonal n (Part.get (order φ) h)] split_ifs with hn · simp [tsub_add_cancel_of_le hn] · simp only [Finset.sum_empty] diff --git a/Mathlib/RingTheory/PowerSeries/Derivative.lean b/Mathlib/RingTheory/PowerSeries/Derivative.lean new file mode 100644 index 0000000000000..466173a395ef2 --- /dev/null +++ b/Mathlib/RingTheory/PowerSeries/Derivative.lean @@ -0,0 +1,173 @@ +/- +Copyright (c) 2023 Richard M. Hill. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Richard M. Hill +-/ +import Mathlib.RingTheory.PowerSeries.Basic +import Mathlib.RingTheory.Derivation.Basic + +/-! +# Definitions + +In this file we define an operation `derivative` (formal differentiation) +on the ring of formal power series in one variable (over an arbitrary commutative semiring). + +Under suitable assumptions, we prove that two power series are equal if their derivatives +are equal and their constant terms are equal. This will give us a simple tool for proving +power series identities. For example, one can easily prove the power series identity +$\exp ( \log (1+X)) = 1+X$ by differentiating twice. + +## Main Definition + +- `PowerSeries.derivative R : Derivation R R⟦X⟧ R⟦X⟧` the formal derivative operation. + This is abbreviated `d⁄dX R`. +-/ + +namespace PowerSeries + +open Polynomial Derivation Nat + +section CommutativeSemiring +variable {R} [CommSemiring R] + +/-- +The formal derivative of a power series in one variable. +This is defined here as a function, but will be packaged as a +derivation `derivative` on `R⟦X⟧`. +-/ +noncomputable def derivativeFun (f : R⟦X⟧) : R⟦X⟧ := mk λ n ↦ coeff R (n + 1) f * (n + 1) + +theorem coeff_derivativeFun (f : R⟦X⟧) (n : ℕ) : + coeff R n f.derivativeFun = coeff R (n + 1) f * (n + 1) := by + rw [derivativeFun, coeff_mk] + +theorem derivativeFun_coe (f : R[X]) : (f : R⟦X⟧).derivativeFun = derivative f := by + ext + rw [coeff_derivativeFun, coeff_coe, coeff_coe, coeff_derivative] + +theorem derivativeFun_add (f g : R⟦X⟧) : + derivativeFun (f + g) = derivativeFun f + derivativeFun g := by + ext + rw [coeff_derivativeFun, map_add, map_add, coeff_derivativeFun, + coeff_derivativeFun, add_mul] + +theorem derivativeFun_C (r : R) : derivativeFun (C R r) = 0 := by + ext n + rw [coeff_derivativeFun, coeff_succ_C, zero_mul, map_zero] + +theorem trunc_derivativeFun (f : R⟦X⟧) (n : ℕ) : + trunc n f.derivativeFun = derivative (trunc (n + 1) f):= by + ext d + rw [coeff_trunc] + split_ifs with h + · have : d + 1 < n + 1 := succ_lt_succ_iff.2 h + rw [coeff_derivativeFun, coeff_derivative, coeff_trunc, if_pos this] + · have : ¬d + 1 < n + 1 := by rwa [succ_lt_succ_iff] + rw [coeff_derivative, coeff_trunc, if_neg this, zero_mul] + +--A special case of `derivativeFun_mul`, used in its proof. +private theorem derivativeFun_coe_mul_coe (f g : R[X]) : derivativeFun (f * g : R⟦X⟧) = + f * derivative g + g * derivative f := by + rw [←coe_mul, derivativeFun_coe, derivative_mul, + add_comm, mul_comm _ g, ←coe_mul, ←coe_mul, Polynomial.coe_add] + +/-- Leibniz rule for formal power series.-/ +theorem derivativeFun_mul (f g : R⟦X⟧) : + derivativeFun (f * g) = f • g.derivativeFun + g • f.derivativeFun := by + ext n + have h₁ : n < n + 1 := lt_succ_self n + have h₂ : n < n + 1 + 1 := Nat.lt_add_right _ _ _ h₁ + rw [coeff_derivativeFun, map_add, coeff_mul_eq_coeff_trunc_mul_trunc _ _ (lt_succ_self _), + smul_eq_mul, smul_eq_mul, coeff_mul_eq_coeff_trunc_mul_trunc₂ g f.derivativeFun h₂ h₁, + coeff_mul_eq_coeff_trunc_mul_trunc₂ f g.derivativeFun h₂ h₁, trunc_derivativeFun, + trunc_derivativeFun, ←map_add, ←derivativeFun_coe_mul_coe, coeff_derivativeFun] + +theorem derivativeFun_one : derivativeFun (1 : R⟦X⟧) = 0 := by + rw [←map_one (C R), derivativeFun_C (1 : R)] + +theorem derivativeFun_smul (r : R) (f : R⟦X⟧) : derivativeFun (r • f) = r • derivativeFun f := by + rw [smul_eq_C_mul, smul_eq_C_mul, derivativeFun_mul, derivativeFun_C, smul_zero, add_zero, + smul_eq_mul] + +variable (R) +/--The formal derivative of a formal power series.-/ +noncomputable def derivative : Derivation R R⟦X⟧ R⟦X⟧ +where + toFun := derivativeFun + map_add' := derivativeFun_add + map_smul' := derivativeFun_smul + map_one_eq_zero' := derivativeFun_one + leibniz' := derivativeFun_mul +/--Abbreviation of `PowerSeries.derivative`, the formal derivative on `R⟦X⟧`.-/ +scoped notation "d⁄dX" => derivative + +variable {R} + +@[simp] theorem derivative_C (r : R) : d⁄dX R (C R r) = 0 := derivativeFun_C r + +theorem coeff_derivative (f : R⟦X⟧) (n : ℕ) : + coeff R n (d⁄dX R f) = coeff R (n + 1) f * (n + 1) := coeff_derivativeFun f n + +theorem derivative_coe (f : R[X]) : d⁄dX R f = Polynomial.derivative f := derivativeFun_coe f + +@[simp] theorem derivative_X : d⁄dX R (X : R⟦X⟧) = 1 := by + ext + rw [coeff_derivative, coeff_one, coeff_X, boole_mul] + simp_rw [add_left_eq_self] + split_ifs with h + · rw [h, cast_zero, zero_add] + · rfl + +theorem trunc_derivative (f : R⟦X⟧) (n : ℕ) : + trunc n (d⁄dX R f) = Polynomial.derivative (trunc (n + 1) f) := by apply trunc_derivativeFun + +theorem trunc_derivative' (f : R⟦X⟧) (n : ℕ) : + trunc (n-1) (d⁄dX R f) = Polynomial.derivative (trunc n f) := by + cases n with + | zero => + simp only [zero_eq, ge_iff_le, tsub_eq_zero_of_le] + rfl + | succ n => + rw [succ_sub_one, trunc_derivative] + +end CommutativeSemiring + +/-In the next lemma, we use `smul_right_inj`, which requires not only `no_smul_divisors ℕ R`, but +also cancellation of addition in `R`. For this reason, the next lemma is stated in the case that `R` +is a `CommRing`.-/ + +/-- If `f` and `g` have the same constant term and derivative, then they are equal.-/ +theorem derivative.ext {R} [CommRing R] [NoZeroSMulDivisors ℕ R] {f g} (hD : d⁄dX R f = d⁄dX R g) + (hc : constantCoeff R f = constantCoeff R g) : f = g := by + ext n + cases n with + | zero => + rw [coeff_zero_eq_constantCoeff, hc] + | succ n => + have equ : coeff R n (d⁄dX R f) = coeff R n (d⁄dX R g) := by rw [hD] + rwa [coeff_derivative, coeff_derivative, ←cast_succ, mul_comm, ←nsmul_eq_mul, + mul_comm, ←nsmul_eq_mul, smul_right_inj] at equ + exact n.succ_ne_zero + +@[simp] theorem derivative_inv {R} [CommRing R] (f : R⟦X⟧ˣ) : + d⁄dX R ↑f⁻¹ = -(↑f⁻¹ : R⟦X⟧) ^ 2 * d⁄dX R f := by + apply Derivation.leibniz_of_mul_eq_one + simp + +@[simp] theorem derivative_invOf {R} [CommRing R] (f : R⟦X⟧) [Invertible f] : + d⁄dX R ⅟f = - ⅟f ^ 2 * d⁄dX R f := by + rw [Derivation.leibniz_invOf, smul_eq_mul] + +/- +The following theorem is stated only in the case that `R` is a field. This is because +there is currently no instance of `Inv R⟦X⟧` for more general base rings `R`. +-/ +@[simp] theorem derivative_inv' {R} [Field R] (f : R⟦X⟧) : d⁄dX R f⁻¹ = -f⁻¹ ^ 2 * d⁄dX R f := by + by_cases constantCoeff R f = 0 + · suffices : f⁻¹ = 0 + rw [this, pow_two, zero_mul, neg_zero, zero_mul, map_zero] + rwa [MvPowerSeries.inv_eq_zero] + apply Derivation.leibniz_of_mul_eq_one + exact PowerSeries.inv_mul_cancel (h := h) + +end PowerSeries diff --git a/Mathlib/RingTheory/ReesAlgebra.lean b/Mathlib/RingTheory/ReesAlgebra.lean index 298b9f6b6dd13..d359e175e4d7f 100644 --- a/Mathlib/RingTheory/ReesAlgebra.lean +++ b/Mathlib/RingTheory/ReesAlgebra.lean @@ -41,7 +41,7 @@ def reesAlgebra : Subalgebra R R[X] where rw [coeff_mul] apply Ideal.sum_mem rintro ⟨j, k⟩ e - rw [← Finset.Nat.mem_antidiagonal.mp e, pow_add] + rw [← Finset.mem_antidiagonal.mp e, pow_add] exact Ideal.mul_mem_mul (hf j) (hg k) one_mem' i := by rw [coeff_one] diff --git a/Mathlib/RingTheory/TensorProduct.lean b/Mathlib/RingTheory/TensorProduct.lean index d9c27f58d5376..f840e76b39925 100644 --- a/Mathlib/RingTheory/TensorProduct.lean +++ b/Mathlib/RingTheory/TensorProduct.lean @@ -733,10 +733,15 @@ protected nonrec def lid : R ⊗[R] A ≃ₐ[R] A := @[simp] theorem lid_toLinearEquiv : (TensorProduct.lid R A).toLinearEquiv = _root_.TensorProduct.lid R A := rfl +variable {R} {A} in @[simp] -theorem lid_tmul (r : R) (a : A) : (TensorProduct.lid R A : R ⊗ A → A) (r ⊗ₜ a) = r • a := rfl +theorem lid_tmul (r : R) (a : A) : TensorProduct.lid R A (r ⊗ₜ a) = r • a := rfl #align algebra.tensor_product.lid_tmul Algebra.TensorProduct.lid_tmul +variable {A} in +@[simp] +theorem lid_symm_apply (a : A) : (TensorProduct.lid R A).symm a = 1 ⊗ₜ a := rfl + variable (S) /-- The base ring is a right identity for the tensor product of algebra, up to algebra isomorphism. @@ -757,6 +762,9 @@ variable {R A} in theorem rid_tmul (r : R) (a : A) : TensorProduct.rid R S A (a ⊗ₜ r) = r • a := rfl #align algebra.tensor_product.rid_tmul Algebra.TensorProduct.rid_tmul +variable {A} in +@[simp] +theorem rid_symm_apply (a : A) : (TensorProduct.rid R S A).symm a = a ⊗ₜ 1 := rfl section @@ -771,12 +779,23 @@ protected def comm : A ⊗[R] B ≃ₐ[R] B ⊗[R] A := @[simp] theorem comm_toLinearEquiv : (Algebra.TensorProduct.comm R A B).toLinearEquiv = _root_.TensorProduct.comm R A B := rfl +variable {A B} in @[simp] theorem comm_tmul (a : A) (b : B) : - (TensorProduct.comm R A B : A ⊗[R] B → B ⊗[R] A) (a ⊗ₜ b) = b ⊗ₜ a := + TensorProduct.comm R A B (a ⊗ₜ b) = b ⊗ₜ a := rfl #align algebra.tensor_product.comm_tmul Algebra.TensorProduct.comm_tmul +variable {A B} in +@[simp] +theorem comm_symm_tmul (a : A) (b : B) : + (TensorProduct.comm R A B).symm (b ⊗ₜ a) = a ⊗ₜ b := + rfl + +theorem comm_symm : + (TensorProduct.comm R A B).symm = TensorProduct.comm R B A := by + ext; rfl + theorem adjoin_tmul_eq_top : adjoin R { t : A ⊗[R] B | ∃ a b, a ⊗ₜ[R] b = t } = ⊤ := top_le_iff.mp <| (top_le_iff.mpr <| span_tmul_eq_top R A B).trans (span_le_adjoin R _) #align algebra.tensor_product.adjoin_tmul_eq_top Algebra.TensorProduct.adjoin_tmul_eq_top @@ -816,10 +835,15 @@ variable {A B C} @[simp] theorem assoc_tmul (a : A) (b : B) (c : C) : - Algebra.TensorProduct.assoc R A B C (a ⊗ₜ b ⊗ₜ c) = a ⊗ₜ (b ⊗ₜ c) := + Algebra.TensorProduct.assoc R A B C ((a ⊗ₜ b) ⊗ₜ c) = a ⊗ₜ (b ⊗ₜ c) := rfl #align algebra.tensor_product.assoc_tmul Algebra.TensorProduct.assoc_tmul +@[simp] +theorem assoc_symm_tmul (a : A) (b : B) (c : C) : + (Algebra.TensorProduct.assoc R A B C).symm (a ⊗ₜ (b ⊗ₜ c)) = (a ⊗ₜ b) ⊗ₜ c := + rfl + end variable {R S A} diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index 10cf5f79d5593..3817273d05bed 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -280,8 +280,7 @@ theorem WfDvdMonoid.of_exists_prime_factors : WfDvdMonoid α := classical refine' RelHomClass.wellFounded - (RelHom.mk _ _ : (DvdNotUnit : α → α → Prop) →r ((· < ·) : ℕ∞ → ℕ∞ → Prop)) - (WithTop.wellFounded_lt Nat.lt_wfRel.wf) + (RelHom.mk _ _ : (DvdNotUnit : α → α → Prop) →r ((· < ·) : ℕ∞ → ℕ∞ → Prop)) wellFounded_lt · intro a by_cases h : a = 0 · exact ⊤ diff --git a/Mathlib/RingTheory/WittVector/IsPoly.lean b/Mathlib/RingTheory/WittVector/IsPoly.lean index 0d70990e0a39a..b106c5583e7ac 100644 --- a/Mathlib/RingTheory/WittVector/IsPoly.lean +++ b/Mathlib/RingTheory/WittVector/IsPoly.lean @@ -328,8 +328,8 @@ theorem bind₁_onePoly_wittPolynomial [hp : Fact p.Prime] (n : ℕ) : AlgHom.map_pow, bind₁_X_right, AlgHom.map_mul] · rw [Finset.mem_range] -- porting note: was `decide` - simp only [add_pos_iff, or_true, not_true, pow_zero, map_one, ge_iff_le, nonpos_iff_eq_zero, - tsub_zero, one_mul, gt_iff_lt, IsEmpty.forall_iff] + intro h + simp only [add_pos_iff, zero_lt_one, or_true, not_true_eq_false] at h #align witt_vector.bind₁_one_poly_witt_polynomial WittVector.bind₁_onePoly_wittPolynomial /-- The function that is constantly one on Witt vectors is a polynomial function. -/ diff --git a/Mathlib/RingTheory/WittVector/Isocrystal.lean b/Mathlib/RingTheory/WittVector/Isocrystal.lean index a66974445c4db..06dc661aff52f 100644 --- a/Mathlib/RingTheory/WittVector/Isocrystal.lean +++ b/Mathlib/RingTheory/WittVector/Isocrystal.lean @@ -197,11 +197,7 @@ instance (m : ℤ) : Isocrystal p k (StandardOneDimIsocrystal p k m) where @[simp] theorem StandardOneDimIsocrystal.frobenius_apply (m : ℤ) (x : StandardOneDimIsocrystal p k m) : - Φ(p, k) x = (p : K(p, k)) ^ m • φ(p, k) x := by - -- Porting note: was just `rfl` - erw [smul_eq_mul] - simp only [map_zpow₀, map_natCast] - rfl + Φ(p, k) x = (p : K(p, k)) ^ m • φ(p, k) x := rfl #align witt_vector.standard_one_dim_isocrystal.frobenius_apply WittVector.StandardOneDimIsocrystal.frobenius_apply end PerfectRing diff --git a/Mathlib/RingTheory/WittVector/MulCoeff.lean b/Mathlib/RingTheory/WittVector/MulCoeff.lean index ab6e5b6171a79..48f1d191b3a68 100644 --- a/Mathlib/RingTheory/WittVector/MulCoeff.lean +++ b/Mathlib/RingTheory/WittVector/MulCoeff.lean @@ -243,7 +243,8 @@ theorem peval_polyOfInterest' (n : ℕ) (x y : 𝕎 k) : x.coeff (n + 1) * y.coeff 0 ^ p ^ (n + 1) := by rw [peval_polyOfInterest] have : (p : k) = 0 := CharP.cast_eq_zero k p - simp only [this, Nat.cast_pow, ne_eq, add_eq_zero, and_false, zero_pow', zero_mul, add_zero] + simp only [this, Nat.cast_pow, ne_eq, add_eq_zero, and_false, zero_pow', zero_mul, add_zero, + not_false_eq_true] have sum_zero_pow_mul_pow_p : ∀ y : 𝕎 k, ∑ x : ℕ in range (n + 1 + 1), (0 : k) ^ x * y.coeff x ^ p ^ (n + 1 - x) = y.coeff 0 ^ p ^ (n + 1) := by intro y diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index acf217ffe7bdc..bda7d41c768df 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1883,13 +1883,13 @@ def toPartENat : Cardinal →+ PartENat where rw [← Nat.cast_add, toNat_cast, Nat.cast_add] · simp_rw [if_neg hy, PartENat.add_top] contrapose! hy - simp only [ne_eq, ite_eq_right_iff, - PartENat.natCast_ne_top, not_forall, exists_prop, and_true] at hy + simp only [ne_eq, ite_eq_right_iff, PartENat.natCast_ne_top, not_forall, exists_prop, + and_true, not_false_eq_true] at hy exact le_add_self.trans_lt hy · simp_rw [if_neg hx, PartENat.top_add] contrapose! hx - simp only [ne_eq, ite_eq_right_iff, - PartENat.natCast_ne_top, not_forall, exists_prop, and_true] at hx + simp only [ne_eq, ite_eq_right_iff, PartENat.natCast_ne_top, not_forall, exists_prop, + and_true, not_false_eq_true] at hx exact le_self_add.trans_lt hx #align cardinal.to_part_enat Cardinal.toPartENat diff --git a/Mathlib/SetTheory/Cardinal/Finite.lean b/Mathlib/SetTheory/Cardinal/Finite.lean index af658391b6c64..4ae92568aecc7 100644 --- a/Mathlib/SetTheory/Cardinal/Finite.lean +++ b/Mathlib/SetTheory/Cardinal/Finite.lean @@ -69,7 +69,7 @@ an equivalence between `α` and `Fin (Nat.card α)`. See also `Finite.equivFin`. def equivFinOfCardPos {α : Type*} (h : Nat.card α ≠ 0) : α ≃ Fin (Nat.card α) := by cases fintypeOrInfinite α · simpa only [card_eq_fintype_card] using Fintype.equivFin α - · simp only [card_eq_zero_of_infinite, ne_eq] at h + · simp only [card_eq_zero_of_infinite, ne_eq, not_true_eq_false] at h #align nat.equiv_fin_of_card_pos Nat.equivFinOfCardPos theorem card_of_subsingleton (a : α) [Subsingleton α] : Nat.card α = 1 := by diff --git a/Mathlib/SetTheory/Lists.lean b/Mathlib/SetTheory/Lists.lean index 9306d7ecadcd2..3b811c7d62078 100644 --- a/Mathlib/SetTheory/Lists.lean +++ b/Mathlib/SetTheory/Lists.lean @@ -369,7 +369,7 @@ def Equiv.decidableMeas : theorem sizeof_pos {b} (l : Lists' α b) : 0 < SizeOf.sizeOf l := by cases l <;> simp only [Lists'.atom.sizeOf_spec, Lists'.nil.sizeOf_spec, Lists'.cons'.sizeOf_spec, - true_or, add_pos_iff] + true_or, add_pos_iff, zero_lt_one] #align lists.sizeof_pos Lists.sizeof_pos theorem lt_sizeof_cons' {b} (a : Lists' α b) (l) : diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 70e2b88d5aeaa..a6836d401410e 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -1003,7 +1003,8 @@ theorem le_of_dvd : ∀ {a b : Ordinal}, b ≠ 0 → a ∣ b → a ≤ b -- Porting note: `Ne` is required. simpa only [mul_one] using mul_le_mul_left' - (one_le_iff_ne_zero.2 fun h : b = 0 => by simp only [h, mul_zero, Ne] at b0) a + (one_le_iff_ne_zero.2 fun h : b = 0 => by + simp only [h, mul_zero, Ne, not_true_eq_false] at b0) a #align ordinal.le_of_dvd Ordinal.le_of_dvd theorem dvd_antisymm {a b : Ordinal} (h₁ : a ∣ b) (h₂ : b ∣ a) : a = b := @@ -2584,12 +2585,11 @@ theorem rank_lt_of_rel (h : r a b) : hwf.rank a < hwf.rank b := #align well_founded.rank_lt_of_rel WellFounded.rank_lt_of_rel theorem rank_strictMono [Preorder α] [WellFoundedLT α] : - StrictMono (rank <| @IsWellFounded.wf α (· < ·) _) := fun _ _ => rank_lt_of_rel _ + StrictMono (rank <| @wellFounded_lt α _ _) := fun _ _ => rank_lt_of_rel _ #align well_founded.rank_strict_mono WellFounded.rank_strictMono theorem rank_strictAnti [Preorder α] [WellFoundedGT α] : - StrictAnti (rank <| @IsWellFounded.wf α (· > ·) _) := fun _ _ => - rank_lt_of_rel <| @IsWellFounded.wf α (· > ·) _ + StrictAnti (rank <| @wellFounded_gt α _ _) := fun _ _ => rank_lt_of_rel wellFounded_gt #align well_founded.rank_strict_anti WellFounded.rank_strictAnti end WellFounded diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 85add9657020c..c7b105ac6ada0 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -167,7 +167,8 @@ instance isWellOrder_out_lt (o : Ordinal) : IsWellOrder o.out.α (· < ·) := namespace Ordinal --- ### Basic properties of the order type +/-! ### Basic properties of the order type -/ + /-- The order type of a well order is an ordinal. -/ def type (r : α → α → Prop) [wo : IsWellOrder α r] : Ordinal := ⟦⟨α, r, wo⟩⟧ @@ -1009,7 +1010,8 @@ theorem sInf_empty : sInf (∅ : Set Ordinal) = 0 := dif_neg Set.not_nonempty_empty #align ordinal.Inf_empty Ordinal.sInf_empty --- ### Successor order properties +/-! ### Successor order properties -/ + private theorem succ_le_iff' {a b : Ordinal} : a + 1 ≤ b ↔ a < b := ⟨lt_of_lt_of_le (inductionOn a fun α r _ => @@ -1584,3 +1586,22 @@ theorem type_fin (n : ℕ) : @type (Fin n) (· < ·) _ = n := by simp #align ordinal.type_fin Ordinal.type_fin end Ordinal + +/-! ### Sorted lists -/ + +theorem List.Sorted.lt_ord_of_lt [LinearOrder α] [IsWellOrder α (· < ·)] {l m : List α} + {o : Ordinal} (hl : l.Sorted (· > ·)) (hm : m.Sorted (· > ·)) (hmltl : m < l) + (hlt : ∀ i ∈ l, Ordinal.typein (· < ·) i < o) : ∀ i ∈ m, Ordinal.typein (· < ·) i < o := by + replace hmltl : List.Lex (· < ·) m l := hmltl + cases l with + | nil => simp at hmltl + | cons a as => + cases m with + | nil => intro i hi; simp at hi + | cons b bs => + intro i hi + suffices h : i ≤ a by refine lt_of_le_of_lt ?_ (hlt a (mem_cons_self a as)); simpa + cases hi with + | head as => exact List.head_le_of_lt hmltl + | tail b hi => exact le_of_lt (lt_of_lt_of_le (List.rel_of_sorted_cons hm _ hi) + (List.head_le_of_lt hmltl)) diff --git a/Mathlib/SetTheory/Ordinal/CantorNormalForm.lean b/Mathlib/SetTheory/Ordinal/CantorNormalForm.lean index 0ce02e0b9ead4..bed9105c56c5b 100644 --- a/Mathlib/SetTheory/Ordinal/CantorNormalForm.lean +++ b/Mathlib/SetTheory/Ordinal/CantorNormalForm.lean @@ -162,11 +162,11 @@ set_option linter.uppercaseLean3 false in /-- The exponents of the Cantor normal form are decreasing. -/ theorem CNF_sorted (b o : Ordinal) : ((CNF b o).map Prod.fst).Sorted (· > ·) := by refine' CNFRec b _ (fun o ho IH ↦ _) o - · simp only [CNF_zero] + · simp only [gt_iff_lt, CNF_zero, map_nil, sorted_nil] · cases' le_or_lt b 1 with hb hb - · simp only [CNF_of_le_one hb ho, map] + · simp only [CNF_of_le_one hb ho, gt_iff_lt, map_cons, map, sorted_singleton] · cases' lt_or_le o b with hob hbo - · simp only [CNF_of_lt ho hob, map] + · simp only [CNF_of_lt ho hob, gt_iff_lt, map_cons, map, sorted_singleton] · rw [CNF_ne_zero ho, map_cons, sorted_cons] refine' ⟨fun a H ↦ _, IH⟩ rw [mem_map] at H diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 0ce5aceecc77e..a1600f206cc20 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -189,7 +189,8 @@ theorem eq_of_cmp_eq : ∀ {o₁ o₂}, cmp o₁ o₂ = Ordering.eq → o₁ = o #align onote.eq_of_cmp_eq ONote.eq_of_cmp_eq protected theorem zero_lt_one : (0 : ONote) < 1 := by - simp only [lt_def, repr, repr_one, opow_zero, one_mul, add_zero, nat_cast_pos] + simp only [lt_def, repr, opow_zero, Nat.succPNat_coe, Nat.cast_one, mul_one, add_zero, + zero_lt_one] #align onote.zero_lt_one ONote.zero_lt_one /-- `NFBelow o b` says that `o` is a normal form ordinal notation @@ -830,9 +831,10 @@ instance nf_opow (o₁ o₂) [NF o₁] [NF o₂] : NF (o₁ ^ o₂) := by haveI := (nf_repr_split' e₂).1 cases' a with a0 n a' · cases' m with m - · by_cases o₂ = 0 <;> simp [(· ^ ·), Pow.pow, pow, opow, opowAux2, *] + · by_cases o₂ = 0 <;> simp only [(· ^ ·), Pow.pow, pow, opow, opowAux2, *] <;> decide · by_cases m = 0 · simp only [(· ^ ·), Pow.pow, pow, opow, opowAux2, *, zero_def] + decide · simp only [(· ^ ·), Pow.pow, pow, opow, opowAux2, mulNat_eq_mul, ofNat, *] infer_instance · simp [(· ^ ·),Pow.pow,pow, opow, opowAux2, e₁, e₂, split_eq_scale_split' e₂] @@ -1098,6 +1100,7 @@ theorem fundamentalSequence_has_prop (o) : FundamentalSequenceProp o (fundamenta eq_self_iff_true, lt_add_iff_pos_right, lt_def, mul_one, Nat.cast_zero, Nat.cast_succ, Nat.succPNat_coe, opow_succ, opow_zero, mul_add_one, PNat.one_coe, succ_zero, true_and_iff, _root_.zero_add, zero_def] + · decide · exact ⟨rfl, inferInstance⟩ · have := opow_pos (repr a') omega_pos refine' diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index 5edcf6c1a2af4..5e50d410af28c 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -8,7 +8,6 @@ import Mathlib.Tactic.Attr.Register import Mathlib.Tactic.Backtrack import Mathlib.Tactic.Basic import Mathlib.Tactic.ByContra -import Mathlib.Tactic.Cache import Mathlib.Tactic.CancelDenoms import Mathlib.Tactic.CancelDenoms.Core import Mathlib.Tactic.Cases @@ -68,7 +67,6 @@ import Mathlib.Tactic.InferParam import Mathlib.Tactic.Inhabit import Mathlib.Tactic.IntervalCases import Mathlib.Tactic.IrreducibleDef -import Mathlib.Tactic.LeftRight import Mathlib.Tactic.LibrarySearch import Mathlib.Tactic.Lift import Mathlib.Tactic.LiftLets @@ -115,12 +113,12 @@ import Mathlib.Tactic.NormNum.Result import Mathlib.Tactic.NthRewrite import Mathlib.Tactic.Observe import Mathlib.Tactic.PPWithUniv +import Mathlib.Tactic.Peel import Mathlib.Tactic.PermuteGoals import Mathlib.Tactic.Polyrith import Mathlib.Tactic.Positivity import Mathlib.Tactic.Positivity.Basic import Mathlib.Tactic.Positivity.Core -import Mathlib.Tactic.PrintPrefix import Mathlib.Tactic.ProjectionNotation import Mathlib.Tactic.Propose import Mathlib.Tactic.ProxyType @@ -135,9 +133,11 @@ import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.Rename import Mathlib.Tactic.RenameBVar import Mathlib.Tactic.Replace +import Mathlib.Tactic.RewriteSearch import Mathlib.Tactic.Rewrites import Mathlib.Tactic.Ring import Mathlib.Tactic.Ring.Basic +import Mathlib.Tactic.Ring.PNat import Mathlib.Tactic.Ring.RingNF import Mathlib.Tactic.RunCmd import Mathlib.Tactic.Sat.FromLRAT diff --git a/Mathlib/Tactic/ApplyFun.lean b/Mathlib/Tactic/ApplyFun.lean index 070658d23e7e2..3aaf64ae18b67 100644 --- a/Mathlib/Tactic/ApplyFun.lean +++ b/Mathlib/Tactic/ApplyFun.lean @@ -105,7 +105,8 @@ def maybeProveInjective (ginj : Expr) (using? : Option Expr) : MetaM Bool := do let err ← mkHasTypeButIsExpectedMsg (← inferType u) (← inferType ginj) throwError "Using clause {err}" -- Try an assumption - try ginj.mvarId!.assumption; return true catch _ => pure () + if ←ginj.mvarId!.assumptionCore then + return true -- Try using that this is an equivalence -- Note: if `f` is itself a metavariable, this can cause it to become an equivalence; -- perhaps making sure `f` is an equivalence would be correct, but maybe users diff --git a/Mathlib/Tactic/Backtrack.lean b/Mathlib/Tactic/Backtrack.lean index e2abf294bbc5d..fe9c907b6219f 100644 --- a/Mathlib/Tactic/Backtrack.lean +++ b/Mathlib/Tactic/Backtrack.lean @@ -43,19 +43,6 @@ def List.tryAllM [Monad m] [Alternative m] (L : List α) (f : α → m β) : m ( return (R.filterMap (fun s => match s with | .inl a => a | _ => none), R.filterMap (fun s => match s with | .inr b => b | _ => none)) -namespace Lean.MVarId - -/-- -Given any tactic that takes a goal, and returns a sequence of alternative outcomes -(each outcome consisting of a list of new subgoals), -we can perform backtracking search by repeatedly applying the tactic. --/ -def firstContinuation (results : MVarId → Nondet MetaM (List MVarId)) - (cont : List MVarId → MetaM α) (g : MVarId) : MetaM α := do - (results g).firstM fun r => observing? do cont r - -end Lean.MVarId - namespace Mathlib.Tactic /-- @@ -70,38 +57,39 @@ structure BacktrackConfig where /-- Maximum recursion depth. -/ maxDepth : Nat := 6 /-- An arbitrary procedure which can be used to modify the list of goals - before each attempt to apply a lemma. + before each attempt to generate alternatives. Called as `proc goals curr`, where `goals` are the original goals for `backtracking`, and `curr` are the current goals. Returning `some l` will replace the current goals with `l` and recurse (consuming one step of maximum depth). - Returning `none` will proceed to applying lemmas without changing goals. + Returning `none` will proceed to generating alternative without changing goals. Failure will cause backtracking. (defaults to `none`) -/ proc : List MVarId → List MVarId → MetaM (Option (List MVarId)) := fun _ _ => pure none - /-- If `suspend g`, then we do not attempt to apply any further lemmas, + /-- If `suspend g`, then we do not consider alternatives for `g`, but return `g` as a new subgoal. (defaults to `false`) -/ suspend : MVarId → MetaM Bool := fun _ => pure false - /-- `discharge g` is called on goals for which no lemmas apply. + /-- `discharge g` is called on goals for which there were no alternatives. If `none` we return `g` as a new subgoal. If `some l`, we replace `g` by `l` in the list of active goals, and recurse. If failure, we backtrack. (defaults to failure) -/ discharge : MVarId → MetaM (Option (List MVarId)) := fun _ => failure - /-- If we solve any "independent" goals, don't fail. -/ + /-- + If we solve any "independent" goals, don't fail. + See `Lean.MVarId.independent?` for the definition of independence. + -/ commitIndependentGoals : Bool := false -def ppMVarIds (gs : List MVarId) : MetaM (List Format) := do - gs.mapM fun g => do ppExpr (← g.getType) - /-- Attempts to solve the `goals`, by recursively calling `alternatives g` on each subgoal that appears. -`alternatives` returns a nondeterministic list of goals -(this is essentially a lazy list of `List MVarId`, -with the extra state required to backtrack in `MetaM`). +`alternatives` returns a nondeterministic list of new subgoals generated from a goal. `backtrack` performs a backtracking search, attempting to close all subgoals. Further flow control options are available via the `Config` argument. + +Returns a list of subgoals which were "suspended" via the `suspend` or `discharge` hooks +in `Config`. In the default configuration, `backtrack` will either return an empty list or fail. -/ partial def backtrack (cfg : BacktrackConfig := {}) (trace : Name := .anonymous) (alternatives : MVarId → Nondet MetaM (List MVarId)) @@ -152,9 +140,9 @@ where run (n+1) gs (g :: acc) else try - -- We attempt to find an expression which can be applied, - -- and for which all resulting sub-goals can be discharged using `run n`. - g.firstContinuation alternatives (fun res => run n (res ++ gs) acc) + -- We attempt to find an alternative, + -- for which all resulting sub-goals can be discharged using `run n`. + (alternatives g).firstM fun r => observing? do run n (r ++ gs) acc catch _ => -- No lemmas could be applied: match (← cfg.discharge g) with @@ -164,8 +152,10 @@ where | some l => (withTraceNode trace (fun _ => return m!"⏬ discharger generated new subgoals") do run n (l ++ gs) acc) - -- A wrapper around `run`, which works on "independent" goals separately first, - -- to reduce backtracking. + /-- + A wrapper around `run`, which works on "independent" goals separately first, + to reduce backtracking. + -/ processIndependentGoals (goals remaining : List MVarId) : MetaM (List MVarId) := do -- Partition the remaining goals into "independent" goals -- (which should be solvable without affecting the solvability of other goals) @@ -193,11 +183,16 @@ where return newSubgoals ++ failed ++ (← (processIndependentGoals goals' ogs <|> pure ogs)) else if !failed.isEmpty then -- If `commitIndependentGoals` is `false`, and we failed on any of the independent goals, - -- the overall failure is inevitable so we can stop here. + -- then overall failure is inevitable so we can stop here. failure else -- Finally, having solved this batch of independent goals, -- recurse (potentially now finding new independent goals). return newSubgoals ++ (← processIndependentGoals goals' ogs) + /-- + Pretty print a list of goals. + -/ + ppMVarIds (gs : List MVarId) : MetaM (List Format) := do + gs.mapM fun g => do ppExpr (← g.getType) end Mathlib.Tactic diff --git a/Mathlib/Tactic/Cache.lean b/Mathlib/Tactic/Cache.lean deleted file mode 100644 index 694382dd9063e..0000000000000 --- a/Mathlib/Tactic/Cache.lean +++ /dev/null @@ -1,182 +0,0 @@ -/- -Copyright (c) 2021 Gabriel Ebner. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Gabriel Ebner --/ -import Lean -import Mathlib.Lean.Meta.DiscrTree - -/-! -# Once-per-file cache for tactics - -This file defines cache data structures for tactics -that are initialized the first time they are accessed. -Since Lean 4 starts one process per file, -these caches are once-per-file -and can for example be used to cache information -about the imported modules. - -The `Cache α` data structure is -the most generic version we define. -It is created using `Cache.mk f` -where `f : MetaM α` performs -the initialization of the cache: -``` -initialize numberOfImports : Cache Nat ← Cache.mk do - (← getEnv).imports.size - --- (does not work in the same module where the cache is defined) -#eval show MetaM Nat from numberOfImports.get -``` - -The `DeclCache α` data structure computes -a fold over the environment's constants: -`DeclCache.mk empty f` constructs such a cache -where `empty : α` and `f : Name → ConstantInfo → α → MetaM α`. -The result of the constants in the imports is cached -between tactic invocations, -while for constants defined in the same file -`f` is evaluated again every time. -This kind of cache can be used e.g. -to populate discrimination trees. --/ - -set_option autoImplicit true - -open Lean Meta - -namespace Mathlib.Tactic - -/-- Once-per-file cache. -/ -def Cache (α : Type) := - IO.Ref <| Sum (MetaM α) <| - Task <| Except Exception α - -instance : Nonempty (Cache α) := - inferInstanceAs <| Nonempty (IO.Ref _) - -/-- Creates a cache with an initialization function. -/ -def Cache.mk (init : MetaM α) : IO (Cache α) := - IO.mkRef <| Sum.inl init - -/-- -Access the cache. -Calling this function for the first time -will initialize the cache with the function -provided in the constructor. --/ -def Cache.get [Monad m] [MonadEnv m] [MonadLog m] [MonadOptions m] [MonadLiftT BaseIO m] - [MonadExcept Exception m] (cache : Cache α) : m α := do - let t ← match ← show BaseIO _ from ST.Ref.get cache with - | Sum.inr t => pure t - | Sum.inl init => - let env ← getEnv - let fileName ← getFileName - let fileMap ← getFileMap - let options ← getOptions -- TODO: sanitize options? - -- Default heartbeats to a reasonable value. - -- otherwise librarySearch times out on mathlib - -- TODO: add customization option - let options := Core.maxHeartbeats.set options <| - options.get? Core.maxHeartbeats.name |>.getD 1000000 - let res ← EIO.asTask do - let metaCtx : Meta.Context := {} - let metaState : Meta.State := {} - let coreCtx : Core.Context := {options, fileName, fileMap} - let coreState : Core.State := {env} - pure (← ((init ‹_›).run ‹_› ‹_›).run ‹_›).1.1 - show BaseIO _ from cache.set (Sum.inr res) - pure res - match t.get with - | Except.ok res => pure res - | Except.error err => throw err - -/-- -Cached fold over the environment's declarations, -where a given function is applied to `α` for every constant. --/ -structure DeclCache (α : Type) where mk' :: - /-- The cached data. -/ - cache : Cache α - /-- Function for adding a declaration from the current file to the cache. -/ - addDecl : Name → ConstantInfo → α → MetaM α - /-- Function for adding a declaration from the library to the cache. - Defaults to the same behaviour as adding a declaration from the current file. -/ - addLibraryDecl : Name → ConstantInfo → α → MetaM α := addDecl -deriving Nonempty - -/-- -Creates a `DeclCache`. -The cached structure `α` is initialized with `empty`, -and then `addLibraryDecl` is called for every imported constant, and the result is cached. -After all imported constants have been added, we run `post`. -When `get` is called, `addDecl` is also called for every constant in the current file. --/ -def DeclCache.mk (profilingName : String) (empty : α) - (addDecl : Name → ConstantInfo → α → MetaM α) - (addLibraryDecl : Name → ConstantInfo → α → MetaM α := addDecl) - (post : α → MetaM α := fun a => pure a) : IO (DeclCache α) := do - let cache ← Cache.mk do - profileitM Exception profilingName (← getOptions) do - let mut a := empty - for (n, c) in (← getEnv).constants.map₁.toList do - a ← addLibraryDecl n c a - return (← post a) - pure { cache := cache, addDecl := addDecl } - -/-- -Access the cache. -Calling this function for the first time -will initialize the cache with the function -provided in the constructor. --/ -def DeclCache.get (cache : DeclCache α) : MetaM α := do - let mut a ← cache.cache.get - for (n, c) in (← getEnv).constants.map₂.toList do - a ← cache.addDecl n c a - return a - -/-- -A type synonym for a `DeclCache` containing a pair of discrimination trees. -The first will store declarations in the current file, -the second will store declarations from imports (and will hopefully be "read-only" after creation). --/ -@[reducible] def DiscrTreeCache (α : Type) : Type := - DeclCache (DiscrTree α true × DiscrTree α true) - -/-- -Build a `DiscrTreeCache`, -from a function that returns a collection of keys and values for each declaration. --/ -def DiscrTreeCache.mk [BEq α] (profilingName : String) - (processDecl : Name → ConstantInfo → MetaM (Array (Array (DiscrTree.Key true) × α))) - (post? : Option (Array α → Array α) := none) - (init : Option (DiscrTree α true) := none) : - IO (DiscrTreeCache α) := - let updateTree := fun name constInfo tree => do - return (← processDecl name constInfo).foldl (fun t (k, v) => t.insertIfSpecific k v) tree - let addDecl := fun name constInfo (tree₁, tree₂) => do - return (← updateTree name constInfo tree₁, tree₂) - let addLibraryDecl := fun name constInfo (tree₁, tree₂) => do - return (tree₁, ← updateTree name constInfo tree₂) - let post := match post? with - | some f => fun (T₁, T₂) => return (T₁, T₂.mapArrays f) - | none => fun T => pure T - match init with - | some t => return ⟨← Cache.mk (pure ({}, t)), addDecl, addLibraryDecl⟩ - | none => DeclCache.mk profilingName ({}, {}) addDecl addLibraryDecl (post := post) - -/-- -Get matches from both the discrimination tree for declarations in the current file, -and for the imports. - -Note that if you are calling this multiple times with the same environment, -it will rebuild the discrimination tree for the current file multiple times, -and it would be more efficient to call `c.get` once, -and then call `DiscrTree.getMatch` multiple times. --/ -def DiscrTreeCache.getMatch (c : DiscrTreeCache α) (e : Expr) : MetaM (Array α) := do - let (locals, imports) ← c.get - -- `DiscrTree.getMatch` returns results in batches, with more specific lemmas coming later. - -- Hence we reverse this list, so we try out more specific lemmas earlier. - return (← locals.getMatch e).reverse ++ (← imports.getMatch e).reverse diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index 45e5bc2ae20a9..6249c27aff145 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -57,7 +57,6 @@ import Mathlib.Tactic.HigherOrder import Mathlib.Tactic.InferParam import Mathlib.Tactic.Inhabit import Mathlib.Tactic.IrreducibleDef -import Mathlib.Tactic.LeftRight import Mathlib.Tactic.LibrarySearch import Mathlib.Tactic.Lift import Mathlib.Tactic.Lint @@ -69,7 +68,6 @@ import Mathlib.Tactic.MkIffOfInductiveProp import Mathlib.Tactic.NthRewrite import Mathlib.Tactic.Observe import Mathlib.Tactic.PermuteGoals -import Mathlib.Tactic.PrintPrefix import Mathlib.Tactic.ProjectionNotation import Mathlib.Tactic.Propose import Mathlib.Tactic.PushNeg diff --git a/Mathlib/Tactic/ComputeDegree.lean b/Mathlib/Tactic/ComputeDegree.lean index d523afc63969e..5ab9de971c149 100644 --- a/Mathlib/Tactic/ComputeDegree.lean +++ b/Mathlib/Tactic/ComputeDegree.lean @@ -448,12 +448,14 @@ elab_rules : tactic | `(tactic| compute_degree $[!%$bang]?) => focus <| withMain -- simplify the left-hand sides, since this is where the degree computations leave -- expressions such as `max (0 * 1) (max (1 + 0 + 3 * 4) (7 * 0))` evalTactic - (← `(tactic| try any_goals conv_lhs => (simp only [Nat.cast_withBot]; norm_num))) + (← `(tactic| try any_goals conv_lhs => + (simp (config := {decide := true}) only [Nat.cast_withBot]; norm_num))) if bang.isSome then let mut false_goals : Array MVarId := #[] let mut new_goals : Array MVarId := #[] for g in ← getGoals do - let gs' ← run g do evalTactic (← `(tactic| try (any_goals norm_num <;> try assumption))) + let gs' ← run g do evalTactic (← + `(tactic| try (any_goals norm_num <;> norm_cast <;> try assumption))) new_goals := new_goals ++ gs'.toArray if ← gs'.anyM fun g' => g'.withContext do return (← g'.getType'').isConstOf ``False then false_goals := false_goals.push g diff --git a/Mathlib/Tactic/Continuity.lean b/Mathlib/Tactic/Continuity.lean index 7993d3206e745..81abbb964af61 100644 --- a/Mathlib/Tactic/Continuity.lean +++ b/Mathlib/Tactic/Continuity.lean @@ -13,7 +13,10 @@ import Mathlib.Algebra.Group.Defs We define the `continuity` tactic using `aesop`. -/ attribute [aesop (rule_sets [Continuous]) unfold norm] Function.comp -attribute [aesop (rule_sets [Continuous]) unfold norm] npowRec +-- FIXME: `npowRec` is an internal implementation detail, +-- and `aesop` certainly should not know about it. +-- If anyone is working on the `continuity` tactic, please try to fix this! +attribute [aesop (rule_sets [Continuous]) norm] npowRec /-- The `continuity` attribute used to tag continuity statements for the `continuity` tactic. -/ diff --git a/Mathlib/Tactic/Convert.lean b/Mathlib/Tactic/Convert.lean index afb24931682c5..a16d50fbadf6f 100644 --- a/Mathlib/Tactic/Convert.lean +++ b/Mathlib/Tactic/Convert.lean @@ -30,8 +30,7 @@ def Lean.MVarId.convert (e : Expr) (sym : Bool) let v ← mkFreshExprMVar (← mkAppM ``Eq (if sym then #[src, tgt] else #[tgt, src])) g.assign (← mkAppM (if sym then ``Eq.mp else ``Eq.mpr) #[v, e]) let m := v.mvarId! - try m.congrN! depth config patterns - catch _ => return [m] + m.congrN! depth config patterns namespace Mathlib.Tactic diff --git a/Mathlib/Tactic/FieldSimp.lean b/Mathlib/Tactic/FieldSimp.lean index 77b7689f2bf1c..cf7853643b95e 100644 --- a/Mathlib/Tactic/FieldSimp.lean +++ b/Mathlib/Tactic/FieldSimp.lean @@ -148,7 +148,7 @@ syntax (name := fieldSimp) "field_simp" (config)? (discharger)? (&" only")? (simpArgs)? (location)? : tactic elab_rules : tactic -| `(tactic| field_simp $[$cfg:config]? $[$dis:discharger]? $[only%$only?]? +| `(tactic| field_simp $[$cfg:config]? $[(discharger := $dis)]? $[only%$only?]? $[$sa:simpArgs]? $[$loc:location]?) => withMainContext do let cfg ← elabSimpConfig (cfg.getD ⟨.missing⟩) .simp let loc := expandOptLocation (mkOptionalNode loc) @@ -156,8 +156,8 @@ elab_rules : tactic let dis ← match dis with | none => pure discharge | some d => do - let ⟨_,d⟩ ← tacticToDischarge d - pure d + let ⟨_, d⟩ ← tacticToDischarge d + pure d let thms0 ← if only?.isSome then simpOnlyBuiltins.foldlM (·.addConst ·) ({} : SimpTheorems) diff --git a/Mathlib/Tactic/Find.lean b/Mathlib/Tactic/Find.lean index af12a88d91db7..0dc003bc008c4 100644 --- a/Mathlib/Tactic/Find.lean +++ b/Mathlib/Tactic/Find.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sebastian Ullrich -/ import Lean -import Mathlib.Tactic.Cache +import Std.Util.Cache /-! # The `#find` command and tactic. diff --git a/Mathlib/Tactic/LeftRight.lean b/Mathlib/Tactic/LeftRight.lean deleted file mode 100644 index da1b102a6865d..0000000000000 --- a/Mathlib/Tactic/LeftRight.lean +++ /dev/null @@ -1,31 +0,0 @@ -/- -Copyright (c) 2022 Siddhartha Gadgil. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Siddhartha Gadgil --/ -import Lean - -namespace Mathlib.Tactic.LeftRight -open Lean Meta Elab Tactic -def leftRightMeta (name : Name) (idx max : Nat) (goal : MVarId) : MetaM (List MVarId) := do - goal.withContext do - goal.checkNotAssigned name - let target ← goal.getType' - matchConstInduct target.getAppFn - (fun _ ↦ throwTacticEx `constructor goal "target is not an inductive datatype") - fun ival us ↦ do - unless ival.ctors.length == max do - throwTacticEx `constructor goal - s!"{name} target applies for inductive types with exactly two constructors" - let ctor := ival.ctors[idx]! - goal.apply <| mkConst ctor us - -end LeftRight - -open LeftRight - -elab "left" : tactic => do - Lean.Elab.Tactic.liftMetaTactic (leftRightMeta `left 0 2) - -elab "right" : tactic => do - Lean.Elab.Tactic.liftMetaTactic (leftRightMeta `right 1 2) diff --git a/Mathlib/Tactic/LibrarySearch.lean b/Mathlib/Tactic/LibrarySearch.lean index 40ff9bee4abfe..e13e8de539cac 100644 --- a/Mathlib/Tactic/LibrarySearch.lean +++ b/Mathlib/Tactic/LibrarySearch.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner, Scott Morrison -/ import Std.Util.Pickle -import Mathlib.Tactic.Cache +import Std.Util.Cache import Mathlib.Tactic.SolveByElim import Std.Data.MLList.Heartbeats +import Mathlib.Lean.Name /-! # Library search @@ -31,6 +32,9 @@ open Lean Meta Std.Tactic.TryThis initialize registerTraceClass `Tactic.librarySearch initialize registerTraceClass `Tactic.librarySearch.lemmas +/-- Configuration for `DiscrTree`. -/ +def discrTreeConfig : WhnfCoreConfig := {} + /-- A "modifier" for a declaration. * `none` indicates the original declaration, @@ -47,17 +51,17 @@ instance : ToString DeclMod where /-- Prepare the discrimination tree entries for a lemma. -/ def processLemma (name : Name) (constInfo : ConstantInfo) : - MetaM (Array (Array (DiscrTree.Key true) × (Name × DeclMod))) := do + MetaM (Array (Array DiscrTree.Key × (Name × DeclMod))) := do if constInfo.isUnsafe then return #[] if ← name.isBlackListed then return #[] withNewMCtxDepth do withReducible do let (_, _, type) ← forallMetaTelescope constInfo.type - let keys ← DiscrTree.mkPath type + let keys ← DiscrTree.mkPath type discrTreeConfig let mut r := #[(keys, (name, .none))] match type.getAppFnArgs with | (``Iff, #[lhs, rhs]) => do - return r.push (← DiscrTree.mkPath rhs, (name, .mp)) - |>.push (← DiscrTree.mkPath lhs, (name, .mpr)) + return r.push (← DiscrTree.mkPath rhs discrTreeConfig, (name, .mp)) + |>.push (← DiscrTree.mkPath lhs discrTreeConfig, (name, .mpr)) | _ => return r /-- Insert a lemma into the discrimination tree. -/ @@ -66,10 +70,10 @@ def processLemma (name : Name) (constInfo : ConstantInfo) : -- `build/lib/MathlibExtras/LibrarySearch.extra` -- so that the cache is rebuilt. def addLemma (name : Name) (constInfo : ConstantInfo) - (lemmas : DiscrTree (Name × DeclMod) true) : MetaM (DiscrTree (Name × DeclMod) true) := do + (lemmas : DiscrTree (Name × DeclMod)) : MetaM (DiscrTree (Name × DeclMod)) := do let mut lemmas := lemmas for (key, value) in ← processLemma name constInfo do - lemmas := lemmas.insertIfSpecific key value + lemmas := lemmas.insertIfSpecific key value discrTreeConfig return lemmas /-- Construct the discrimination tree of all lemmas. -/ @@ -98,7 +102,7 @@ Retrieve the current current of lemmas. initialize librarySearchLemmas : DiscrTreeCache (Name × DeclMod) ← unsafe do let path ← cachePath if (← path.pathExists) then - let (d, _r) ← unpickle (DiscrTree (Name × DeclMod) true) path + let (d, _r) ← unpickle (DiscrTree (Name × DeclMod)) path -- We can drop the `CompactedRegion` value; we do not plan to free it DiscrTreeCache.mk "apply?: using cache" processLemma (init := some d) else diff --git a/Mathlib/Tactic/Measurability.lean b/Mathlib/Tactic/Measurability.lean index b02c3499798bb..8cc7bbf740ff3 100644 --- a/Mathlib/Tactic/Measurability.lean +++ b/Mathlib/Tactic/Measurability.lean @@ -15,7 +15,10 @@ We define the `measurability` tactic using `aesop`. -/ open Lean.Parser.Tactic (config) attribute [aesop (rule_sets [Measurable]) unfold norm] Function.comp -attribute [aesop (rule_sets [Measurable]) unfold norm] npowRec +-- FIXME: `npowRec` is an internal implementation detail, +-- and `aesop` certainly should not know about it. +-- If anyone is working on the `measurability` tactic, please try to fix this! +attribute [aesop (rule_sets [Measurable]) norm] npowRec /-- The `measurability` attribute used to tag continuity statements for the `measurability` tactic. -/ diff --git a/Mathlib/Tactic/MkIffOfInductiveProp.lean b/Mathlib/Tactic/MkIffOfInductiveProp.lean index 939cda6d74fd6..6457562047a28 100644 --- a/Mathlib/Tactic/MkIffOfInductiveProp.lean +++ b/Mathlib/Tactic/MkIffOfInductiveProp.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, David Renshaw -/ import Lean +import Std.Tactic.LeftRight import Mathlib.Lean.Meta import Mathlib.Tactic.Basic -import Mathlib.Tactic.LeftRight /-! # mk_iff_of_inductive_prop @@ -35,12 +35,14 @@ Fails if `n < m`. -/ private def select (m n : Nat) (goal : MVarId) : MetaM MVarId := match m,n with | 0, 0 => pure goal - | 0, (_ + 1) => do let [new_goal] ← Mathlib.Tactic.LeftRight.leftRightMeta `left 0 2 goal - | throwError "expected only one new goal" - pure new_goal - | (m + 1), (n + 1) => do let [new_goal] ← Mathlib.Tactic.LeftRight.leftRightMeta `right 1 2 goal - | throwError "expected only one new goal" - select m n new_goal + | 0, (_ + 1) => do + let [new_goal] ← Std.Tactic.NthConstructor.nthConstructor `left 0 (some 2) goal + | throwError "expected only one new goal" + pure new_goal + | (m + 1), (n + 1) => do + let [new_goal] ← Std.Tactic.NthConstructor.nthConstructor `right 1 (some 2) goal + | throwError "expected only one new goal" + select m n new_goal | _, _ => failure /-- `compactRelation bs as_ps`: Produce a relation of the form: @@ -131,7 +133,7 @@ structure Shape : Type where while proving the iff theorem, and a proposition representing the constructor. -/ def constrToProp (univs : List Level) (params : List Expr) (idxs : List Expr) (c : Name) : - MetaM (Shape × Expr) := + MetaM (Shape × Expr) := do let type := (← getConstInfo c).instantiateTypeLevelParams univs let type' ← Meta.forallBoundedTelescope type (params.length) fun fvars ty ↦ do pure $ ty.replaceFVars fvars params.toArray diff --git a/Mathlib/Tactic/NormNum/Core.lean b/Mathlib/Tactic/NormNum/Core.lean index ddb35ca06aa60..6ebf7ed63e1d7 100644 --- a/Mathlib/Tactic/NormNum/Core.lean +++ b/Mathlib/Tactic/NormNum/Core.lean @@ -51,22 +51,25 @@ def mkNormNumExt (n : Name) : ImportM NormNumExt := do /-- Each `norm_num` extension is labelled with a collection of patterns which determine the expressions to which it should be applied. -/ -abbrev Entry := Array (Array (DiscrTree.Key true)) × Name +abbrev Entry := Array (Array DiscrTree.Key) × Name /-- The state of the `norm_num` extension environment -/ structure NormNums where /-- The tree of `norm_num` extensions. -/ - tree : DiscrTree NormNumExt true := {} + tree : DiscrTree NormNumExt := {} /-- Erased `norm_num`s. -/ erased : PHashSet Name := {} deriving Inhabited +/-- Configuration for `DiscrTree`. -/ +def discrTreeConfig : WhnfCoreConfig := {} + /-- Environment extensions for `norm_num` declarations -/ initialize normNumExt : ScopedEnvExtension Entry (Entry × NormNumExt) NormNums ← -- we only need this to deduplicate entries in the DiscrTree have : BEq NormNumExt := ⟨fun _ _ ↦ false⟩ /- Insert `v : NormNumExt` into the tree `dt` on all key sequences given in `kss`. -/ - let insert kss v dt := kss.foldl (fun dt ks ↦ dt.insertCore ks v) dt + let insert kss v dt := kss.foldl (fun dt ks ↦ dt.insertCore ks v discrTreeConfig) dt registerScopedEnvExtension { mkInitial := pure {} ofOLeanEntry := fun _ e@(_, n) ↦ return (e, ← mkNormNumExt n) @@ -84,7 +87,7 @@ def derive {α : Q(Type u)} (e : Q($α)) (post := false) : MetaM (Result e) := d profileitM Exception "norm_num" (← getOptions) do let s ← saveState let normNums := normNumExt.getState (← getEnv) - let arr ← normNums.tree.getMatch e + let arr ← normNums.tree.getMatch e discrTreeConfig for ext in arr do if (bif post then ext.post else ext.pre) && ! normNums.erased.contains ext.name then try @@ -177,7 +180,7 @@ initialize registerBuiltinAttribute { let e ← elabTerm stx none let (_, _, e) ← lambdaMetaTelescope (← mkLambdaFVars (← getLCtx).getFVars e) return e - DiscrTree.mkPath e + DiscrTree.mkPath e discrTreeConfig normNumExt.add ((keys, declName), ext) kind | _ => throwUnsupportedSyntax erase := fun declName => do diff --git a/Mathlib/Tactic/NormNum/Inv.lean b/Mathlib/Tactic/NormNum/Inv.lean index aa6a0874a5d23..98a7f08c57fb8 100644 --- a/Mathlib/Tactic/NormNum/Inv.lean +++ b/Mathlib/Tactic/NormNum/Inv.lean @@ -107,7 +107,7 @@ theorem isRat_inv_pos {α} [DivisionRing α] [CharZero α] {a : α} {n d : ℕ} IsRat a (.ofNat (Nat.succ n)) d → IsRat a⁻¹ (.ofNat d) (Nat.succ n) := by rintro ⟨_, rfl⟩ have := invertibleOfNonzero (α := α) (Nat.cast_ne_zero.2 (Nat.succ_ne_zero n)) - refine ⟨this, by simp⟩ + exact ⟨this, by simp⟩ theorem isRat_inv_one {α} [DivisionRing α] : {a : α} → IsNat a (nat_lit 1) → IsNat a⁻¹ (nat_lit 1) diff --git a/Mathlib/Tactic/NormNum/LegendreSymbol.lean b/Mathlib/Tactic/NormNum/LegendreSymbol.lean index db5269773b3c2..8417a02e44a34 100644 --- a/Mathlib/Tactic/NormNum/LegendreSymbol.lean +++ b/Mathlib/Tactic/NormNum/LegendreSymbol.lean @@ -124,6 +124,7 @@ theorem jacobiSymNat.odd_even (a b c : ℕ) (r : ℤ) (ha : a % 2 = 1) (hb : b % jacobiSymNat a b = r := by have ha' : legendreSym 2 a = 1 := by simp only [legendreSym.mod 2 a, Int.ofNat_mod_ofNat, ha] + decide rcases eq_or_ne c 0 with (rfl | hc') · rw [← hr, Nat.eq_zero_of_dvd_of_div_eq_zero (Nat.dvd_of_mod_eq_zero hb) hc] · haveI : NeZero c := ⟨hc'⟩ diff --git a/Mathlib/Tactic/NormNum/OfScientific.lean b/Mathlib/Tactic/NormNum/OfScientific.lean index 7b564e68174e0..58fe46e88ef09 100644 --- a/Mathlib/Tactic/NormNum/OfScientific.lean +++ b/Mathlib/Tactic/NormNum/OfScientific.lean @@ -50,7 +50,7 @@ to rat casts if the scientific notation is inherited from the one for rationals. haveI' : $e =Q OfScientific.ofScientific $m $b $exp := ⟨⟩ haveI' lh : @OfScientific.ofScientific $α $σα =Q (fun m s e ↦ (Rat.ofScientific m s e : $α)) := ⟨⟩ match b with - | ~q(true) => + | ~q(true) => let rme ← derive (q(mkRat $m (10 ^ $exp)) : Q($α)) let some ⟨q, n, d, p⟩ := rme.toRat' dα | failure return .isRat' dα q n d q(isRat_ofScientific_of_true $σα $lh $p) diff --git a/Mathlib/Tactic/NormNum/Result.lean b/Mathlib/Tactic/NormNum/Result.lean index e1545d810c1c5..2b6685e01ffa8 100644 --- a/Mathlib/Tactic/NormNum/Result.lean +++ b/Mathlib/Tactic/NormNum/Result.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Invertible.GroupWithZero import Mathlib.Data.Sigma.Basic import Mathlib.Data.Nat.Basic import Mathlib.Data.Int.Cast.Basic +import Qq.MetaM /-! ## The `Result` type for `norm_num` diff --git a/Mathlib/Tactic/Peel.lean b/Mathlib/Tactic/Peel.lean new file mode 100644 index 0000000000000..717fa7fcb05cc --- /dev/null +++ b/Mathlib/Tactic/Peel.lean @@ -0,0 +1,245 @@ +/- +Copyright (c) 2023 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import Mathlib.Order.Filter.Basic +import Mathlib.Tactic.Basic + +/-! +# The `peel` tactic + +`peel h with h' idents*` tries to apply `forall_imp` (or `Exists.imp`, or `Filter.Eventually.mp`, +`Filter.Frequently.mp` and `Filter.eventually_of_forall`) with the argument `h` and uses `idents*` +to introduce variables with the supplied names, giving the "peeled" argument the name `h'`. + +One can provide a numeric argument as in `peel 4 h` which will peel 4 quantifiers off +the expressions automatically name any variables not specifically named in the `with` clause. + +In addition, the user may supply a term `e` via `... using e` in order to close the goal +immediately. In particular, `peel h using e` is equivalent to `peel h; exact e`. The `using` syntax +may be paired with any of the other features of `peel`. +-/ + +namespace Mathlib.Tactic.Peel +open Lean Expr Meta Elab Tactic + +/-- +Peels matching quantifiers off of a given term and the goal and introduces the relevant variables. + +- `peel e` peels all quantifiers (at reducible transparency), + using `this` for the name of the peeled hypothesis. +- `peel e with h` is `peel e` but names the peeled hypothesis `h`. + If `h` is `_` then uses `this` for the name of the peeled hypothesis. +- `peel n e` peels `n` quantifiers (at default transparency). +- `peel n e with h x y z ...` peels `n` quantifiers, names the peeled hypothesis `h`, + and uses `x`, `y`, `z`, and so on to name the introduced variables; these names may be `_`. + If `h` is `_` then uses `this` for the name of the peeled hypothesis. + The length of the list of variables does not need to equal `n`. +- `peel e with h x₁ ... xₙ` is `peel n e with h x₁ ... xₙ`. + +There are also variants that apply to an iff in the goal: +- `peel n` peels `n` quantifiers in an iff. +- `peel with x₁ ... xₙ` peels `n` quantifiers in an iff and names them. + +Given `p q : ℕ → Prop`, `h : ∀ x, p x`, and a goal `⊢ : ∀ x, q x`, the tactic `peel h with h' x` +will introduce `x : ℕ`, `h' : p x` into the context and the new goal will be `⊢ q x`. This works +with `∃`, as well as `∀ᶠ` and `∃ᶠ`, and it can even be applied to a sequence of quantifiers. Note +that this is a logically weaker setup, so using this tactic is not always feasible. + +For a more complex example, given a hypothesis and a goal: +``` +h : ∀ ε > (0 : ℝ), ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℝ) < ε +⊢ ∀ ε > (0 : ℝ), ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℝ) ≤ ε +``` +(which differ only in `<`/`≤`), applying `peel h with h_peel ε hε N n hn` will yield a tactic state: +``` +h : ∀ ε > (0 : ℝ), ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℝ) < ε +ε : ℝ +hε : 0 < ε +N n : ℕ +hn : N ≤ n +h_peel : 1 / (n + 1 : ℝ) < ε +⊢ 1 / (n + 1 : ℝ) ≤ ε +``` +and the goal can be closed with `exact h_peel.le`. +Note that in this example, `h` and the goal are logically equivalent statements, but `peel` +*cannot* be immediately applied to show that the goal implies `h`. + +In addition, `peel` supports goals of the form `(∀ x, p x) ↔ ∀ x, q x`, or likewise for any +other quantifier. In this case, there is no hypothesis or term to supply, but otherwise the syntax +is the same. So for such goals, the syntax is `peel 1` or `peel with x`, and after which the +resulting goal is `p x ↔ q x`. The `congr!` tactic can also be applied to goals of this form using +`congr! 1 with x`. While `congr!` applies congruence lemmas in general, `peel` can be relied upon +to only apply to outermost quantifiers. + +Finally, the user may supply a term `e` via `... using e` in order to close the goal +immediately. In particular, `peel h using e` is equivalent to `peel h; exact e`. The `using` syntax +may be paired with any of the other features of `peel`. + +This tactic works by repeatedly applying lemmas such as `forall_imp`, `Exists.imp`, +`Filter.Eventually.mp`, `Filter.Frequently.mp`, and `Filter.eventually_of_forall`. +-/ +syntax (name := peel) + "peel" (num)? (ppSpace colGt term)? + (" with" (ppSpace colGt (ident <|> hole))+)? (usingArg)? : tactic + +private lemma and_imp_left_of_imp_imp {p q r : Prop} (h : r → p → q) : r ∧ p → r ∧ q := by tauto + +private theorem eventually_imp {α : Type*} {p q : α → Prop} {f : Filter α} + (hq : ∀ (x : α), p x → q x) (hp : ∀ᶠ (x : α) in f, p x) : ∀ᶠ (x : α) in f, q x := + Filter.Eventually.mp hp (Filter.eventually_of_forall hq) + +private theorem frequently_imp {α : Type*} {p q : α → Prop} {f : Filter α} + (hq : ∀ (x : α), p x → q x) (hp : ∃ᶠ (x : α) in f, p x) : ∃ᶠ (x : α) in f, q x := + Filter.Frequently.mp hp (Filter.eventually_of_forall hq) + +private theorem eventually_congr {α : Type*} {p q : α → Prop} {f : Filter α} + (hq : ∀ (x : α), p x ↔ q x) : (∀ᶠ (x : α) in f, p x) ↔ ∀ᶠ (x : α) in f, q x := by + congr! 2; exact hq _ + +private theorem frequently_congr {α : Type*} {p q : α → Prop} {f : Filter α} + (hq : ∀ (x : α), p x ↔ q x) : (∃ᶠ (x : α) in f, p x) ↔ ∃ᶠ (x : α) in f, q x := by + congr! 2; exact hq _ + +/-- The list of constants that are regarded as being quantifiers. -/ +def quantifiers : List Name := + [``Exists, ``And, ``Filter.Eventually, ``Filter.Frequently] + +/-- If `unfold` is false then do `whnfR`, otherwise unfold everything that's not a quantifier, +according to the `quantifiers` list. -/ +def whnfQuantifier (p : Expr) (unfold : Bool) : MetaM Expr := do + if unfold then + whnfHeadPred p fun e => + if let .const n .. := e.getAppFn then + return !(n ∈ quantifiers) + else + return false + else + whnfR p + +/-- Throws an error saying `ty` and `target` could not be matched up. -/ +def throwPeelError {α : Type} (ty target : Expr) : MetaM α := + throwError "Tactic 'peel' could not match quantifiers in{indentD ty}\nand{indentD target}" + +/-- If `f` is a lambda then use its binding name to generate a new hygienic name, +and otherwise choose a new hygienic name. -/ +def mkFreshBinderName (f : Expr) : MetaM Name := + mkFreshUserName (if let .lam n .. := f then n else `a) + +/-- Applies a "peel theorem" with two main arguments, where the first is the new goal +and the second can be filled in using `e`. Then it intros two variables with the +provided names. + +If, for example, `goal : ∃ y : α, q y` and `thm := Exists.imp`, the metavariable returned has +type `q x` where `x : α` has been introduced into the context. -/ +def applyPeelThm (thm : Name) (goal : MVarId) + (e : Expr) (ty target : Expr) (n : Name) (n' : Name) : + MetaM (FVarId × List MVarId) := do + let new_goal :: ge :: _ ← goal.applyConst thm <|> throwPeelError ty target + | throwError "peel: internal error" + ge.assignIfDefeq e <|> throwPeelError ty target + let (fvars, new_goal) ← new_goal.introN 2 [n, n'] + return (fvars[1]!, [new_goal]) + +/-- This is the core to the `peel` tactic. + +It tries to match `e` and `goal` as quantified statements (using `∀` and the quantifiers in +the `quantifiers` list), then applies "peel theorems" using `applyPeelThm`. + +We treat `∧` as a quantifier for sake of dealing with quantified statements +like `∃ δ > (0 : ℝ), q δ`, which is notation for `∃ δ, δ > (0 : ℝ) ∧ q δ`. -/ +def peelCore (goal : MVarId) (e : Expr) (n? : Option Name) (n' : Name) (unfold : Bool) : + MetaM (FVarId × List MVarId) := goal.withContext do + let ty ← whnfQuantifier (← inferType e) unfold + let target ← whnfQuantifier (← goal.getType) unfold + if ty.isForall && target.isForall then + applyPeelThm ``forall_imp goal e ty target (← n?.getDM (mkFreshUserName target.bindingName!)) n' + else if ty.getAppFn.isConst + && ty.getAppNumArgs == target.getAppNumArgs + && ty.getAppFn == target.getAppFn then + match target.getAppFnArgs with + | (``Exists, #[_, p]) => + applyPeelThm ``Exists.imp goal e ty target (← n?.getDM (mkFreshBinderName p)) n' + | (``And, #[_, _]) => + applyPeelThm ``and_imp_left_of_imp_imp goal e ty target (← n?.getDM (mkFreshUserName `p)) n' + | (``Filter.Eventually, #[_, p, _]) => + applyPeelThm ``eventually_imp goal e ty target (← n?.getDM (mkFreshBinderName p)) n' + | (``Filter.Frequently, #[_, p, _]) => + applyPeelThm ``frequently_imp goal e ty target (← n?.getDM (mkFreshBinderName p)) n' + | _ => throwPeelError ty target + else + throwPeelError ty target + +/-- Given a list `l` of names, this peels `num` quantifiers off of the expression `e` and +the main goal and introduces variables with the provided names until the list of names is exhausted. +Note: the name `n?` (with default `this`) is used for the name of the expression `e` with +quantifiers peeled. -/ +def peelArgs (e : Expr) (num : Nat) (l : List Name) (n? : Option Name) (unfold : Bool := true) : + TacticM Unit := do + match num with + | 0 => return + | num + 1 => + let fvarId ← liftMetaTacticAux (peelCore · e l.head? (n?.getD `this) unfold) + peelArgs (.fvar fvarId) num l.tail n? + unless num == 0 do + if let some mvarId ← observing? do (← getMainGoal).clear fvarId then + replaceMainGoal [mvarId] + +/-- Similar to `peelArgs` but peels arbitrarily many quantifiers. Returns whether or not +any quantifiers were peeled. -/ +partial def peelUnbounded (e : Expr) (n? : Option Name) (unfold : Bool := false) : + TacticM Bool := do + let fvarId? ← observing? <| liftMetaTacticAux (peelCore · e none (n?.getD `this) unfold) + if let some fvarId := fvarId? then + let peeled ← peelUnbounded (.fvar fvarId) n? + if peeled then + if let some mvarId ← observing? do (← getMainGoal).clear fvarId then + replaceMainGoal [mvarId] + return true + else + return false + +/-- Peel off a single quantifier from an `↔`. -/ +def peelIffAux : TacticM Unit := do + evalTactic (← `(tactic| focus + first | apply forall_congr' + | apply exists_congr + | apply eventually_congr + | apply frequently_congr + | apply and_congr_right + | fail "failed to apply a quantifier congruence lemma.")) + +/-- Peel off quantifiers from an `↔` and assign the names given in `l` to the introduced +variables. -/ +def peelArgsIff (l : List Name) : TacticM Unit := withMainContext do + match l with + | [] => pure () + | h :: hs => + peelIffAux + let goal ← getMainGoal + let (_, new_goal) ← goal.intro h + replaceMainGoal [new_goal] + peelArgsIff hs + +elab_rules : tactic + | `(tactic| peel $[$num?:num]? $e:term $[with $n? $l?*]?) => withMainContext do + /- we use `elabTermForApply` instead of `elabTerm` so that terms passed to `peel` can contain + quantifiers with implicit bound variables without causing errors or requiring `@`. -/ + let e ← elabTermForApply e false + let n? := n?.bind fun n => if n.raw.isIdent then pure n.raw.getId else none + let l := (l?.getD #[]).map getNameOfIdent' |>.toList + -- If num is not present and if there are any provided variable names, + -- use the number of variable names. + let num? := num?.map (·.getNat) <|> if l.isEmpty then none else l.length + if let some num := num? then + peelArgs e num l n? + else + unless ← peelUnbounded e n? do + throwPeelError (← inferType e) (← getMainTarget) + | `(tactic| peel $n:num) => peelArgsIff <| .replicate n.getNat `_ + | `(tactic| peel with $args*) => peelArgsIff (args.map getNameOfIdent').toList + +macro_rules + | `(tactic| peel $[$n:num]? $[$e:term]? $[with $h*]? using $u:term) => + `(tactic| peel $[$n:num]? $[$e:term]? $[with $h*]?; exact $u) diff --git a/Mathlib/Tactic/Positivity/Basic.lean b/Mathlib/Tactic/Positivity/Basic.lean index 26c7373b0183c..51b724cdf2dc1 100644 --- a/Mathlib/Tactic/Positivity/Basic.lean +++ b/Mathlib/Tactic/Positivity/Basic.lean @@ -7,6 +7,9 @@ import Std.Lean.Parser import Mathlib.Data.Int.Order.Basic import Mathlib.Data.Int.CharZero import Mathlib.Data.Nat.Factorial.Basic +import Mathlib.Data.Rat.Order +import Mathlib.Data.Rat.Cast.CharZero +import Mathlib.Data.Rat.Cast.Order import Mathlib.Tactic.Positivity.Core import Mathlib.Tactic.HaveI import Mathlib.Algebra.GroupPower.Order @@ -460,6 +463,26 @@ def evalIntCast : PositivityExt where eval {u α} _zα _pα e := do | .none => pure .none +/-- Extension for Rat.cast. -/ +@[positivity Rat.cast _] +def evalRatCast : PositivityExt where eval {u α} _zα _pα e := do + let _rα : Q(RatCast $α) ← synthInstanceQ (q(RatCast $α)) + let ~q(Rat.cast ($a : ℚ)) := e | throwError "not Rat.cast" + let zα' : Q(Zero ℚ) := q(inferInstance) + let pα' : Q(PartialOrder ℚ) := q(inferInstance) + match ← core zα' pα' a with + | .positive pa => + let _oα ← synthInstanceQ (q(LinearOrderedField $α) : Q(Type u)) + pure (.positive (q((Rat.cast_pos (K := $α)).mpr $pa) : Expr)) + | .nonnegative pa => + let _oα ← synthInstanceQ (q(LinearOrderedField $α) : Q(Type u)) + pure (.nonnegative (q((Rat.cast_nonneg (K := $α)).mpr $pa) : Expr)) + | .nonzero pa => + let _oα ← synthInstanceQ (q(DivisionRing $α) : Q(Type u)) + let _cα ← synthInstanceQ (q(CharZero $α) : Q(Prop)) + pure (.nonzero (q((Rat.cast_ne_zero (α := $α)).mpr $pa) : Expr)) + | .none => pure .none + /-- Extension for Nat.succ. -/ @[positivity Nat.succ _] def evalNatSucc : PositivityExt where eval {_u _α} _zα _pα e := do @@ -468,6 +491,12 @@ def evalNatSucc : PositivityExt where eval {_u _α} _zα _pα e := do /-- Extension for Nat.factorial. -/ @[positivity Nat.factorial _] -def evalFactorial : PositivityExt where eval {_ _} _ _ e := do - let .app _ (a : Q(Nat)) ← whnfR e | throwError "not Nat.factorial" +def evalFactorial : PositivityExt where eval {_ _} _ _ (e : Q(ℕ)) := do + let ~q(Nat.factorial $a) := e | throwError "failed to match Nat.factorial" pure (.positive (q(Nat.factorial_pos $a) : Expr)) + +/-- Extension for Nat.ascFactorial. -/ +@[positivity Nat.ascFactorial _ _] +def evalAscFactorial : PositivityExt where eval {_ _} _ _ (e : Q(ℕ)) := do + let ~q(Nat.ascFactorial $n $k) := e | throwError "failed to match Nat.ascFactorial" + pure (.positive (q(Nat.ascFactorial_pos $n $k) : Expr)) diff --git a/Mathlib/Tactic/Positivity/Core.lean b/Mathlib/Tactic/Positivity/Core.lean index 060f6d5e9addb..755eb8cbe93ee 100644 --- a/Mathlib/Tactic/Positivity/Core.lean +++ b/Mathlib/Tactic/Positivity/Core.lean @@ -62,16 +62,19 @@ def mkPositivityExt (n : Name) : ImportM PositivityExt := do let { env, opts, .. } ← read IO.ofExcept <| unsafe env.evalConstCheck PositivityExt opts ``PositivityExt n +/-- Configuration for `DiscrTree`. -/ +def discrTreeConfig : WhnfCoreConfig := {} + /-- Each `positivity` extension is labelled with a collection of patterns which determine the expressions to which it should be applied. -/ -abbrev Entry := Array (Array (DiscrTree.Key true)) × Name +abbrev Entry := Array (Array DiscrTree.Key) × Name /-- Environment extensions for `positivity` declarations -/ initialize positivityExt : PersistentEnvExtension Entry (Entry × PositivityExt) - (List Entry × DiscrTree PositivityExt true) ← + (List Entry × DiscrTree PositivityExt) ← -- we only need this to deduplicate entries in the DiscrTree have : BEq PositivityExt := ⟨fun _ _ => false⟩ - let insert kss v dt := kss.foldl (fun dt ks => dt.insertCore ks v) dt + let insert kss v dt := kss.foldl (fun dt ks => dt.insertCore ks v discrTreeConfig) dt registerPersistentEnvExtension { mkInitial := pure ([], {}) addImportedFn := fun s => do @@ -101,7 +104,7 @@ initialize registerBuiltinAttribute { let e ← elabTerm stx none let (_, _, e) ← lambdaMetaTelescope (← mkLambdaFVars (← getLCtx).getFVars e) return e - DiscrTree.mkPath e + DiscrTree.mkPath e discrTreeConfig setEnv <| positivityExt.addEntry env ((keys, declName), ext) | _ => throwUnsupportedSyntax } @@ -287,7 +290,7 @@ def orElse (t₁ : Strictness zα pα e) (t₂ : MetaM (Strictness zα pα e)) : def core (e : Q($α)) : MetaM (Strictness zα pα e) := do let mut result := .none trace[Tactic.positivity] "trying to prove positivity of {e}" - for ext in ← (positivityExt.getState (← getEnv)).2.getMatch e do + for ext in ← (positivityExt.getState (← getEnv)).2.getMatch e discrTreeConfig do try result ← orElse result <| ext.eval zα pα e catch err => diff --git a/Mathlib/Tactic/PrintPrefix.lean b/Mathlib/Tactic/PrintPrefix.lean deleted file mode 100644 index 08fabad6e9429..0000000000000 --- a/Mathlib/Tactic/PrintPrefix.lean +++ /dev/null @@ -1,60 +0,0 @@ -/- -Copyright (c) 2021 Shing Tak Lam. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Shing Tak Lam, Daniel Selsam, Mario Carneiro --/ -import Lean.Elab.Command - -open Lean Meta Elab Command - -namespace Lean -namespace Meta - -structure FindOptions where - stage1 : Bool := true - checkPrivate : Bool := false - -def findCore (ϕ : ConstantInfo → MetaM Bool) (opts : FindOptions := {}) : - MetaM (Array ConstantInfo) := do - let matches_ ← if !opts.stage1 then pure #[] else - (← getEnv).constants.map₁.foldM (init := #[]) check - (← getEnv).constants.map₂.foldlM (init := matches_) check -where - check matches_ name cinfo := do - if opts.checkPrivate || !isPrivateName name then - if ← ϕ cinfo then pure $ matches_.push cinfo else pure matches_ - else pure matches_ - -def find (msg : String) - (ϕ : ConstantInfo → MetaM Bool) (opts : FindOptions := {}) : TermElabM String := do - let cinfos ← findCore ϕ opts - let cinfos := cinfos.qsort fun p q ↦ p.name.lt q.name - let mut msg := msg - for cinfo in cinfos do - msg := msg ++ s!"{cinfo.name} : {← Meta.ppExpr cinfo.type}\n" - pure msg - -end Meta - -namespace Elab.Command - -syntax (name := printPrefix) "#print prefix " ident : command - -/-- -The command `#print prefix foo` will print all definitions that start with -the namespace `foo`. --/ -@[command_elab printPrefix] def elabPrintPrefix : CommandElab -| `(#print prefix%$tk $name:ident) => do - let nameId := name.getId - liftTermElabM do - let mut msg ← find "" fun cinfo ↦ pure $ nameId.isPrefixOf cinfo.name - if msg.isEmpty then - if let [name] ← resolveGlobalConst name then - msg ← find msg fun cinfo ↦ pure $ name.isPrefixOf cinfo.name - if !msg.isEmpty then - logInfoAt tk msg -| _ => throwUnsupportedSyntax - -end Elab.Command -end Lean diff --git a/Mathlib/Tactic/Propose.lean b/Mathlib/Tactic/Propose.lean index 4099f6aafbd6a..f438066a003a6 100644 --- a/Mathlib/Tactic/Propose.lean +++ b/Mathlib/Tactic/Propose.lean @@ -7,7 +7,7 @@ import Mathlib.Lean.Expr.Basic import Mathlib.Lean.Meta import Mathlib.Lean.Meta.Basic import Mathlib.Lean.Meta.DiscrTree -import Mathlib.Tactic.Cache +import Std.Util.Cache import Mathlib.Tactic.Core import Mathlib.Tactic.SolveByElim import Mathlib.Tactic.TryThis @@ -42,7 +42,10 @@ open Lean Meta Std.Tactic.TryThis initialize registerTraceClass `Tactic.propose -initialize proposeLemmas : DeclCache (DiscrTree Name true) ← +/-- Configuration for `DiscrTree`. -/ +def discrTreeConfig : WhnfCoreConfig := {} + +initialize proposeLemmas : DeclCache (DiscrTree Name) ← DeclCache.mk "have?: init cache" {} fun name constInfo lemmas => do if constInfo.isUnsafe then return lemmas if ← name.isBlackListed then return lemmas @@ -50,7 +53,8 @@ initialize proposeLemmas : DeclCache (DiscrTree Name true) ← let (mvars, _, _) ← forallMetaTelescope constInfo.type let mut lemmas := lemmas for m in mvars do - lemmas := lemmas.insertIfSpecific (← DiscrTree.mkPath (← inferType m)) name + let path ← DiscrTree.mkPath (← inferType m) discrTreeConfig + lemmas := lemmas.insertIfSpecific path name discrTreeConfig pure lemmas /-- Shortcut for calling `solveByElim`. -/ @@ -73,11 +77,11 @@ We look up candidate lemmas from a discrimination tree using the first such expr Returns an array of pairs, containing the names of found lemmas and the resulting application. -/ -def propose (lemmas : DiscrTree Name s) (type : Expr) (required : Array Expr) +def propose (lemmas : DiscrTree Name) (type : Expr) (required : Array Expr) (solveByElimDepth := 15) : MetaM (Array (Name × Expr)) := do guard !required.isEmpty let ty ← whnfR (← instantiateMVars (← inferType required[0]!)) - let candidates ← lemmas.getMatch ty + let candidates ← lemmas.getMatch ty discrTreeConfig candidates.filterMapM fun lem : Name => try trace[Tactic.propose] "considering {lem}" diff --git a/Mathlib/Tactic/PushNeg.lean b/Mathlib/Tactic/PushNeg.lean index a4fd41b57d5ec..b5875a12442a3 100644 --- a/Mathlib/Tactic/PushNeg.lean +++ b/Mathlib/Tactic/PushNeg.lean @@ -9,6 +9,7 @@ import Mathlib.Lean.Expr import Mathlib.Logic.Basic import Mathlib.Init.Order.Defs import Mathlib.Tactic.Conv +import Mathlib.Init.Set set_option autoImplicit true @@ -35,6 +36,17 @@ theorem not_lt_eq (a b : β) : (¬ (a < b)) = (b ≤ a) := propext not_lt theorem not_ge_eq (a b : β) : (¬ (a ≥ b)) = (a < b) := propext not_le theorem not_gt_eq (a b : β) : (¬ (a > b)) = (a ≤ b) := propext not_lt +theorem not_nonempty_eq (s : Set γ) : (¬ s.Nonempty) = (s = ∅) := by + have A : ∀ (x : γ), ¬(x ∈ (∅ : Set γ)) := fun x ↦ id + simp only [Set.Nonempty, not_exists, eq_iff_iff] + exact ⟨fun h ↦ Set.ext (fun x ↦ by simp only [h x, false_iff, A]), fun h ↦ by rwa [h]⟩ + +theorem ne_empty_eq_nonempty (s : Set γ) : (s ≠ ∅) = s.Nonempty := by + rw [ne_eq, ← not_nonempty_eq s, not_not] + +theorem empty_ne_eq_nonempty (s : Set γ) : (∅ ≠ s) = s.Nonempty := by + rw [ne_comm, ne_empty_eq_nonempty] + /-- Make `push_neg` use `not_and_or` rather than the default `not_and`. -/ register_option push_neg.use_distrib : Bool := { defValue := false @@ -71,7 +83,19 @@ def transformNegationStep (e : Expr) : SimpM (Option Simp.Step) := do return mkSimpStep (mkAnd (mkNot p) (mkNot q)) (←mkAppM ``not_or_eq #[p, q]) | (``Iff, #[p, q]) => return mkSimpStep (mkOr (mkAnd p (mkNot q)) (mkAnd (mkNot p) q)) (←mkAppM ``not_iff #[p, q]) - | (``Eq, #[_ty, e₁, e₂]) => + | (``Eq, #[ty, e₁, e₂]) => + if ty.isAppOfArity ``Set 1 then + -- test if equality is of the form `s = ∅`, and negate it to `s.Nonempty` + if e₂.isAppOfArity ``EmptyCollection.emptyCollection 2 then + let thm ← mkAppM ``ne_empty_eq_nonempty #[e₁] + let some (_, _, rhs) := (← inferType thm).eq? | return none + return mkSimpStep rhs thm + -- test if equality is of the form `∅ = s`, and negate it to `s.Nonempty` + if e₁.isAppOfArity ``EmptyCollection.emptyCollection 2 then + let thm ← mkAppM ``empty_ne_eq_nonempty #[e₂] + let some (_, _, rhs) := (← inferType thm).eq? | return none + return mkSimpStep rhs thm + -- negate `a = b` to `a ≠ b` return Simp.Step.visit { expr := ← mkAppM ``Ne #[e₁, e₂] } | (``Ne, #[_ty, e₁, e₂]) => return mkSimpStep (← mkAppM ``Eq #[e₁, e₂]) (← mkAppM ``not_ne_eq #[e₁, e₂]) @@ -79,6 +103,11 @@ def transformNegationStep (e : Expr) : SimpM (Option Simp.Step) := do | (``LT.lt, #[_ty, _inst, e₁, e₂]) => handleIneq e₁ e₂ ``not_lt_eq | (``GE.ge, #[_ty, _inst, e₁, e₂]) => handleIneq e₁ e₂ ``not_ge_eq | (``GT.gt, #[_ty, _inst, e₁, e₂]) => handleIneq e₁ e₂ ``not_gt_eq + | (``Set.Nonempty, #[_ty, e]) => + -- negate `s.Nonempty` to `s = ∅` + let thm ← mkAppM ``not_nonempty_eq #[e] + let some (_, _, rhs) := (← inferType thm).eq? | return none + return mkSimpStep rhs thm | (``Exists, #[_, .lam n typ bo bi]) => return mkSimpStep (.forallE n typ (mkNot bo) bi) (← mkAppM ``not_exists_eq #[.lam n typ bo bi]) @@ -108,7 +137,7 @@ partial def transformNegation (e : Expr) : SimpM Simp.Step := do /-- Common entry point to `push_neg` as a conv. -/ def pushNegCore (tgt : Expr) : MetaM Simp.Result := do let myctx : Simp.Context := - { config := { eta := true, zeta := false }, + { config := { eta := true, zeta := false, proj := false }, simpTheorems := #[ ] congrTheorems := (← getSimpCongrTheorems) } (·.1) <$> Simp.main tgt myctx (methods := { pre := transformNegation }) diff --git a/Mathlib/Tactic/Relation/Rfl.lean b/Mathlib/Tactic/Relation/Rfl.lean index c81e3a024e888..5fe045dfabe5b 100644 --- a/Mathlib/Tactic/Relation/Rfl.lean +++ b/Mathlib/Tactic/Relation/Rfl.lean @@ -41,7 +41,7 @@ def _root_.Lean.MVarId.liftReflToEq (mvarId : MVarId) : MetaM MVarId := do if rel.isAppOf `Eq then -- No need to lift Eq to Eq return mvarId - for lem in ← (Std.Tactic.reflExt.getState (← getEnv)).getMatch rel do + for lem in ← (Std.Tactic.reflExt.getState (← getEnv)).getMatch rel Std.Tactic.reflExt.config do let res ← observing? do -- First create an equality relating the LHS and RHS -- and reduce the goal to proving that LHS is related to LHS. diff --git a/Mathlib/Tactic/Relation/Trans.lean b/Mathlib/Tactic/Relation/Trans.lean index fd92d2cea3d62..df0b91d4958f8 100644 --- a/Mathlib/Tactic/Relation/Trans.lean +++ b/Mathlib/Tactic/Relation/Trans.lean @@ -21,11 +21,14 @@ open Lean Meta Elab initialize registerTraceClass `Tactic.trans +/-- Discrimation tree settings for the `trans` extension. -/ +def transExt.config : WhnfCoreConfig := {} + /-- Environment extension storing transitivity lemmas -/ initialize transExt : - SimpleScopedEnvExtension (Name × Array (DiscrTree.Key true)) (DiscrTree Name true) ← + SimpleScopedEnvExtension (Name × Array DiscrTree.Key) (DiscrTree Name) ← registerSimpleScopedEnvExtension { - addEntry := fun dt (n, ks) ↦ dt.insertCore ks n + addEntry := fun dt (n, ks) ↦ dt.insertCore ks n transExt.config initial := {} } @@ -43,7 +46,7 @@ initialize registerBuiltinAttribute { let some xyHyp := xs.pop.back? | fail let .app (.app _ _) _ ← inferType yzHyp | fail let .app (.app _ _) _ ← inferType xyHyp | fail - let key ← withReducible <| DiscrTree.mkPath rel + let key ← withReducible <| DiscrTree.mkPath rel transExt.config transExt.add (decl, key) kind } @@ -129,7 +132,9 @@ elab "trans" t?:(ppSpace colGt term)? : tactic => withMainContext do let t'? ← t?.mapM (elabTermWithHoles · ty (← getMainTag)) let s ← saveState trace[Tactic.trans]"trying homogeneous case" - for lem in (← (transExt.getState (← getEnv)).getUnify rel).push ``Trans.simple do + let lemmas := + (← (transExt.getState (← getEnv)).getUnify rel transExt.config).push ``Trans.simple + for lem in lemmas do trace[Tactic.trans]"trying lemma {lem}" try liftMetaTactic fun g ↦ do @@ -147,7 +152,7 @@ elab "trans" t?:(ppSpace colGt term)? : tactic => withMainContext do trace[Tactic.trans]"trying heterogeneous case" let t'? ← t?.mapM (elabTermWithHoles · none (← getMainTag)) let s ← saveState - for lem in (← (transExt.getState (← getEnv)).getUnify rel).push + for lem in (← (transExt.getState (← getEnv)).getUnify rel transExt.config).push ``HEq.trans |>.push ``HEq.trans do try liftMetaTactic fun g ↦ do diff --git a/Mathlib/Tactic/Replace.lean b/Mathlib/Tactic/Replace.lean index 95f70094f52df..317e30c22613f 100644 --- a/Mathlib/Tactic/Replace.lean +++ b/Mathlib/Tactic/Replace.lean @@ -49,7 +49,5 @@ elab_rules : tactic let name := optBinderIdent.name n let hId? := (← getLCtx).findFromUserName? name |>.map fun d ↦ d.fvarId match hId? with - | some hId => - try replaceMainGoal [goal1, ← goal2.clear hId] - catch | _ => pure () + | some hId => replaceMainGoal [goal1, (← observing? <| goal2.clear hId).getD goal2] | none => replaceMainGoal [goal1, goal2] diff --git a/Mathlib/Tactic/RewriteSearch.lean b/Mathlib/Tactic/RewriteSearch.lean new file mode 100644 index 0000000000000..ba11a43575115 --- /dev/null +++ b/Mathlib/Tactic/RewriteSearch.lean @@ -0,0 +1,312 @@ +/- +Copyright (c) 2023 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ +import Mathlib.Data.List.EditDistance.Estimator +import Mathlib.Data.MLList.BestFirst +import Mathlib.Data.Nat.Interval + +/-! +# The `rw_search` tactic + +`rw_search` attempts to solve an equality goal +by repeatedly rewriting using lemmas from the library. + +If no solution is found, +the best sequence of rewrites found before `maxHeartbeats` elapses is returned. + +The search is a best-first search, minimising the Levenshtein edit distance between +the pretty-printed expressions on either side of the equality. +(The strings are tokenized at spaces, +separating delimiters `(`, `)`, `[`, `]`, and `,` into their own tokens.) + +The implementation avoids completely computing edit distances where possible, +only computing lower bounds sufficient to decide which path to take in the search. + +# Future improvements + +We could call `simp` as an atomic step of rewriting. + +The edit distance heuristic could be replaced by something else. + +No effort has been made to choose the best tokenization scheme, and this should be investigated. +Moreover, the Levenshtein distance function is customizable with different weights for each token, +and it would be interesting to try optimizing these +(or dynamically updating them, +adding weight to tokens that persistently appear on one side of the equation but not the other.) + +The `rw_search` tactic will rewrite by local hypotheses, +but will not use local hypotheses to discharge side conditions. +This limitation would need to be resolved in the `rw?` tactic first. + +-/ + +set_option autoImplicit true + +namespace Mathlib.Tactic.RewriteSearch + +open Lean Meta +open Mathlib.Tactic.Rewrites + +initialize registerTraceClass `rw_search +initialize registerTraceClass `rw_search.detail + +/-- Separate a string into a list of strings by pulling off initial `(` or `]` characters, +and pulling off terminal `)`, `]`, or `,` characters. -/ +partial def splitDelimiters (s : String) : List String := + let rec /-- Pull off leading delimiters. -/ auxStart front pre := + let head := s.get front + if head = '(' || head = '[' then + auxStart (s.next front) (head.toString :: pre) + else + (front, pre) + let rec /-- Pull off trailing delimiters. -/ auxEnd back suff := + let last := s.get back + if last = ')' || last = ']' || last = ',' then + auxEnd (s.prev back) (last.toString :: suff) + else + (back, suff) + let (frontAfterStart, pre) := auxStart 0 [] + let (backAfterEnd, suff) := auxEnd (s.prev s.endPos) [] + pre.reverse ++ [s.extract frontAfterStart (s.next backAfterEnd)] ++ suff + +/-- +Tokenize a string at whitespace, and then pull off delimiters. +-/ +-- `rw_search` seems to return better results if we tokenize +-- rather than just breaking the string into characters, +-- but more extensive testing would be great! +-- Tokenizing of course makes the edit distance calculations cheaper, +-- because the lists are significantly shorter. +-- On the other hand, because we currently use unweighted edit distance, +-- this makes it "cheap" to change one identifier for another. +def tokenize (e : Expr) : MetaM (List String) := do + let s := (← ppExpr e).pretty + return s.splitOn.map splitDelimiters |>.join + +/-- +Data structure containing the history of a rewrite search. +-/ +structure SearchNode where mk' :: + /-- The lemmas used so far. -/ + -- The first component is the index amongst successful rewrites at the previous node. + history : Array (Nat × Expr × Bool) + /-- The metavariable context after rewriting. + We carry this around so the search can safely backtrack. -/ + mctx : MetavarContext + /-- The current goal. -/ + goal : MVarId + /-- The type of the current goal. -/ + type : Expr + /-- The pretty printed current goal. -/ + ppGoal : String + /-- The tokenization of the left-hand-side of the current goal. -/ + lhs : List String + /-- The tokenization of the right-hand-side of the current goal. -/ + rhs : List String + /-- Whether the current goal can be closed by `rfl` (or `none` if this hasn't been test yet). -/ + rfl? : Option Bool := none + /-- The edit distance between the tokenizations of the two sides + (or `none` if this hasn't been computed yet). -/ + dist? : Option Nat := none +namespace SearchNode + +/-- +What is the cost for changing a token? +`Levenshtein.defaultCost` just uses constant cost `1` for any token. + +It may be interesting to try others. +the only one I've experimented with so far is `Levenshtein.stringLogLengthCost`, +which performs quite poorly! +-/ +def editCost : Levenshtein.Cost String String Nat := Levenshtein.defaultCost + +/-- Check whether a goal can be solved by `rfl`, and fill in the `SearchNode.rfl?` field. -/ +def compute_rfl? (n : SearchNode) : MetaM SearchNode := withMCtx n.mctx do + if (← try? n.goal.applyRfl).isSome then + pure { n with mctx := ← getMCtx, rfl? := some true } + else + pure { n with rfl? := some false } + +/-- Fill in the `SearchNode.dist?` field with the edit distance between the two sides. -/ +def compute_dist? (n : SearchNode) : SearchNode := + match n.dist? with + | some _ => n + | none => + { n with dist? := some (levenshtein editCost n.lhs n.rhs) } + +/-- Represent a search node as string, solely for debugging. -/ +def toString (n : SearchNode) : MetaM String := do + let n := n.compute_dist? + let tac ← match n.history.back? with + | some (_, e, true) => do let pp ← ppExpr e; pure s!"rw [← {pp}]" + | some (_, e, false) => do let pp ← ppExpr e; pure s!"rw [{pp}]" + | none => pure "" + return s!"depth: {n.history.size}\n" ++ + s!"history: {n.history.map fun p => hash p % 10000}\n" ++ + tac ++ "\n" ++ + "-- " ++ n.ppGoal ++ "\n" ++ + s!"distance: {n.dist?.get!}+{n.history.size}, {n.ppGoal.length}" + +/-- Construct a `SearchNode`. -/ +def mk (history : Array (Nat × Expr × Bool)) (goal : MVarId) (ctx : Option MetavarContext := none) : + MetaM (Option SearchNode) := goal.withContext do + let type ← whnfR (← instantiateMVars (← goal.getType)) + match type.eq? with + | none => return none + | some (_, lhs, rhs) => + let lhsTokens ← tokenize lhs + let rhsTokens ← tokenize rhs + let r := + { history := history + mctx := ← ctx.getDM getMCtx + goal := goal + type := type + ppGoal := (← ppExpr type).pretty + lhs := lhsTokens + rhs := rhsTokens } + return some r + +/-- Construct an initial `SearchNode` from a goal. -/ +def init (goal : MVarId) : MetaM (Option SearchNode) := mk #[] goal +/-- Add an additional step to the `SearchNode` history. -/ +def push (n : SearchNode) (expr : Expr) (symm : Bool) (k : Nat) (g : MVarId) + (ctx : Option MetavarContext := none) : MetaM (Option SearchNode) := + mk (n.history.push (k, expr, symm)) g ctx + +/-- Report the index of the most recently applied lemma, in the ordering returned by `rw?`. -/ +def lastIdx (n : SearchNode) : Nat := + match n.history.back? with + | some (k, _) => k + | none => 0 + +instance : Ord SearchNode where + compare := compareOn fun n => toLex (toLex (n.ppGoal.length, n.lastIdx), n.ppGoal) + +/-- +A somewhat arbitrary penalty function. +Note that `n.lastIdx` penalizes using later lemmas from a particular call to `rw?` at a node, +but once we have moved on to the next node these penalties are "forgiven". + +(You might in interpret this as encouraging +the algorithm to "trust" the ordering provided by `rw?`.) + +I tried out a various (positive) linear combinations of +`.history.size`, `.lastIdx`, and `.ppGoal.length` (and also the `.log2`s of these). +* `.lastIdx.log2` is quite good, and the best coefficient is around 1. +* `.lastIdx / 10` is almost as good. +* `.history.size` makes things worse (similarly with `.log2`). +* `.ppGoal.length` makes little difference (similarly with `.log2`). +Here testing consisting of running the current `rw_search` test suite, +rejecting values for which any failed, and trying to minimize the run time reported by +```shell +lake build && \ +time (lake env lean test/RewriteSearch/Basic.lean; \ + lake env lean test/RewriteSearch/Polynomial.lean) +``` + +With a larger test suite it might be worth running this minimization again, +and considering other penalty functions. + +(If you do this, please choose a penalty function which is in the interior of the region +where the test suite works. +I think it would be a bad idea to optimize the run time at the expense of fragility.) +-/ +def penalty (n : SearchNode) : Nat := n.lastIdx.log2 + n.ppGoal.length.log2 + +/-- The priority function for search is Levenshtein distance plus a penalty. -/ +abbrev prio (n : SearchNode) : Thunk Nat := + (Thunk.pure n.penalty) + (Thunk.mk fun _ => levenshtein editCost n.lhs n.rhs) +/-- We can obtain lower bounds, and improve them, for the Levenshtein distance. -/ +abbrev estimator (n : SearchNode) : Type := + Estimator.trivial n.penalty × LevenshteinEstimator editCost n.lhs n.rhs + +/-- Given a `RewriteResult` from the `rw?` tactic, create a new `SearchNode` with the new goal. -/ +def rewrite (n : SearchNode) (r : Rewrites.RewriteResult) (k : Nat) : + MetaM (Option SearchNode) := + withMCtx r.mctx do + let goal' ← n.goal.replaceTargetEq r.result.eNew r.result.eqProof + n.push r.expr r.symm k goal' (← getMCtx) + +/-- +Given a pair of `DiscrTree` trees +indexing all rewrite lemmas in the imported files and the current file, +try rewriting the current goal in the `SearchNode` by one of them, +returning a `MLList MetaM SearchNode`, i.e. a lazy list of next possible goals. +-/ +def rewrites (hyps : Array (Expr × Bool × Nat)) + (lemmas : DiscrTree (Name × Bool × Nat) × DiscrTree (Name × Bool × Nat)) + (forbidden : NameSet := ∅) (n : SearchNode) : MLList MetaM SearchNode := .squash fun _ => do + if ← isTracingEnabledFor `rw_search then do + trace[rw_search] "searching:\n{← toString n}" + return rewritesCore hyps lemmas n.mctx n.goal n.type forbidden + |>.enum + |>.filterMapM fun ⟨k, r⟩ => do n.rewrite r k + +/-- +Perform best first search on the graph of rewrites from the specified `SearchNode`. +-/ +def search (n : SearchNode) + (stopAtRfl := true) (stopAtDistZero := true) + (forbidden : NameSet := ∅) (maxQueued : Option Nat := none) : + MLList MetaM SearchNode := .squash fun _ => do + let hyps ← localHypotheses + let lemmas ← rewriteLemmas.get + let search := bestFirstSearchCore (maxQueued := maxQueued) + (β := String) (removeDuplicatesBy? := some SearchNode.ppGoal) + prio estimator (rewrites hyps lemmas forbidden) n + let search ← + if ← isTracingEnabledFor `rw_search then do + pure <| search.mapM fun n => do trace[rw_search] "{← toString n}"; pure n + else + pure search + let search := if stopAtRfl then + search.mapM compute_rfl? |>.takeUpToFirst fun n => n.rfl? = some true + else + search + return if stopAtDistZero then + search.map (fun r => r.compute_dist?) |>.takeUpToFirst (fun r => r.dist? = some 0) + else + search + +end SearchNode + +open Elab Tactic + +/-- +`rw_search` attempts to solve an equality goal +by repeatedly rewriting using lemmas from the library. + +If no solution is found, +the best sequence of rewrites found before `maxHeartbeats` elapses is returned. + +The search is a best-first search, minimising the Levenshtein edit distance between +the pretty-printed expressions on either side of the equality. +(The strings are tokenized at spaces, +separating delimiters `(`, `)`, `[`, `]`, and `,` into their own tokens.) + +You can use `rw_search [-my_lemma, -my_theorem]` +to prevent `rw_search` from using the names theorems. +-/ +syntax "rw_search" (forbidden)? : tactic + +elab_rules : tactic | + `(tactic| rw_search%$tk $[[ $[-$forbidden],* ]]?) => withMainContext do + let forbidden : NameSet := + ((forbidden.getD #[]).map Syntax.getId).foldl (init := ∅) fun s n => s.insert n + let .some init ← SearchNode.init (← getMainGoal) | throwError "Goal is not an equality." + let results := init.search (forbidden := forbidden) |>.whileAtLeastHeartbeatsPercent 20 + let results ← results.force + let min ← match results with + | [] => failure + | h :: t => + pure <| t.foldl (init := h) fun r₁ r₂ => + if r₁.rfl? = some true then r₁ + else if r₂.rfl? = some true then r₂ + else if r₁.dist?.getD 0 ≤ r₂.dist?.getD 0 then r₁ else r₂ + setMCtx min.mctx + replaceMainGoal [min.goal] + let type? := if min.rfl? = some true then none else some (← min.goal.getType) + addRewriteSuggestion tk (min.history.toList.map (·.2)) + type? (origSpan? := ← getRef) diff --git a/Mathlib/Tactic/Rewrites.lean b/Mathlib/Tactic/Rewrites.lean index b84c2348b1be5..33573fa80c66f 100644 --- a/Mathlib/Tactic/Rewrites.lean +++ b/Mathlib/Tactic/Rewrites.lean @@ -8,7 +8,7 @@ import Std.Data.MLList.Heartbeats import Std.Tactic.Relation.Rfl import Mathlib.Data.MLList.Dedup import Mathlib.Lean.Meta.DiscrTree -import Mathlib.Tactic.Cache +import Std.Util.Cache import Mathlib.Lean.Meta import Mathlib.Tactic.TryThis import Mathlib.Control.Basic @@ -53,9 +53,12 @@ def forwardWeight := 2 /-- Weight to multiply the "specificity" of a rewrite lemma by when rewriting backwards. -/ def backwardWeight := 1 +/-- Configuration for `DiscrTree`. -/ +def discrTreeConfig : WhnfCoreConfig := {} + /-- Prepare the discrimination tree entries for a lemma. -/ def processLemma (name : Name) (constInfo : ConstantInfo) : - MetaM (Array (Array (DiscrTree.Key true) × (Name × Bool × Nat))) := do + MetaM (Array (Array DiscrTree.Key × (Name × Bool × Nat))) := do if constInfo.isUnsafe then return #[] if ← name.isBlackListed then return #[] -- We now remove some injectivity lemmas which are not useful to rewrite by. @@ -69,8 +72,8 @@ def processLemma (name : Name) (constInfo : ConstantInfo) : match type.getAppFnArgs with | (``Eq, #[_, lhs, rhs]) | (``Iff, #[lhs, rhs]) => do - let lhsKey ← DiscrTree.mkPath lhs - let rhsKey ← DiscrTree.mkPath rhs + let lhsKey ← DiscrTree.mkPath lhs discrTreeConfig + let rhsKey ← DiscrTree.mkPath rhs discrTreeConfig return #[(lhsKey, (name, false, forwardWeight * lhsKey.size)), (rhsKey, (name, true, backwardWeight * rhsKey.size))] | _ => return #[] @@ -82,11 +85,12 @@ def localHypotheses (except : List FVarId := []) : MetaM (Array (Expr × Bool × for h in r do if except.contains h.fvarId! then continue let (_, _, type) ← forallMetaTelescopeReducing (← inferType h) + let type ← whnfR type match type.getAppFnArgs with | (``Eq, #[_, lhs, rhs]) | (``Iff, #[lhs, rhs]) => do - let lhsKey : Array (DiscrTree.Key true) ← DiscrTree.mkPath lhs - let rhsKey : Array (DiscrTree.Key true) ← DiscrTree.mkPath rhs + let lhsKey : Array DiscrTree.Key ← DiscrTree.mkPath lhs discrTreeConfig + let rhsKey : Array DiscrTree.Key ← DiscrTree.mkPath rhs discrTreeConfig result := result.push (h, false, forwardWeight * lhsKey.size) |>.push (h, true, backwardWeight * rhsKey.size) | _ => pure () @@ -98,10 +102,10 @@ def localHypotheses (except : List FVarId := []) : MetaM (Array (Expr × Bool × -- `build/lib/MathlibExtras/Rewrites.extra` -- so that the cache is rebuilt. def addLemma (name : Name) (constInfo : ConstantInfo) - (lemmas : DiscrTree (Name × Bool × Nat) true) : MetaM (DiscrTree (Name × Bool × Nat) true) := do + (lemmas : DiscrTree (Name × Bool × Nat)) : MetaM (DiscrTree (Name × Bool × Nat)) := do let mut lemmas := lemmas for (key, value) in ← processLemma name constInfo do - lemmas := lemmas.insertIfSpecific key value + lemmas := lemmas.insertIfSpecific key value discrTreeConfig return lemmas /-- Construct the discrimination tree of all lemmas. -/ @@ -130,7 +134,7 @@ Retrieve the current cache of lemmas. initialize rewriteLemmas : DiscrTreeCache (Name × Bool × Nat) ← unsafe do let path ← cachePath if (← path.pathExists) then - let (d, _r) ← unpickle (DiscrTree (Name × Bool × Nat) true) path + let (d, _r) ← unpickle (DiscrTree (Name × Bool × Nat)) path -- We can drop the `CompactedRegion` value; we do not plan to free it DiscrTreeCache.mk "rw?: using cache" processLemma (init := some d) else @@ -213,12 +217,14 @@ See also `rewrites` for a more convenient interface. -- because `MLList.squash` executes lazily, -- so there is no opportunity for `← getMCtx` to record the context at the call site. def rewritesCore (hyps : Array (Expr × Bool × Nat)) - (lemmas : DiscrTree (Name × Bool × Nat) s × DiscrTree (Name × Bool × Nat) s) - (ctx : MetavarContext) (goal : MVarId) (target : Expr) (side : SideConditions := .solveByElim) : + (lemmas : DiscrTree (Name × Bool × Nat) × DiscrTree (Name × Bool × Nat)) + (ctx : MetavarContext) (goal : MVarId) (target : Expr) + (forbidden : NameSet := ∅) + (side : SideConditions := .solveByElim) : MLList MetaM RewriteResult := MLList.squash fun _ => do -- Get all lemmas which could match some subexpression - let candidates := (← lemmas.1.getSubexpressionMatches target) - ++ (← lemmas.2.getSubexpressionMatches target) + let candidates := (← lemmas.1.getSubexpressionMatches target discrTreeConfig) + ++ (← lemmas.2.getSubexpressionMatches target discrTreeConfig) -- Sort them by our preferring weighting -- (length of discriminant key, doubled for the forward implication) @@ -230,6 +236,7 @@ def rewritesCore (hyps : Array (Expr × Bool × Nat)) let mut backward : NameSet := ∅ let mut deduped := #[] for (l, s, w) in candidates do + if forbidden.contains l then continue if s then if ¬ backward.contains l then deduped := deduped.push (l, s, w) @@ -276,11 +283,11 @@ Find lemmas which can rewrite the goal, and deduplicate based on pretty-printed Note that this builds a `HashMap` containing the results, and so may consume significant memory. -/ def rewritesDedup (hyps : Array (Expr × Bool × Nat)) - (lemmas : DiscrTree (Name × Bool × Nat) s × DiscrTree (Name × Bool × Nat) s) - (mctx : MetavarContext) - (goal : MVarId) (target : Expr) (side : SideConditions := .solveByElim) : - MLList MetaM RewriteResult := MLList.squash fun _ => do - return rewritesCore hyps lemmas mctx goal target side + (lemmas : DiscrTree (Name × Bool × Nat) × DiscrTree (Name × Bool × Nat)) + (mctx : MetavarContext) (goal : MVarId) (target : Expr) + (forbidden : NameSet := ∅) (side : SideConditions := .solveByElim) : + MLList MetaM RewriteResult := + rewritesCore hyps lemmas mctx goal target forbidden side -- Don't report duplicate results. -- (TODO: a config flag to disable this, -- if distinct-but-pretty-print-the-same results are desirable?) @@ -289,11 +296,12 @@ def rewritesDedup (hyps : Array (Expr × Bool × Nat)) /-- Find lemmas which can rewrite the goal. -/ def rewrites (hyps : Array (Expr × Bool × Nat)) - (lemmas : DiscrTree (Name × Bool × Nat) s × DiscrTree (Name × Bool × Nat) s) + (lemmas : DiscrTree (Name × Bool × Nat) × DiscrTree (Name × Bool × Nat)) (goal : MVarId) (target : Expr) - (side : SideConditions := .solveByElim) (stopAtRfl : Bool := false) (max : Nat := 20) + (forbidden : NameSet := ∅) (side : SideConditions := .solveByElim) + (stopAtRfl : Bool := false) (max : Nat := 20) (leavePercentHeartbeats : Nat := 10) : MetaM (List RewriteResult) := do - let results ← rewritesDedup hyps lemmas (← getMCtx) goal target side + let results ← rewritesDedup hyps lemmas (← getMCtx) goal target forbidden side -- Stop if we find a rewrite after which `with_reducible rfl` would succeed. |>.mapM RewriteResult.computeRfl -- TODO could simply not compute this if `stopAtRfl` is False |>.takeUpToFirst (fun r => stopAtRfl && r.rfl? = some true) @@ -311,19 +319,29 @@ def rewrites (hyps : Array (Expr × Bool × Nat)) open Lean.Parser.Tactic +/-- +Syntax for excluding some names, e.g. `[-my_lemma, -my_theorem]`. +-/ +-- TODO: allow excluding local hypotheses. +syntax forbidden := " [" (("-" ident),*,?) "]" + /-- `rw?` tries to find a lemma which can rewrite the goal. `rw?` should not be left in proofs; it is a search tool, like `apply?`. Suggestions are printed as `rw [h]` or `rw [←h]`. + +You can use `rw? [-my_lemma, -my_theorem]` to prevent `rw?` using the named lemmas. -/ -syntax (name := rewrites') "rw?" (ppSpace location)? : tactic +syntax (name := rewrites') "rw?" (ppSpace location)? (forbidden)? : tactic open Elab.Tactic Elab Tactic in elab_rules : tactic | - `(tactic| rw?%$tk $[$loc]?) => do + `(tactic| rw?%$tk $[$loc]? $[[ $[-$forbidden],* ]]?) => do let lems ← rewriteLemmas.get + let forbidden : NameSet := + ((forbidden.getD #[]).map Syntax.getId).foldl (init := ∅) fun s n => s.insert n reportOutOfHeartbeats `rewrites tk let goal ← getMainGoal -- TODO fix doc of core to say that * fails only if all failed @@ -333,7 +351,7 @@ elab_rules : tactic | if a.isImplementationDetail then return let target ← instantiateMVars (← f.getType) let hyps ← localHypotheses (except := [f]) - let results ← rewrites hyps lems goal target (stopAtRfl := false) + let results ← rewrites hyps lems goal target (stopAtRfl := false) forbidden reportOutOfHeartbeats `rewrites tk if results.isEmpty then throwError "Could not find any lemmas which can rewrite the hypothesis { @@ -349,7 +367,7 @@ elab_rules : tactic | do withMainContext do let target ← instantiateMVars (← goal.getType) let hyps ← localHypotheses - let results ← rewrites hyps lems goal target (stopAtRfl := true) + let results ← rewrites hyps lems goal target (stopAtRfl := true) forbidden reportOutOfHeartbeats `rewrites tk if results.isEmpty then throwError "Could not find any lemmas which can rewrite the goal" diff --git a/Mathlib/Tactic/Ring.lean b/Mathlib/Tactic/Ring.lean index 1cedfb08c1db5..1a19f9b955fc7 100644 --- a/Mathlib/Tactic/Ring.lean +++ b/Mathlib/Tactic/Ring.lean @@ -1,2 +1,3 @@ import Mathlib.Tactic.Ring.Basic import Mathlib.Tactic.Ring.RingNF +import Mathlib.Tactic.Ring.PNat diff --git a/Mathlib/Tactic/Ring/Basic.lean b/Mathlib/Tactic/Ring/Basic.lean index d2efa09403bc8..fc46d76e953e7 100644 --- a/Mathlib/Tactic/Ring/Basic.lean +++ b/Mathlib/Tactic/Ring/Basic.lean @@ -1086,6 +1086,29 @@ partial def eval {u} {α : Q(Type u)} (sα : Q(CommSemiring $α)) | _ => els | _, _, _ => els +/-- `CSLift α β` is a typeclass used by `ring` for lifting operations from `α` +(which is not a commutative semiring) into a commutative semiring `β` by using an injective map +`lift : α → β`. -/ +class CSLift (α : Type u) (β : outParam (Type u)) where + /-- `lift` is the "canonical injection" from `α` to `β` -/ + lift : α → β + /-- `lift` is an injective function -/ + inj : Function.Injective lift + +/-- `CSLiftVal a b` means that `b = lift a`. This is used by `ring` to construct an expression `b` +from the input expression `a`, and then run the usual ring algorithm on `b`. -/ +class CSLiftVal {α} {β : outParam (Type u)} [CSLift α β] (a : α) (b : outParam β) : Prop where + /-- The output value `b` is equal to the lift of `a`. This can be supplied by the default + instance which sets `b := lift a`, but `ring` will treat this as an atom so it is more useful + when there are other instances which distribute addition or multiplication. -/ + eq : b = CSLift.lift a + +instance (priority := low) {α β} [CSLift α β] (a : α) : CSLiftVal a (CSLift.lift a) := ⟨rfl⟩ + +theorem of_lift {α β} [inst : CSLift α β] {a b : α} {a' b' : β} + [h1 : CSLiftVal a a'] [h2 : CSLiftVal b b'] (h : a' = b') : a = b := + inst.2 <| by rwa [← h1.1, ← h2.1] + open Lean Parser.Tactic Elab Command Elab.Tactic Meta Qq theorem of_eq (_ : (a : R) = c) (_ : b = c) : a = b := by subst_vars; rfl @@ -1104,17 +1127,39 @@ def proveEq (g : MVarId) : AtomM Unit := do let .sort u ← whnf (← inferType α) | unreachable! let v ← try u.dec catch _ => throwError "not a type{indentExpr α}" have α : Q(Type v) := α + let sα ← + try Except.ok <$> synthInstanceQ q(CommSemiring $α) + catch e => pure (.error e) have e₁ : Q($α) := e₁; have e₂ : Q($α) := e₂ - let sα ← synthInstanceQ (q(CommSemiring $α) : Q(Type v)) - let c ← mkCache sα - profileitM Exception "ring" (← getOptions) do - let ⟨a, va, pa⟩ ← eval sα c e₁ - let ⟨b, vb, pb⟩ ← eval sα c e₂ - unless va.eq vb do - let g ← mkFreshExprMVar (← (← ringCleanupRef.get) q($a = $b)) - throwError "ring failed, ring expressions not equal\n{g.mvarId!}" - let pb : Q($e₂ = $a) := pb - g.assign q(of_eq $pa $pb) + let eq ← match sα with + | .ok sα => ringCore sα e₁ e₂ + | .error e => + let β ← mkFreshExprMVarQ q(Type v) + let e₁' ← mkFreshExprMVarQ q($β) + let e₂' ← mkFreshExprMVarQ q($β) + let (sβ, (pf : Q($e₁' = $e₂' → $e₁ = $e₂))) ← try + let _l ← synthInstanceQ q(CSLift $α $β) + let sβ ← synthInstanceQ q(CommSemiring $β) + let _ ← synthInstanceQ q(CSLiftVal $e₁ $e₁') + let _ ← synthInstanceQ q(CSLiftVal $e₂ $e₂') + pure (sβ, q(of_lift (a := $e₁) (b := $e₂))) + catch _ => throw e + pure q($pf $(← ringCore sβ e₁' e₂')) + g.assign eq +where + /-- The core of `proveEq` takes expressions `e₁ e₂ : α` where `α` is a `CommSemiring`, + and returns a proof that they are equal (or fails). -/ + ringCore {v : Level} {α : Q(Type v)} (sα : Q(CommSemiring $α)) + (e₁ e₂ : Q($α)) : AtomM Q($e₁ = $e₂) := do + let c ← mkCache sα + profileitM Exception "ring" (← getOptions) do + let ⟨a, va, pa⟩ ← eval sα c e₁ + let ⟨b, vb, pb⟩ ← eval sα c e₂ + unless va.eq vb do + let g ← mkFreshExprMVar (← (← ringCleanupRef.get) q($a = $b)) + throwError "ring failed, ring expressions not equal\n{g.mvarId!}" + let pb : Q($e₂ = $a) := pb + return q(of_eq $pa $pb) /-- Tactic for solving equations of *commutative* (semi)rings, diff --git a/Mathlib/Tactic/Ring/PNat.lean b/Mathlib/Tactic/Ring/PNat.lean new file mode 100644 index 0000000000000..25291a31ef2f1 --- /dev/null +++ b/Mathlib/Tactic/Ring/PNat.lean @@ -0,0 +1,41 @@ +/- +Copyright (c) 2023 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import Mathlib.Tactic.Ring.Basic +import Mathlib.Data.PNat.Basic + +/-! +# Additional instances for `ring` over `PNat` + +This adds some instances which enable `ring` to work on `PNat` even though it is not a commutative +semiring, by lifting to `Nat`. +-/ + +namespace Mathlib.Tactic.Ring + +instance : CSLift ℕ+ Nat where + lift := PNat.val + inj := PNat.coe_injective + +-- FIXME: this `no_index` seems to be in the wrong place, but +-- #synth CSLiftVal (3 : ℕ+) _ doesn't work otherwise +instance {n} : CSLiftVal (no_index (OfNat.ofNat (n+1)) : ℕ+) (n + 1) := ⟨rfl⟩ + +instance {n h} : CSLiftVal (Nat.toPNat n h) n := ⟨rfl⟩ + +instance {n} : CSLiftVal (Nat.succPNat n) (n + 1) := ⟨rfl⟩ + +instance {n} : CSLiftVal (Nat.toPNat' n) (n.pred + 1) := ⟨rfl⟩ + +instance {n k} : CSLiftVal (PNat.divExact n k) (n.div k + 1) := ⟨rfl⟩ + +instance {n n' k k'} [h1 : CSLiftVal (n : ℕ+) n'] [h2 : CSLiftVal (k : ℕ+) k'] : + CSLiftVal (n + k) (n' + k') := ⟨by simp [h1.1, h2.1, CSLift.lift]⟩ + +instance {n n' k k'} [h1 : CSLiftVal (n : ℕ+) n'] [h2 : CSLiftVal (k : ℕ+) k'] : + CSLiftVal (n * k) (n' * k') := ⟨by simp [h1.1, h2.1, CSLift.lift]⟩ + +instance {n n' k} [h1 : CSLiftVal (n : ℕ+) n'] : + CSLiftVal (n ^ k) (n' ^ k) := ⟨by simp [h1.1, CSLift.lift]⟩ diff --git a/Mathlib/Tactic/Simps/Basic.lean b/Mathlib/Tactic/Simps/Basic.lean index fa5048aa5dd6a..5778f4bf8cffe 100644 --- a/Mathlib/Tactic/Simps/Basic.lean +++ b/Mathlib/Tactic/Simps/Basic.lean @@ -4,17 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Mathlib.Init.Data.Nat.Notation -import Mathlib.Lean.Message -import Mathlib.Lean.Expr.Basic import Mathlib.Data.KVMap import Mathlib.Tactic.Simps.NotationClass import Std.Classes.Dvd import Std.Data.String.Basic import Std.Util.LibraryNote -import Mathlib.Tactic.RunCmd -- not necessary, but useful for debugging import Mathlib.Lean.Linter import Std.Data.List.Count +import Mathlib.Lean.Expr.Basic /-! # Simps attribute @@ -369,7 +366,7 @@ structure ProjectionData where target of the first projection is a structure with at least two projections). The composition of these projections is required to be definitionally equal to the provided Expression. -/ - projNrs : List ℕ + projNrs : List Nat /-- A boolean specifying whether `simp` lemmas are generated for this projection by default. -/ isDefault : Bool /-- A boolean specifying whether this projection is written as prefix. -/ @@ -469,7 +466,7 @@ def projectionsInfo (l : List ProjectionData) (pref : String) (str : Name) : Mes /-- Find the indices of the projections that need to be applied to elaborate `$e.$projName`. Example: If `e : α ≃+ β` and ``projName = `invFun`` then this returns `[0, 1]`, because the first projection of `MulEquiv` is `toEquiv` and the second projection of `Equiv` is `invFun`. -/ -def findProjectionIndices (strName projName : Name) : MetaM (List ℕ) := do +def findProjectionIndices (strName projName : Name) : MetaM (List Nat) := do let env ← getEnv let .some baseStr := findField? env strName projName | throwError "{strName} has no field {projName} in parent structure" @@ -481,8 +478,8 @@ def findProjectionIndices (strName projName : Name) : MetaM (List ℕ) := do return allProjs.map (env.getProjectionFnInfo? · |>.get!.i) /-- Auxiliary function of `getCompositeOfProjections`. -/ -partial def getCompositeOfProjectionsAux - (proj : String) (e : Expr) (pos : Array ℕ) (args : Array Expr) : MetaM (Expr × Array ℕ) := do +partial def getCompositeOfProjectionsAux (proj : String) (e : Expr) (pos : Array Nat) + (args : Array Expr) : MetaM (Expr × Array Nat) := do let env ← getEnv let .const structName _ := (← whnf (←inferType e)).getAppFn | throwError "{e} doesn't have a structure as type" @@ -515,7 +512,7 @@ partial def getCompositeOfProjectionsAux we will be able to generate the "projection" `λ {A} (f : gradedFun A) (x : A i) (y : A j) ↦ ↑(↑(f.toFun i j) x) y`, which projection notation cannot do. -/ -def getCompositeOfProjections (structName : Name) (proj : String) : MetaM (Expr × Array ℕ) := do +def getCompositeOfProjections (structName : Name) (proj : String) : MetaM (Expr × Array Nat) := do let strExpr ← mkConstWithLevelParams structName let type ← inferType strExpr forallTelescopeReducing type fun typeArgs _ ↦ @@ -995,7 +992,7 @@ partial def headStructureEtaReduce (e : Expr) : MetaM Expr := do -/ partial def addProjections (nm : Name) (type lhs rhs : Expr) (args : Array Expr) (mustBeStr : Bool) (cfg : Config) - (todo : List (String × Syntax)) (toApply : List ℕ) : MetaM (Array Name) := do + (todo : List (String × Syntax)) (toApply : List Nat) : MetaM (Array Name) := do -- we don't want to unfold non-reducible definitions (like `set`) to apply more arguments trace[simps.debug] "Type of the Expression before normalizing: {type}" withTransparency cfg.typeMd <| forallTelescopeReducing type fun typeArgs tgt ↦ withDefault do diff --git a/Mathlib/Tactic/Simps/NotationClass.lean b/Mathlib/Tactic/Simps/NotationClass.lean index 1d326fad8cb73..5c16cb4328b21 100644 --- a/Mathlib/Tactic/Simps/NotationClass.lean +++ b/Mathlib/Tactic/Simps/NotationClass.lean @@ -4,10 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Std.Lean.NameMapAttribute -import Mathlib.Lean.Expr.Basic import Lean.Elab.Exception -import Qq.MetaM +import Std.Lean.NameMapAttribute +import Std.Lean.Expr import Std.Tactic.Lint /-! @@ -36,7 +35,7 @@ in the file where we declare `@[simps]`. For further documentation, see `Tactic. arguments of classes that use the projection. -/ syntax (name := notation_class) "notation_class" "*"? (ppSpace ident)? (ppSpace ident)? : attr -open Lean Meta Elab Term Qq +open Lean Meta Elab Term namespace Simps @@ -119,7 +118,7 @@ initialize notationClassAttr : NameMapExtension AutomaticProjectionData ← do match (← getEnv).find? findArgs with | none => throwError "no such declaration {findArgs}" | some declInfo => - unless ← MetaM.run' <| isDefEq declInfo.type q(findArgType) do + unless ← MetaM.run' <| isDefEq declInfo.type (mkConst ``findArgType) do throwError "declaration {findArgs} has wrong type" ext.add projName ⟨src, coercion.isNone, findArgs⟩ | _ => throwUnsupportedSyntax } diff --git a/Mathlib/Tactic/ToAdditive.lean b/Mathlib/Tactic/ToAdditive.lean index ca165fbfe65cd..0217a19889152 100644 --- a/Mathlib/Tactic/ToAdditive.lean +++ b/Mathlib/Tactic/ToAdditive.lean @@ -462,7 +462,7 @@ def expand (e : Expr) : MetaM Expr := do trace[to_additive_detail] "expanded {e} to {e'}" return .continue e' if e != e₂ then - trace[to_additive_detail] "expand:\nBefore: {e}\nAfter: {e₂}" + trace[to_additive_detail] "expand:\nBefore: {e}\nAfter: {e₂}" return e₂ /-- Reorder pi-binders. See doc of `reorderAttr` for the interpretation of the argument -/ @@ -951,7 +951,7 @@ partial def applyAttributes (stx : Syntax) (rawAttrs : Array Syntax) (thisAttr s let env ← getEnv match getAttributeImpl env attr.name with | Except.error errMsg => throwError errMsg - | Except.ok attrImpl => + | Except.ok attrImpl => let runAttr := do attrImpl.add src attr.stx attr.kind attrImpl.add tgt attr.stx attr.kind diff --git a/Mathlib/Tactic/Widget/Calc.lean b/Mathlib/Tactic/Widget/Calc.lean index 78e20c9f607bc..634be8df3c637 100644 --- a/Mathlib/Tactic/Widget/Calc.lean +++ b/Mathlib/Tactic/Widget/Calc.lean @@ -91,7 +91,7 @@ def suggestSteps (pos : Array Lean.SubExpr.GoalsLocation) (goalType : Expr) (par else s!"_ {relStr} {newLhsStr} := by sorry\n{spc}_ {relStr} {newRhsStr} := by sorry\n" ++ s!"{spc}_ {relStr} {rhsStr} := by sorry" - | false, true => + | false, true => if params.isFirst then s!"{lhsStr} {relStr} {newRhsStr} := by sorry\n{spc}_ {relStr} {rhsStr} := by sorry" else diff --git a/Mathlib/Tactic/Zify.lean b/Mathlib/Tactic/Zify.lean index e2a709ca938bb..178055cdcc419 100644 --- a/Mathlib/Tactic/Zify.lean +++ b/Mathlib/Tactic/Zify.lean @@ -5,15 +5,15 @@ Authors: Moritz Doll, Mario Carneiro, Robert Y. Lewis -/ import Mathlib.Tactic.Basic import Mathlib.Tactic.NormCast -import Mathlib.Data.Int.Basic +import Mathlib.Tactic.Attr.Register /-! # `zify` tactic -The `zify` tactic is used to shift propositions from `ℕ` to `ℤ`. -This is often useful since `ℤ` has well-behaved subtraction. +The `zify` tactic is used to shift propositions from `Nat` to `Int`. +This is often useful since `Int` has well-behaved subtraction. ``` -example (a b c x y z : ℕ) (h : ¬ x*y*z < 0) : c < a + 3*b := by +example (a b c x y z : Nat) (h : ¬ x*y*z < 0) : c < a + 3*b := by zify zify at h /- @@ -31,10 +31,10 @@ open Lean.Parser.Tactic open Lean.Elab.Tactic /-- -The `zify` tactic is used to shift propositions from `ℕ` to `ℤ`. -This is often useful since `ℤ` has well-behaved subtraction. +The `zify` tactic is used to shift propositions from `Nat` to `Int`. +This is often useful since `Int` has well-behaved subtraction. ``` -example (a b c x y z : ℕ) (h : ¬ x*y*z < 0) : c < a + 3*b := by +example (a b c x y z : Nat) (h : ¬ x*y*z < 0) : c < a + 3*b := by zify zify at h /- @@ -45,16 +45,18 @@ example (a b c x y z : ℕ) (h : ¬ x*y*z < 0) : c < a + 3*b := by `zify` can be given extra lemmas to use in simplification. This is especially useful in the presence of nat subtraction: passing `≤` arguments will allow `push_cast` to do more work. ``` -example (a b c : ℕ) (h : a - b < c) (hab : b ≤ a) : false := by +example (a b c : Nat) (h : a - b < c) (hab : b ≤ a) : false := by zify [hab] at h /- h : ↑a - ↑b < ↑c -/ ``` `zify` makes use of the `@[zify_simps]` attribute to move propositions, -and the `push_cast` tactic to simplify the `ℤ`-valued expressions. -`zify` is in some sense dual to the `lift` tactic. `lift (z : ℤ) to ℕ` will change the type of an -integer `z` (in the supertype) to `ℕ` (the subtype), given a proof that `z ≥ 0`; -propositions concerning `z` will still be over `ℤ`. `zify` changes propositions about `ℕ` (the -subtype) to propositions about `ℤ` (the supertype), without changing the type of any variable. +and the `push_cast` tactic to simplify the `Int`-valued expressions. +`zify` is in some sense dual to the `lift` tactic. +`lift (z : Int) to Nat` will change the type of an +integer `z` (in the supertype) to `Nat` (the subtype), given a proof that `z ≥ 0`; +propositions concerning `z` will still be over `Int`. +`zify` changes propositions about `Nat` (the subtype) to propositions about `Int` (the supertype), +without changing the type of any variable. -/ syntax (name := zify) "zify" (simpArgs)? (location)? : tactic @@ -90,11 +92,11 @@ def zifyProof (simpArgs : Option (Syntax.TSepArray `Lean.Parser.Tactic.simpStar let (r, _) ← simp prop ctx_result.ctx applySimpResultToProp' proof prop r -@[zify_simps] lemma nat_cast_eq (a b : ℕ) : a = b ↔ (a : ℤ) = (b : ℤ) := Int.ofNat_inj.symm -@[zify_simps] lemma nat_cast_le (a b : ℕ) : a ≤ b ↔ (a : ℤ) ≤ (b : ℤ) := Int.ofNat_le.symm -@[zify_simps] lemma nat_cast_lt (a b : ℕ) : a < b ↔ (a : ℤ) < (b : ℤ) := Int.ofNat_lt.symm -@[zify_simps] lemma nat_cast_ne (a b : ℕ) : a ≠ b ↔ (a : ℤ) ≠ (b : ℤ) := by - simp only [ne_eq, Int.cast_eq_cast_iff_Nat] -@[zify_simps] lemma nat_cast_dvd (a b : ℕ) : a ∣ b ↔ (a : ℤ) ∣ (b : ℤ) := Int.ofNat_dvd.symm +@[zify_simps] lemma nat_cast_eq (a b : Nat) : a = b ↔ (a : Int) = (b : Int) := Int.ofNat_inj.symm +@[zify_simps] lemma nat_cast_le (a b : Nat) : a ≤ b ↔ (a : Int) ≤ (b : Int) := Int.ofNat_le.symm +@[zify_simps] lemma nat_cast_lt (a b : Nat) : a < b ↔ (a : Int) < (b : Int) := Int.ofNat_lt.symm +@[zify_simps] lemma nat_cast_ne (a b : Nat) : a ≠ b ↔ (a : Int) ≠ (b : Int) := + not_congr Int.ofNat_inj.symm +@[zify_simps] lemma nat_cast_dvd (a b : Nat) : a ∣ b ↔ (a : Int) ∣ (b : Int) := Int.ofNat_dvd.symm -- TODO: is it worth adding lemmas for Prime and Coprime as well? -- Doing so in this file would require adding imports. diff --git a/Mathlib/Testing/SlimCheck/Gen.lean b/Mathlib/Testing/SlimCheck/Gen.lean index 64214d33bca39..7da83f3ed8198 100644 --- a/Mathlib/Testing/SlimCheck/Gen.lean +++ b/Mathlib/Testing/SlimCheck/Gen.lean @@ -55,7 +55,7 @@ def choose (α : Type u) [Preorder α] [BoundedRandom α] (lo hi : α) (h : lo lemma chooseNatLt_aux {lo hi : Nat} (a : Nat) (h : Nat.succ lo ≤ a ∧ a ≤ hi) : lo ≤ Nat.pred a ∧ Nat.pred a < hi := - And.intro (Nat.le_pred_of_lt (Nat.lt_of_succ_le h.left)) <| + And.intro (Nat.le_sub_one_of_lt (Nat.lt_of_succ_le h.left)) <| show a.pred.succ ≤ hi by rw [Nat.succ_pred_eq_of_pos] exact h.right diff --git a/Mathlib/Testing/SlimCheck/Testable.lean b/Mathlib/Testing/SlimCheck/Testable.lean index 79b992c257dfb..b96dc6deddd5f 100644 --- a/Mathlib/Testing/SlimCheck/Testable.lean +++ b/Mathlib/Testing/SlimCheck/Testable.lean @@ -413,16 +413,16 @@ instance LE.printableProp [Repr α] [LE α] {x y : α} : PrintableProp (x ≤ y) instance LT.printableProp [Repr α] [LT α] {x y : α} : PrintableProp (x < y) where printProp := s!"{repr x} < {repr y}" -instance And.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x ∧ y) where +instance And.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x ∧ y) where printProp := s!"{printProp x} ∧ {printProp y}" -instance Or.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x ∨ y) where +instance Or.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x ∨ y) where printProp := s!"{printProp x} ∨ {printProp y}" -instance Iff.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x ↔ y) where +instance Iff.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x ↔ y) where printProp := s!"{printProp x} ↔ {printProp y}" -instance Imp.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x → y) where +instance Imp.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x → y) where printProp := s!"{printProp x} → {printProp y}" instance Not.printableProp [PrintableProp x] : PrintableProp (¬x) where diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 339e31b443991..83180a8028e4d 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -694,11 +694,11 @@ instance (S : Subgroup G) : TopologicalGroup S := end Subgroup -/-- The (topological-space) closure of a subgroup of a space `M` with `ContinuousMul` is +/-- The (topological-space) closure of a subgroup of a topological group is itself a subgroup. -/ @[to_additive - "The (topological-space) closure of an additive subgroup of a space `M` with - `ContinuousAdd` is itself an additive subgroup."] + "The (topological-space) closure of an additive subgroup of an additive topological group is + itself an additive subgroup."] def Subgroup.topologicalClosure (s : Subgroup G) : Subgroup G := { s.toSubmonoid.topologicalClosure with carrier := _root_.closure (s : Set G) @@ -797,6 +797,11 @@ def Subgroup.commGroupTopologicalClosure [T2Space G] (s : Subgroup G) #align subgroup.comm_group_topological_closure Subgroup.commGroupTopologicalClosure #align add_subgroup.add_comm_group_topological_closure AddSubgroup.addCommGroupTopologicalClosure +variable (G) in +@[to_additive] +lemma Subgroup.coe_topologicalClosure_bot : + ((⊥ : Subgroup G).topologicalClosure : Set G) = _root_.closure ({1} : Set G) := by simp + @[to_additive exists_nhds_half_neg] theorem exists_nhds_split_inv {s : Set G} (hs : s ∈ 𝓝 (1 : G)) : ∃ V ∈ 𝓝 (1 : G), ∀ v ∈ V, ∀ w ∈ V, v / w ∈ s := by @@ -1378,7 +1383,7 @@ end ContinuousConstSMulOp section TopologicalGroup -variable [TopologicalSpace α] [Group α] [TopologicalGroup α] {s t : Set α} +variable [TopologicalSpace G] [Group G] [TopologicalGroup G] {s t : Set G} @[to_additive] theorem IsOpen.div_left (ht : IsOpen t) : IsOpen (s / t) := by @@ -1413,7 +1418,7 @@ theorem subset_interior_div : interior s / interior t ⊆ interior (s / t) := #align subset_interior_sub subset_interior_sub @[to_additive] -theorem IsOpen.mul_closure (hs : IsOpen s) (t : Set α) : s * closure t = s * t := by +theorem IsOpen.mul_closure (hs : IsOpen s) (t : Set G) : s * closure t = s * t := by refine' (mul_subset_iff.2 fun a ha b hb => _).antisymm (mul_subset_mul_left subset_closure) rw [mem_closure_iff] at hb have hbU : b ∈ s⁻¹ * {a * b} := ⟨a⁻¹, a * b, Set.inv_mem_inv.2 ha, rfl, inv_mul_cancel_left _ _⟩ @@ -1423,20 +1428,20 @@ theorem IsOpen.mul_closure (hs : IsOpen s) (t : Set α) : s * closure t = s * t #align is_open.add_closure IsOpen.add_closure @[to_additive] -theorem IsOpen.closure_mul (ht : IsOpen t) (s : Set α) : closure s * t = s * t := by +theorem IsOpen.closure_mul (ht : IsOpen t) (s : Set G) : closure s * t = s * t := by rw [← inv_inv (closure s * t), mul_inv_rev, inv_closure, ht.inv.mul_closure, mul_inv_rev, inv_inv, inv_inv] #align is_open.closure_mul IsOpen.closure_mul #align is_open.closure_add IsOpen.closure_add @[to_additive] -theorem IsOpen.div_closure (hs : IsOpen s) (t : Set α) : s / closure t = s / t := by +theorem IsOpen.div_closure (hs : IsOpen s) (t : Set G) : s / closure t = s / t := by simp_rw [div_eq_mul_inv, inv_closure, hs.mul_closure] #align is_open.div_closure IsOpen.div_closure #align is_open.sub_closure IsOpen.sub_closure @[to_additive] -theorem IsOpen.closure_div (ht : IsOpen t) (s : Set α) : closure s / t = s / t := by +theorem IsOpen.closure_div (ht : IsOpen t) (s : Set G) : closure s / t = s / t := by simp_rw [div_eq_mul_inv, ht.inv.closure_mul] #align is_open.closure_div IsOpen.closure_div #align is_open.closure_sub IsOpen.closure_sub @@ -1452,8 +1457,8 @@ theorem IsClosed.mul_right_of_isCompact (ht : IsClosed t) (hs : IsCompact s) : exact IsClosed.smul_left_of_isCompact ht (hs.image continuous_op) @[to_additive] -theorem QuotientGroup.isClosedMap_coe {H : Subgroup α} (hH : IsCompact (H : Set α)) : - IsClosedMap ((↑) : α → α ⧸ H) := by +theorem QuotientGroup.isClosedMap_coe {H : Subgroup G} (hH : IsCompact (H : Set G)) : + IsClosedMap ((↑) : G → G ⧸ H) := by intro t ht rw [← quotientMap_quotient_mk'.isClosed_preimage] convert ht.mul_right_of_isCompact hH @@ -1461,6 +1466,57 @@ theorem QuotientGroup.isClosedMap_coe {H : Subgroup α} (hH : IsCompact (H : Set rw [iUnion_subtype, ← iUnion_mul_right_image] rfl +@[to_additive] +lemma subset_mul_closure_one (s : Set G) : s ⊆ s * (closure {1} : Set G) := by + have : s ⊆ s * ({1} : Set G) := by simpa using Subset.rfl + exact this.trans (smul_subset_smul_left subset_closure) + +@[to_additive] +lemma IsCompact.mul_closure_one_eq_closure {K : Set G} (hK : IsCompact K) : + K * (closure {1} : Set G) = closure K := by + apply Subset.antisymm ?_ ?_ + · calc + K * (closure {1} : Set G) ⊆ closure K * (closure {1} : Set G) := + smul_subset_smul_right subset_closure + _ ⊆ closure (K * ({1} : Set G)) := smul_set_closure_subset _ _ + _ = closure K := by simp + · have : IsClosed (K * (closure {1} : Set G)) := + IsClosed.smul_left_of_isCompact isClosed_closure hK + rw [IsClosed.closure_subset_iff this] + exact subset_mul_closure_one K + +@[to_additive] +lemma IsClosed.mul_closure_one_eq {F : Set G} (hF : IsClosed F) : + F * (closure {1} : Set G) = F := by + refine Subset.antisymm ?_ (subset_mul_closure_one F) + calc + F * (closure {1} : Set G) = closure F * closure ({1} : Set G) := by rw [hF.closure_eq] + _ ⊆ closure (F * ({1} : Set G)) := smul_set_closure_subset _ _ + _ = F := by simp [hF.closure_eq] + +@[to_additive] +lemma compl_mul_closure_one_eq {t : Set G} (ht : t * (closure {1} : Set G) = t) : + tᶜ * (closure {1} : Set G) = tᶜ := by + refine Subset.antisymm ?_ (subset_mul_closure_one tᶜ) + rintro - ⟨x, g, hx, hg, rfl⟩ + by_contra H + have : x ∈ t * (closure {1} : Set G) := by + rw [← Subgroup.coe_topologicalClosure_bot G] at hg ⊢ + simp only [smul_eq_mul, mem_compl_iff, not_not] at H + exact ⟨x * g, g⁻¹, H, Subgroup.inv_mem _ hg, by simp⟩ + rw [ht] at this + exact hx this + +@[to_additive] +lemma compl_mul_closure_one_eq_iff {t : Set G} : + tᶜ * (closure {1} : Set G) = tᶜ ↔ t * (closure {1} : Set G) = t := + ⟨fun h ↦ by simpa using compl_mul_closure_one_eq h, fun h ↦ compl_mul_closure_one_eq h⟩ + +@[to_additive] +lemma IsOpen.mul_closure_one_eq {U : Set G} (hU : IsOpen U) : + U * (closure {1} : Set G) = U := + compl_mul_closure_one_eq_iff.1 (hU.isClosed_compl.mul_closure_one_eq) + end TopologicalGroup section FilterMul @@ -1726,6 +1782,59 @@ theorem exists_isCompact_isClosed_subset_isCompact_nhds_one rcases exists_mem_nhds_isClosed_subset L1 with ⟨K, hK, K_closed, KL⟩ exact ⟨K, Lcomp.of_isClosed_subset K_closed KL, K_closed, KL, hK⟩ +/-- If a point in a topological group has a compact neighborhood, then the group is +locally compact. -/ +@[to_additive] +theorem IsCompact.locallyCompactSpace_of_mem_nhds_of_group {K : Set G} (hK : IsCompact K) {x : G} + (h : K ∈ 𝓝 x) : LocallyCompactSpace G := by + refine ⟨fun y n hn ↦ ?_⟩ + have A : (y * x⁻¹) • K ∈ 𝓝 y := by + rw [← preimage_smul_inv] + exact (continuous_const_smul _).continuousAt.preimage_mem_nhds (by simpa using h) + rcases exists_mem_nhds_isClosed_subset (inter_mem A hn) with ⟨L, hL, L_closed, LK⟩ + refine ⟨L, hL, LK.trans (inter_subset_right _ _), ?_⟩ + exact (hK.smul (y * x⁻¹)).of_isClosed_subset L_closed (LK.trans (inter_subset_left _ _)) + +-- The next instance creates a loop between weakly locally compact space and locally compact space +-- for topological groups. Hopefully, it shouldn't create problems. +/-- A topological group which is weakly locally compact is automatically locally compact. -/ +@[to_additive] +instance (priority := 90) instLocallyCompactSpaceOfWeaklyOfGroup [WeaklyLocallyCompactSpace G] : + LocallyCompactSpace G := by + rcases exists_compact_mem_nhds (1 : G) with ⟨K, K_comp, hK⟩ + exact K_comp.locallyCompactSpace_of_mem_nhds_of_group hK + +/-- If a function defined on a topological group has a support contained in a +compact set, then either the function is trivial or the group is locally compact. -/ +@[to_additive + "If a function defined on a topological additive group has a support contained in a compact + set, then either the function is trivial or the group is locally compact."] +theorem eq_zero_or_locallyCompactSpace_of_support_subset_isCompact_of_group + [TopologicalSpace α] [Zero α] [T1Space α] + {f : G → α} {k : Set G} (hk : IsCompact k) (hf : support f ⊆ k) (h'f : Continuous f) : + f = 0 ∨ LocallyCompactSpace G := by + by_cases h : ∀ x, f x = 0 + · apply Or.inl + ext x + exact h x + apply Or.inr + push_neg at h + obtain ⟨x, hx⟩ : ∃ x, f x ≠ 0 := h + have : k ∈ 𝓝 x := + mem_of_superset (h'f.isOpen_support.mem_nhds hx) hf + exact IsCompact.locallyCompactSpace_of_mem_nhds_of_group hk this + +/-- If a function defined on a topological group has compact support, then either +the function is trivial or the group is locally compact. -/ +@[to_additive + "If a function defined on a topological additive group has compact support, + then either the function is trivial or the group is locally compact."] +theorem HasCompactSupport.eq_zero_or_locallyCompactSpace_of_group + [TopologicalSpace α] [Zero α] [T1Space α] + {f : G → α} (hf : HasCompactSupport f) (h'f : Continuous f) : + f = 0 ∨ LocallyCompactSpace G := + eq_zero_or_locallyCompactSpace_of_support_subset_isCompact_of_group hf (subset_tsupport f) h'f + /-- In a locally compact group, any neighborhood of the identity contains a compact closed neighborhood of the identity, even without separation assumptions on the space. -/ @[to_additive @@ -1750,6 +1859,18 @@ theorem exists_isCompact_isClosed_nhds_one [WeaklyLocallyCompactSpace G] : let ⟨K, Kcl, Kcomp, _, K1⟩ := exists_isCompact_isClosed_subset_isCompact_nhds_one Lcomp L1 ⟨K, Kcl, Kcomp, K1⟩ +/-- A quotient of a locally compact group is locally compact. -/ +@[to_additive] +instance [LocallyCompactSpace G] (N : Subgroup G) : LocallyCompactSpace (G ⧸ N) := by + refine ⟨fun x n hn ↦ ?_⟩ + let π := ((↑) : G → G ⧸ N) + have C : Continuous π := continuous_coinduced_rng + obtain ⟨y, rfl⟩ : ∃ y, π y = x := Quot.exists_rep x + have : π ⁻¹' n ∈ 𝓝 y := preimage_nhds_coinduced hn + rcases local_compact_nhds this with ⟨s, s_mem, hs, s_comp⟩ + exact ⟨π '' s, (QuotientGroup.isOpenMap_coe N).image_mem_nhds s_mem, mapsTo'.mp hs, + s_comp.image C⟩ + end section diff --git a/Mathlib/Topology/Algebra/Group/TopologicalAbelianization.lean b/Mathlib/Topology/Algebra/Group/TopologicalAbelianization.lean new file mode 100644 index 0000000000000..5a27525a78e01 --- /dev/null +++ b/Mathlib/Topology/Algebra/Group/TopologicalAbelianization.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2023 María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: María Inés de Frutos-Fernández +-/ +import Mathlib.GroupTheory.Abelianization +import Mathlib.Topology.Algebra.Group.Basic + +/-! +# The topological abelianization of a group. + +This file defines the topological abelianization of a topological group. + +## Main definitions + +* `TopologicalAbelianization`: defines the topological abelianization of a group `G` as the quotient + of `G` by the topological closure of its commutator subgroup.. + +## Main results +- `instNormalCommutatorClosure` : the topological closure of the commutator of a topological group + `G` is a normal subgroup. + +## Tags +group, topological abelianization + +-/ + +variable (G : Type*) [Group G] [TopologicalSpace G] [TopologicalGroup G] + +instance instNormalCommutatorClosure : (commutator G).topologicalClosure.Normal := + Subgroup.is_normal_topologicalClosure (commutator G) + +/-- The topological abelianization of `absoluteGaloisGroup`, that is, the quotient of + `absoluteGaloisGroup` by the topological closure of its commutator subgroup. -/ +abbrev TopologicalAbelianization := G ⧸ Subgroup.topologicalClosure (commutator G) + +local notation "G_ab" => TopologicalAbelianization + +namespace TopologicalAbelianization + +instance commGroup : CommGroup (G_ab G) where + mul_comm := fun x y => + Quotient.inductionOn₂' x y fun a b => + Quotient.sound' <| + QuotientGroup.leftRel_apply.mpr <| by + have h : (a * b)⁻¹ * (b * a) = ⁅b⁻¹, a⁻¹⁆ := by group + rw [h] + exact Subgroup.le_topologicalClosure _ (Subgroup.commutator_mem_commutator + (Subgroup.mem_top b⁻¹) (Subgroup.mem_top a⁻¹)) + __ : Group (G_ab G) := inferInstance + +end TopologicalAbelianization diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean b/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean index 5e4522deae7aa..0c125c0f1b8a5 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean @@ -184,47 +184,57 @@ We prove two versions of the Cauchy product formula. The first one is involving `Nat` subtraction. In order to avoid `Nat` subtraction, we also provide `tsum_mul_tsum_eq_tsum_sum_antidiagonal`, where the `n`-th term is a sum over all pairs `(k, l)` such that `k+l=n`, which corresponds to the -`Finset` `Finset.Nat.antidiagonal n` +`Finset` `Finset.antidiagonal n`. +This in fact allows us to generalize to any type satisfying `[Finset.HasAntidiagonal A]` -/ section CauchyProduct -variable [TopologicalSpace α] [NonUnitalNonAssocSemiring α] {f g : ℕ → α} +section HasAntidiagonal +variable {A : Type*} [AddCommMonoid A] [HasAntidiagonal A] +variable [TopologicalSpace α] [NonUnitalNonAssocSemiring α] {f g : A → α} /- The family `(k, l) : ℕ × ℕ ↦ f k * g l` is summable if and only if the family -`(n, k, l) : Σ (n : ℕ), Nat.antidiagonal n ↦ f k * g l` is summable. -/ +`(n, k, l) : Σ (n : ℕ), antidiagonal n ↦ f k * g l` is summable. -/ theorem summable_mul_prod_iff_summable_mul_sigma_antidiagonal : - (Summable fun x : ℕ × ℕ => f x.1 * g x.2) ↔ - Summable fun x : Σn : ℕ, Nat.antidiagonal n => f (x.2 : ℕ × ℕ).1 * g (x.2 : ℕ × ℕ).2 := - Nat.sigmaAntidiagonalEquivProd.summable_iff.symm + (Summable fun x : A × A => f x.1 * g x.2) ↔ + Summable fun x : Σn : A, antidiagonal n => f (x.2 : A × A).1 * g (x.2 : A × A).2 := + Finset.sigmaAntidiagonalEquivProd.summable_iff.symm #align summable_mul_prod_iff_summable_mul_sigma_antidiagonal summable_mul_prod_iff_summable_mul_sigma_antidiagonal variable [T3Space α] [TopologicalSemiring α] theorem summable_sum_mul_antidiagonal_of_summable_mul - (h : Summable fun x : ℕ × ℕ => f x.1 * g x.2) : - Summable fun n => ∑ kl in Nat.antidiagonal n, f kl.1 * g kl.2 := by + (h : Summable fun x : A × A => f x.1 * g x.2) : + Summable fun n => ∑ kl in antidiagonal n, f kl.1 * g kl.2 := by rw [summable_mul_prod_iff_summable_mul_sigma_antidiagonal] at h conv => congr; ext; rw [← Finset.sum_finset_coe, ← tsum_fintype] exact h.sigma' fun n => (hasSum_fintype _).summable #align summable_sum_mul_antidiagonal_of_summable_mul summable_sum_mul_antidiagonal_of_summable_mul /-- The **Cauchy product formula** for the product of two infinites sums indexed by `ℕ`, expressed -by summing on `Finset.Nat.antidiagonal`. +by summing on `Finset.antidiagonal`. See also `tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm` if `f` and `g` are absolutely summable. -/ theorem tsum_mul_tsum_eq_tsum_sum_antidiagonal (hf : Summable f) (hg : Summable g) - (hfg : Summable fun x : ℕ × ℕ => f x.1 * g x.2) : - ((∑' n, f n) * ∑' n, g n) = ∑' n, ∑ kl in Nat.antidiagonal n, f kl.1 * g kl.2 := by + (hfg : Summable fun x : A × A => f x.1 * g x.2) : + ((∑' n, f n) * ∑' n, g n) = ∑' n, ∑ kl in antidiagonal n, f kl.1 * g kl.2 := by conv_rhs => congr; ext; rw [← Finset.sum_finset_coe, ← tsum_fintype] - rw [tsum_mul_tsum hf hg hfg, ← Nat.sigmaAntidiagonalEquivProd.tsum_eq (_ : ℕ × ℕ → α)] + rw [tsum_mul_tsum hf hg hfg, ← sigmaAntidiagonalEquivProd.tsum_eq (_ : A × A → α)] exact tsum_sigma' (fun n => (hasSum_fintype _).summable) (summable_mul_prod_iff_summable_mul_sigma_antidiagonal.mp hfg) #align tsum_mul_tsum_eq_tsum_sum_antidiagonal tsum_mul_tsum_eq_tsum_sum_antidiagonal +end HasAntidiagonal + +section Nat + +variable [TopologicalSpace α] [NonUnitalNonAssocSemiring α] {f g : ℕ → α} +variable [T3Space α] [TopologicalSemiring α] + theorem summable_sum_mul_range_of_summable_mul (h : Summable fun x : ℕ × ℕ => f x.1 * g x.2) : Summable fun n => ∑ k in range (n + 1), f k * g (n - k) := by simp_rw [← Nat.sum_antidiagonal_eq_sum_range_succ fun k l => f k * g l] @@ -243,4 +253,6 @@ theorem tsum_mul_tsum_eq_tsum_sum_range (hf : Summable f) (hg : Summable g) exact tsum_mul_tsum_eq_tsum_sum_antidiagonal hf hg hfg #align tsum_mul_tsum_eq_tsum_sum_range tsum_mul_tsum_eq_tsum_sum_range +end Nat + end CauchyProduct diff --git a/Mathlib/Topology/Algebra/MulAction.lean b/Mathlib/Topology/Algebra/MulAction.lean index 826236de0c29d..c7f813cf9d6b1 100644 --- a/Mathlib/Topology/Algebra/MulAction.lean +++ b/Mathlib/Topology/Algebra/MulAction.lean @@ -140,6 +140,26 @@ instance MulOpposite.continuousSMul : ContinuousSMul M Xᵐᵒᵖ := #align mul_opposite.has_continuous_smul MulOpposite.continuousSMul #align add_opposite.has_continuous_vadd AddOpposite.continuousVAdd +@[to_additive] +lemma IsCompact.smul_set {k : Set M} {u : Set X} (hk : IsCompact k) (hu : IsCompact u) : + IsCompact (k • u) := by + rw [← Set.image_smul_prod] + exact IsCompact.image (hk.prod hu) continuous_smul + +@[to_additive] +lemma smul_set_closure_subset (K : Set M) (L : Set X) : + closure K • closure L ⊆ closure (K • L) := by + rintro - ⟨x, y, hx, hy, rfl⟩ + apply mem_closure_iff_nhds.2 (fun u hu ↦ ?_) + have A : (fun p ↦ p.fst • p.snd) ⁻¹' u ∈ 𝓝 (x, y) := + (continuous_smul.continuousAt (x := (x, y))).preimage_mem_nhds hu + obtain ⟨a, ha, b, hb, hab⟩ : + ∃ a, a ∈ 𝓝 x ∧ ∃ b, b ∈ 𝓝 y ∧ a ×ˢ b ⊆ (fun p ↦ p.fst • p.snd) ⁻¹' u := + mem_nhds_prod_iff.1 A + obtain ⟨x', ⟨x'a, x'K⟩⟩ : Set.Nonempty (a ∩ K) := mem_closure_iff_nhds.1 hx a ha + obtain ⟨y', ⟨y'b, y'L⟩⟩ : Set.Nonempty (b ∩ L) := mem_closure_iff_nhds.1 hy b hb + exact ⟨x' • y', hab (Set.mk_mem_prod x'a y'b), Set.smul_mem_smul x'K y'L⟩ + end SMul section Monoid diff --git a/Mathlib/Topology/Algebra/Order/LeftRight.lean b/Mathlib/Topology/Algebra/Order/LeftRight.lean index 04ebb944a582b..861f49609cb13 100644 --- a/Mathlib/Topology/Algebra/Order/LeftRight.lean +++ b/Mathlib/Topology/Algebra/Order/LeftRight.lean @@ -27,6 +27,44 @@ left continuous, right continuous open Set Filter Topology +section Preorder + +variable {α : Type*} [TopologicalSpace α] [Preorder α] + +lemma frequently_lt_nhds (a : α) [NeBot (𝓝[<] a)] : ∃ᶠ x in 𝓝 a, x < a := + frequently_iff_neBot.2 ‹_› + +lemma frequently_gt_nhds (a : α) [NeBot (𝓝[>] a)] : ∃ᶠ x in 𝓝 a, a < x := + frequently_iff_neBot.2 ‹_› + +theorem Filter.Eventually.exists_lt {a : α} [NeBot (𝓝[<] a)] {p : α → Prop} + (h : ∀ᶠ x in 𝓝 a, p x) : ∃ b < a, p b := + ((frequently_lt_nhds a).and_eventually h).exists +#align filter.eventually.exists_lt Filter.Eventually.exists_lt + +theorem Filter.Eventually.exists_gt {a : α} [NeBot (𝓝[>] a)] {p : α → Prop} + (h : ∀ᶠ x in 𝓝 a, p x) : ∃ b > a, p b := + ((frequently_gt_nhds a).and_eventually h).exists +#align filter.eventually.exists_gt Filter.Eventually.exists_gt + +theorem nhdsWithin_Ici_neBot {a b : α} (H₂ : a ≤ b) : NeBot (𝓝[Ici a] b) := + nhdsWithin_neBot_of_mem H₂ +#align nhds_within_Ici_ne_bot nhdsWithin_Ici_neBot + +instance nhdsWithin_Ici_self_neBot (a : α) : NeBot (𝓝[≥] a) := + nhdsWithin_Ici_neBot (le_refl a) +#align nhds_within_Ici_self_ne_bot nhdsWithin_Ici_self_neBot + +theorem nhdsWithin_Iic_neBot {a b : α} (H : a ≤ b) : NeBot (𝓝[Iic b] a) := + nhdsWithin_neBot_of_mem H +#align nhds_within_Iic_ne_bot nhdsWithin_Iic_neBot + +instance nhdsWithin_Iic_self_neBot (a : α) : NeBot (𝓝[≤] a) := + nhdsWithin_Iic_neBot (le_refl a) +#align nhds_within_Iic_self_ne_bot nhdsWithin_Iic_self_neBot + +end Preorder + section PartialOrder variable {α β : Type*} [TopologicalSpace α] [PartialOrder α] [TopologicalSpace β] diff --git a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean index 42c2ceeec7339..13bdb361e1f54 100644 --- a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean +++ b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov +Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov, Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.BigOperators.Order @@ -12,44 +12,62 @@ import Mathlib.Order.Filter.Archimedean import Mathlib.Order.Filter.CountableInter import Mathlib.Topology.Order.Basic -#align_import topology.algebra.order.liminf_limsup from "leanprover-community/mathlib"@"ffde2d8a6e689149e44fd95fa862c23a57f8c780" +#align_import topology.algebra.order.liminf_limsup from "leanprover-community/mathlib"@"ce64cd319bb6b3e82f31c2d38e79080d377be451" /-! # Lemmas about liminf and limsup in an order topology. + +## Main declarations + +* `BoundedLENhdsClass`: Typeclass stating that neighborhoods are eventually bounded above. +* `BoundedGENhdsClass`: Typeclass stating that neighborhoods are eventually bounded below. + +## Implementation notes + +The same lemmas are true in `ℝ`, `ℝ × ℝ`, `ι → ℝ`, `EuclideanSpace ι ℝ`. To avoid code +duplication, we provide an ad hoc axiomatisation of the properties we need. -/ open Filter TopologicalSpace -open Topology Classical +open scoped Topology Classical universe u v -variable {α : Type u} {β : Type v} +variable {ι α β R S : Type*} {π : ι → Type*} -section LiminfLimsup +/-- Ad hoc typeclass stating that neighborhoods are eventually bounded above. -/ +class BoundedLENhdsClass (α : Type*) [Preorder α] [TopologicalSpace α] : Prop where + isBounded_le_nhds (a : α) : (𝓝 a).IsBounded (· ≤ ·) +#align bounded_le_nhds_class BoundedLENhdsClass -section OrderClosedTopology +/-- Ad hoc typeclass stating that neighborhoods are eventually bounded below. -/ +class BoundedGENhdsClass (α : Type*) [Preorder α] [TopologicalSpace α] : Prop where + isBounded_ge_nhds (a : α) : (𝓝 a).IsBounded (· ≥ ·) +#align bounded_ge_nhds_class BoundedGENhdsClass -variable [SemilatticeSup α] [TopologicalSpace α] [OrderTopology α] +section Preorder +variable [Preorder α] [Preorder β] [TopologicalSpace α] [TopologicalSpace β] + +section BoundedLENhdsClass +variable [BoundedLENhdsClass α] [BoundedLENhdsClass β] {f : Filter ι} {u : ι → α} {a : α} theorem isBounded_le_nhds (a : α) : (𝓝 a).IsBounded (· ≤ ·) := - (isTop_or_exists_gt a).elim (fun h ↦ ⟨a, eventually_of_forall h⟩) fun ⟨b, hb⟩ ↦ - ⟨b, ge_mem_nhds hb⟩ + BoundedLENhdsClass.isBounded_le_nhds _ #align is_bounded_le_nhds isBounded_le_nhds -theorem Filter.Tendsto.isBoundedUnder_le {f : Filter β} {u : β → α} {a : α} - (h : Tendsto u f (𝓝 a)) : f.IsBoundedUnder (· ≤ ·) u := +theorem Filter.Tendsto.isBoundedUnder_le (h : Tendsto u f (𝓝 a)) : f.IsBoundedUnder (· ≤ ·) u := (isBounded_le_nhds a).mono h #align filter.tendsto.is_bounded_under_le Filter.Tendsto.isBoundedUnder_le -theorem Filter.Tendsto.bddAbove_range_of_cofinite {u : β → α} {a : α} +theorem Filter.Tendsto.bddAbove_range_of_cofinite [IsDirected α (· ≤ ·)] (h : Tendsto u cofinite (𝓝 a)) : BddAbove (Set.range u) := h.isBoundedUnder_le.bddAbove_range_of_cofinite #align filter.tendsto.bdd_above_range_of_cofinite Filter.Tendsto.bddAbove_range_of_cofinite -theorem Filter.Tendsto.bddAbove_range {u : ℕ → α} {a : α} (h : Tendsto u atTop (𝓝 a)) : - BddAbove (Set.range u) := +theorem Filter.Tendsto.bddAbove_range [IsDirected α (· ≤ ·)] {u : ℕ → α} + (h : Tendsto u atTop (𝓝 a)) : BddAbove (Set.range u) := h.isBoundedUnder_le.bddAbove_range #align filter.tendsto.bdd_above_range Filter.Tendsto.bddAbove_range @@ -57,33 +75,47 @@ theorem isCobounded_ge_nhds (a : α) : (𝓝 a).IsCobounded (· ≥ ·) := (isBounded_le_nhds a).isCobounded_flip #align is_cobounded_ge_nhds isCobounded_ge_nhds -theorem Filter.Tendsto.isCoboundedUnder_ge {f : Filter β} {u : β → α} {a : α} [NeBot f] - (h : Tendsto u f (𝓝 a)) : f.IsCoboundedUnder (· ≥ ·) u := +theorem Filter.Tendsto.isCoboundedUnder_ge [NeBot f] (h : Tendsto u f (𝓝 a)) : + f.IsCoboundedUnder (· ≥ ·) u := h.isBoundedUnder_le.isCobounded_flip #align filter.tendsto.is_cobounded_under_ge Filter.Tendsto.isCoboundedUnder_ge -end OrderClosedTopology +instance : BoundedGENhdsClass αᵒᵈ := ⟨@isBounded_le_nhds α _ _ _⟩ + +instance Prod.instBoundedLENhdsClass : BoundedLENhdsClass (α × β) := by + refine ⟨fun x ↦ ?_⟩ + obtain ⟨a, ha⟩ := isBounded_le_nhds x.1 + obtain ⟨b, hb⟩ := isBounded_le_nhds x.2 + rw [← @Prod.mk.eta _ _ x, nhds_prod_eq] + exact ⟨(a, b), ha.prod_mk hb⟩ + +instance Pi.instBoundedLENhdsClass [Finite ι] [∀ i, Preorder (π i)] [∀ i, TopologicalSpace (π i)] + [∀ i, BoundedLENhdsClass (π i)] : BoundedLENhdsClass (∀ i, π i) := by + refine' ⟨fun x ↦ _⟩ + rw [nhds_pi] + choose f hf using fun i ↦ isBounded_le_nhds (x i) + exact ⟨f, eventually_pi hf⟩ -section OrderClosedTopology +end BoundedLENhdsClass -variable [SemilatticeInf α] [TopologicalSpace α] [OrderTopology α] +section BoundedGENhdsClass +variable [BoundedGENhdsClass α] [BoundedGENhdsClass β] {f : Filter ι} {u : ι → α} {a : α} theorem isBounded_ge_nhds (a : α) : (𝓝 a).IsBounded (· ≥ ·) := - isBounded_le_nhds (α := αᵒᵈ) a + BoundedGENhdsClass.isBounded_ge_nhds _ #align is_bounded_ge_nhds isBounded_ge_nhds -theorem Filter.Tendsto.isBoundedUnder_ge {f : Filter β} {u : β → α} {a : α} - (h : Tendsto u f (𝓝 a)) : f.IsBoundedUnder (· ≥ ·) u := +theorem Filter.Tendsto.isBoundedUnder_ge (h : Tendsto u f (𝓝 a)) : f.IsBoundedUnder (· ≥ ·) u := (isBounded_ge_nhds a).mono h #align filter.tendsto.is_bounded_under_ge Filter.Tendsto.isBoundedUnder_ge -theorem Filter.Tendsto.bddBelow_range_of_cofinite {u : β → α} {a : α} +theorem Filter.Tendsto.bddBelow_range_of_cofinite [IsDirected α (· ≥ ·)] (h : Tendsto u cofinite (𝓝 a)) : BddBelow (Set.range u) := h.isBoundedUnder_ge.bddBelow_range_of_cofinite #align filter.tendsto.bdd_below_range_of_cofinite Filter.Tendsto.bddBelow_range_of_cofinite -theorem Filter.Tendsto.bddBelow_range {u : ℕ → α} {a : α} (h : Tendsto u atTop (𝓝 a)) : - BddBelow (Set.range u) := +theorem Filter.Tendsto.bddBelow_range [IsDirected α (· ≥ ·)] {u : ℕ → α} + (h : Tendsto u atTop (𝓝 a)) : BddBelow (Set.range u) := h.isBoundedUnder_ge.bddBelow_range #align filter.tendsto.bdd_below_range Filter.Tendsto.bddBelow_range @@ -91,12 +123,50 @@ theorem isCobounded_le_nhds (a : α) : (𝓝 a).IsCobounded (· ≤ ·) := (isBounded_ge_nhds a).isCobounded_flip #align is_cobounded_le_nhds isCobounded_le_nhds -theorem Filter.Tendsto.isCoboundedUnder_le {f : Filter β} {u : β → α} {a : α} [NeBot f] - (h : Tendsto u f (𝓝 a)) : f.IsCoboundedUnder (· ≤ ·) u := +theorem Filter.Tendsto.isCoboundedUnder_le [NeBot f] (h : Tendsto u f (𝓝 a)) : + f.IsCoboundedUnder (· ≤ ·) u := h.isBoundedUnder_ge.isCobounded_flip #align filter.tendsto.is_cobounded_under_le Filter.Tendsto.isCoboundedUnder_le -end OrderClosedTopology +instance : BoundedLENhdsClass αᵒᵈ := ⟨@isBounded_ge_nhds α _ _ _⟩ + +instance Prod.instBoundedGENhdsClass : BoundedGENhdsClass (α × β) := + ⟨(Prod.instBoundedLENhdsClass (α := αᵒᵈ) (β := βᵒᵈ)).isBounded_le_nhds⟩ + +instance Pi.instBoundedGENhdsClass [Finite ι] [∀ i, Preorder (π i)] [∀ i, TopologicalSpace (π i)] + [∀ i, BoundedGENhdsClass (π i)] : BoundedGENhdsClass (∀ i, π i) := + ⟨(Pi.instBoundedLENhdsClass (π := fun i ↦ (π i)ᵒᵈ)).isBounded_le_nhds⟩ + +end BoundedGENhdsClass + +-- See note [lower instance priority] +instance (priority := 100) OrderTop.to_BoundedLENhdsClass [OrderTop α] : BoundedLENhdsClass α := + ⟨fun _a ↦ isBounded_le_of_top⟩ +#align order_top.to_bounded_le_nhds_class OrderTop.to_BoundedLENhdsClass + +-- See note [lower instance priority] +instance (priority := 100) OrderBot.to_BoundedGENhdsClass [OrderBot α] : BoundedGENhdsClass α := + ⟨fun _a ↦ isBounded_ge_of_bot⟩ +#align order_bot.to_bounded_ge_nhds_class OrderBot.to_BoundedGENhdsClass + +-- See note [lower instance priority] +instance (priority := 100) OrderTopology.to_BoundedLENhdsClass [IsDirected α (· ≤ ·)] + [OrderTopology α] : BoundedLENhdsClass α := + ⟨fun a ↦ + ((isTop_or_exists_gt a).elim fun h ↦ ⟨a, eventually_of_forall h⟩) <| + Exists.imp fun _b ↦ ge_mem_nhds⟩ +#align order_topology.to_bounded_le_nhds_class OrderTopology.to_BoundedLENhdsClass + +-- See note [lower instance priority] +instance (priority := 100) OrderTopology.to_BoundedGENhdsClass [IsDirected α (· ≥ ·)] + [OrderTopology α] : BoundedGENhdsClass α := + ⟨fun a ↦ ((isBot_or_exists_lt a).elim fun h ↦ ⟨a, eventually_of_forall h⟩) <| + Exists.imp fun _b ↦ le_mem_nhds⟩ +#align order_topology.to_bounded_ge_nhds_class OrderTopology.to_BoundedGENhdsClass + +end Preorder + +section LiminfLimsup section ConditionallyCompleteLinearOrder @@ -258,7 +328,7 @@ end LiminfLimsup section Monotone -variable {ι R S : Type*} {F : Filter ι} [NeBot F] +variable {F : Filter ι} [NeBot F] [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] [ConditionallyCompleteLinearOrder S] [TopologicalSpace S] [OrderTopology S] @@ -303,7 +373,7 @@ theorem Antitone.map_limsSup_of_continuousAt {F : Filter R} [NeBot F] {f : R → simpa [IsBot] using not_bot obtain ⟨m, l_m, m_lt⟩ : (Set.Ioo l F.limsSup).Nonempty := by contrapose! h' - refine' ⟨l, l_lt, by rwa [Set.not_nonempty_iff_eq_empty] at h'⟩ + exact ⟨l, l_lt, h'⟩ have B : F.liminf f ≤ f m := by apply liminf_le_of_frequently_le _ _ · apply (frequently_lt_of_lt_limsSup cobdd m_lt).mono @@ -399,7 +469,7 @@ open Topology open Filter Set -variable {ι : Type*} {R : Type*} [CompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] +variable [CompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] theorem iInf_eq_of_forall_le_of_tendsto {x : R} {as : ι → R} (x_le : ∀ i, x ≤ as i) {F : Filter ι} [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : ⨅ i, as i = x := by @@ -412,7 +482,7 @@ theorem iSup_eq_of_forall_le_of_tendsto {x : R} {as : ι → R} (le_x : ∀ i, a iInf_eq_of_forall_le_of_tendsto (R := Rᵒᵈ) le_x as_lim #align supr_eq_of_forall_le_of_tendsto iSup_eq_of_forall_le_of_tendsto -theorem iUnion_Ici_eq_Ioi_of_lt_of_tendsto {ι : Type*} (x : R) {as : ι → R} (x_lt : ∀ i, x < as i) +theorem iUnion_Ici_eq_Ioi_of_lt_of_tendsto (x : R) {as : ι → R} (x_lt : ∀ i, x < as i) {F : Filter ι} [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : ⋃ i : ι, Ici (as i) = Ioi x := by have obs : x ∉ range as := by @@ -426,7 +496,7 @@ theorem iUnion_Ici_eq_Ioi_of_lt_of_tendsto {ι : Type*} (x : R) {as : ι → R} exact iUnion_Ici_eq_Ioi_iInf obs #align Union_Ici_eq_Ioi_of_lt_of_tendsto iUnion_Ici_eq_Ioi_of_lt_of_tendsto -theorem iUnion_Iic_eq_Iio_of_lt_of_tendsto {ι : Type*} (x : R) {as : ι → R} (lt_x : ∀ i, as i < x) +theorem iUnion_Iic_eq_Iio_of_lt_of_tendsto (x : R) {as : ι → R} (lt_x : ∀ i, as i < x) {F : Filter ι} [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : ⋃ i : ι, Iic (as i) = Iio x := iUnion_Ici_eq_Ioi_of_lt_of_tendsto (R := Rᵒᵈ) x lt_x as_lim @@ -512,10 +582,7 @@ theorem limsup_eq_tendsto_sum_indicator_atTop (R : Type*) [StrictOrderedSemiring end Indicator section LiminfLimsupAddSub - -set_option autoImplicit true - -variable {R : Type*} [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] +variable [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] /-- `liminf (c + xᵢ) = c + liminf xᵢ`. -/ lemma limsup_const_add (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index e576ceb1b6702..6e49b0b9bb7c6 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -848,7 +848,7 @@ scoped[Topology] notation "𝓝" => nhds scoped[Topology] notation "𝓝[" s "] " x:100 => nhdsWithin x s /-- Notation for the filter of punctured neighborhoods of a point. -/ -scoped[Topology] notation "𝓝[≠] " x:100 => nhdsWithin x {x}ᶜ +scoped[Topology] notation "𝓝[≠] " x:100 => nhdsWithin x (@singleton _ (Set _) instSingletonSet x)ᶜ /-- Notation for the filter of right neighborhoods of a point. -/ scoped[Topology] notation "𝓝[≥] " x:100 => nhdsWithin x (Set.Ici x) @@ -1082,6 +1082,10 @@ theorem tendsto_nhds_of_eventually_eq {f : β → α} {a : α} (h : ∀ᶠ x in Tendsto f l (𝓝 a) := tendsto_const_nhds.congr' (.symm h) +theorem Filter.EventuallyEq.tendsto {f : β → α} {a : α} (hf : f =ᶠ[l] fun _ ↦ a) : + Tendsto f l (𝓝 a) := + tendsto_nhds_of_eventually_eq hf + /-! ### Cluster points @@ -1609,10 +1613,12 @@ theorem IsOpen.preimage {f : α → β} (hf : Continuous f) {s : Set β} (h : Is hf.isOpen_preimage s h #align is_open.preimage IsOpen.preimage -theorem Continuous.congr {f g : α → β} (h : Continuous f) (h' : ∀ x, f x = g x) : Continuous g := by - convert h - ext - rw [h'] +theorem continuous_congr {f g : α → β} (h : ∀ x, f x = g x) : + Continuous f ↔ Continuous g := + .of_eq <| congrArg _ <| funext h + +theorem Continuous.congr {f g : α → β} (h : Continuous f) (h' : ∀ x, f x = g x) : Continuous g := + continuous_congr h' |>.mp h #align continuous.congr Continuous.congr /-- A function between topological spaces is continuous at a point `x₀` diff --git a/Mathlib/Topology/Bornology/Basic.lean b/Mathlib/Topology/Bornology/Basic.lean index fc3438f419ed3..34b1b3c40a9af 100644 --- a/Mathlib/Topology/Bornology/Basic.lean +++ b/Mathlib/Topology/Bornology/Basic.lean @@ -173,6 +173,12 @@ theorem isBounded_empty : IsBounded (∅ : Set α) := by exact univ_mem #align bornology.is_bounded_empty Bornology.isBounded_empty +theorem nonempty_of_not_isBounded (h : ¬IsBounded s) : s.Nonempty := by + rw [nonempty_iff_ne_empty] + rintro rfl + exact h isBounded_empty +#align metric.nonempty_of_unbounded Bornology.nonempty_of_not_isBounded + @[simp] theorem isBounded_singleton : IsBounded ({x} : Set α) := by rw [isBounded_def] diff --git a/Mathlib/Topology/Bornology/Constructions.lean b/Mathlib/Topology/Bornology/Constructions.lean index bf91b61dc9731..df634267185ba 100644 --- a/Mathlib/Topology/Bornology/Constructions.lean +++ b/Mathlib/Topology/Bornology/Constructions.lean @@ -89,7 +89,7 @@ theorem isBounded_prod : IsBounded (s ×ˢ t) ↔ s = ∅ ∨ t = ∅ ∨ IsBoun theorem isBounded_prod_self : IsBounded (s ×ˢ s) ↔ IsBounded s := by rcases s.eq_empty_or_nonempty with (rfl | hs); · simp - exact (isBounded_prod_of_nonempty (hs.prod hs)).trans (and_self_iff _) + exact (isBounded_prod_of_nonempty (hs.prod hs)).trans and_self_iff #align bornology.is_bounded_prod_self Bornology.isBounded_prod_self /-! diff --git a/Mathlib/Topology/Category/CompHaus/EffectiveEpi.lean b/Mathlib/Topology/Category/CompHaus/EffectiveEpi.lean index 86f8f0f541019..27794821af15e 100644 --- a/Mathlib/Topology/Category/CompHaus/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/CompHaus/EffectiveEpi.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Adam Topaz -/ -import Mathlib.CategoryTheory.Sites.Coherent +import Mathlib.CategoryTheory.Sites.RegularExtensive import Mathlib.Topology.Category.CompHaus.Limits /-! @@ -18,7 +18,7 @@ In this file, we show that the following are all equivalent: - The family `π` is jointly surjective. This is the main result of this file, which can be found in `CompHaus.effectiveEpiFamily_tfae` -As a consequence, we also show that `CompHaus` is precoherent. +As a consequence, we also show that `CompHaus` is precoherent and preregular. # Projects @@ -192,6 +192,7 @@ theorem effectiveEpiFamily_of_jointly_surjective open EffectiveEpiFamily +-- TODO: prove this for `Type*` open List in theorem effectiveEpiFamily_tfae {α : Type} [Fintype α] {B : CompHaus.{u}} @@ -247,4 +248,13 @@ lemma effectiveEpi_iff_surjective {X Y : CompHaus} (f : X ⟶ Y) : rw [← epi_iff_surjective] exact effectiveEpi_iff_epi (fun _ _ ↦ (effectiveEpiFamily_tfae _ _).out 0 1) f +instance : Preregular CompHaus where + exists_fac := by + intro X Y Z f π hπ + refine ⟨pullback f π, pullback.fst f π, ?_, pullback.snd f π, (pullback.condition _ _).symm⟩ + rw [CompHaus.effectiveEpi_iff_surjective] at hπ ⊢ + intro y + obtain ⟨z,hz⟩ := hπ (f y) + exact ⟨⟨(y, z), hz.symm⟩, rfl⟩ + end CompHaus diff --git a/Mathlib/Topology/Category/CompHaus/Limits.lean b/Mathlib/Topology/Category/CompHaus/Limits.lean index c79397d1e2425..870f4295b05f4 100644 --- a/Mathlib/Topology/Category/CompHaus/Limits.lean +++ b/Mathlib/Topology/Category/CompHaus/Limits.lean @@ -6,6 +6,8 @@ Authors: Adam Topaz import Mathlib.Topology.Category.CompHaus.Basic import Mathlib.CategoryTheory.Limits.Shapes.Pullbacks +import Mathlib.CategoryTheory.Extensive +import Mathlib.CategoryTheory.Limits.Preserves.Finite /-! @@ -24,7 +26,7 @@ namespace CompHaus universe u -open CategoryTheory +open CategoryTheory Limits section Pullbacks @@ -230,6 +232,18 @@ lemma finiteCoproduct.ι_desc_apply {B : CompHaus} {π : (a : α) → X a ⟶ B} simp only [ι_desc] -- `elementwise` should work here, but doesn't +instance : PreservesFiniteCoproducts compHausToTop := by + refine ⟨fun J hJ ↦ ⟨fun {F} ↦ ?_⟩⟩ + suffices : PreservesColimit (Discrete.functor (F.obj ∘ Discrete.mk)) compHausToTop + · exact preservesColimitOfIsoDiagram _ Discrete.natIsoFunctor.symm + apply preservesColimitOfPreservesColimitCocone (CompHaus.finiteCoproduct.isColimit _) + exact TopCat.sigmaCofanIsColimit _ + +instance : FinitaryExtensive CompHaus := + have := fullyFaithfulReflectsLimits compHausToTop + have := fullyFaithfulReflectsColimits compHausToTop + finitaryExtensive_of_preserves_and_reflects compHausToTop + end FiniteCoproducts end CompHaus diff --git a/Mathlib/Topology/Category/Profinite/EffectiveEpi.lean b/Mathlib/Topology/Category/Profinite/EffectiveEpi.lean index 748d71c99c3c8..f5dbb4bc4deba 100644 --- a/Mathlib/Topology/Category/Profinite/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/Profinite/EffectiveEpi.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Jon Eugster. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson, Boris Bolvig Kjær, Jon Eugster, Sina Hazratpour -/ -import Mathlib.CategoryTheory.Sites.Coherent +import Mathlib.CategoryTheory.Sites.RegularExtensive import Mathlib.Topology.Category.Profinite.Limits /-! @@ -15,6 +15,9 @@ In this file, we show that the following are all equivalent: - The induced map `∐ X ⟶ B` is epimorphic. - The family `π` is jointly surjective. +As a consequence, we show (see `effectiveEpi_iff_surjective`) that all epimorphisms in `Profinite`  +are effective, and that `Profinite` is preregular. + ## Main results - `Profinite.effectiveEpiFamily_tfae`: characterise being an effective epimorphic family. @@ -269,6 +272,15 @@ lemma effectiveEpi_iff_surjective {X Y : Profinite} (f : X ⟶ Y) : rw [← epi_iff_surjective] exact effectiveEpi_iff_epi (fun _ _ ↦ (effectiveEpiFamily_tfae _ _).out 0 1) f +instance : Preregular Profinite where + exists_fac := by + intro X Y Z f π hπ + refine ⟨pullback f π, pullback.fst f π, ?_, pullback.snd f π, (pullback.condition _ _).symm⟩ + rw [Profinite.effectiveEpi_iff_surjective] at hπ ⊢ + intro y + obtain ⟨z,hz⟩ := hπ (f y) + exact ⟨⟨(y, z), hz.symm⟩, rfl⟩ + end JointlySurjective section Coherent diff --git a/Mathlib/Topology/Category/Profinite/Limits.lean b/Mathlib/Topology/Category/Profinite/Limits.lean index 1e969c392828b..bdafcd02463b7 100644 --- a/Mathlib/Topology/Category/Profinite/Limits.lean +++ b/Mathlib/Topology/Category/Profinite/Limits.lean @@ -5,6 +5,7 @@ Authors: Adam Topaz -/ import Mathlib.Topology.Category.Profinite.Basic import Mathlib.CategoryTheory.Limits.Shapes.Pullbacks +import Mathlib.Topology.Category.CompHaus.Limits /-! @@ -24,7 +25,7 @@ namespace Profinite universe u -open CategoryTheory +open CategoryTheory Limits section Pullbacks @@ -215,6 +216,18 @@ lemma finiteCoproduct.ι_desc_apply {B : Profinite} {π : (a : α) → X a ⟶ B change (ι X a ≫ desc X π) _ = _ simp only [ι_desc] +instance : PreservesFiniteCoproducts profiniteToCompHaus := by + refine ⟨fun J hJ ↦ ⟨fun {F} ↦ ?_⟩⟩ + suffices : PreservesColimit (Discrete.functor (F.obj ∘ Discrete.mk)) profiniteToCompHaus + · exact preservesColimitOfIsoDiagram _ Discrete.natIsoFunctor.symm + apply preservesColimitOfPreservesColimitCocone (Profinite.finiteCoproduct.isColimit _) + exact CompHaus.finiteCoproduct.isColimit _ + +instance : FinitaryExtensive Profinite := + have := fullyFaithfulReflectsLimits profiniteToCompHaus + have := fullyFaithfulReflectsColimits profiniteToCompHaus + finitaryExtensive_of_preserves_and_reflects profiniteToCompHaus + end FiniteCoproducts end Profinite diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean new file mode 100644 index 0000000000000..812d2351b1a05 --- /dev/null +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -0,0 +1,1831 @@ +/- +Copyright (c) 2023 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.Algebra.Category.ModuleCat.Free +import Mathlib.Topology.Category.Profinite.CofilteredLimit +import Mathlib.Topology.Category.Profinite.Product +import Mathlib.Topology.LocallyConstant.Algebra + +/-! + +# Nöbeling's theorem + +This file proves Nöbeling's theorem. + +## Main result + +- `Nobeling`: For `S : Profinite`, the `ℤ`-module `LocallyConstant S ℤ` is free. + +## Proof idea + +We follow the proof of theorem 5.4 in [scholze2019condensed], in which the idea is to embed `S` in +a product of `I` copies of `Bool` for some sufficiently large `I`, and then to choose a +well-ordering on `I` and use ordinal induction over that well-order. Here we can let `I` be +the set of clopen subsets of `S` since `S` is totally separated. + +The above means it suffices to prove the following statement: For a closed subset `C` of `I → Bool`, +the `ℤ`-module `LocallyConstant C ℤ` is free. + +For `i : I`, let `e C i : LocallyConstant C ℤ` denote the map `fun f ↦ (if f.val i then 1 else 0)`. + +The basis will consist of products `e C iᵣ * ⋯ * e C i₁` with `iᵣ > ⋯ > i₁` which cannot be written +as linear combinations of lexicographically smaller products. We call this set `GoodProducts C` + +What is proved by ordinal induction is that this set is linearly independent. The fact that it +spans can be proved directly. + +## References + +- [scholze2019condensed], Theorem 5.4. +-/ + +universe u + +namespace Profinite + +namespace NobelingProof + +variable {I : Type u} [Inhabited I] [LinearOrder I] [IsWellOrder I (·<·)] (C : Set (I → Bool)) + +open Profinite ContinuousMap CategoryTheory Limits Opposite Submodule + +section Projections +/-! + +## Projection maps + +The purpose of this section is twofold. + +Firstly, in the proof that the set `GoodProducts C` spans the whole module `LocallyConstant C ℤ`, +we need to project `C` down to finite discrete subsets and write `C` as a cofiltered limit of those. + +Secondly, in the inductive argument, we need to project `C` down to "smaller" sets satisfying the +inductive hypothesis. + +In this section we define the relevant projection maps and prove some compatibility results. + +### Main definitions + +* Let `J : I → Prop`. Then `Proj J : (I → Bool) → (I → Bool)` is the projection mapping everything + that satisfies `J i` to itself, and everything else to `false`. + +* The image of `C` under `Proj J` is denoted `π C J` and the corresponding map `C → π C J` is called + `ProjRestrict`. If `J` implies `K` we have a map `ProjRestricts : π C K → π C J`. + +* `spanCone_isLimit` establishes that when `C` is compact, it can be written as a limit of its + images under the maps `Proj (· ∈ s)` where `s : Finset I`. +-/ + +variable (J K L : I → Prop) [∀ i, Decidable (J i)] [∀ i, Decidable (K i)] [∀ i, Decidable (L i)] + +/-- +The projection mapping everything that satisfies `J i` to itself, and everything else to `false` +-/ +def Proj : (I → Bool) → (I → Bool) := + fun c i ↦ if J i then c i else false + +@[simp] +theorem continuous_proj : + Continuous (Proj J : (I → Bool) → (I → Bool)) := by + dsimp [Proj] + apply continuous_pi + intro i + split + · apply continuous_apply + · apply continuous_const + +/-- The image of `Proj π J` -/ +def π : Set (I → Bool) := (Proj J) '' C + +/-- The restriction of `Proj π J` to a subset, mapping to its image. -/ +@[simps!] +def ProjRestrict : C → π C J := + Set.MapsTo.restrict (Proj J) _ _ (Set.mapsTo_image _ _) + +@[simp] +theorem continuous_projRestrict : Continuous (ProjRestrict C J) := + Continuous.restrict _ (continuous_proj _) + +theorem proj_eq_self {x : I → Bool} (h : ∀ i, x i ≠ false → J i) : Proj J x = x := by + ext i + simp only [Proj, ite_eq_left_iff] + contrapose! + simpa only [ne_comm] using h i + +theorem proj_prop_eq_self (hh : ∀ i x, x ∈ C → x i ≠ false → J i) : π C J = C := by + ext x + refine ⟨fun ⟨y, hy, h⟩ ↦ ?_, fun h ↦ ⟨x, h, ?_⟩⟩ + · rwa [← h, proj_eq_self]; exact (hh · y hy) + · rw [proj_eq_self]; exact (hh · x h) + +theorem proj_comp_of_subset (h : ∀ i, J i → K i) : (Proj J ∘ Proj K) = + (Proj J : (I → Bool) → (I → Bool)) := by + ext x i; dsimp [Proj]; aesop + +theorem proj_eq_of_subset (h : ∀ i, J i → K i) : π (π C K) J = π C J := by + ext x + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · obtain ⟨y, ⟨z, hz, rfl⟩, rfl⟩ := h + refine ⟨z, hz, (?_ : _ = (Proj J ∘ Proj K) z)⟩ + rw [proj_comp_of_subset J K h] + · obtain ⟨y, hy, rfl⟩ := h + dsimp [π] + rw [← Set.image_comp] + refine ⟨y, hy, ?_⟩ + rw [proj_comp_of_subset J K h] + +variable {J K L} + +/-- A variant of `ProjRestrict` with domain of the form `π C K` -/ +@[simps!] +def ProjRestricts (h : ∀ i, J i → K i) : π C K → π C J := + Homeomorph.setCongr (proj_eq_of_subset C J K h) ∘ ProjRestrict (π C K) J + +@[simp] +theorem continuous_projRestricts (h : ∀ i, J i → K i) : Continuous (ProjRestricts C h) := + Continuous.comp (Homeomorph.continuous _) (continuous_projRestrict _ _) + +theorem surjective_projRestricts (h : ∀ i, J i → K i) : Function.Surjective (ProjRestricts C h) := + (Homeomorph.surjective _).comp (Set.surjective_mapsTo_image_restrict _ _) + +variable (J) in +theorem projRestricts_eq_id : ProjRestricts C (fun i (h : J i) ↦ h) = id := by + ext ⟨x, y, hy, rfl⟩ i + simp (config := { contextual := true }) only [π, Proj, ProjRestricts_coe, id_eq, if_true] + +theorem projRestricts_eq_comp (hJK : ∀ i, J i → K i) (hKL : ∀ i, K i → L i) : + ProjRestricts C hJK ∘ ProjRestricts C hKL = ProjRestricts C (fun i ↦ hKL i ∘ hJK i) := by + ext x i + simp only [π, Proj, Function.comp_apply, ProjRestricts_coe] + aesop + +theorem projRestricts_comp_projRestrict (h : ∀ i, J i → K i) : + ProjRestricts C h ∘ ProjRestrict C K = ProjRestrict C J := by + ext x i + simp only [π, Proj, Function.comp_apply, ProjRestricts_coe, ProjRestrict_coe] + aesop + +variable (J) + +/-- The objectwise map in the isomorphism `spanFunctor ≅ Profinite.indexFunctor`. -/ +def iso_map : C(π C J, (IndexFunctor.obj C J)) := + ⟨fun x ↦ ⟨fun i ↦ x.val i.val, by + rcases x with ⟨x, y, hy, rfl⟩ + refine ⟨y, hy, ?_⟩ + ext ⟨i, hi⟩ + simp [precomp, Proj, hi]⟩, by + refine Continuous.subtype_mk (continuous_pi fun i ↦ ?_) _ + exact (continuous_apply i.val).comp continuous_subtype_val⟩ + +lemma iso_map_bijective : Function.Bijective (iso_map C J) := by + refine ⟨fun a b h ↦ ?_, fun a ↦ ?_⟩ + · ext i + rw [Subtype.ext_iff] at h + by_cases hi : J i + · exact congr_fun h ⟨i, hi⟩ + · rcases a with ⟨_, c, hc, rfl⟩ + rcases b with ⟨_, d, hd, rfl⟩ + simp only [Proj, if_neg hi] + · refine ⟨⟨fun i ↦ if hi : J i then a.val ⟨i, hi⟩ else false, ?_⟩, ?_⟩ + · rcases a with ⟨_, y, hy, rfl⟩ + exact ⟨y, hy, rfl⟩ + · ext i + exact dif_pos i.prop + +variable {C} (hC : IsCompact C) + +/-- +For a given compact subset `C` of `I → Bool`, `spanFunctor` is the functor from the poset of finsets +of `I` to `Profinite`, sending a finite subset set `J` to the image of `C` under the projection +`Proj J`. +-/ +noncomputable +def spanFunctor [∀ (s : Finset I) (i : I), Decidable (i ∈ s)] : + (Finset I)ᵒᵖ ⥤ Profinite.{u} where + obj s := @Profinite.of (π C (· ∈ (unop s))) _ + (by rw [← isCompact_iff_compactSpace]; exact hC.image (continuous_proj _)) _ _ + map h := ⟨(ProjRestricts C (leOfHom h.unop)), continuous_projRestricts _ _⟩ + map_id J := by simp only [projRestricts_eq_id C (· ∈ (unop J))]; rfl + map_comp _ _ := by dsimp; congr; dsimp; rw [projRestricts_eq_comp] + +/-- The limit cone on `spanFunctor` with point `C`. -/ +noncomputable +def spanCone [∀ (s : Finset I) (i : I), Decidable (i ∈ s)] : Cone (spanFunctor hC) where + pt := @Profinite.of C _ (by rwa [← isCompact_iff_compactSpace]) _ _ + π := + { app := fun s ↦ ⟨ProjRestrict C (· ∈ unop s), continuous_projRestrict _ _⟩ + naturality := by + intro X Y h + simp only [Functor.const_obj_obj, Homeomorph.setCongr, Homeomorph.homeomorph_mk_coe, + Functor.const_obj_map, Category.id_comp, ← projRestricts_comp_projRestrict C + (leOfHom h.unop)] + rfl } + +/-- `spanCone` is a limit cone. -/ +noncomputable +def spanCone_isLimit [∀ (s : Finset I) (i : I), Decidable (i ∈ s)] [DecidableEq I] : + CategoryTheory.Limits.IsLimit (spanCone hC) := by + refine (IsLimit.postcomposeHomEquiv (NatIso.ofComponents + (fun s ↦ (Profinite.isoOfBijective _ (iso_map_bijective C (· ∈ unop s)))) ?_) (spanCone hC)) + (IsLimit.ofIsoLimit (indexCone_isLimit hC) (Cones.ext (Iso.refl _) ?_)) + · intro ⟨s⟩ ⟨t⟩ ⟨⟨⟨f⟩⟩⟩ + ext x + have : iso_map C (· ∈ t) ∘ ProjRestricts C f = IndexFunctor.map C f ∘ iso_map C (· ∈ s) := by + ext _ i; exact dif_pos i.prop + exact congr_fun this x + · intro ⟨s⟩ + ext x + have : iso_map C (· ∈ s) ∘ ProjRestrict C (· ∈ s) = IndexFunctor.π_app C (· ∈ s) := by + ext _ i; exact dif_pos i.prop + erw [← this] + rfl + +end Projections + +section Products +/-! + +## Defining the basis + +Our proposed basis consists of products `e C iᵣ * ⋯ * e C i₁` with `iᵣ > ⋯ > i₁` which cannot be +written as linear combinations of lexicographically smaller products. See below for the definition +of `e`. + +### Main definitions + +* For `i : I`, we let `e C i : LocallyConstant C ℤ` denote the map + `fun f ↦ (if f.val i then 1 else 0)`. + +* `Products I` is the type of lists of decreasing elements of `I`, so a typical element is + `[i₁, i₂,..., iᵣ]` with `i₁ > i₂ > ... > iᵣ`. + +* `Products.eval C` is the `C`-evaluation of a list. It takes a term `[i₁, i₂,..., iᵣ] : Products I` + and returns the actual product `e C i₁ ··· e C iᵣ : LocallyConstant C ℤ`. + +* `GoodProducts C` is the set of `Products I` such that their `C`-evaluation cannot be written as + a linear combination of evaluations of lexicographically smaller lists. + +### Main results + +* `Products.evalFacProp` and `Products.evalFacProps` establish the fact that `Products.eval`  + interacts nicely with the projection maps from the previous section. + +* `GoodProducts.span_iff_products`: the good products span `LocallyConstant C ℤ` iff all the + products span `LocallyConstant C ℤ`. + +-/ + +/-- +`e C i` is the locally constant map from `C : Set (I → Bool)` to `ℤ` sending `f` to 1 if +`f.val i = true`, and 0 otherwise. +-/ +def e (i : I) : LocallyConstant C ℤ where + toFun := fun f ↦ (if f.val i then 1 else 0) + isLocallyConstant := by + rw [IsLocallyConstant.iff_continuous] + exact (continuous_of_discreteTopology (f := fun (a : Bool) ↦ (if a then (1 : ℤ) else 0))).comp + ((continuous_apply i).comp continuous_subtype_val) + +/-- +`Products I` is the type of lists of decreasing elements of `I`, so a typical element is +`[i₁, i₂, ...]` with `i₁ > i₂ > ...`. We order `Products I` lexicographically, so `[] < [i₁, ...]`, +and `[i₁, i₂, ...] < [j₁, j₂, ...]` if either `i₁ < j₁`, or `i₁ = j₁` and `[i₂, ...] < [j₂, ...]`. + +Terms `m = [i₁, i₂, ..., iᵣ]` of this type will be used to represent products of the form +`e C i₁ ··· e C iᵣ : LocallyConstant C ℤ` . The function associated to `m` is `m.eval`. +-/ +def Products (I : Type*) [LinearOrder I] := {l : List I // l.Chain' (·>·)} + +namespace Products + +instance : LinearOrder (Products I) := + inferInstanceAs (LinearOrder {l : List I // l.Chain' (·>·)}) + +@[simp] +theorem lt_iff_lex_lt (l m : Products I) : l < m ↔ List.Lex (·<·) l.val m.val := by + cases l; cases m; rw [Subtype.mk_lt_mk]; exact Iff.rfl + +instance : IsWellFounded (Products I) (·<·) := by + have : (· < · : Products I → _ → _) = (fun l m ↦ List.Lex (·<·) l.val m.val) := by + ext; exact lt_iff_lex_lt _ _ + rw [this] + dsimp [Products] + rw [(by rfl : (·>· : I → _) = flip (·<·))] + infer_instance + +/-- The evaluation `e C i₁ ··· e C iᵣ : C → ℤ` of a formal product `[i₁, i₂, ..., iᵣ]`. -/ +def eval (l : Products I) := (l.1.map (e C)).prod + +/-- +The predicate on products which we prove picks out a basis of `LocallyConstant C ℤ`. We call such a +product "good". +-/ +def isGood (l : Products I) : Prop := + l.eval C ∉ Submodule.span ℤ ((Products.eval C) '' {m | m < l}) + +theorem rel_head!_of_mem {i : I} {l : Products I} (hi : i ∈ l.val) : i ≤ l.val.head! := + List.Sorted.le_head! (List.chain'_iff_pairwise.mp l.prop) hi + +theorem head!_le_of_lt {q l : Products I} (h : q < l) (hq : q.val ≠ []) : + q.val.head! ≤ l.val.head! := + List.head!_le_of_lt l.val q.val h hq + +end Products + +/-- The set of good products. -/ +def GoodProducts := {l : Products I | l.isGood C} + +namespace GoodProducts + +/-- Evaluation of good products. -/ +def eval (l : {l : Products I // l.isGood C}) : LocallyConstant C ℤ := + Products.eval C l.1 + +theorem injective : Function.Injective (eval C) := by + intro ⟨a, ha⟩ ⟨b, hb⟩ h + dsimp [eval] at h + rcases lt_trichotomy a b with (h'|rfl|h') + · exfalso; apply hb; rw [← h] + exact Submodule.subset_span ⟨a, h', rfl⟩ + · rfl + · exfalso; apply ha; rw [h] + exact Submodule.subset_span ⟨b, ⟨h',rfl⟩⟩ + +/-- The image of the good products in the module `LocallyConstant C ℤ`. -/ +def range := Set.range (GoodProducts.eval C) + +/-- The type of good products is equivalent to its image. -/ +noncomputable +def equiv_range : GoodProducts C ≃ range C := + Equiv.ofInjective (eval C) (injective C) + +theorem equiv_toFun_eq_eval : (equiv_range C).toFun = Set.rangeFactorization (eval C) := rfl + +theorem linearIndependent_iff_range : LinearIndependent ℤ (GoodProducts.eval C) ↔ + LinearIndependent ℤ (fun (p : range C) ↦ p.1) := by + rw [← @Set.rangeFactorization_eq _ _ (GoodProducts.eval C), ← equiv_toFun_eq_eval C] + exact linearIndependent_equiv (equiv_range C) + +end GoodProducts + +namespace Products + +theorem eval_eq (l : Products I) (x : C) : + l.eval C x = if ∀ i, i ∈ l.val → (x.val i = true) then 1 else 0 := by + change LocallyConstant.evalMonoidHom x (l.eval C) = _ + rw [eval, map_list_prod] + split_ifs with h + · simp only [List.map_map] + apply List.prod_eq_one + simp only [List.mem_map, Function.comp_apply] + rintro _ ⟨i, hi, rfl⟩ + exact if_pos (h i hi) + · simp only [List.map_map, List.prod_eq_zero_iff, List.mem_map, Function.comp_apply] + push_neg at h + convert h with i + dsimp [LocallyConstant.evalMonoidHom, e] + simp only [ite_eq_right_iff, one_ne_zero] + +theorem evalFacProp {l : Products I} (J : I → Prop) + (h : ∀ a, a ∈ l.val → J a) [∀ j, Decidable (J j)] : + l.eval (π C J) ∘ ProjRestrict C J = l.eval C := by + ext x + dsimp [ProjRestrict] + rw [Products.eval_eq, Products.eval_eq] + congr + apply forall_congr; intro i + apply forall_congr; intro hi + simp [h i hi, Proj] + +theorem evalFacProps {l : Products I} (J K : I → Prop) + (h : ∀ a, a ∈ l.val → J a) [∀ j, Decidable (J j)] [∀ j, Decidable (K j)] + (hJK : ∀ i, J i → K i) : + l.eval (π C J) ∘ ProjRestricts C hJK = l.eval (π C K) := by + have : l.eval (π C J) ∘ Homeomorph.setCongr (proj_eq_of_subset C J K hJK) = + l.eval (π (π C K) J) := by + ext; simp [Homeomorph.setCongr, Products.eval_eq] + rw [ProjRestricts, ← Function.comp.assoc, this, ← evalFacProp (π C K) J h] + +theorem prop_of_isGood {l : Products I} (J : I → Prop) [∀ j, Decidable (J j)] + (h : l.isGood (π C J)) : ∀ a, a ∈ l.val → J a := by + intro i hi + by_contra h' + apply h + suffices : eval (π C J) l = 0 + · rw [this] + exact Submodule.zero_mem _ + ext ⟨_, ⟨_, ⟨_, rfl⟩⟩⟩ + rw [eval_eq, if_neg] + · rfl + · intro h + specialize h i hi + simp only [Proj, Bool.ite_eq_true_distrib, if_false_right_eq_and] at h + exact h' h.1 + +/-- An arbitrary product `e C i₁ * e C i₂ * ... * e C iᵣ` is in the ℤ-span of the good products. -/ +theorem eval_mem_span_goodProducts (l : Products I) : + l.eval C ∈ span ℤ (Set.range (GoodProducts.eval C)) := by + let L : Products I → Prop := fun m ↦ m.eval C ∈ span ℤ (Set.range (GoodProducts.eval C)) + suffices L l by assumption + apply IsWellFounded.induction (·<· : Products I → Products I → Prop) + intro l h + dsimp + by_cases hl : l.isGood C + · apply subset_span + exact ⟨⟨l, hl⟩, rfl⟩ + · simp only [Products.isGood, not_not] at hl + suffices : Products.eval C '' {m | m < l} ⊆ span ℤ (Set.range (GoodProducts.eval C)) + · rw [← span_le] at this + exact this hl + rintro a ⟨m, hm, rfl⟩ + exact h m hm + +end Products + +/-- The good products span `LocallyConstant C ℤ` if and only all the products do. -/ +theorem GoodProducts.span_iff_products : ⊤ ≤ span ℤ (Set.range (eval C)) ↔ + ⊤ ≤ span ℤ (Set.range (Products.eval C)) := by + refine ⟨fun h ↦ le_trans h (span_mono (fun a ⟨b, hb⟩ ↦ ⟨b.val, hb⟩)), fun h ↦ le_trans h ?_⟩ + rw [span_le] + rintro f ⟨l, rfl⟩ + let L : Products I → Prop := fun m ↦ m.eval C ∈ span ℤ (Set.range (GoodProducts.eval C)) + suffices L l by assumption + apply IsWellFounded.induction (·<· : Products I → Products I → Prop) + intro l h + dsimp + by_cases hl : l.isGood C + · apply subset_span + exact ⟨⟨l, hl⟩, rfl⟩ + · simp only [Products.isGood, not_not] at hl + suffices : Products.eval C '' {m | m < l} ⊆ span ℤ (Set.range (GoodProducts.eval C)) + · rw [← span_le] at this + exact this hl + rintro a ⟨m, hm, rfl⟩ + exact h m hm + +end Products + +section Span +/-! +## The good products span + +Most of the argument is developing an API for `π C (· ∈ s)` when `s : Finset I`; then the image +of `C` is finite with the discrete topology. In this case, there is a direct argument that the good +products span. The general result is deduced from this. + +### Main theorems + +* `GoodProducts.spanFin` : The good products span the locally constant functions on `π C (· ∈ s)` + if `s` is finite. + +* `GoodProducts.span` : The good products span `LocallyConstant C ℤ` for every closed subset `C`. +-/ + +section Fin + +variable (s : Finset I) + +/-- The `ℤ`-linear map induced by precomposition of the projection `C → π C (· ∈ s)`. -/ +noncomputable +def πJ : LocallyConstant (π C (· ∈ s)) ℤ →ₗ[ℤ] LocallyConstant C ℤ := + LocallyConstant.comapₗ ℤ _ (continuous_projRestrict C (· ∈ s)) + +theorem eval_eq_πJ (l : Products I) (hl : l.isGood (π C (· ∈ s))) : + l.eval C = πJ C s (l.eval (π C (· ∈ s))) := by + ext f + simp only [πJ, LocallyConstant.comapₗ, LinearMap.coe_mk, AddHom.coe_mk, + (continuous_projRestrict C (· ∈ s)), LocallyConstant.coe_comap, Function.comp_apply] + exact (congr_fun (Products.evalFacProp C (· ∈ s) (Products.prop_of_isGood C (· ∈ s) hl)) _).symm + +/-- `π C (· ∈ s)` is finite for a finite set `s`. -/ +noncomputable +instance : Fintype (π C (· ∈ s)) := by + let f : π C (· ∈ s) → (s → Bool) := fun x j ↦ x.val j.val + refine Fintype.ofInjective f ?_ + intro ⟨_, x, hx, rfl⟩ ⟨_, y, hy, rfl⟩ h + ext i + by_cases hi : i ∈ s + · exact congrFun h ⟨i, hi⟩ + · simp only [Proj, if_neg hi] + + +open Classical in +/-- The Kronecker delta as a locally constant map from `π C (· ∈ s)` to `ℤ`. -/ +noncomputable +def spanFinBasis (x : π C (· ∈ s)) : LocallyConstant (π C (· ∈ s)) ℤ where + toFun := fun y ↦ if y = x then 1 else 0 + isLocallyConstant := + haveI : DiscreteTopology (π C (· ∈ s)) := discrete_of_t1_of_finite + IsLocallyConstant.of_discrete _ + +open Classical in +theorem spanFinBasis.span : ⊤ ≤ Submodule.span ℤ (Set.range (spanFinBasis C s)) := by + intro f _ + rw [Finsupp.mem_span_range_iff_exists_finsupp] + use Finsupp.onFinset (Finset.univ) f.toFun (fun _ _ ↦ Finset.mem_univ _) + ext x + change LocallyConstant.evalₗ ℤ x _ = _ + simp only [zsmul_eq_mul, map_finsupp_sum, LocallyConstant.evalₗ_apply, + LocallyConstant.coe_mul, Pi.mul_apply, spanFinBasis, LocallyConstant.coe_mk, mul_ite, mul_one, + mul_zero, Finsupp.sum_ite_eq, Finsupp.mem_support_iff, ne_eq, ite_not] + split_ifs with h <;> [exact h.symm; rfl] + +/-- +A certain explicit list of locally constant maps. The theorem `factors_prod_eq_basis` shows that the +product of the elements in this list is the delta function `spanFinBasis C s x`. +-/ +def factors (x : π C (· ∈ s)) : List (LocallyConstant (π C (· ∈ s)) ℤ) := + List.map (fun i ↦ if x.val i = true then e (π C (· ∈ s)) i else (1 - (e (π C (· ∈ s)) i))) + (s.sort (·≥·)) + +theorem list_prod_apply (x : C) (l : List (LocallyConstant C ℤ)) : + l.prod x = (l.map (LocallyConstant.evalMonoidHom x)).prod := by + rw [← map_list_prod (LocallyConstant.evalMonoidHom x) l] + rfl + +theorem factors_prod_eq_basis_of_eq {x y : (π C fun x ↦ x ∈ s)} (h : y = x) : + (factors C s x).prod y = 1 := by + rw [list_prod_apply (π C (· ∈ s)) y _] + apply List.prod_eq_one + simp only [h, List.mem_map, LocallyConstant.evalMonoidHom, factors] + rintro _ ⟨a, ⟨b, _, rfl⟩, rfl⟩ + dsimp + split_ifs with hh + · rw [e, LocallyConstant.coe_mk, if_pos hh] + · rw [LocallyConstant.sub_apply, e, LocallyConstant.coe_mk, LocallyConstant.coe_mk, if_neg hh] + simp only [LocallyConstant.toFun_eq_coe, LocallyConstant.coe_one, Pi.one_apply, sub_zero] + +theorem e_mem_of_eq_true {x : (π C (· ∈ s))} {a : I} (hx : x.val a = true) : + e (π C (· ∈ s)) a ∈ factors C s x := by + rcases x with ⟨_, z, hz, rfl⟩ + simp only [factors, List.mem_map, Finset.mem_sort] + refine ⟨a, ⟨?_, if_pos hx⟩⟩ + simp only [Proj, Bool.ite_eq_true_distrib, if_false_right_eq_and] at hx + exact hx.1 + +theorem one_sub_e_mem_of_false {x y : (π C (· ∈ s))} {a : I} (ha : y.val a = true) + (hx : x.val a = false) : 1 - e (π C (· ∈ s)) a ∈ factors C s x := by + simp only [factors, List.mem_map, Finset.mem_sort] + use a + simp only [hx, ite_false, and_true] + rcases y with ⟨_, z, hz, rfl⟩ + simp only [Proj, Bool.ite_eq_true_distrib, if_false_right_eq_and] at ha + exact ha.1 + +theorem factors_prod_eq_basis_of_ne {x y : (π C (· ∈ s))} (h : y ≠ x) : + (factors C s x).prod y = 0 := by + rw [list_prod_apply (π C (· ∈ s)) y _] + apply List.prod_eq_zero + simp only [List.mem_map] + obtain ⟨a, ha⟩ : ∃ a, y.val a ≠ x.val a + · contrapose! h; ext; apply h + cases hx : x.val a + · rw [hx, ne_eq, Bool.not_eq_false] at ha + refine ⟨1 - (e (π C (· ∈ s)) a), ⟨one_sub_e_mem_of_false _ _ ha hx, ?_⟩⟩ + rw [e, LocallyConstant.evalMonoidHom_apply, LocallyConstant.sub_apply, + LocallyConstant.coe_one, Pi.one_apply, LocallyConstant.coe_mk, if_pos ha, sub_self] + · refine ⟨e (π C (· ∈ s)) a, ⟨e_mem_of_eq_true _ _ hx, ?_⟩⟩ + rw [hx] at ha + rw [LocallyConstant.evalMonoidHom_apply, e, LocallyConstant.coe_mk, if_neg ha] + +/-- If `s` is finite, the product of the elements of the list `factors C s x` +is the delta function at `x`. -/ +theorem factors_prod_eq_basis (x : π C (· ∈ s)) : + (factors C s x).prod = spanFinBasis C s x := by + ext y + dsimp [spanFinBasis] + split_ifs with h <;> [exact factors_prod_eq_basis_of_eq _ _ h; + exact factors_prod_eq_basis_of_ne _ _ h] + +theorem GoodProducts.finsupp_sum_mem_span_eval {a : I} {as : List I} + (ha : List.Chain' (· > ·) (a :: as)) {c : Products I →₀ ℤ} + (hc : (c.support : Set (Products I)) ⊆ {m | m.val ≤ as}) : + (Finsupp.sum c fun a_1 b ↦ e (π C (· ∈ s)) a * b • Products.eval (π C (· ∈ s)) a_1) ∈ + Submodule.span ℤ (Products.eval (π C (· ∈ s)) '' {m | m.val ≤ a :: as}) := by + apply Submodule.finsupp_sum_mem + intro m hm + have hsm := (LinearMap.mulLeft ℤ (e (π C (· ∈ s)) a)).map_smul + dsimp at hsm + rw [hsm] + apply Submodule.smul_mem + apply Submodule.subset_span + have hmas : m.val ≤ as + · apply hc + simpa only [Finset.mem_coe, Finsupp.mem_support_iff] using hm + refine ⟨⟨a :: m.val, ha.cons_of_le m.prop hmas⟩, ⟨List.cons_le_cons a hmas, ?_⟩⟩ + simp only [Products.eval, List.map, List.prod_cons] + +/-- If `s` is a finite subset of `I`, then the good products span. -/ +theorem GoodProducts.spanFin : ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (· ∈ s)))) := by + rw [span_iff_products] + refine le_trans (spanFinBasis.span C s) ?_ + rw [Submodule.span_le] + rintro _ ⟨x, rfl⟩ + rw [← factors_prod_eq_basis] + let l := s.sort (·≥·) + dsimp [factors] + suffices : l.Chain' (·>·) → (l.map (fun i ↦ if x.val i = true then e (π C (· ∈ s)) i + else (1 - (e (π C (· ∈ s)) i)))).prod ∈ + Submodule.span ℤ ((Products.eval (π C (· ∈ s))) '' {m | m.val ≤ l}) + · exact Submodule.span_mono (Set.image_subset_range _ _) (this (Finset.sort_sorted_gt _).chain') + induction l with + | nil => + intro _ + apply Submodule.subset_span + exact ⟨⟨[], List.chain'_nil⟩,⟨Or.inl rfl, rfl⟩⟩ + | cons a as ih => + rw [List.map_cons, List.prod_cons] + intro ha + specialize ih (by rw [List.chain'_cons'] at ha; exact ha.2) + rw [Finsupp.mem_span_image_iff_total] at ih + simp only [Finsupp.mem_supported, Finsupp.total_apply] at ih + obtain ⟨c, hc, hc'⟩ := ih + rw [← hc']; clear hc' + have hmap := fun g ↦ map_finsupp_sum (LinearMap.mulLeft ℤ (e (π C (· ∈ s)) a)) c g + dsimp at hmap ⊢ + split_ifs + · rw [hmap] + exact finsupp_sum_mem_span_eval _ _ ha hc + · ring_nf + rw [hmap] + apply Submodule.add_mem + · apply Submodule.neg_mem + exact finsupp_sum_mem_span_eval _ _ ha hc + · apply Submodule.finsupp_sum_mem + intro m hm + apply Submodule.smul_mem + apply Submodule.subset_span + refine ⟨m, ⟨?_, rfl⟩⟩ + simp only [Set.mem_setOf_eq] + have hmas : m.val ≤ as := + hc (by simpa only [Finset.mem_coe, Finsupp.mem_support_iff] using hm) + refine le_trans hmas ?_ + cases as with + | nil => exact (List.nil_lt_cons a []).le + | cons b bs => + apply le_of_lt + rw [List.chain'_cons] at ha + have hlex := List.lt.head bs (b :: bs) ha.1 + exact (List.lt_iff_lex_lt _ _).mp hlex + +end Fin + +theorem fin_comap_jointlySurjective + (hC : IsClosed C) + (f : LocallyConstant C ℤ) : ∃ (s : Finset I) + (g : LocallyConstant (π C (· ∈ s)) ℤ), f = g.comap (ProjRestrict C (· ∈ s)) := by + obtain ⟨J, g, h⟩ := @Profinite.exists_locallyConstant (Finset I)ᵒᵖ _ _ _ + (spanCone hC.isCompact) _ + (spanCone_isLimit hC.isCompact) f + exact ⟨(Opposite.unop J), g, h⟩ + +/-- The good products span all of `LocallyConstant C ℤ` if `C` is closed. -/ +theorem GoodProducts.span (hC : IsClosed C) : + ⊤ ≤ Submodule.span ℤ (Set.range (eval C)) := by + rw [span_iff_products] + intro f _ + obtain ⟨K, f', rfl⟩ : ∃ K f', f = πJ C K f' := fin_comap_jointlySurjective C hC f + refine Submodule.span_mono ?_ <| Submodule.apply_mem_span_image_of_mem_span (πJ C K) <| + spanFin C K (Submodule.mem_top : f' ∈ ⊤) + rintro l ⟨y, ⟨m, rfl⟩, rfl⟩ + exact ⟨m.val, eval_eq_πJ C K m.val m.prop⟩ + +end Span + +section Ordinal +/-! + +## Relating elements of the well-order `I` with ordinals + +We choose a well-ordering on `I`. This amounts to regarding `I` as an ordinal, and as such it +can be regarded as the set of all strictly smaller ordinals, allowing to apply ordinal induction. + +### Main definitions + +* `ord I i` is the term `i` of `I` regarded as an ordinal. + +* `term I ho` is a sufficiently small ordinal regarded as a term of `I`. + +* `contained C o` is a predicate saying that `C` is "small" enough in relation to the ordinal `o` + to satisfy the inductive hypothesis. + +* `P I` is the predicate on ordinals about linear independence of good products, which the rest of + this file is spent on proving by induction. +-/ + +variable (I) + +/-- A term of `I` regarded as an ordinal. -/ +def ord (i : I) : Ordinal := Ordinal.typein ((·<·) : I → I → Prop) i + +/-- An ordinal regarded as a term of `I`. -/ +noncomputable +def term {o : Ordinal} (ho : o < Ordinal.type ((·<·) : I → I → Prop)) : I := + Ordinal.enum ((·<·) : I → I → Prop) o ho + +variable {I} + +theorem term_ord_aux {i : I} (ho : ord I i < Ordinal.type ((·<·) : I → I → Prop)) : + term I ho = i := by + simp only [term, ord, Ordinal.enum_typein] + +@[simp] +theorem ord_term_aux {o : Ordinal} (ho : o < Ordinal.type ((·<·) : I → I → Prop)) : + ord I (term I ho) = o := by + simp only [ord, term, Ordinal.typein_enum] + +theorem ord_term {o : Ordinal} (ho : o < Ordinal.type ((·<·) : I → I → Prop)) (i : I) : + ord I i = o ↔ term I ho = i := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · subst h + exact term_ord_aux ho + · subst h + exact ord_term_aux ho + +/-- A predicate saying that `C` is "small" enough to satisfy the inductive hypothesis. -/ +def contained (o : Ordinal) : Prop := ∀ f, f ∈ C → ∀ (i : I), f i = true → ord I i < o + +variable (I) in +/-- +The predicate on ordinals which we prove by induction, see `GoodProducts.P0`, +`GoodProducts.Plimit` and `GoodProducts.linearIndependentAux` in the section `Induction` below +-/ +def P (o : Ordinal) : Prop := + o ≤ Ordinal.type (·<· : I → I → Prop) → + (∀ (C : Set (I → Bool)), IsClosed C → contained C o → + LinearIndependent ℤ (GoodProducts.eval C)) + +theorem Products.prop_of_isGood_of_contained {l : Products I} (o : Ordinal) (h : l.isGood C) + (hsC : contained C o) (i : I) (hi : i ∈ l.val) : ord I i < o := by + by_contra h' + apply h + suffices : eval C l = 0 + · simp [this, Submodule.zero_mem] + ext x + simp only [eval_eq, LocallyConstant.coe_zero, Pi.zero_apply, ite_eq_right_iff, one_ne_zero] + contrapose! h' + exact hsC x.val x.prop i (h'.1 i hi) + +end Ordinal + +section Zero +/-! + +## The zero case of the induction + +In this case, we have `contained C 0` which means that `C` is either empty or a singleton. +-/ + +instance : Subsingleton (LocallyConstant (∅ : Set (I → Bool)) ℤ) := + subsingleton_iff.mpr (fun _ _ ↦ LocallyConstant.ext isEmptyElim) + +instance : IsEmpty { l // Products.isGood (∅ : Set (I → Bool)) l } := + isEmpty_iff.mpr fun ⟨l, hl⟩ ↦ hl <| by + rw [subsingleton_iff.mp inferInstance (Products.eval ∅ l) 0] + exact Submodule.zero_mem _ + +theorem GoodProducts.linearIndependentEmpty : + LinearIndependent ℤ (eval (∅ : Set (I → Bool))) := linearIndependent_empty_type + +/-- The empty list as a `Products` -/ +def Products.nil : Products I := ⟨[], by simp only [List.chain'_nil]⟩ + +theorem Products.lt_nil_empty : { m : Products I | m < Products.nil } = ∅ := by + ext ⟨m, hm⟩ + refine ⟨fun h ↦ ?_, by tauto⟩ + simp only [Set.mem_setOf_eq, lt_iff_lex_lt, nil, List.Lex.not_nil_right] at h + +instance {α : Type*} [TopologicalSpace α] [Inhabited α] : Nontrivial (LocallyConstant α ℤ) := by + refine ⟨0, 1, fun h ↦ ?_⟩ + apply @zero_ne_one ℤ + exact FunLike.congr_fun h default + +theorem Products.isGood_nil : Products.isGood ({fun _ ↦ false} : Set (I → Bool)) Products.nil:= by + intro h + simp only [Products.lt_nil_empty, Products.eval, List.map, List.prod_nil, Set.image_empty, + Submodule.span_empty, Submodule.mem_bot, one_ne_zero] at h + +theorem Products.span_nil_eq_top : + Submodule.span ℤ (eval ({fun _ ↦ false} : Set (I → Bool)) '' {nil}) = ⊤ := by + rw [Set.image_singleton, eq_top_iff] + intro f _ + rw [Submodule.mem_span_singleton] + refine ⟨f default, ?_⟩ + simp only [eval, List.map, List.prod_nil, zsmul_eq_mul, mul_one] + ext x + obtain rfl : x = default := by simp only [Set.default_coe_singleton, eq_iff_true_of_subsingleton] + rfl + +/-- There is a unique `GoodProducts` for the singleton `{fun _ ↦ false}`. -/ +noncomputable +instance : Unique { l // Products.isGood ({fun _ ↦ false} : Set (I → Bool)) l } where + default := ⟨Products.nil, Products.isGood_nil⟩ + uniq := by + intro ⟨⟨l, hl⟩, hll⟩ + ext + apply Subtype.ext + apply (List.Lex.nil_left_or_eq_nil l (r := (·<·))).resolve_left + intro _ + apply hll + have he : {Products.nil} ⊆ {m | m < ⟨l,hl⟩} + · simpa only [Products.nil, Products.lt_iff_lex_lt, Set.singleton_subset_iff, Set.mem_setOf_eq] + apply Submodule.span_mono (Set.image_subset _ he) + rw [Products.span_nil_eq_top] + exact Submodule.mem_top + +instance (α : Type*) [TopologicalSpace α] : NoZeroSMulDivisors ℤ (LocallyConstant α ℤ) := by + constructor + intro c f h + rw [or_iff_not_imp_left] + intro hc + ext x + apply mul_right_injective₀ hc + simp [LocallyConstant.ext_iff] at h ⊢ + exact h x + +theorem GoodProducts.linearIndependentSingleton : + LinearIndependent ℤ (eval ({fun _ ↦ false} : Set (I → Bool))) := by + refine linearIndependent_unique (eval ({fun _ ↦ false} : Set (I → Bool))) ?_ + simp only [eval, Products.eval, List.map, List.prod_nil, ne_eq, one_ne_zero, not_false_eq_true] + +end Zero + +section Maps +/-! + +## `ℤ`-linear maps induced by projections + +We define injective `ℤ`-linear maps between modules of the form `LocallyConstant C ℤ` induced by +precomposition with the projections defined in the section `Projections`. + +### Main definitions + +* `πs` and `πs'` are the `ℤ`-linear maps corresponding to `ProjRestrict` and `ProjRestricts`  + respectively. + +### Main result + +* We prove that `πs` and `πs'` interact well with `Products.eval` and the main application is the + theorem `isGood_mono` which says that the property `isGood` is "monotone" on ordinals. +-/ + +theorem contained_eq_proj (o : Ordinal) (h : contained C o) : + C = π C (ord I · < o) := by + have := proj_prop_eq_self C (ord I · < o) + simp [π, Bool.not_eq_false] at this + exact (this (fun i x hx ↦ h x hx i)).symm + +theorem isClosed_proj (o : Ordinal) (hC : IsClosed C) : IsClosed (π C (ord I · < o)) := + (continuous_proj (ord I · < o)).isClosedMap C hC + +theorem contained_proj (o : Ordinal) : contained (π C (ord I · < o)) o := by + intro x ⟨_, ⟨_, h⟩⟩ j hj + dsimp [Proj] at h + simp only [← congr_fun h j, Bool.ite_eq_true_distrib, if_false_right_eq_and] at hj + exact hj.1 + +/-- The `ℤ`-linear map induced by precomposition of the projection `C → π C (ord I · < o)`. -/ +@[simps!] +noncomputable +def πs (o : Ordinal) : LocallyConstant (π C (ord I · < o)) ℤ →ₗ[ℤ] LocallyConstant C ℤ := + LocallyConstant.comapₗ ℤ (ProjRestrict C (ord I · < o)) (continuous_projRestrict _ _) + +theorem coe_πs (o : Ordinal) (f : LocallyConstant (π C (ord I · < o)) ℤ) : + πs C o f = f ∘ ProjRestrict C (ord I · < o) := by + simp only [πs_apply, continuous_projRestrict, LocallyConstant.coe_comap] + +theorem injective_πs (o : Ordinal) : Function.Injective (πs C o) := + LocallyConstant.comap_injective _ (continuous_projRestrict _ _) + (Set.surjective_mapsTo_image_restrict _ _) + +/-- The `ℤ`-linear map induced by precomposition of the projection + `π C (ord I · < o₂) → π C (ord I · < o₁)` for `o₁ ≤ o₂`. -/ +@[simps!] +noncomputable +def πs' {o₁ o₂ : Ordinal} (h : o₁ ≤ o₂) : + LocallyConstant (π C (ord I · < o₁)) ℤ →ₗ[ℤ] LocallyConstant (π C (ord I · < o₂)) ℤ := + LocallyConstant.comapₗ ℤ (ProjRestricts C (fun _ hh ↦ lt_of_lt_of_le hh h)) + (continuous_projRestricts _ _) + +theorem coe_πs' {o₁ o₂ : Ordinal} (h : o₁ ≤ o₂) (f : LocallyConstant (π C (ord I · < o₁)) ℤ) : + (πs' C h f).toFun = f.toFun ∘ (ProjRestricts C (fun _ hh ↦ lt_of_lt_of_le hh h)) := by + simp only [πs'_apply, LocallyConstant.toFun_eq_coe, continuous_projRestricts, + LocallyConstant.coe_comap] + +theorem injective_πs' {o₁ o₂ : Ordinal} (h : o₁ ≤ o₂) : Function.Injective (πs' C h) := + LocallyConstant.comap_injective _ (continuous_projRestricts _ _) (surjective_projRestricts _ _) + +namespace Products + +theorem lt_ord_of_lt {l m : Products I} {o : Ordinal} (h₁ : m < l) + (h₂ : ∀ i ∈ l.val, ord I i < o) : ∀ i ∈ m.val, ord I i < o := + List.Sorted.lt_ord_of_lt (List.chain'_iff_pairwise.mp l.2) (List.chain'_iff_pairwise.mp m.2) h₁ h₂ + +theorem eval_πs {l : Products I} {o : Ordinal} (hlt : ∀ i ∈ l.val, ord I i < o) : + πs C o (l.eval (π C (ord I · < o))) = l.eval C := by + simpa only [← LocallyConstant.coe_inj, πs_apply, continuous_projRestrict, + LocallyConstant.coe_comap] using evalFacProp C (ord I · < o) hlt + +theorem eval_πs' {l : Products I} {o₁ o₂ : Ordinal} (h : o₁ ≤ o₂) + (hlt : ∀ i ∈ l.val, ord I i < o₁) : + πs' C h (l.eval (π C (ord I · < o₁))) = l.eval (π C (ord I · < o₂)) := by + rw [← LocallyConstant.coe_inj, ← LocallyConstant.toFun_eq_coe] + simp only [πs'_apply, LocallyConstant.toFun_eq_coe, continuous_projRestricts, + LocallyConstant.coe_comap] + exact evalFacProps C (fun (i : I) ↦ ord I i < o₁) (fun (i : I) ↦ ord I i < o₂) hlt _ + +theorem eval_πs_image {l : Products I} {o : Ordinal} + (hl : ∀ i ∈ l.val, ord I i < o) : eval C '' { m | m < l } = + (πs C o) '' (eval (π C (ord I · < o)) '' { m | m < l }) := by + ext f + simp only [Set.mem_image, Set.mem_setOf_eq, exists_exists_and_eq_and] + apply exists_congr; intro m + apply and_congr_right; intro hm + rw [eval_πs C (lt_ord_of_lt hm hl)] + +theorem eval_πs_image' {l : Products I} {o₁ o₂ : Ordinal} (h : o₁ ≤ o₂) + (hl : ∀ i ∈ l.val, ord I i < o₁) : eval (π C (ord I · < o₂)) '' { m | m < l } = + (πs' C h) '' (eval (π C (ord I · < o₁)) '' { m | m < l }) := by + ext f + simp only [Set.mem_image, Set.mem_setOf_eq, exists_exists_and_eq_and] + apply exists_congr; intro m + apply and_congr_right; intro hm + rw [eval_πs' C h (lt_ord_of_lt hm hl)] + +theorem head_lt_ord_of_isGood {l : Products I} {o : Ordinal} + (h : l.isGood (π C (ord I · < o))) (hn : l.val ≠ []) : ord I (l.val.head!) < o := + prop_of_isGood C (ord I · < o) h l.val.head! (List.head!_mem_self hn) + +/-- +If `l` is good w.r.t. `π C (ord I · < o₁)` and `o₁ ≤ o₂`, then it is good w.r.t. +`π C (ord I · < o₂)` +-/ +theorem isGood_mono {l : Products I} {o₁ o₂ : Ordinal} (h : o₁ ≤ o₂) + (hl : l.isGood (π C (ord I · < o₁))) : l.isGood (π C (ord I · < o₂)) := by + intro hl' + apply hl + rwa [eval_πs_image' C h (prop_of_isGood C _ hl), ← eval_πs' C h (prop_of_isGood C _ hl), + Submodule.apply_mem_span_image_iff_mem_span (injective_πs' C h)] at hl' + +end Products + +end Maps + +section Limit +/-! + +## The limit case of the induction + +We relate linear independence in `LocallyConstant (π C (ord I · < o')) ℤ` with linear independence +in `LocallyConstant C ℤ`, where `contained C o` and `o' < o`. + +When `o` is a limit ordinal, we prove that the good products in `LocallyConstant C ℤ` are linearly +independent if and only if a certain directed union is linearly independent. Each term in this +directed union is in bijection with the good products w.r.t. `π C (ord I · < o')` for an ordinal +`o' < o`, and these are linearly independent by the inductive hypothesis. + +### Main definitions + +* `GoodProducts.smaller` is the image of good products coming from a smaller ordinal. + +* `GoodProducts.range_equiv`: The image of the `GoodProducts` in `C` is equivalent to the union of + `smaller C o'` over all ordinals `o' < o`. + +### Main results + +* `Products.limitOrdinal`: for `o` a limit ordinal such that `contained C o`, a product `l` is good + w.r.t. `C` iff it there exists an ordinal `o' < o` such that `l` is good w.r.t. + `π C (ord I · < o')`. + +* `GoodProducts.linearIndependent_iff_union_smaller` is the result mentioned above, that the good + products are linearly independent iff a directed union is. +-/ + +namespace GoodProducts + +/-- +The image of the `GoodProducts` for `π C (ord I · < o)` in `LocallyConstant C ℤ`. The name `smaller` +refers to the setting in which we will use this, when we are mapping in `GoodProducts` from a +smaller set, i.e. when `o` is a smaller ordinal than the one `C` is "contained" in. +-/ +def smaller (o : Ordinal) : Set (LocallyConstant C ℤ) := + (πs C o) '' (range (π C (ord I · < o))) + +/-- +The map from the image of the `GoodProducts` in `LocallyConstant (π C (ord I · < o)) ℤ` to +`smaller C o` +-/ +noncomputable +def range_equiv_smaller_toFun (o : Ordinal) (x : range (π C (ord I · < o))) : smaller C o := + ⟨πs C o ↑x, x.val, x.property, rfl⟩ + +theorem range_equiv_smaller_toFun_bijective (o : Ordinal) : + Function.Bijective (range_equiv_smaller_toFun C o) := by + dsimp [range_equiv_smaller_toFun] + refine ⟨fun a b hab ↦ ?_, fun ⟨a, b, hb⟩ ↦ ?_⟩ + · ext1 + simp only [Subtype.mk.injEq] at hab + exact injective_πs C o hab + · use ⟨b, hb.1⟩ + simpa only [Subtype.mk.injEq] using hb.2 + +/-- +The equivalence from the image of the `GoodProducts` in `LocallyConstant (π C (ord I · < o)) ℤ` to +`smaller C o` +-/ +noncomputable +def range_equiv_smaller (o : Ordinal) : range (π C (ord I · < o)) ≃ smaller C o := + Equiv.ofBijective (range_equiv_smaller_toFun C o) (range_equiv_smaller_toFun_bijective C o) + +theorem smaller_factorization (o : Ordinal) : + (fun (p : smaller C o) ↦ p.1) ∘ (range_equiv_smaller C o).toFun = + (πs C o) ∘ (fun (p : range (π C (ord I · < o))) ↦ p.1) := by rfl + +theorem linearIndependent_iff_smaller (o : Ordinal) : + LinearIndependent ℤ (GoodProducts.eval (π C (ord I · < o))) ↔ + LinearIndependent ℤ (fun (p : smaller C o) ↦ p.1) := by + rw [GoodProducts.linearIndependent_iff_range, + ← LinearMap.linearIndependent_iff (πs C o) + (LinearMap.ker_eq_bot_of_injective (injective_πs _ _)), ← smaller_factorization C o] + exact linearIndependent_equiv _ + +theorem smaller_mono {o₁ o₂ : Ordinal} (h : o₁ ≤ o₂) : smaller C o₁ ⊆ smaller C o₂ := by + rintro f ⟨g, hg, rfl⟩ + simp only [smaller, Set.mem_image] + use πs' C h g + obtain ⟨⟨l, gl⟩, rfl⟩ := hg + refine ⟨?_, ?_⟩ + · use ⟨l, Products.isGood_mono C h gl⟩ + ext x + rw [eval, ← Products.eval_πs' _ h (Products.prop_of_isGood C _ gl), eval] + · rw [← LocallyConstant.coe_inj, coe_πs C o₂, ← LocallyConstant.toFun_eq_coe, coe_πs', + Function.comp.assoc, projRestricts_comp_projRestrict C _, coe_πs] + rfl + +end GoodProducts + +variable {o : Ordinal} (ho : o.IsLimit) (hsC : contained C o) + +theorem Products.limitOrdinal (l : Products I) : l.isGood (π C (ord I · < o)) ↔ + ∃ (o' : Ordinal), o' < o ∧ l.isGood (π C (ord I · < o')) := by + refine ⟨fun h ↦ ?_, fun ⟨o', ⟨ho', hl⟩⟩ ↦ isGood_mono C (le_of_lt ho') hl⟩ + use Finset.sup l.val.toFinset (fun a ↦ Order.succ (ord I a)) + have ha : ⊥ < o := by rw [Ordinal.bot_eq_zero, Ordinal.pos_iff_ne_zero]; exact ho.1 + have hslt : Finset.sup l.val.toFinset (fun a ↦ Order.succ (ord I a)) < o + · simp only [Finset.sup_lt_iff ha, List.mem_toFinset] + exact fun b hb ↦ ho.2 _ (prop_of_isGood C (ord I · < o) h b hb) + refine ⟨hslt, fun he ↦ h ?_⟩ + have hlt : ∀ i ∈ l.val, ord I i < Finset.sup l.val.toFinset (fun a ↦ Order.succ (ord I a)) + · intro i hi + simp only [Finset.lt_sup_iff, List.mem_toFinset, Order.lt_succ_iff] + exact ⟨i, hi, le_rfl⟩ + rwa [eval_πs_image' C (le_of_lt hslt) hlt, ← eval_πs' C (le_of_lt hslt) hlt, + Submodule.apply_mem_span_image_iff_mem_span (injective_πs' C _)] + +theorem GoodProducts.union : range C = ⋃ (e : {o' // o' < o}), (smaller C e.val) := by + ext p + simp only [smaller, range, Set.mem_iUnion, Set.mem_image, Set.mem_range, Subtype.exists] + refine ⟨fun hp ↦ ?_, fun hp ↦ ?_⟩ + · obtain ⟨l, hl, rfl⟩ := hp + rw [contained_eq_proj C o hsC, Products.limitOrdinal C ho] at hl + obtain ⟨o', ho'⟩ := hl + refine ⟨o', ho'.1, eval (π C (ord I · < o')) ⟨l, ho'.2⟩, ⟨l, ho'.2, rfl⟩, ?_⟩ + exact Products.eval_πs C (Products.prop_of_isGood C _ ho'.2) + · obtain ⟨o', h, _, ⟨l, hl, rfl⟩, rfl⟩ := hp + refine ⟨l, ?_, (Products.eval_πs C (Products.prop_of_isGood C _ hl)).symm⟩ + rw [contained_eq_proj C o hsC] + exact Products.isGood_mono C (le_of_lt h) hl + +/-- +The image of the `GoodProducts` in `C` is equivalent to the union of `smaller C o'` over all +ordinals `o' < o`. +-/ +def GoodProducts.range_equiv : range C ≃ ⋃ (e : {o' // o' < o}), (smaller C e.val) := + Equiv.Set.ofEq (union C ho hsC) + +theorem GoodProducts.range_equiv_factorization : + (fun (p : ⋃ (e : {o' // o' < o}), (smaller C e.val)) ↦ p.1) ∘ (range_equiv C ho hsC).toFun = + (fun (p : range C) ↦ (p.1 : LocallyConstant C ℤ)) := rfl + +theorem GoodProducts.linearIndependent_iff_union_smaller {o : Ordinal} (ho : o.IsLimit) + (hsC : contained C o) : LinearIndependent ℤ (GoodProducts.eval C) ↔ + LinearIndependent ℤ (fun (p : ⋃ (e : {o' // o' < o}), (smaller C e.val)) ↦ p.1) := by + rw [GoodProducts.linearIndependent_iff_range, ← range_equiv_factorization C ho hsC] + exact linearIndependent_equiv (range_equiv C ho hsC) + +end Limit + +section Successor +/-! + +## The successor case in the induction + +Here we assume that `o` is an ordinal such that `contained C (o+1)` and `o < I`. The element in `I` +corresponding to `o` is called `term I ho`, but in this informal docstring we refer to it simply as +`o`. + +This section follows the proof in [scholze2019condensed] quite closely. A translation of the +notation there is as follows: + +``` +[scholze2019condensed] | This file +`S₀` |`C0` +`S₁` |`C1` +`\overline{S}` |`π C (ord I · < o) +`\overline{S}'` |`C'` +The left map in the exact sequence |`πs` +The right map in the exact sequence |`Linear_CC'` +``` + +When comparing the proof of the successor case in Theorem 5.4 in [scholze2019condensed] with this +proof, one should read the phrase "is a basis" as "is linearly independent". Also, the short exact +sequence in [scholze2019condensed] is only proved to be left exact here (indeed, that is enough +since we are only proving linear independence). + +This section is split into two sections. The first one, `ExactSequence` defines the left exact +sequence mentioned in the previous paragraph (see `succ_mono` and `succ_exact`). It corresponds to +the penultimate paragraph of the proof in [scholze2019condensed]. The second one, `GoodProducts` +corresponds to the last paragraph in the proof in [scholze2019condensed]. + +### Main definitions + +The main definitions in the section `ExactSequence` are all just notation explained in the table +above. + +The main definitions in the section `GoodProducts` are as follows: + +* `MaxProducts`: the set of good products that contain the ordinal `o` (since we have + `contained C (o+1)`, these all start with `o`). + +* `GoodProducts.sum_equiv`: the equivalence between `GoodProducts C` and the disjoint union of + `MaxProducts C` and `GoodProducts (π C (ord I · < o))`. + +### Main results + +* The main results in the section `ExactSequence` are `succ_mono` and `succ_exact` which together + say that the secuence given by `πs` and `Linear_CC'` is left exact: + ``` + f g + 0 --→ LocallyConstant (π C (ord I · < o)) ℤ --→ LocallyConstant C ℤ --→ LocallyConstant C' ℤ + ``` + where `f` is `πs` and `g` is `Linear_CC'`. + +The main results in the section `GoodProducts` are as follows: + +* `Products.max_eq_eval` says that the linear map on the right in the exact sequence, i.e. + `Linear_CC'`, takes the evaluation of a term of `MaxProducts` to the evaluation of the + corresponding list with the leading `o` removed. + +* `GoodProducts.maxTail_isGood` says that removing the leading `o` from a term of `MaxProducts C`  + yields a list which `isGood` with respect to `C'`. +-/ + +variable {o : Ordinal} (hC : IsClosed C) (hsC : contained C (Order.succ o)) + (ho : o < Ordinal.type (·<· : I → I → Prop)) + +section ExactSequence + +/-- The subset of `C` consisting of those elements whose `o`-th entry is `false`. -/ +def C0 := C ∩ {f | f (term I ho) = false} + +/-- The subset of `C` consisting of those elements whose `o`-th entry is `true`. -/ +def C1 := C ∩ {f | f (term I ho) = true} + +theorem isClosed_C0 : IsClosed (C0 C ho) := by + refine hC.inter ?_ + have h : Continuous (fun (f : I → Bool) ↦ f (term I ho)) := continuous_apply (term I ho) + exact IsClosed.preimage h (s := {false}) (isClosed_discrete _) + +theorem isClosed_C1 : IsClosed (C1 C ho) := by + refine hC.inter ?_ + have h : Continuous (fun (f : I → Bool) ↦ f (term I ho)) := continuous_apply (term I ho) + exact IsClosed.preimage h (s := {true}) (isClosed_discrete _) + +theorem contained_C1 : contained (π (C1 C ho) (ord I · < o)) o := + contained_proj _ _ + +theorem union_C0C1_eq : (C0 C ho) ∪ (C1 C ho) = C := by + ext x + simp only [C0, C1, Set.mem_union, Set.mem_inter_iff, Set.mem_setOf_eq, + ← and_or_left, and_iff_left_iff_imp, Bool.dichotomy (x (term I ho)), implies_true] + +/-- +The intersection of `C0` and the projection of `C1`. We will apply the inductive hypothesis to +this set. +-/ +def C' := C0 C ho ∩ π (C1 C ho) (ord I · < o) + +theorem isClosed_C' : IsClosed (C' C ho) := + IsClosed.inter (isClosed_C0 _ hC _) (isClosed_proj _ _ (isClosed_C1 _ hC _)) + +theorem contained_C' : contained (C' C ho) o := fun f hf i hi ↦ contained_C1 C ho f hf.2 i hi + +variable (o) + +/-- Swapping the `o`-th coordinate to `true`. -/ +noncomputable +def SwapTrue : (I → Bool) → I → Bool := + fun f i ↦ if ord I i = o then true else f i + +theorem continuous_swapTrue : + Continuous (SwapTrue o : (I → Bool) → I → Bool) := by + dsimp [SwapTrue] + apply continuous_pi + intro i + apply Continuous.comp' + · apply continuous_bot + · apply continuous_apply + +variable {o} + +theorem swapTrue_mem_C1 (f : π (C1 C ho) (ord I · < o)) : + SwapTrue o f.val ∈ C1 C ho := by + obtain ⟨f, g, hg, rfl⟩ := f + convert hg + dsimp [SwapTrue] + ext i + split_ifs with h + · rw [ord_term ho] at h + simpa only [← h] using hg.2.symm + · simp only [Proj, ite_eq_left_iff, not_lt, @eq_comm _ false, ← Bool.not_eq_true] + specialize hsC g hg.1 i + intro h' + contrapose! hsC + exact ⟨hsC, Order.succ_le_of_lt (h'.lt_of_ne' h)⟩ + +/-- The first way to map `C'` into `C`. -/ +def CC'₀ : C' C ho → C := fun g ↦ ⟨g.val,g.prop.1.1⟩ + +/-- The second way to map `C'` into `C`. -/ +noncomputable +def CC'₁ : C' C ho → C := + fun g ↦ ⟨SwapTrue o g.val, (swapTrue_mem_C1 C hsC ho ⟨g.val,g.prop.2⟩).1⟩ + +theorem continuous_CC'₀ : Continuous (CC'₀ C ho) := Continuous.subtype_mk continuous_subtype_val _ + +theorem continuous_CC'₁ : Continuous (CC'₁ C hsC ho) := + Continuous.subtype_mk (Continuous.comp (continuous_swapTrue o) continuous_subtype_val) _ + +/-- The `ℤ`-linear map induced by precomposing with `CC'₀` -/ +noncomputable +def Linear_CC'₀ : LocallyConstant C ℤ →ₗ[ℤ] LocallyConstant (C' C ho) ℤ := + LocallyConstant.comapₗ ℤ (CC'₀ C ho) (continuous_CC'₀ C ho) + +/-- The `ℤ`-linear map induced by precomposing with `CC'₁` -/ +noncomputable +def Linear_CC'₁ : LocallyConstant C ℤ →ₗ[ℤ] LocallyConstant (C' C ho) ℤ := + LocallyConstant.comapₗ ℤ (CC'₁ C hsC ho) (continuous_CC'₁ C hsC ho) + +/-- The difference between `Linear_CC'₁` and `Linear_CC'₀`. -/ +noncomputable +def Linear_CC' : LocallyConstant C ℤ →ₗ[ℤ] LocallyConstant (C' C ho) ℤ := + Linear_CC'₁ C hsC ho - Linear_CC'₀ C ho + +theorem CC_comp_zero : ∀ y, (Linear_CC' C hsC ho) ((πs C o) y) = 0 := by + intro y + ext x + dsimp [Linear_CC', Linear_CC'₀, Linear_CC'₁, LocallyConstant.sub_apply] + simp only [continuous_CC'₀, continuous_CC'₁, LocallyConstant.coe_comap, continuous_projRestrict, + Function.comp_apply, sub_eq_zero] + congr 1 + ext i + dsimp [CC'₀, CC'₁, ProjRestrict, Proj] + apply if_ctx_congr Iff.rfl _ (fun _ ↦ rfl) + simp only [SwapTrue, ite_eq_right_iff] + intro h₁ h₂ + exact (h₁.ne h₂).elim + +theorem C0_projOrd {x : I → Bool} (hx : x ∈ C0 C ho) : Proj (ord I · < o) x = x := by + ext i + simp only [Proj, Set.mem_setOf, ite_eq_left_iff, not_lt] + intro hi + rw [le_iff_lt_or_eq] at hi + cases' hi with hi hi + · specialize hsC x hx.1 i + rw [← not_imp_not] at hsC + simp only [not_lt, Bool.not_eq_true, Order.succ_le_iff] at hsC + exact (hsC hi).symm + · simp only [C0, Set.mem_inter_iff, Set.mem_setOf_eq] at hx + rw [eq_comm, ord_term ho] at hi + rw [← hx.2, hi] + +theorem C1_projOrd {x : I → Bool} (hx : x ∈ C1 C ho) : SwapTrue o (Proj (ord I · < o) x) = x := by + ext i + dsimp [SwapTrue, Proj] + split_ifs with hi h + · rw [ord_term ho] at hi + rw [← hx.2, hi] + · rfl + · simp only [not_lt] at h + have h' : o < ord I i := lt_of_le_of_ne h (Ne.symm hi) + specialize hsC x hx.1 i + rw [← not_imp_not] at hsC + simp only [not_lt, Bool.not_eq_true, Order.succ_le_iff] at hsC + exact (hsC h').symm + +open Classical in +theorem CC_exact {f : LocallyConstant C ℤ} (hf : Linear_CC' C hsC ho f = 0) : + ∃ y, πs C o y = f := by + dsimp [Linear_CC', Linear_CC'₀, Linear_CC'₁] at hf + simp only [sub_eq_zero, ← LocallyConstant.coe_inj, LocallyConstant.coe_comap, + continuous_CC'₀, continuous_CC'₁] at hf + let C₀C : C0 C ho → C := fun x ↦ ⟨x.val, x.prop.1⟩ + have h₀ : Continuous C₀C := Continuous.subtype_mk continuous_induced_dom _ + let C₁C : π (C1 C ho) (ord I · < o) → C := + fun x ↦ ⟨SwapTrue o x.val, (swapTrue_mem_C1 C hsC ho x).1⟩ + have h₁ : Continuous C₁C := Continuous.subtype_mk + ((continuous_swapTrue o).comp continuous_subtype_val) _ + refine ⟨LocallyConstant.piecewise' ?_ (isClosed_C0 C hC ho) + (isClosed_proj _ o (isClosed_C1 C hC ho)) (f.comap C₀C) (f.comap C₁C) ?_, ?_⟩ + · rintro _ ⟨y, hyC, rfl⟩ + simp only [Set.mem_union, Set.mem_setOf_eq, Set.mem_univ, iff_true] + rw [← union_C0C1_eq C ho] at hyC + refine hyC.imp (fun hyC ↦ ?_) (fun hyC ↦ ⟨y, hyC, rfl⟩) + rwa [C0_projOrd C hsC ho hyC] + · intro x hx + simpa only [h₀, h₁, LocallyConstant.coe_comap] using (congrFun hf ⟨x, hx⟩).symm + · ext ⟨x, hx⟩ + rw [← union_C0C1_eq C ho] at hx + cases' hx with hx₀ hx₁ + · have hx₀' : ProjRestrict C (ord I · < o) ⟨x, hx⟩ = x + · simpa only [ProjRestrict, Set.MapsTo.val_restrict_apply] using C0_projOrd C hsC ho hx₀ + simp only [hx₀', hx₀, h₀, LocallyConstant.piecewise'_apply_left, LocallyConstant.coe_comap, + Function.comp_apply, πs_apply, continuous_projRestrict] + · have hx₁' : (ProjRestrict C (ord I · < o) ⟨x, hx⟩).val ∈ π (C1 C ho) (ord I · < o) + · simpa only [ProjRestrict, Set.MapsTo.val_restrict_apply] using ⟨x, hx₁, rfl⟩ + simp only [πs_apply, continuous_projRestrict, LocallyConstant.coe_comap, Function.comp_apply, + hx₁', LocallyConstant.piecewise'_apply_right, h₁] + congr + exact C1_projOrd C hsC ho hx₁ + +variable (o) in +theorem succ_mono : CategoryTheory.Mono (ModuleCat.ofHom (πs C o)) := by + rw [ModuleCat.mono_iff_injective] + exact injective_πs _ _ + +theorem succ_exact : + CategoryTheory.Exact (ModuleCat.ofHom (πs C o)) (ModuleCat.ofHom (Linear_CC' C hsC ho)) := by + rw [ModuleCat.exact_iff] + ext f + rw [LinearMap.mem_ker, LinearMap.mem_range] + refine ⟨fun ⟨y, hy⟩ ↦ ?_, fun hf ↦ ?_⟩ + · simpa only [ModuleCat.ofHom, ← hy] using CC_comp_zero _ _ _ y + · exact CC_exact _ hC _ ho hf + +end ExactSequence + +section GoodProducts + +namespace GoodProducts + +/-- +The `GoodProducts` in `C` that contain `o` (they necessarily start with `o`, see +`GoodProducts.head!_eq_o_of_maxProducts`) +-/ +def MaxProducts : Set (Products I) := {l | l.isGood C ∧ term I ho ∈ l.val} + +theorem union_succ : GoodProducts C = GoodProducts (π C (ord I · < o)) ∪ MaxProducts C ho := by + ext l + simp only [GoodProducts, MaxProducts, Set.mem_union, Set.mem_setOf_eq] + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · by_cases hh : term I ho ∈ l.val + · exact Or.inr ⟨h, hh⟩ + · left + intro he + apply h + have h' := Products.prop_of_isGood_of_contained C _ h hsC + simp only [Order.lt_succ_iff] at h' + simp only [not_imp_not] at hh + have hh' : ∀ a ∈ l.val, ord I a < o + · intro a ha + refine (h' a ha).lt_of_ne ?_ + rw [ne_eq, ord_term ho a] + rintro rfl + contradiction + rwa [Products.eval_πs_image C hh', ← Products.eval_πs C hh', + Submodule.apply_mem_span_image_iff_mem_span (injective_πs _ _)] + · refine h.elim (fun hh ↦ ?_) And.left + have := Products.isGood_mono C (Order.lt_succ o).le hh + rwa [contained_eq_proj C (Order.succ o) hsC] + +/-- The inclusion map from the sum of `GoodProducts (π C (ord I · < o))` and + `(MaxProducts C ho)` to `Products I`. -/ +def sum_to : (GoodProducts (π C (ord I · < o))) ⊕ (MaxProducts C ho) → Products I := + Sum.elim Subtype.val Subtype.val + +theorem injective_sum_to : Function.Injective (sum_to C ho) := by + refine Function.Injective.sum_elim Subtype.val_injective Subtype.val_injective + (fun ⟨a,ha⟩ ⟨b,hb⟩ ↦ (fun (hab : a = b) ↦ ?_)) + rw [← hab] at hb + have ha' := Products.prop_of_isGood C _ ha (term I ho) hb.2 + simp only [ord_term_aux, lt_self_iff_false] at ha' + +theorem sum_to_range : + Set.range (sum_to C ho) = GoodProducts (π C (ord I · < o)) ∪ MaxProducts C ho := by + have h : Set.range (sum_to C ho) = _ ∪ _ := Set.Sum.elim_range _ _; rw [h]; congr<;> ext l + · exact ⟨fun ⟨m,hm⟩ ↦ by rw [← hm]; exact m.prop, fun hl ↦ ⟨⟨l,hl⟩, rfl⟩⟩ + · exact ⟨fun ⟨m,hm⟩ ↦ by rw [← hm]; exact m.prop, fun hl ↦ ⟨⟨l,hl⟩, rfl⟩⟩ + +/-- The equivalence from the sum of `GoodProducts (π C (ord I · < o))` and + `(MaxProducts C ho)` to `GoodProducts C`. -/ +noncomputable +def sum_equiv : GoodProducts (π C (ord I · < o)) ⊕ (MaxProducts C ho) ≃ GoodProducts C := + calc _ ≃ Set.range (sum_to C ho) := Equiv.ofInjective (sum_to C ho) (injective_sum_to C ho) + _ ≃ _ := Equiv.Set.ofEq <| by rw [sum_to_range C ho, union_succ C hsC ho] + +theorem sum_equiv_comp_eval_eq_elim : eval C ∘ (sum_equiv C hsC ho).toFun = + (Sum.elim (fun (l : GoodProducts (π C (ord I · < o))) ↦ Products.eval C l.1) + (fun (l : MaxProducts C ho) ↦ Products.eval C l.1)) := by + ext ⟨_,_⟩ <;> [rfl; rfl] + +/-- Let + +`N := LocallyConstant (π C (ord I · < o)) ℤ` + +`M := LocallyConstant C ℤ` + +`P := LocallyConstant (C' C ho) ℤ` + +`ι := GoodProducts (π C (ord I · < o))` + +`ι' := GoodProducts (C' C ho')` + +`v : ι → N := GoodProducts.eval (π C (ord I · < o))` + +Then `SumEval C ho` is the map `u` in the diagram below. It is linearly independent if and only if +`GoodProducts.eval C` is, see `linearIndependent_iff_sum`. The top row is the exact sequence given +by `succ_exact` and `succ_mono`. The left square commutes by `GoodProducts.square_commutes`. +``` +0 --→ N --→ M --→ P + ↑ ↑ ↑ + v| u| | + ι → ι ⊕ ι' ← ι' +``` +-/ +def SumEval : GoodProducts (π C (ord I · < o)) ⊕ MaxProducts C ho → + LocallyConstant C ℤ := + Sum.elim (fun l ↦ l.1.eval C) (fun l ↦ l.1.eval C) + +theorem linearIndependent_iff_sum : + LinearIndependent ℤ (eval C) ↔ LinearIndependent ℤ (SumEval C ho) := by + rw [← linearIndependent_equiv (sum_equiv C hsC ho), SumEval, + ← sum_equiv_comp_eval_eq_elim C hsC ho] + exact Iff.rfl + +theorem span_sum : Set.range (eval C) = Set.range (Sum.elim + (fun (l : GoodProducts (π C (ord I · < o))) ↦ Products.eval C l.1) + (fun (l : MaxProducts C ho) ↦ Products.eval C l.1)) := by + rw [← sum_equiv_comp_eval_eq_elim C hsC ho, Equiv.toFun_as_coe, + EquivLike.range_comp (e := sum_equiv C hsC ho)] + + +theorem square_commutes : SumEval C ho ∘ Sum.inl = + ModuleCat.ofHom (πs C o) ∘ eval (π C (ord I · < o)) := by + ext l + dsimp [SumEval] + rw [← Products.eval_πs C (Products.prop_of_isGood _ _ l.prop)] + rfl + +end GoodProducts + +theorem swapTrue_eq_true (x : I → Bool) : SwapTrue o x (term I ho) = true := by + simp only [SwapTrue, ord_term_aux, ite_true] + +theorem mem_C'_eq_false : ∀ x, x ∈ C' C ho → x (term I ho) = false := by + rintro x ⟨_, y, _, rfl⟩ + simp only [Proj, ord_term_aux, lt_self_iff_false, ite_false] + +/-- `List.tail` as a `Products`. -/ +def Products.Tail (l : Products I) : Products I := + ⟨l.val.tail, List.Chain'.tail l.prop⟩ + +theorem Products.max_eq_o_cons_tail (l : Products I) (hl : l.val ≠ []) + (hlh : l.val.head! = term I ho) : l.val = term I ho :: l.Tail.val := by + rw [← List.cons_head!_tail hl, hlh] + rfl + +theorem Products.max_eq_o_cons_tail' (l : Products I) (hl : l.val ≠ []) + (hlh : l.val.head! = term I ho) (hlc : List.Chain' (·>·) (term I ho :: l.Tail.val)) : + l = ⟨term I ho :: l.Tail.val, hlc⟩ := by + simp_rw [← max_eq_o_cons_tail ho l hl hlh] + rfl + +theorem GoodProducts.head!_eq_o_of_maxProducts (l : ↑(MaxProducts C ho)) : + l.val.val.head! = term I ho := by + rw [eq_comm, ← ord_term ho] + have hm := l.prop.2 + have := Products.prop_of_isGood_of_contained C _ l.prop.1 hsC l.val.val.head! + (List.head!_mem_self (List.ne_nil_of_mem hm)) + simp only [Order.lt_succ_iff] at this + refine eq_of_le_of_not_lt this (not_lt.mpr ?_) + have h : ord I (term I ho) ≤ ord I l.val.val.head! + · simp only [← ord_term_aux, ord, Ordinal.typein_le_typein, not_lt] + exact Products.rel_head!_of_mem hm + rwa [ord_term_aux] at h + +theorem GoodProducts.max_eq_o_cons_tail (l : MaxProducts C ho) : + l.val.val = (term I ho) :: l.val.Tail.val := + Products.max_eq_o_cons_tail ho l.val (List.ne_nil_of_mem l.prop.2) + (head!_eq_o_of_maxProducts _ hsC ho l) + +theorem Products.evalCons {l : List I} {a : I} + (hla : (a::l).Chain' (·>·)) : Products.eval C ⟨a::l,hla⟩ = + (e C a) * Products.eval C ⟨l,List.Chain'.sublist hla (List.tail_sublist (a::l))⟩ := by + simp only [eval._eq_1, List.map, List.prod_cons] + +theorem Products.max_eq_eval (l : Products I) (hl : l.val ≠ []) + (hlh : l.val.head! = term I ho) : + Linear_CC' C hsC ho (l.eval C) = l.Tail.eval (C' C ho) := by + have hlc : ((term I ho) :: l.Tail.val).Chain' (·>·) + · rw [← max_eq_o_cons_tail ho l hl hlh]; exact l.prop + rw [max_eq_o_cons_tail' ho l hl hlh hlc, Products.evalCons] + ext x + simp only [Linear_CC', Linear_CC'₀, Linear_CC'₁, LocallyConstant.comapₗ, Subtype.coe_eta, + LinearMap.sub_apply, LinearMap.coe_mk, AddHom.coe_mk, LocallyConstant.sub_apply, + continuous_CC'₀, continuous_CC'₁, LocallyConstant.coe_comap, LocallyConstant.coe_mul, + Function.comp_apply, Pi.mul_apply] + rw [CC'₁, CC'₀, Products.eval_eq, Products.eval_eq, Products.eval_eq] + simp only [mul_ite, mul_one, mul_zero] + have hi' : ∀ i, i ∈ l.Tail.val → (x.val i = SwapTrue o x.val i) + · intro i hi + simp only [SwapTrue, @eq_comm _ (x.val i), ite_eq_right_iff, ord_term ho] + rintro rfl + exact ((List.Chain.rel hlc hi).ne rfl).elim + have H : + (∀ i, i ∈ l.Tail.val → (x.val i = true)) = (∀ i, i ∈ l.Tail.val → (SwapTrue o x.val i = true)) + · apply forall_congr; intro i; apply forall_congr; intro hi; rw [hi' i hi] + simp only [H] + split_ifs with h₁ h₂ h₃ <;> dsimp [e] + · rw [if_pos (swapTrue_eq_true _ _), if_neg] + · rfl + · simp [mem_C'_eq_false C ho x x.prop, Bool.coe_false] + · push_neg at h₂; obtain ⟨i, hi⟩ := h₂; exfalso; rw [hi' i hi.1] at hi; exact hi.2 (h₁ i hi.1) + · push_neg at h₁; obtain ⟨i, hi⟩ := h₁; exfalso; rw [← hi' i hi.1] at hi; exact hi.2 (h₃ i hi.1) + · rfl + +namespace GoodProducts + +theorem max_eq_eval (l : MaxProducts C ho) : + Linear_CC' C hsC ho (l.val.eval C) = l.val.Tail.eval (C' C ho) := + Products.max_eq_eval _ _ _ _ (List.ne_nil_of_mem l.prop.2) + (head!_eq_o_of_maxProducts _ hsC ho l) + +theorem max_eq_eval_unapply : + (Linear_CC' C hsC ho) ∘ (fun (l : MaxProducts C ho) ↦ Products.eval C l.val) = + (fun l ↦ l.val.Tail.eval (C' C ho)) := by + ext1 l + exact max_eq_eval _ _ _ _ + +theorem chain'_cons_of_lt (l : MaxProducts C ho) + (q : Products I) (hq : q < l.val.Tail) : + List.Chain' (fun x x_1 ↦ x > x_1) (term I ho :: q.val) := by + rw [List.chain'_iff_pairwise] + simp only [gt_iff_lt, List.pairwise_cons] + refine ⟨fun a ha ↦ lt_of_le_of_lt (Products.rel_head!_of_mem ha) ?_, + List.chain'_iff_pairwise.mp q.prop⟩ + refine lt_of_le_of_lt (Products.head!_le_of_lt hq (q.val.ne_nil_of_mem ha)) ?_ + by_cases hM : l.val.Tail.val = [] + · rw [Products.lt_iff_lex_lt, hM] at hq + simp only [List.Lex.not_nil_right] at hq + · have := l.val.prop + rw [max_eq_o_cons_tail C hsC ho l, List.chain'_iff_pairwise] at this + exact List.rel_of_pairwise_cons this (List.head!_mem_self hM) + +theorem good_lt_maxProducts (q : GoodProducts (π C (ord I · < o))) + (l : MaxProducts C ho) : List.Lex (·<·) q.val.val l.val.val := by + by_cases h : q.val.val = [] + · rw [h, max_eq_o_cons_tail C hsC ho l] + exact List.Lex.nil + · rw [← List.cons_head!_tail h, max_eq_o_cons_tail C hsC ho l] + apply List.Lex.rel + rw [← Ordinal.typein_lt_typein (·<·)] + simp only [term, Ordinal.typein_enum] + exact Products.prop_of_isGood C _ q.prop q.val.val.head! (List.head!_mem_self h) + +/-- +Removing the leading `o` from a term of `MaxProducts C` yields a list which `isGood` with respect to +`C'`. +-/ +theorem maxTail_isGood (l : MaxProducts C ho) + (h₁: ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (ord I · < o))))) : + l.val.Tail.isGood (C' C ho) := by + -- Write `l.Tail` as a linear combination of smaller products: + intro h + rw [Finsupp.mem_span_image_iff_total, ← max_eq_eval C hsC ho] at h + obtain ⟨m, ⟨hmmem, hmsum⟩⟩ := h + rw [Finsupp.total_apply] at hmsum + + -- Write the image of `l` under `Linear_CC'` as `Linear_CC'` applied to the linear combination + -- above, with leading `term I ho`'s added to each term: + have : (Linear_CC' C hsC ho) (l.val.eval C) = (Linear_CC' C hsC ho) + (Finsupp.sum m fun i a ↦ a • ((term I ho :: i.1).map (e C)).prod) + · rw [← hmsum] + simp only [map_finsupp_sum] + apply Finsupp.sum_congr + intro q hq + rw [LinearMap.map_smul] + rw [Finsupp.mem_supported] at hmmem + have hx'' : q < l.val.Tail := hmmem hq + have : ∃ (p : Products I), p.val ≠ [] ∧ p.val.head! = term I ho ∧ q = p.Tail := + ⟨⟨term I ho :: q.val, chain'_cons_of_lt C hsC ho l q hx''⟩, + ⟨List.cons_ne_nil _ _, by simp only [List.head!_cons], + by simp only [Products.Tail, List.tail_cons, Subtype.coe_eta]⟩⟩ + obtain ⟨p, hp⟩ := this + rw [hp.2.2, ← Products.max_eq_eval C hsC ho p hp.1 hp.2.1] + dsimp [Products.eval] + rw [Products.max_eq_o_cons_tail ho p hp.1 hp.2.1] + rfl + have hse := succ_exact C hC hsC ho + rw [ModuleCat.exact_iff] at hse + dsimp [ModuleCat.ofHom] at hse + + -- Rewrite `this` using exact sequence manipulations to conclude that a term is in the range of + -- the linear map `πs`: + rw [← LinearMap.sub_mem_ker_iff, ← hse] at this + obtain ⟨(n : LocallyConstant (π C (ord I · < o)) ℤ), hn⟩ := this + rw [eq_sub_iff_add_eq] at hn + have hn' := h₁ (Submodule.mem_top : n ∈ ⊤) + rw [Finsupp.mem_span_range_iff_exists_finsupp] at hn' + obtain ⟨w,hc⟩ := hn' + rw [← hc, map_finsupp_sum] at hn + apply l.prop.1 + rw [← hn] + + -- Now we just need to prove that a sum of two terms belongs to a span: + apply Submodule.add_mem + · apply Submodule.finsupp_sum_mem + intro q _ + erw [LinearMap.map_smul (fₗ := πs C o) (c := w q) (x := eval (π C (ord I · < o)) q)] + apply Submodule.smul_mem + apply Submodule.subset_span + dsimp only [eval] + rw [Products.eval_πs C (Products.prop_of_isGood _ _ q.prop)] + refine ⟨q.val, ⟨?_, rfl⟩⟩ + simp only [Products.lt_iff_lex_lt, Set.mem_setOf_eq] + exact good_lt_maxProducts C hsC ho q l + · apply Submodule.finsupp_sum_mem + intro q hq + apply Submodule.smul_mem + apply Submodule.subset_span + rw [Finsupp.mem_supported] at hmmem + rw [← Finsupp.mem_support_iff] at hq + refine ⟨⟨term I ho :: q.val, chain'_cons_of_lt C hsC ho l q (hmmem hq)⟩, ⟨?_, rfl⟩⟩ + simp only [Products.lt_iff_lex_lt, Set.mem_setOf_eq] + rw [max_eq_o_cons_tail C hsC ho l] + exact List.Lex.cons ((Products.lt_iff_lex_lt q l.val.Tail).mp (hmmem hq)) + +/-- Given `l : MaxProducts C ho`, its `Tail` is a `GoodProducts (C' C ho)`. -/ +noncomputable +def MaxToGood + (h₁: ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (ord I · < o))))) : + MaxProducts C ho → GoodProducts (C' C ho) := + fun l ↦ ⟨l.val.Tail, maxTail_isGood C hC hsC ho l h₁⟩ + +theorem maxToGood_injective + (h₁: ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (ord I · < o))))) : + (MaxToGood C hC hsC ho h₁).Injective := by + intro m n h + apply Subtype.ext ∘ Subtype.ext + rw [Subtype.ext_iff] at h + dsimp [MaxToGood] at h + rw [max_eq_o_cons_tail C hsC ho m, max_eq_o_cons_tail C hsC ho n, h] + +theorem linearIndependent_comp_of_eval + (h₁: ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (ord I · < o))))) : + LinearIndependent ℤ (eval (C' C ho)) → + LinearIndependent ℤ (ModuleCat.ofHom (Linear_CC' C hsC ho) ∘ SumEval C ho ∘ Sum.inr) := by + dsimp [SumEval, ModuleCat.ofHom] + erw [max_eq_eval_unapply C hsC ho] + intro h + let f := MaxToGood C hC hsC ho h₁ + have hf : f.Injective := maxToGood_injective C hC hsC ho h₁ + have hh : (fun l ↦ Products.eval (C' C ho) l.val.Tail) = eval (C' C ho) ∘ f := rfl + rw [hh] + exact h.comp f hf + +end GoodProducts + +end GoodProducts + +end Successor + +section Induction +/-! + +## The induction + +Here we put together the results of the sections `Zero`, `Limit` and `Successor` to prove the +predicate `P I o` holds for all ordinals `o`, and conclude with the main result: + +* `GoodProducts.linearIndependent` which says that `GoodProducts C` is linearly independent when `C` + is closed. + +We also define + +* `GoodProducts.Basis` which uses `GoodProducts.linearIndependent` and `GoodProducts.span` to + define a basis for `LocallyConstant C ℤ`  +-/ + +theorem GoodProducts.P0 : P I 0 := fun _ C _ hsC ↦ by + have : C ⊆ {(fun _ ↦ false)} := fun c hc ↦ by + ext x; exact Bool.eq_false_iff.mpr (fun ht ↦ (Ordinal.not_lt_zero (ord I x)) (hsC c hc x ht)) + rw [Set.subset_singleton_iff_eq] at this + cases this + · subst C + exact linearIndependentEmpty + · subst C + exact linearIndependentSingleton + +theorem GoodProducts.Plimit (o : Ordinal) (ho : Ordinal.IsLimit o) : + (∀ (o' : Ordinal), o' < o → P I o') → P I o := by + intro h hho C hC hsC + rw [linearIndependent_iff_union_smaller C ho hsC] + exact linearIndependent_iUnion_of_directed + (Monotone.directed_le fun _ _ h ↦ GoodProducts.smaller_mono C h) fun ⟨o', ho'⟩ ↦ + (linearIndependent_iff_smaller _ _).mp (h o' ho' (le_of_lt (lt_of_lt_of_le ho' hho)) + (π C (ord I · < o')) (isClosed_proj _ _ hC) (contained_proj _ _)) + +theorem GoodProducts.linearIndependentAux (μ : Ordinal) : P I μ := by + refine Ordinal.limitRecOn μ P0 (fun o h ho C hC hsC ↦ ?_) + (fun o ho h ↦ (GoodProducts.Plimit o ho (fun o' ho' ↦ (h o' ho')))) + have ho' : o < Ordinal.type (·<· : I → I → Prop) := + lt_of_lt_of_le (Order.lt_succ _) ho + rw [linearIndependent_iff_sum C hsC ho'] + refine ModuleCat.linearIndependent_leftExact ?_ ?_ (succ_mono C o) (succ_exact C hC hsC ho') + (square_commutes C ho') + · exact h (le_of_lt ho') (π C (ord I · < o)) (isClosed_proj C o hC) (contained_proj C o) + · exact linearIndependent_comp_of_eval C hC hsC ho' (span (π C (ord I · < o)) + (isClosed_proj C o hC)) (h (le_of_lt ho') (C' C ho') (isClosed_C' C hC ho') + (contained_C' C ho')) + +theorem GoodProducts.linearIndependent (hC : IsClosed C) : + LinearIndependent ℤ (GoodProducts.eval C) := + GoodProducts.linearIndependentAux (Ordinal.type (·<· : I → I → Prop)) (le_refl _) + C hC (fun _ _ _ _ ↦ Ordinal.typein_lt_type _ _) + +/-- `GoodProducts C` as a `ℤ`-basis for `LocallyConstant C ℤ`. -/ +noncomputable +def GoodProducts.Basis (hC : IsClosed C) : + Basis (GoodProducts C) ℤ (LocallyConstant C ℤ) := + Basis.mk (GoodProducts.linearIndependent C hC) (GoodProducts.span C hC) + +end Induction + +variable {S : Profinite} {ι : S → I → Bool} (hι : ClosedEmbedding ι) + +/-- +Given a profinite set `S` and a closed embedding `S → (I → Bool)`, the `ℤ`-module +`LocallyConstant C ℤ` is free. +-/ +theorem Nobeling_aux : Module.Free ℤ (LocallyConstant S ℤ) := Module.Free.of_equiv' + (Module.Free.of_basis <| GoodProducts.Basis _ hι.closed_range) (LocallyConstant.congrLeftₗ ℤ + (Homeomorph.ofEmbedding ι hι.toEmbedding)).symm + +end NobelingProof + +variable (S : Profinite.{u}) + +open Classical in +/-- The embedding `S → (I → Bool)` where `I` is the set of clopens of `S`. -/ +noncomputable +def Nobeling.ι : S → ({C : Set S // IsClopen C} → Bool) := fun s C => decide (s ∈ C.1) + +open Classical in +/-- The map `Nobeling.ι` is a closed embedding. -/ +theorem Nobeling.embedding : ClosedEmbedding (Nobeling.ι S) := by + apply Continuous.closedEmbedding + · dsimp [ι] + refine continuous_pi ?_ + intro C + rw [← IsLocallyConstant.iff_continuous] + refine ((IsLocallyConstant.tfae _).out 0 3).mpr ?_ + rintro ⟨⟩ + · refine IsClopen.isOpen (isClopen_compl_iff.mp ?_) + convert C.2 + ext x + simp only [Set.mem_compl_iff, Set.mem_preimage, Set.mem_singleton_iff, + decide_eq_false_iff_not, not_not] + · refine IsClopen.isOpen ?_ + convert C.2 + ext x + simp only [Set.mem_preimage, Set.mem_singleton_iff, decide_eq_true_eq] + · intro a b h + by_contra hn + obtain ⟨C, hC, hh⟩ := exists_clopen_of_totally_separated hn + apply hh.2 ∘ of_decide_eq_true + dsimp [ι] at h + rw [← congr_fun h ⟨C, hC⟩] + exact decide_eq_true hh.1 + +end Profinite + +open Profinite NobelingProof + +/-- Nöbeling's theorem: the `ℤ`-module `LocallyConstant S ℤ` is free for every `S : Profinite` -/ +instance LocallyConstant.freeOfProfinite (S : Profinite.{u}) : + Module.Free ℤ (LocallyConstant S ℤ) := + @Nobeling_aux {C : Set S // IsClopen C} ⟨⟨∅, isClopen_empty⟩⟩ + (IsWellOrder.linearOrder WellOrderingRel) WellOrderingRel.isWellOrder + S (Nobeling.ι S) (Nobeling.embedding S) diff --git a/Mathlib/Topology/Category/Stonean/Basic.lean b/Mathlib/Topology/Category/Stonean/Basic.lean index 5eabeca8333c6..773840ef269d8 100644 --- a/Mathlib/Topology/Category/Stonean/Basic.lean +++ b/Mathlib/Topology/Category/Stonean/Basic.lean @@ -240,6 +240,17 @@ instance (X : Stonean) : Projective (toProfinite.obj X) where ext exact congr_fun h.right _ +/-- Every Stonean space is projective in `Stonean`. -/ +instance (X : Stonean) : Projective X where + factors := by + intro B C φ f _ + haveI : ExtremallyDisconnected X.compHaus.toTop := X.extrDisc + have hf : Function.Surjective f := by rwa [← Stonean.epi_iff_surjective] + obtain ⟨f', h⟩ := CompactT2.ExtremallyDisconnected.projective φ.continuous f.continuous hf + use ⟨f', h.left⟩ + ext + exact congr_fun h.right _ + end Stonean namespace CompHaus diff --git a/Mathlib/Topology/Category/Stonean/EffectiveEpi.lean b/Mathlib/Topology/Category/Stonean/EffectiveEpi.lean index e3f18d04043db..ba5b104aed8b9 100644 --- a/Mathlib/Topology/Category/Stonean/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/Stonean/EffectiveEpi.lean @@ -15,6 +15,9 @@ In this file, we show that the following are all equivalent: - The induced map `∐ X ⟶ B` is epimorphic. - The family `π` is jointly surjective. +As a consequence, we show (see `effectiveEpi_iff_surjective`) that all epimorphisms in `Stonean`  +are effective, and that `Stonean` is preregular. + ## Main results - `Stonean.effectiveEpiFamily_tfae`: characterise being an effective epimorphic family. - `Stonean.instPrecoherent`: `Stonean` is precoherent. @@ -170,6 +173,11 @@ lemma effectiveEpi_iff_surjective {X Y : Stonean} (f : X ⟶ Y) : rw [← epi_iff_surjective] exact effectiveEpi_iff_epi (fun _ _ ↦ (effectiveEpiFamily_tfae _ _).out 0 1) f +instance : Preregular Stonean where + exists_fac := by + intro X Y Z f π hπ + exact ⟨X, 𝟙 X, inferInstance, Projective.factors f π⟩ + end JointlySurjective section Coherent diff --git a/Mathlib/Topology/Category/Stonean/Limits.lean b/Mathlib/Topology/Category/Stonean/Limits.lean index a476b2167ffd9..1501553ee523f 100644 --- a/Mathlib/Topology/Category/Stonean/Limits.lean +++ b/Mathlib/Topology/Category/Stonean/Limits.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Adam Topaz, Dagur Asgeirsson, Filippo A. E. Nuccio, Riccardo Brasca -/ import Mathlib.Topology.Category.Stonean.Basic +import Mathlib.Topology.Category.CompHaus.Limits /-! # Explicit (co)limits in the category of Stonean spaces @@ -74,17 +75,17 @@ lemma finiteCoproduct.hom_ext {B : Stonean.{u}} (f g : finiteCoproduct X ⟶ B) @[simps] def finiteCoproduct.cocone (F : Discrete α ⥤ Stonean) : Cocone F where - pt := finiteCoproduct F.obj - ι := Discrete.natTrans fun a => finiteCoproduct.ι F.obj a + pt := finiteCoproduct (F.obj ∘ Discrete.mk) + ι := Discrete.natTrans fun a => finiteCoproduct.ι (F.obj ∘ Discrete.mk) a.as /-- The explicit finite coproduct cocone is a colimit cocone. -/ @[simps] def finiteCoproduct.isColimit (F : Discrete α ⥤ Stonean) : IsColimit (finiteCoproduct.cocone F) where - desc := fun s => finiteCoproduct.desc _ fun a => s.ι.app a + desc := fun s => finiteCoproduct.desc _ fun a => s.ι.app ⟨a⟩ fac := fun s ⟨a⟩ => finiteCoproduct.ι_desc _ _ _ uniq := fun s m hm => finiteCoproduct.hom_ext _ _ _ fun a => by - specialize hm a + specialize hm ⟨a⟩ ext t apply_fun (fun q => q t) at hm exact hm @@ -140,6 +141,13 @@ lemma Sigma.openEmbedding_ι {α : Type} [Fintype α] (Z : α → Stonean.{u}) ( simp only [coproductIsoCoproduct, colimit.comp_coconePointUniqueUpToIso_inv, finiteCoproduct.explicitCocone_pt, finiteCoproduct.explicitCocone_ι, Discrete.natTrans_app] +instance : PreservesFiniteCoproducts Stonean.toCompHaus := by + refine ⟨fun J hJ ↦ ⟨fun {F} ↦ ?_⟩⟩ + suffices : PreservesColimit (Discrete.functor (F.obj ∘ Discrete.mk)) Stonean.toCompHaus + · exact preservesColimitOfIsoDiagram _ Discrete.natIsoFunctor.symm + apply preservesColimitOfPreservesColimitCocone (Stonean.finiteCoproduct.isColimit _) + exact CompHaus.finiteCoproduct.isColimit _ + end FiniteCoproducts end Stonean @@ -298,6 +306,40 @@ def pullbackIsoPullback : Stonean.pullback f hi ≅ end Isos +/-- The forgetful from `Stonean` to `TopCat` creates pullbacks along open embeddings -/ +noncomputable +def createsPullbacksOfOpenEmbedding : + CreatesLimit (cospan f i) (Stonean.toCompHaus ⋙ compHausToTop) := +createsLimitOfFullyFaithfulOfIso (Stonean.pullback f hi) (by + refine (@TopCat.isoOfHomeo (TopCat.of _) (TopCat.of _) + (TopCat.pullbackHomeoPreimage f f.2 i hi.toEmbedding)).symm ≪≫ ?_ + refine ?_ ≪≫ Limits.lim.mapIso (diagramIsoCospan _).symm + exact (TopCat.pullbackConeIsLimit f i).conePointUniqueUpToIso (limit.isLimit _)) + +instance : HasPullbacksOfInclusions Stonean where + hasPullbackInl f := by + apply (config := { allowSynthFailures := true }) hasPullback_symmetry + apply Stonean.HasPullbackOpenEmbedding + apply Stonean.Sigma.openEmbedding_ι + +noncomputable +instance : PreservesPullbacksOfInclusions Stonean.toCompHaus.{u} where + preservesPullbackInl := by + intros X Y Z f + apply (config := { allowSynthFailures := true }) preservesPullbackSymmetry + have : OpenEmbedding (coprod.inl : X ⟶ X ⨿ Y) := Stonean.Sigma.openEmbedding_ι _ _ + have := Stonean.createsPullbacksOfOpenEmbedding f this + refine @preservesLimitOfReflectsOfPreserves _ _ _ _ _ _ _ _ _ Stonean.toCompHaus + compHausToTop inferInstance ?_ + apply (config := { allowSynthFailures := true }) ReflectsLimitsOfShape.reflectsLimit + apply (config := { allowSynthFailures := true }) ReflectsLimitsOfSize.reflectsLimitsOfShape + exact reflectsLimitsOfSizeShrink _ + +instance : FinitaryExtensive Stonean.{u} := + have := fullyFaithfulReflectsLimits Stonean.toCompHaus + have := fullyFaithfulReflectsColimits Stonean.toCompHaus + finitaryExtensive_of_preserves_and_reflects Stonean.toCompHaus + end Stonean end Pullback diff --git a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean index 0d21234da04dd..de8a5354a414c 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean @@ -175,6 +175,31 @@ theorem range_pullback_to_prod {X Y Z : TopCat} (f : X ⟶ Z) (g : Y ⟶ Z) : aesop_cat #align Top.range_pullback_to_prod TopCat.range_pullback_to_prod +/-- The pullback along an embedding is (isomorphic to) the preimage. -/ +noncomputable +def pullbackHomeoPreimage + {X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] + (f : X → Z) (hf : Continuous f) (g : Y → Z) (hg : Embedding g) : + { p : X × Y // f p.1 = g p.2 } ≃ₜ f ⁻¹' Set.range g where + toFun := fun x ↦ ⟨x.1.1, _, x.2.symm⟩ + invFun := fun x ↦ ⟨⟨x.1, Exists.choose x.2⟩, (Exists.choose_spec x.2).symm⟩ + left_inv := by + intro x + ext <;> dsimp + apply hg.inj + convert x.prop + exact Exists.choose_spec (p := fun y ↦ g y = f (↑x : X × Y).1) _ + right_inv := fun x ↦ rfl + continuous_toFun := by + apply Continuous.subtype_mk + exact continuous_fst.comp continuous_subtype_val + continuous_invFun := by + apply Continuous.subtype_mk + refine continuous_prod_mk.mpr ⟨continuous_subtype_val, hg.toInducing.continuous_iff.mpr ?_⟩ + convert hf.comp continuous_subtype_val + ext x + exact Exists.choose_spec x.2 + theorem inducing_pullback_to_prod {X Y Z : TopCat.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) : Inducing <| ⇑(prod.lift pullback.fst pullback.snd : pullback f g ⟶ X ⨯ Y) := ⟨by simp [prod_topology, pullback_topology, induced_compose, ← coe_comp]⟩ diff --git a/Mathlib/Topology/Clopen.lean b/Mathlib/Topology/Clopen.lean index 1cc0da0544c9f..a0fc25ade8dd3 100644 --- a/Mathlib/Topology/Clopen.lean +++ b/Mathlib/Topology/Clopen.lean @@ -16,15 +16,15 @@ open Set Filter Topology TopologicalSpace Classical universe u v -variable {α : Type u} {β : Type v} {ι : Type*} {π : ι → Type*} +variable {X : Type u} {Y : Type v} {ι : Type*} -variable [TopologicalSpace α] [TopologicalSpace β] {s t : Set α} +variable [TopologicalSpace X] [TopologicalSpace Y] {s t : Set X} section Clopen -- porting note: todo: redefine as `IsClosed s ∧ IsOpen s` /-- A set is clopen if it is both open and closed. -/ -def IsClopen (s : Set α) : Prop := +def IsClopen (s : Set X) : Prop := IsOpen s ∧ IsClosed s #align is_clopen IsClopen @@ -34,7 +34,7 @@ protected theorem IsClopen.isOpen (hs : IsClopen s) : IsOpen s := hs.1 protected theorem IsClopen.isClosed (hs : IsClopen s) : IsClosed s := hs.2 #align is_clopen.is_closed IsClopen.isClosed -theorem isClopen_iff_frontier_eq_empty {s : Set α} : IsClopen s ↔ frontier s = ∅ := by +theorem isClopen_iff_frontier_eq_empty : IsClopen s ↔ frontier s = ∅ := by rw [IsClopen, ← closure_eq_iff_isClosed, ← interior_eq_iff_isOpen, frontier, diff_eq_empty] refine' ⟨fun h => (h.2.trans h.1.symm).subset, fun h => _⟩ exact ⟨interior_subset.antisymm (subset_closure.trans h), @@ -44,115 +44,112 @@ theorem isClopen_iff_frontier_eq_empty {s : Set α} : IsClopen s ↔ frontier s alias ⟨IsClopen.frontier_eq, _⟩ := isClopen_iff_frontier_eq_empty #align is_clopen.frontier_eq IsClopen.frontier_eq -theorem IsClopen.union {s t : Set α} (hs : IsClopen s) (ht : IsClopen t) : IsClopen (s ∪ t) := +theorem IsClopen.union (hs : IsClopen s) (ht : IsClopen t) : IsClopen (s ∪ t) := ⟨hs.1.union ht.1, hs.2.union ht.2⟩ #align is_clopen.union IsClopen.union -theorem IsClopen.inter {s t : Set α} (hs : IsClopen s) (ht : IsClopen t) : IsClopen (s ∩ t) := +theorem IsClopen.inter (hs : IsClopen s) (ht : IsClopen t) : IsClopen (s ∩ t) := ⟨hs.1.inter ht.1, hs.2.inter ht.2⟩ #align is_clopen.inter IsClopen.inter -@[simp] theorem isClopen_empty : IsClopen (∅ : Set α) := ⟨isOpen_empty, isClosed_empty⟩ +@[simp] theorem isClopen_empty : IsClopen (∅ : Set X) := ⟨isOpen_empty, isClosed_empty⟩ #align is_clopen_empty isClopen_empty -@[simp] theorem isClopen_univ : IsClopen (univ : Set α) := ⟨isOpen_univ, isClosed_univ⟩ +@[simp] theorem isClopen_univ : IsClopen (univ : Set X) := ⟨isOpen_univ, isClosed_univ⟩ #align is_clopen_univ isClopen_univ -theorem IsClopen.compl {s : Set α} (hs : IsClopen s) : IsClopen sᶜ := +theorem IsClopen.compl (hs : IsClopen s) : IsClopen sᶜ := ⟨hs.2.isOpen_compl, hs.1.isClosed_compl⟩ #align is_clopen.compl IsClopen.compl @[simp] -theorem isClopen_compl_iff {s : Set α} : IsClopen sᶜ ↔ IsClopen s := +theorem isClopen_compl_iff : IsClopen sᶜ ↔ IsClopen s := ⟨fun h => compl_compl s ▸ IsClopen.compl h, IsClopen.compl⟩ #align is_clopen_compl_iff isClopen_compl_iff -theorem IsClopen.diff {s t : Set α} (hs : IsClopen s) (ht : IsClopen t) : IsClopen (s \ t) := +theorem IsClopen.diff (hs : IsClopen s) (ht : IsClopen t) : IsClopen (s \ t) := hs.inter ht.compl #align is_clopen.diff IsClopen.diff -theorem IsClopen.prod {s : Set α} {t : Set β} (hs : IsClopen s) (ht : IsClopen t) : - IsClopen (s ×ˢ t) := +theorem IsClopen.prod {t : Set Y} (hs : IsClopen s) (ht : IsClopen t) : IsClopen (s ×ˢ t) := ⟨hs.1.prod ht.1, hs.2.prod ht.2⟩ #align is_clopen.prod IsClopen.prod -theorem isClopen_iUnion_of_finite {β : Type*} [Finite β] {s : β → Set α} (h : ∀ i, IsClopen (s i)) : +theorem isClopen_iUnion_of_finite [Finite Y] {s : Y → Set X} (h : ∀ i, IsClopen (s i)) : IsClopen (⋃ i, s i) := ⟨isOpen_iUnion (forall_and.1 h).1, isClosed_iUnion_of_finite (forall_and.1 h).2⟩ #align is_clopen_Union isClopen_iUnion_of_finite -theorem Set.Finite.isClopen_biUnion {β : Type*} {s : Set β} {f : β → Set α} (hs : s.Finite) +theorem Set.Finite.isClopen_biUnion {s : Set Y} {f : Y → Set X} (hs : s.Finite) (h : ∀ i ∈ s, IsClopen <| f i) : IsClopen (⋃ i ∈ s, f i) := ⟨isOpen_biUnion fun i hi => (h i hi).1, hs.isClosed_biUnion fun i hi => (h i hi).2⟩ #align is_clopen_bUnion Set.Finite.isClopen_biUnion -theorem isClopen_biUnion_finset {β : Type*} {s : Finset β} {f : β → Set α} +theorem isClopen_biUnion_finset {s : Finset Y} {f : Y → Set X} (h : ∀ i ∈ s, IsClopen <| f i) : IsClopen (⋃ i ∈ s, f i) := - s.finite_toSet.isClopen_biUnion h + s.finite_toSet.isClopen_biUnion h #align is_clopen_bUnion_finset isClopen_biUnion_finset -theorem isClopen_iInter_of_finite {β : Type*} [Finite β] {s : β → Set α} (h : ∀ i, IsClopen (s i)) : +theorem isClopen_iInter_of_finite [Finite Y] {s : Y → Set X} (h : ∀ i, IsClopen (s i)) : IsClopen (⋂ i, s i) := ⟨isOpen_iInter_of_finite (forall_and.1 h).1, isClosed_iInter (forall_and.1 h).2⟩ #align is_clopen_Inter isClopen_iInter_of_finite -theorem Set.Finite.isClopen_biInter {β : Type*} {s : Set β} (hs : s.Finite) {f : β → Set α} +theorem Set.Finite.isClopen_biInter {s : Set Y} (hs : s.Finite) {f : Y → Set X} (h : ∀ i ∈ s, IsClopen (f i)) : IsClopen (⋂ i ∈ s, f i) := ⟨hs.isOpen_biInter fun i hi => (h i hi).1, isClosed_biInter fun i hi => (h i hi).2⟩ #align is_clopen_bInter Set.Finite.isClopen_biInter -theorem isClopen_biInter_finset {β : Type*} {s : Finset β} {f : β → Set α} +theorem isClopen_biInter_finset {s : Finset Y} {f : Y → Set X} (h : ∀ i ∈ s, IsClopen (f i)) : IsClopen (⋂ i ∈ s, f i) := s.finite_toSet.isClopen_biInter h #align is_clopen_bInter_finset isClopen_biInter_finset -theorem IsClopen.preimage {s : Set β} (h : IsClopen s) {f : α → β} (hf : Continuous f) : +theorem IsClopen.preimage {s : Set Y} (h : IsClopen s) {f : X → Y} (hf : Continuous f) : IsClopen (f ⁻¹' s) := ⟨h.1.preimage hf, h.2.preimage hf⟩ #align is_clopen.preimage IsClopen.preimage -theorem ContinuousOn.preimage_clopen_of_clopen {f : α → β} {s : Set α} {t : Set β} +theorem ContinuousOn.preimage_clopen_of_clopen {f : X → Y} {s : Set X} {t : Set Y} (hf : ContinuousOn f s) (hs : IsClopen s) (ht : IsClopen t) : IsClopen (s ∩ f ⁻¹' t) := ⟨ContinuousOn.preimage_open_of_open hf hs.1 ht.1, ContinuousOn.preimage_closed_of_closed hf hs.2 ht.2⟩ #align continuous_on.preimage_clopen_of_clopen ContinuousOn.preimage_clopen_of_clopen /-- The intersection of a disjoint covering by two open sets of a clopen set will be clopen. -/ -theorem isClopen_inter_of_disjoint_cover_clopen {Z a b : Set α} (h : IsClopen Z) (cover : Z ⊆ a ∪ b) - (ha : IsOpen a) (hb : IsOpen b) (hab : Disjoint a b) : IsClopen (Z ∩ a) := by +theorem isClopen_inter_of_disjoint_cover_clopen {s a b : Set X} (h : IsClopen s) (cover : s ⊆ a ∪ b) + (ha : IsOpen a) (hb : IsOpen b) (hab : Disjoint a b) : IsClopen (s ∩ a) := by refine' ⟨IsOpen.inter h.1 ha, _⟩ - have : IsClosed (Z ∩ bᶜ) := IsClosed.inter h.2 (isClosed_compl_iff.2 hb) + have : IsClosed (s ∩ bᶜ) := IsClosed.inter h.2 (isClosed_compl_iff.2 hb) convert this using 1 - refine' (inter_subset_inter_right Z hab.subset_compl_right).antisymm _ + refine' (inter_subset_inter_right s hab.subset_compl_right).antisymm _ rintro x ⟨hx₁, hx₂⟩ exact ⟨hx₁, by simpa [not_mem_of_mem_compl hx₂] using cover hx₁⟩ #align is_clopen_inter_of_disjoint_cover_clopen isClopen_inter_of_disjoint_cover_clopen @[simp] -theorem isClopen_discrete [DiscreteTopology α] (x : Set α) : IsClopen x := +theorem isClopen_discrete [DiscreteTopology X] (s : Set X) : IsClopen s := ⟨isOpen_discrete _, isClosed_discrete _⟩ #align is_clopen_discrete isClopen_discrete -- porting note: new lemma -theorem isClopen_range_inl : IsClopen (range (Sum.inl : α → α ⊕ β)) := +theorem isClopen_range_inl : IsClopen (range (Sum.inl : X → X ⊕ Y)) := ⟨isOpen_range_inl, isClosed_range_inl⟩ -- porting note: new lemma -theorem isClopen_range_inr : IsClopen (range (Sum.inr : β → α ⊕ β)) := +theorem isClopen_range_inr : IsClopen (range (Sum.inr : Y → X ⊕ Y)) := ⟨isOpen_range_inr, isClosed_range_inr⟩ -theorem isClopen_range_sigmaMk {ι : Type*} {σ : ι → Type*} [∀ i, TopologicalSpace (σ i)] {i : ι} : - IsClopen (Set.range (@Sigma.mk ι σ i)) := +theorem isClopen_range_sigmaMk {X : ι → Type*} [∀ i, TopologicalSpace (X i)] {i : ι} : + IsClopen (Set.range (@Sigma.mk ι X i)) := ⟨openEmbedding_sigmaMk.open_range, closedEmbedding_sigmaMk.closed_range⟩ #align clopen_range_sigma_mk isClopen_range_sigmaMk -protected theorem QuotientMap.isClopen_preimage {f : α → β} (hf : QuotientMap f) {s : Set β} : +protected theorem QuotientMap.isClopen_preimage {f : X → Y} (hf : QuotientMap f) {s : Set Y} : IsClopen (f ⁻¹' s) ↔ IsClopen s := and_congr hf.isOpen_preimage hf.isClosed_preimage #align quotient_map.is_clopen_preimage QuotientMap.isClopen_preimage -variable {X : Type*} [TopologicalSpace X] - theorem continuous_boolIndicator_iff_clopen (U : Set X) : Continuous U.boolIndicator ↔ IsClopen U := by constructor diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 201af69174a83..c479f7091e622 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -144,6 +144,29 @@ theorem isCompact_iff_ultrafilter_le_nhds : alias ⟨IsCompact.ultrafilter_le_nhds, _⟩ := isCompact_iff_ultrafilter_le_nhds #align is_compact.ultrafilter_le_nhds IsCompact.ultrafilter_le_nhds +theorem isCompact_iff_ultrafilter_le_nhds' : + IsCompact s ↔ ∀ f : Ultrafilter α, s ∈ f → ∃ a ∈ s, ↑f ≤ 𝓝 a := by + simp only [isCompact_iff_ultrafilter_le_nhds, le_principal_iff, Ultrafilter.mem_coe] + +alias ⟨IsCompact.ultrafilter_le_nhds', _⟩ := isCompact_iff_ultrafilter_le_nhds' + +/-- If a compact set belongs to a filter and this filter has a unique cluster point `y` in this set, +then the filter is less than or equal to `𝓝 y`. -/ +lemma IsCompact.le_nhds_of_unique_clusterPt (hs : IsCompact s) {l : Filter α} {y : α} + (hmem : s ∈ l) (h : ∀ x ∈ s, ClusterPt x l → x = y) : l ≤ 𝓝 y := by + refine le_iff_ultrafilter.2 fun f hf ↦ ?_ + rcases hs.ultrafilter_le_nhds' f (hf hmem) with ⟨x, hxs, hx⟩ + convert ← hx + exact h x hxs (.mono (.of_le_nhds hx) hf) + +/-- If values of `f : β → α` belong to a compact set `s` eventually along a filter `l` +and `y` is a unique `MapClusterPt` for `f` along `l` in `s`, +then `f` tends to `𝓝 y` along `l`. -/ +lemma IsCompact.tendsto_nhds_of_unique_mapClusterPt {l : Filter β} {y : α} {f : β → α} + (hs : IsCompact s) (hmem : ∀ᶠ x in l, f x ∈ s) (h : ∀ x ∈ s, MapClusterPt x l f → x = y) : + Tendsto f l (𝓝 y) := + hs.le_nhds_of_unique_clusterPt (mem_map.2 hmem) h + /-- For every open directed cover of a compact set, there exists a single element of the cover which itself includes the set. -/ theorem IsCompact.elim_directed_cover {ι : Type v} [hι : Nonempty ι] (hs : IsCompact s) @@ -249,8 +272,8 @@ theorem LocallyFinite.finite_nonempty_inter_compact {ι : Type*} {f : ι → Set theorem IsCompact.inter_iInter_nonempty {s : Set α} {ι : Type v} (hs : IsCompact s) (Z : ι → Set α) (hZc : ∀ i, IsClosed (Z i)) (hsZ : ∀ t : Finset ι, (s ∩ ⋂ i ∈ t, Z i).Nonempty) : (s ∩ ⋂ i, Z i).Nonempty := by - simp only [nonempty_iff_ne_empty] at hsZ ⊢ - apply mt (hs.elim_finite_subfamily_closed Z hZc); push_neg; exact hsZ + contrapose! hsZ + exact hs.elim_finite_subfamily_closed Z hZc hsZ #align is_compact.inter_Inter_nonempty IsCompact.inter_iInter_nonempty /-- Cantor's intersection theorem: @@ -401,8 +424,8 @@ theorem Set.Subsingleton.isCompact {s : Set α} (hs : s.Subsingleton) : IsCompac -- porting note: golfed a proof instead of fixing it theorem Set.Finite.isCompact_biUnion {s : Set ι} {f : ι → Set α} (hs : s.Finite) (hf : ∀ i ∈ s, IsCompact (f i)) : IsCompact (⋃ i ∈ s, f i) := - isCompact_iff_ultrafilter_le_nhds.2 <| fun l hl => by - rw [le_principal_iff, Ultrafilter.mem_coe, Ultrafilter.finite_biUnion_mem_iff hs] at hl + isCompact_iff_ultrafilter_le_nhds'.2 <| fun l hl => by + rw [Ultrafilter.finite_biUnion_mem_iff hs] at hl rcases hl with ⟨i, his, hi⟩ rcases (hf i his).ultrafilter_le_nhds _ (le_principal_iff.2 hi) with ⟨x, hxi, hlx⟩ exact ⟨x, mem_iUnion₂.2 ⟨i, his, hxi⟩, hlx⟩ @@ -690,6 +713,12 @@ theorem cluster_point_of_compact [CompactSpace α] (f : Filter α) [NeBot f] : by simpa using isCompact_univ (show f ≤ 𝓟 univ by simp) #align cluster_point_of_compact cluster_point_of_compact +nonrec theorem Ultrafilter.le_nhds_lim [CompactSpace α] (F : Ultrafilter α) : ↑F ≤ 𝓝 F.lim := by + rcases isCompact_univ.ultrafilter_le_nhds F (by simp) with ⟨x, -, h⟩ + exact le_nhds_lim ⟨x, h⟩ +set_option linter.uppercaseLean3 false in +#align ultrafilter.le_nhds_Lim Ultrafilter.le_nhds_lim + theorem CompactSpace.elim_nhds_subcover [CompactSpace α] (U : α → Set α) (hU : ∀ x, U x ∈ 𝓝 x) : ∃ t : Finset α, ⋃ x ∈ t, U x = ⊤ := by obtain ⟨t, -, s⟩ := IsCompact.elim_nhds_subcover isCompact_univ U fun x _ => hU x @@ -708,6 +737,20 @@ theorem IsClosed.isCompact [CompactSpace α] {s : Set α} (h : IsClosed s) : IsC isCompact_univ.of_isClosed_subset h (subset_univ _) #align is_closed.is_compact IsClosed.isCompact +/-- If a filter has a unique cluster point `y` in a compact topological space, +then the filter is less than or equal to `𝓝 y`. -/ +lemma le_nhds_of_unique_clusterPt [CompactSpace α] {l : Filter α} {y : α} + (h : ∀ x, ClusterPt x l → x = y) : l ≤ 𝓝 y := + isCompact_univ.le_nhds_of_unique_clusterPt univ_mem fun x _ ↦ h x + +/-- If `y` is a unique `MapClusterPt` for `f` along `l` +and the codomain of `f` is a compact space, +then `f` tends to `𝓝 y` along `l`. -/ +lemma tendsto_nhds_of_unique_mapClusterPt [CompactSpace α] {l : Filter β} {y : α} {f : β → α} + (h : ∀ x, MapClusterPt x l f → x = y) : + Tendsto f l (𝓝 y) := + le_nhds_of_unique_clusterPt h + /-- `α` is a noncompact topological space if it is not a compact space. -/ class NoncompactSpace (α : Type*) [TopologicalSpace α] : Prop where /-- In a noncompact space, `Set.univ` is not a compact set. -/ @@ -833,6 +876,11 @@ theorem isClosedMap_snd_of_compactSpace {X : Type*} [TopologicalSpace X] [Compac exact hs ⟨hU trivial, hzV⟩ hzs #align is_closed_proj_of_is_compact isClosedMap_snd_of_compactSpace +/-- If `Y` is a compact topological space, then `Prod.fst : X × Y → X` is a closed map. -/ +theorem isClosedMap_fst_of_compactSpace {X Y : Type*} [TopologicalSpace X] + [TopologicalSpace Y] [CompactSpace Y] : IsClosedMap (Prod.fst : X × Y → X) := + isClosedMap_snd_of_compactSpace.comp isClosedMap_swap + theorem exists_subset_nhds_of_compactSpace [CompactSpace α] {ι : Type*} [Nonempty ι] {V : ι → Set α} (hV : Directed (· ⊇ ·) V) (hV_closed : ∀ i, IsClosed (V i)) {U : Set α} (hU : ∀ x ∈ ⋂ i, V i, U ∈ 𝓝 x) : ∃ i, V i ⊆ U := @@ -912,13 +960,12 @@ protected theorem ClosedEmbedding.compactSpace [h : CompactSpace β] {f : α → theorem IsCompact.prod {s : Set α} {t : Set β} (hs : IsCompact s) (ht : IsCompact t) : IsCompact (s ×ˢ t) := by - rw [isCompact_iff_ultrafilter_le_nhds] at hs ht ⊢ + rw [isCompact_iff_ultrafilter_le_nhds'] at hs ht ⊢ intro f hfs - rw [le_principal_iff] at hfs obtain ⟨a : α, sa : a ∈ s, ha : map Prod.fst f.1 ≤ 𝓝 a⟩ := - hs (f.map Prod.fst) (le_principal_iff.2 <| mem_map.2 <| mem_of_superset hfs fun x => And.left) + hs (f.map Prod.fst) (mem_map.2 <| mem_of_superset hfs fun x => And.left) obtain ⟨b : β, tb : b ∈ t, hb : map Prod.snd f.1 ≤ 𝓝 b⟩ := - ht (f.map Prod.snd) (le_principal_iff.2 <| mem_map.2 <| mem_of_superset hfs fun x => And.right) + ht (f.map Prod.snd) (mem_map.2 <| mem_of_superset hfs fun x => And.right) rw [map_le_iff_le_comap] at ha hb refine' ⟨⟨a, b⟩, ⟨sa, tb⟩, _⟩ rw [nhds_prod_eq]; exact le_inf ha hb diff --git a/Mathlib/Topology/Compactness/LocallyCompact.lean b/Mathlib/Topology/Compactness/LocallyCompact.lean index d3ab6876a6068..fdb99c4e8e14d 100644 --- a/Mathlib/Topology/Compactness/LocallyCompact.lean +++ b/Mathlib/Topology/Compactness/LocallyCompact.lean @@ -164,7 +164,9 @@ instance (priority := 100) [LocallyCompactSpace X] : WeaklyLocallyCompactSpace X /-- In a locally compact space, for every containment `K ⊆ U` of a compact set `K` in an open set `U`, there is a compact neighborhood `L` such that `K ⊆ L ⊆ U`: equivalently, there is a - compact `L` such that `K ⊆ interior L` and `L ⊆ U`. -/ + compact `L` such that `K ⊆ interior L` and `L ⊆ U`. + See also `exists_compact_closed_between`, in which one guarantees additionally that `L` is closed + if the space is regular. -/ theorem exists_compact_between [hX : LocallyCompactSpace X] {K U : Set X} (hK : IsCompact K) (hU : IsOpen U) (h_KU : K ⊆ U) : ∃ L, IsCompact L ∧ K ⊆ interior L ∧ L ⊆ U := by choose V hVc hxV hKV using fun x : K => exists_compact_subset hU (h_KU x.2) @@ -207,9 +209,3 @@ protected theorem IsOpen.locallyCompactSpace [LocallyCompactSpace X] {s : Set X} LocallyCompactSpace s := hs.openEmbedding_subtype_val.locallyCompactSpace #align is_open.locally_compact_space IsOpen.locallyCompactSpace - -nonrec theorem Ultrafilter.le_nhds_lim [CompactSpace X] (F : Ultrafilter X) : ↑F ≤ 𝓝 F.lim := by - rcases isCompact_univ.ultrafilter_le_nhds F (by simp) with ⟨x, -, h⟩ - exact le_nhds_lim ⟨x, h⟩ -set_option linter.uppercaseLean3 false in -#align ultrafilter.le_nhds_Lim Ultrafilter.le_nhds_lim diff --git a/Mathlib/Topology/Compactness/SigmaCompact.lean b/Mathlib/Topology/Compactness/SigmaCompact.lean index 3e357ff279bb2..c6be308e14da6 100644 --- a/Mathlib/Topology/Compactness/SigmaCompact.lean +++ b/Mathlib/Topology/Compactness/SigmaCompact.lean @@ -17,24 +17,24 @@ open Set Filter Topology TopologicalSpace Classical universe u v -variable {α : Type u} {β : Type v} {ι : Type*} {π : ι → Type*} -variable [TopologicalSpace α] [TopologicalSpace β] {s t : Set α} +variable {X : Type*} {Y : Type*} {ι : Type*} +variable [TopologicalSpace X] [TopologicalSpace Y] {s t : Set X} -/-- A subset `s ⊆ α` is called **σ-compact** if it is the union of countably many compact sets. -/ -def IsSigmaCompact (s : Set α) : Prop := - ∃ K : ℕ → Set α, (∀ n, IsCompact (K n)) ∧ ⋃ n, K n = s +/-- A subset `s ⊆ X` is called **σ-compact** if it is the union of countably many compact sets. -/ +def IsSigmaCompact (s : Set X) : Prop := + ∃ K : ℕ → Set X, (∀ n, IsCompact (K n)) ∧ ⋃ n, K n = s /-- Compact sets are σ-compact. -/ -lemma IsCompact.isSigmaCompact {s : Set α} (hs : IsCompact s) : IsSigmaCompact s := +lemma IsCompact.isSigmaCompact {s : Set X} (hs : IsCompact s) : IsSigmaCompact s := ⟨fun _ => s, fun _ => hs, iUnion_const _⟩ /-- The empty set is σ-compact. -/ @[simp] -lemma isSigmaCompact_empty : IsSigmaCompact (∅ : Set α) := +lemma isSigmaCompact_empty : IsSigmaCompact (∅ : Set X) := IsCompact.isSigmaCompact isCompact_empty /-- Countable unions of compact sets are σ-compact. -/ -lemma isSigmaCompact_iUnion_of_isCompact {ι : Type*} [hι : Countable ι] (s : ι → Set α) +lemma isSigmaCompact_iUnion_of_isCompact [hι : Countable ι] (s : ι → Set X) (hcomp : ∀ i, IsCompact (s i)) : IsSigmaCompact (⋃ i, s i) := by rcases isEmpty_or_nonempty ι · simp only [iUnion_of_empty, isSigmaCompact_empty] @@ -43,14 +43,14 @@ lemma isSigmaCompact_iUnion_of_isCompact {ι : Type*} [hι : Countable ι] (s : exact ⟨s ∘ f, fun n ↦ hcomp (f n), Function.Surjective.iUnion_comp hf _⟩ /-- Countable unions of compact sets are σ-compact. -/ -lemma isSigmaCompact_sUnion_of_isCompact {S : Set (Set α)} (hc : Set.Countable S) - (hcomp : ∀ (s : Set α), s ∈ S → IsCompact s) : IsSigmaCompact (⋃₀ S) := by +lemma isSigmaCompact_sUnion_of_isCompact {S : Set (Set X)} (hc : Set.Countable S) + (hcomp : ∀ (s : Set X), s ∈ S → IsCompact s) : IsSigmaCompact (⋃₀ S) := by have : Countable S := countable_coe_iff.mpr hc rw [sUnion_eq_iUnion] apply isSigmaCompact_iUnion_of_isCompact _ (fun ⟨s, hs⟩ ↦ hcomp s hs) /-- Countable unions of σ-compact sets are σ-compact. -/ -lemma isSigmaCompact_iUnion {ι : Type*} [Countable ι] (s : ι → Set α) +lemma isSigmaCompact_iUnion [Countable ι] (s : ι → Set X) (hcomp : ∀ i, IsSigmaCompact (s i)) : IsSigmaCompact (⋃ i, s i) := by -- Choose a decomposition s_i = ⋃ K_i,j for each i. choose K hcomp hcov using fun i ↦ hcomp i @@ -64,13 +64,13 @@ lemma isSigmaCompact_iUnion {ι : Type*} [Countable ι] (s : ι → Set α) exact isSigmaCompact_iUnion_of_isCompact K.uncurry fun x ↦ (hcomp x.1 x.2) /-- Countable unions of σ-compact sets are σ-compact. -/ -lemma isSigmaCompact_sUnion (S : Set (Set α)) (hc : Set.Countable S) - (hcomp : ∀ s : S, IsSigmaCompact s (α := α)) : IsSigmaCompact (⋃₀ S) := by +lemma isSigmaCompact_sUnion (S : Set (Set X)) (hc : Set.Countable S) + (hcomp : ∀ s : S, IsSigmaCompact s (X := X)) : IsSigmaCompact (⋃₀ S) := by have : Countable S := countable_coe_iff.mpr hc apply sUnion_eq_iUnion.symm ▸ isSigmaCompact_iUnion _ hcomp /-- Countable unions of σ-compact sets are σ-compact. -/ -lemma isSigmaCompact_biUnion {ι : Type*} {s : Set ι} {S : ι → Set α} (hc : Set.Countable s) +lemma isSigmaCompact_biUnion {s : Set ι} {S : ι → Set X} (hc : Set.Countable s) (hcomp : ∀ (i : ι), i ∈ s → IsSigmaCompact (S i)) : IsSigmaCompact (⋃ (i : ι) (_ : i ∈ s), S i) := by have : Countable ↑s := countable_coe_iff.mpr hc @@ -78,7 +78,7 @@ lemma isSigmaCompact_biUnion {ι : Type*} {s : Set ι} {S : ι → Set α} (hc : exact isSigmaCompact_iUnion _ (fun ⟨i', hi'⟩ ↦ hcomp i' hi') /-- A closed subset of a σ-compact set is σ-compact. -/ -lemma IsSigmaCompact.of_isClosed_subset {s t : Set α} (ht : IsSigmaCompact t) +lemma IsSigmaCompact.of_isClosed_subset {s t : Set X} (ht : IsSigmaCompact t) (hs : IsClosed s) (h : s ⊆ t) : IsSigmaCompact s := by rcases ht with ⟨K, hcompact, hcov⟩ refine ⟨(fun n ↦ s ∩ (K n)), fun n ↦ (hcompact n).inter_left hs, ?_⟩ @@ -86,19 +86,19 @@ lemma IsSigmaCompact.of_isClosed_subset {s t : Set α} (ht : IsSigmaCompact t) exact inter_eq_left.mpr h /-- If `s` is σ-compact and `f` is continuous on `s`, `f(s)` is σ-compact.-/ -lemma IsSigmaCompact.image_of_continuousOn {f : α → β} {s : Set α} (hs : IsSigmaCompact s) +lemma IsSigmaCompact.image_of_continuousOn {f : X → Y} {s : Set X} (hs : IsSigmaCompact s) (hf : ContinuousOn f s) : IsSigmaCompact (f '' s) := by rcases hs with ⟨K, hcompact, hcov⟩ refine ⟨fun n ↦ f '' K n, ?_, hcov.symm ▸ image_iUnion.symm⟩ exact fun n ↦ (hcompact n).image_of_continuousOn (hf.mono (hcov.symm ▸ subset_iUnion K n)) /-- If `s` is σ-compact and `f` continuous, `f(s)` is σ-compact. -/ -lemma IsSigmaCompact.image {f : α → β} (hf : Continuous f) {s : Set α} (hs : IsSigmaCompact s) : +lemma IsSigmaCompact.image {f : X → Y} (hf : Continuous f) {s : Set X} (hs : IsSigmaCompact s) : IsSigmaCompact (f '' s) := hs.image_of_continuousOn hf.continuousOn /-- If `f : X → Y` is `Inducing`, the image `f '' s` of a set `s` is σ-compact if and only `s` is σ-compact. -/ -lemma Inducing.isSigmaCompact_iff {f : α → β} {s : Set α} +lemma Inducing.isSigmaCompact_iff {f : X → Y} {s : Set X} (hf : Inducing f) : IsSigmaCompact s ↔ IsSigmaCompact (f '' s) := by constructor · exact fun h ↦ h.image hf.continuous @@ -119,132 +119,132 @@ lemma Inducing.isSigmaCompact_iff {f : α → β} {s : Set α} /-- If `f : X → Y` is an `Embedding`, the image `f '' s` of a set `s` is σ-compact if and only `s` is σ-compact. -/ -lemma Embedding.isSigmaCompact_iff {f : α → β} {s : Set α} +lemma Embedding.isSigmaCompact_iff {f : X → Y} {s : Set X} (hf : Embedding f) : IsSigmaCompact s ↔ IsSigmaCompact (f '' s) := hf.toInducing.isSigmaCompact_iff /-- Sets of subtype are σ-compact iff the image under a coercion is. -/ -lemma Subtype.isSigmaCompact_iff {p : α → Prop} {s : Set { a // p a }} : - IsSigmaCompact s ↔ IsSigmaCompact ((↑) '' s : Set α) := +lemma Subtype.isSigmaCompact_iff {p : X → Prop} {s : Set { a // p a }} : + IsSigmaCompact s ↔ IsSigmaCompact ((↑) '' s : Set X) := embedding_subtype_val.isSigmaCompact_iff /-- A σ-compact space is a space that is the union of a countable collection of compact subspaces. Note that a locally compact separable T₂ space need not be σ-compact. The sequence can be extracted using `compactCovering`. -/ -class SigmaCompactSpace (α : Type*) [TopologicalSpace α] : Prop where +class SigmaCompactSpace (X : Type*) [TopologicalSpace X] : Prop where /-- In a σ-compact space, `Set.univ` is a σ-compact set. -/ - isSigmaCompact_univ : IsSigmaCompact (univ : Set α) + isSigmaCompact_univ : IsSigmaCompact (univ : Set X) #align sigma_compact_space SigmaCompactSpace /-- A topological space is σ-compact iff `univ` is σ-compact. -/ -lemma isSigmaCompact_univ_iff : IsSigmaCompact (univ : Set α) ↔ SigmaCompactSpace α := +lemma isSigmaCompact_univ_iff : IsSigmaCompact (univ : Set X) ↔ SigmaCompactSpace X := ⟨fun h => ⟨h⟩, fun h => h.1⟩ /-- In a σ-compact space, `univ` is σ-compact. -/ -lemma isSigmaCompact_univ [h : SigmaCompactSpace α] : IsSigmaCompact (univ : Set α) := +lemma isSigmaCompact_univ [h : SigmaCompactSpace X] : IsSigmaCompact (univ : Set X) := isSigmaCompact_univ_iff.mpr h /-- A topological space is σ-compact iff there exists a countable collection of compact subspaces that cover the entire space. -/ lemma SigmaCompactSpace_iff_exists_compact_covering : - SigmaCompactSpace α ↔ ∃ K : ℕ → Set α, (∀ n, IsCompact (K n)) ∧ ⋃ n, K n = univ := by + SigmaCompactSpace X ↔ ∃ K : ℕ → Set X, (∀ n, IsCompact (K n)) ∧ ⋃ n, K n = univ := by rw [← isSigmaCompact_univ_iff, IsSigmaCompact] -lemma SigmaCompactSpace.exists_compact_covering [h : SigmaCompactSpace α] : - ∃ K : ℕ → Set α, (∀ n, IsCompact (K n)) ∧ ⋃ n, K n = univ := +lemma SigmaCompactSpace.exists_compact_covering [h : SigmaCompactSpace X] : + ∃ K : ℕ → Set X, (∀ n, IsCompact (K n)) ∧ ⋃ n, K n = univ := SigmaCompactSpace_iff_exists_compact_covering.mp h /-- If `X` is σ-compact, `im f` is σ-compact. -/ -lemma isSigmaCompact_range {f : α → β} (hf : Continuous f) [SigmaCompactSpace α] : +lemma isSigmaCompact_range {f : X → Y} (hf : Continuous f) [SigmaCompactSpace X] : IsSigmaCompact (range f) := image_univ ▸ isSigmaCompact_univ.image hf /-- A subset `s` is σ-compact iff `s` (with the subspace topology) is a σ-compact space. -/ -lemma isSigmaCompact_iff_isSigmaCompact_univ {s : Set α} : +lemma isSigmaCompact_iff_isSigmaCompact_univ {s : Set X} : IsSigmaCompact s ↔ IsSigmaCompact (univ : Set s) := by rw [Subtype.isSigmaCompact_iff, image_univ, Subtype.range_coe] -lemma isSigmaCompact_iff_sigmaCompactSpace {s : Set α} : +lemma isSigmaCompact_iff_sigmaCompactSpace {s : Set X} : IsSigmaCompact s ↔ SigmaCompactSpace s := isSigmaCompact_iff_isSigmaCompact_univ.trans isSigmaCompact_univ_iff -- see Note [lower instance priority] -instance (priority := 200) CompactSpace.sigma_compact [CompactSpace α] : SigmaCompactSpace α := +instance (priority := 200) CompactSpace.sigma_compact [CompactSpace X] : SigmaCompactSpace X := ⟨⟨fun _ => univ, fun _ => isCompact_univ, iUnion_const _⟩⟩ #align compact_space.sigma_compact CompactSpace.sigma_compact -theorem SigmaCompactSpace.of_countable (S : Set (Set α)) (Hc : S.Countable) - (Hcomp : ∀ s ∈ S, IsCompact s) (HU : ⋃₀ S = univ) : SigmaCompactSpace α := +theorem SigmaCompactSpace.of_countable (S : Set (Set X)) (Hc : S.Countable) + (Hcomp : ∀ s ∈ S, IsCompact s) (HU : ⋃₀ S = univ) : SigmaCompactSpace X := ⟨(exists_seq_cover_iff_countable ⟨_, isCompact_empty⟩).2 ⟨S, Hc, Hcomp, HU⟩⟩ #align sigma_compact_space.of_countable SigmaCompactSpace.of_countable -- see Note [lower instance priority] instance (priority := 100) sigmaCompactSpace_of_locally_compact_second_countable - [LocallyCompactSpace α] [SecondCountableTopology α] : SigmaCompactSpace α := by - choose K hKc hxK using fun x : α => exists_compact_mem_nhds x + [LocallyCompactSpace X] [SecondCountableTopology X] : SigmaCompactSpace X := by + choose K hKc hxK using fun x : X => exists_compact_mem_nhds x rcases countable_cover_nhds hxK with ⟨s, hsc, hsU⟩ refine' SigmaCompactSpace.of_countable _ (hsc.image K) (ball_image_iff.2 fun x _ => hKc x) _ rwa [sUnion_image] #align sigma_compact_space_of_locally_compact_second_countable sigmaCompactSpace_of_locally_compact_second_countable -- porting note: doesn't work on the same line -variable (α) -variable [SigmaCompactSpace α] +variable (X) +variable [SigmaCompactSpace X] open SigmaCompactSpace /-- A choice of compact covering for a `σ`-compact space, chosen to be monotone. -/ -def compactCovering : ℕ → Set α := +def compactCovering : ℕ → Set X := Accumulate exists_compact_covering.choose #align compact_covering compactCovering -theorem isCompact_compactCovering (n : ℕ) : IsCompact (compactCovering α n) := +theorem isCompact_compactCovering (n : ℕ) : IsCompact (compactCovering X n) := isCompact_accumulate (Classical.choose_spec SigmaCompactSpace.exists_compact_covering).1 n #align is_compact_compact_covering isCompact_compactCovering -theorem iUnion_compactCovering : ⋃ n, compactCovering α n = univ := by +theorem iUnion_compactCovering : ⋃ n, compactCovering X n = univ := by rw [compactCovering, iUnion_accumulate] exact (Classical.choose_spec SigmaCompactSpace.exists_compact_covering).2 #align Union_compact_covering iUnion_compactCovering @[mono] -theorem compactCovering_subset ⦃m n : ℕ⦄ (h : m ≤ n) : compactCovering α m ⊆ compactCovering α n := +theorem compactCovering_subset ⦃m n : ℕ⦄ (h : m ≤ n) : compactCovering X m ⊆ compactCovering X n := monotone_accumulate h #align compact_covering_subset compactCovering_subset -variable {α} +variable {X} -theorem exists_mem_compactCovering (x : α) : ∃ n, x ∈ compactCovering α n := - iUnion_eq_univ_iff.mp (iUnion_compactCovering α) x +theorem exists_mem_compactCovering (x : X) : ∃ n, x ∈ compactCovering X n := + iUnion_eq_univ_iff.mp (iUnion_compactCovering X) x #align exists_mem_compact_covering exists_mem_compactCovering -instance [SigmaCompactSpace β] : SigmaCompactSpace (α × β) := - ⟨⟨fun n => compactCovering α n ×ˢ compactCovering β n, fun _ => +instance [SigmaCompactSpace Y] : SigmaCompactSpace (X × Y) := + ⟨⟨fun n => compactCovering X n ×ˢ compactCovering Y n, fun _ => (isCompact_compactCovering _ _).prod (isCompact_compactCovering _ _), by - simp only [iUnion_prod_of_monotone (compactCovering_subset α) (compactCovering_subset β), + simp only [iUnion_prod_of_monotone (compactCovering_subset X) (compactCovering_subset Y), iUnion_compactCovering, univ_prod_univ]⟩⟩ -instance [Finite ι] [∀ i, TopologicalSpace (π i)] [∀ i, SigmaCompactSpace (π i)] : - SigmaCompactSpace (∀ i, π i) := by - refine' ⟨⟨fun n => Set.pi univ fun i => compactCovering (π i) n, - fun n => isCompact_univ_pi fun i => isCompact_compactCovering (π i) _, _⟩⟩ +instance [Finite ι] {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, SigmaCompactSpace (X i)] : + SigmaCompactSpace (∀ i, X i) := by + refine' ⟨⟨fun n => Set.pi univ fun i => compactCovering (X i) n, + fun n => isCompact_univ_pi fun i => isCompact_compactCovering (X i) _, _⟩⟩ rw [iUnion_univ_pi_of_monotone] · simp only [iUnion_compactCovering, pi_univ] - · exact fun i => compactCovering_subset (π i) + · exact fun i => compactCovering_subset (X i) -instance [SigmaCompactSpace β] : SigmaCompactSpace (Sum α β) := - ⟨⟨fun n => Sum.inl '' compactCovering α n ∪ Sum.inr '' compactCovering β n, fun n => - ((isCompact_compactCovering α n).image continuous_inl).union - ((isCompact_compactCovering β n).image continuous_inr), +instance [SigmaCompactSpace Y] : SigmaCompactSpace (Sum X Y) := + ⟨⟨fun n => Sum.inl '' compactCovering X n ∪ Sum.inr '' compactCovering Y n, fun n => + ((isCompact_compactCovering X n).image continuous_inl).union + ((isCompact_compactCovering Y n).image continuous_inr), by simp only [iUnion_union_distrib, ← image_iUnion, iUnion_compactCovering, image_univ, range_inl_union_range_inr]⟩⟩ -instance [Countable ι] [∀ i, TopologicalSpace (π i)] [∀ i, SigmaCompactSpace (π i)] : - SigmaCompactSpace (Σi, π i) := by +instance [Countable ι] {X : ι → Type*} [∀ i, TopologicalSpace (X i)] + [∀ i, SigmaCompactSpace (X i)] : SigmaCompactSpace (Σi, X i) := by cases isEmpty_or_nonempty ι · infer_instance · rcases exists_surjective_nat ι with ⟨f, hf⟩ - refine' ⟨⟨fun n => ⋃ k ≤ n, Sigma.mk (f k) '' compactCovering (π (f k)) n, fun n => _, _⟩⟩ + refine' ⟨⟨fun n => ⋃ k ≤ n, Sigma.mk (f k) '' compactCovering (X (f k)) n, fun n => _, _⟩⟩ · refine' (finite_le_nat _).isCompact_biUnion fun k _ => _ exact (isCompact_compactCovering _ _).image continuous_sigmaMk · simp only [iUnion_eq_univ_iff, Sigma.forall, mem_iUnion] @@ -254,34 +254,34 @@ instance [Countable ι] [∀ i, TopologicalSpace (π i)] [∀ i, SigmaCompactSpa refine' ⟨max k n, k, le_max_left _ _, mem_image_of_mem _ _⟩ exact compactCovering_subset _ (le_max_right _ _) hn -protected theorem ClosedEmbedding.sigmaCompactSpace {e : β → α} (he : ClosedEmbedding e) : - SigmaCompactSpace β := - ⟨⟨fun n => e ⁻¹' compactCovering α n, fun n => +protected theorem ClosedEmbedding.sigmaCompactSpace {e : Y → X} (he : ClosedEmbedding e) : + SigmaCompactSpace Y := + ⟨⟨fun n => e ⁻¹' compactCovering X n, fun n => he.isCompact_preimage (isCompact_compactCovering _ _), by rw [← preimage_iUnion, iUnion_compactCovering, preimage_univ]⟩⟩ #align closed_embedding.sigma_compact_space ClosedEmbedding.sigmaCompactSpace -- porting note: new lemma -theorem IsClosed.sigmaCompactSpace {s : Set α} (hs : IsClosed s) : SigmaCompactSpace s := +theorem IsClosed.sigmaCompactSpace {s : Set X} (hs : IsClosed s) : SigmaCompactSpace s := (closedEmbedding_subtype_val hs).sigmaCompactSpace -instance [SigmaCompactSpace β] : SigmaCompactSpace (ULift.{u} β) := +instance [SigmaCompactSpace Y] : SigmaCompactSpace (ULift.{u} Y) := ULift.closedEmbedding_down.sigmaCompactSpace -/-- If `α` is a `σ`-compact space, then a locally finite family of nonempty sets of `α` can have +/-- If `X` is a `σ`-compact space, then a locally finite family of nonempty sets of `X` can have only countably many elements, `Set.Countable` version. -/ -protected theorem LocallyFinite.countable_univ {ι : Type*} {f : ι → Set α} (hf : LocallyFinite f) +protected theorem LocallyFinite.countable_univ {f : ι → Set X} (hf : LocallyFinite f) (hne : ∀ i, (f i).Nonempty) : (univ : Set ι).Countable := by - have := fun n => hf.finite_nonempty_inter_compact (isCompact_compactCovering α n) + have := fun n => hf.finite_nonempty_inter_compact (isCompact_compactCovering X n) refine (countable_iUnion fun n => (this n).countable).mono fun i _ => ?_ rcases hne i with ⟨x, hx⟩ - rcases iUnion_eq_univ_iff.1 (iUnion_compactCovering α) x with ⟨n, hn⟩ + rcases iUnion_eq_univ_iff.1 (iUnion_compactCovering X) x with ⟨n, hn⟩ exact mem_iUnion.2 ⟨n, x, hx, hn⟩ #align locally_finite.countable_univ LocallyFinite.countable_univ -/-- If `f : ι → Set α` is a locally finite covering of a σ-compact topological space by nonempty +/-- If `f : ι → Set X` is a locally finite covering of a σ-compact topological space by nonempty sets, then the index type `ι` is encodable. -/ -protected noncomputable def LocallyFinite.encodable {ι : Type*} {f : ι → Set α} +protected noncomputable def LocallyFinite.encodable {ι : Type*} {f : ι → Set X} (hf : LocallyFinite f) (hne : ∀ i, (f i).Nonempty) : Encodable ι := @Encodable.ofEquiv _ _ (hf.countable_univ hne).toEncodable (Equiv.Set.univ _).symm #align locally_finite.encodable LocallyFinite.encodable @@ -289,13 +289,13 @@ protected noncomputable def LocallyFinite.encodable {ι : Type*} {f : ι → Set /-- In a topological space with sigma compact topology, if `f` is a function that sends each point `x` of a closed set `s` to a neighborhood of `x` within `s`, then for some countable set `t ⊆ s`, the neighborhoods `f x`, `x ∈ t`, cover the whole set `s`. -/ -theorem countable_cover_nhdsWithin_of_sigma_compact {f : α → Set α} {s : Set α} (hs : IsClosed s) +theorem countable_cover_nhdsWithin_of_sigma_compact {f : X → Set X} {s : Set X} (hs : IsClosed s) (hf : ∀ x ∈ s, f x ∈ 𝓝[s] x) : ∃ (t : _) (_ : t ⊆ s), t.Countable ∧ s ⊆ ⋃ x ∈ t, f x := by simp only [nhdsWithin, mem_inf_principal] at hf choose t ht hsub using fun n => - ((isCompact_compactCovering α n).inter_right hs).elim_nhds_subcover _ fun x hx => hf x hx.right + ((isCompact_compactCovering X n).inter_right hs).elim_nhds_subcover _ fun x hx => hf x hx.right refine' - ⟨⋃ n, (t n : Set α), iUnion_subset fun n x hx => (ht n x hx).2, + ⟨⋃ n, (t n : Set X), iUnion_subset fun n x hx => (ht n x hx).2, countable_iUnion fun n => (t n).countable_toSet, fun x hx => mem_iUnion₂.2 _⟩ rcases exists_mem_compactCovering x with ⟨n, hn⟩ rcases mem_iUnion₂.1 (hsub n ⟨hn, hx⟩) with ⟨y, hyt : y ∈ t n, hyf : x ∈ s → x ∈ f y⟩ @@ -305,8 +305,8 @@ theorem countable_cover_nhdsWithin_of_sigma_compact {f : α → Set α} {s : Set /-- In a topological space with sigma compact topology, if `f` is a function that sends each point `x` to a neighborhood of `x`, then for some countable set `s`, the neighborhoods `f x`, `x ∈ s`, cover the whole space. -/ -theorem countable_cover_nhds_of_sigma_compact {f : α → Set α} (hf : ∀ x, f x ∈ 𝓝 x) : - ∃ s : Set α, s.Countable ∧ ⋃ x ∈ s, f x = univ := by +theorem countable_cover_nhds_of_sigma_compact {f : X → Set X} (hf : ∀ x, f x ∈ 𝓝 x) : + ∃ s : Set X, s.Countable ∧ ⋃ x ∈ s, f x = univ := by simp only [← nhdsWithin_univ] at hf rcases countable_cover_nhdsWithin_of_sigma_compact isClosed_univ fun x _ => hf x with ⟨s, -, hsc, hsU⟩ @@ -337,13 +337,13 @@ structure CompactExhaustion (X : Type*) [TopologicalSpace X] where namespace CompactExhaustion -instance : @RelHomClass (CompactExhaustion α) ℕ (Set α) LE.le HasSubset.Subset where +instance : @RelHomClass (CompactExhaustion X) ℕ (Set X) LE.le HasSubset.Subset where coe := toFun coe_injective' | ⟨_, _, _, _⟩, ⟨_, _, _, _⟩, rfl => rfl map_rel f _ _ h := monotone_nat_of_le_succ (fun n ↦ (f.subset_interior_succ' n).trans interior_subset) h -variable (K : CompactExhaustion α) +variable (K : CompactExhaustion X) @[simp] theorem toFun_eq_coe : K.toFun = K := rfl @@ -372,25 +372,25 @@ theorem iUnion_eq : ⋃ n, K n = univ := K.iUnion_eq' #align compact_exhaustion.Union_eq CompactExhaustion.iUnion_eq -theorem exists_mem (x : α) : ∃ n, x ∈ K n := +theorem exists_mem (x : X) : ∃ n, x ∈ K n := iUnion_eq_univ_iff.1 K.iUnion_eq x #align compact_exhaustion.exists_mem CompactExhaustion.exists_mem /-- The minimal `n` such that `x ∈ K n`. -/ -protected noncomputable def find (x : α) : ℕ := +protected noncomputable def find (x : X) : ℕ := Nat.find (K.exists_mem x) #align compact_exhaustion.find CompactExhaustion.find -theorem mem_find (x : α) : x ∈ K (K.find x) := +theorem mem_find (x : X) : x ∈ K (K.find x) := Nat.find_spec (K.exists_mem x) #align compact_exhaustion.mem_find CompactExhaustion.mem_find -theorem mem_iff_find_le {x : α} {n : ℕ} : x ∈ K n ↔ K.find x ≤ n := +theorem mem_iff_find_le {x : X} {n : ℕ} : x ∈ K n ↔ K.find x ≤ n := ⟨fun h => Nat.find_min' (K.exists_mem x) h, fun h => K.subset h <| K.mem_find x⟩ #align compact_exhaustion.mem_iff_find_le CompactExhaustion.mem_iff_find_le /-- Prepend the empty set to a compact exhaustion `K n`. -/ -def shiftr : CompactExhaustion α where +def shiftr : CompactExhaustion X where toFun n := Nat.casesOn n ∅ K isCompact' n := Nat.casesOn n isCompact_empty K.isCompact subset_interior_succ' n := Nat.casesOn n (empty_subset _) K.subset_interior_succ @@ -398,11 +398,11 @@ def shiftr : CompactExhaustion α where #align compact_exhaustion.shiftr CompactExhaustion.shiftr @[simp] -theorem find_shiftr (x : α) : K.shiftr.find x = K.find x + 1 := +theorem find_shiftr (x : X) : K.shiftr.find x = K.find x + 1 := Nat.find_comp_succ _ _ (not_mem_empty _) #align compact_exhaustion.find_shiftr CompactExhaustion.find_shiftr -theorem mem_diff_shiftr_find (x : α) : x ∈ K.shiftr (K.find x + 1) \ K.shiftr (K.find x) := +theorem mem_diff_shiftr_find (x : X) : x ∈ K.shiftr (K.find x + 1) \ K.shiftr (K.find x) := ⟨K.mem_find _, mt K.shiftr.mem_iff_find_le.1 <| by simp only [find_shiftr, not_le, Nat.lt_succ_self]⟩ #align compact_exhaustion.mem_diff_shiftr_find CompactExhaustion.mem_diff_shiftr_find @@ -424,8 +424,8 @@ noncomputable def choice (X : Type*) [TopologicalSpace X] [WeaklyLocallyCompactS exact iUnion_mono' fun n => ⟨n + 1, subset_union_right _ _⟩ #align compact_exhaustion.choice CompactExhaustion.choice -noncomputable instance [LocallyCompactSpace α] [SigmaCompactSpace α] : - Inhabited (CompactExhaustion α) := - ⟨CompactExhaustion.choice α⟩ +noncomputable instance [LocallyCompactSpace X] [SigmaCompactSpace X] : + Inhabited (CompactExhaustion X) := + ⟨CompactExhaustion.choice X⟩ end CompactExhaustion diff --git a/Mathlib/Topology/CompletelyRegular.lean b/Mathlib/Topology/CompletelyRegular.lean new file mode 100644 index 0000000000000..653f17818480f --- /dev/null +++ b/Mathlib/Topology/CompletelyRegular.lean @@ -0,0 +1,97 @@ +/- +Copyright (c) 2023 Matias Heikkilä. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Matias Heikkilä +-/ +import Mathlib.Topology.UrysohnsLemma +import Mathlib.Topology.UnitInterval + +/-! +# Completely regular topological spaces. + +This file defines `CompletelyRegularSpace` and `T35Space`. + +## Main definitions + +* `CompletelyRegularSpace`: A completely regular space `X` is such that each closed set `K ⊆ X` + and a point `x ∈ Kᶜ`, there is a continuous function `f` from `X` to the unit interval, so that + `f x = 0` and `f k = 1` for all `k ∈ K`. A completely regular space is a regular space, and a + normal space is a completely regular space. +* `T35Space`: A T₃.₅ space is a completely regular space that is also T₁. A T₃.₅ space is a T₃ + space and a T₄ space is a T₃.₅ space. + +## Main results + +### Completely regular spaces + +* `CompletelyRegularSpace.regularSpace`: A completely regular space is a regular space. +* `NormalSpace.completelyRegularSpace`: A normal space is a completely regular space. + +### T₃.₅ spaces + +* `T35Space.t3Space`: A T₃.₅ is a T₃ space. +* `T4Space.t35Space`: A T₄ space is a T₃.₅ space. + +## Implementation notes + +The present definition `CompletelyRegularSpace` is a slight modification of the one given in +[russell1974]. There it's assumed that any point `x ∈ Kᶜ` is separated from the closed set `K` by a +continuous *real* valued function `f` (as opposed to `f` being unit-interval-valued). This can be +converted to the present definition by replacing a real-valued `f` by `h ∘ g ∘ f`, with +`g : x ↦ max(x, 0)` and `h : x ↦ min(x, 1)`. Some sources (e.g. [russell1974]) also assume that a +completely regular space is T₁. Here a completely regular space that is also T₁ is called a T₃.₅ +space. + +## References + +* [Russell C. Walker, *The Stone-Čech Compactification*][russell1974] +-/ + +universe u + +noncomputable section + +open Set Topology Filter unitInterval + +variable {X : Type u} [TopologicalSpace X] [T1Space X] + +/-- A space is completely regular if points can be separated from closed sets via + continuous functions to the unit interval. -/ +@[mk_iff completelyRegularSpace_iff] +class CompletelyRegularSpace (X : Type u) [TopologicalSpace X] : Prop where + completely_regular : ∀ (x : X), ∀ (K : Set X) (_: IsClosed K), x ∉ K → + ∃ f : X → I, Continuous f ∧ f x = 0 ∧ EqOn f 1 K + +instance CompletelyRegularSpace.RegularSpace [CompletelyRegularSpace X] : RegularSpace X := by + rw [regularSpace_iff] + intro s a hs ha + obtain ⟨f, cf, hf, hhf⟩ := CompletelyRegularSpace.completely_regular a s hs ha + apply disjoint_of_map (f := f) + apply Disjoint.mono (cf.tendsto_nhdsSet_nhds hhf) cf.continuousAt + exact disjoint_nhds_nhds.mpr (hf.symm ▸ zero_ne_one).symm + +instance NormalSpace.completelyRegularSpace [NormalSpace X] : CompletelyRegularSpace X := by + rw [completelyRegularSpace_iff] + intro x K hK hx + have cx : IsClosed {x} := by apply T1Space.t1 + have d : Disjoint {x} K := by rwa [Set.disjoint_iff, subset_empty_iff, singleton_inter_eq_empty] + let ⟨⟨f, cf⟩, hfx, hfK, hficc⟩ := exists_continuous_zero_one_of_closed cx hK d + let g : X → I := fun x => ⟨f x, hficc x⟩ + have cg : Continuous g := cf.subtype_mk hficc + have hgx : g x = 0 := Subtype.ext (hfx (mem_singleton_iff.mpr (Eq.refl x))) + have hgK : EqOn g 1 K := fun k hk => Subtype.ext (hfK hk) + exact ⟨g, cg, hgx, hgK⟩ + +/-- A T₃.₅ space is a completely regular space that is also T1. -/ +@[mk_iff t35Space_iff] +class T35Space (X : Type u) [TopologicalSpace X] extends T1Space X, CompletelyRegularSpace X : Prop + +instance T35Space.t3space [T35Space X] : T3Space X := by + have : T0Space X := by apply T1Space.t0Space + have : RegularSpace X := by apply CompletelyRegularSpace.RegularSpace + exact {} + +instance T4Space.t35Space [T4Space X] : T35Space X := by + have : T1Space X := by apply T2Space.t1Space + have : CompletelyRegularSpace X := by apply NormalSpace.completelyRegularSpace + exact {} diff --git a/Mathlib/Topology/Connected/Basic.lean b/Mathlib/Topology/Connected/Basic.lean index 24ec718879d21..4f2b3c80dd0e3 100644 --- a/Mathlib/Topology/Connected/Basic.lean +++ b/Mathlib/Topology/Connected/Basic.lean @@ -977,7 +977,7 @@ lemma PreconnectedSpace.induction₂ [PreconnectedSpace α] (P : α → α → P P x y := by refine PreconnectedSpace.induction₂' P (fun z ↦ ?_) h' x y filter_upwards [h z] with a ha - refine ⟨ha, h'' ha⟩ + exact ⟨ha, h'' ha⟩ /-- In a preconnected set, given a transitive relation `P`, if `P x y` and `P y x` are true for `y` close enough to `x`, then `P x y` holds for all `x, y`. This is a version of the fact @@ -994,7 +994,7 @@ lemma IsPreconnected.induction₂' {s : Set α} (hs : IsPreconnected s) (P : α have Z := h x hx rwa [nhdsWithin_eq_map_subtype_coe] at Z · rintro ⟨a, ha⟩ ⟨b, hb⟩ ⟨c, hc⟩ hab hbc - exact h' a b c ha hb hc hab hbc + exact h' a b c ha hb hc hab hbc /-- In a preconnected set, if a symmetric transitive relation `P x y` is true for `y` close enough to `x`, then it holds for all `x, y`. This is a version of the fact that, if an equivalence @@ -1017,7 +1017,6 @@ theorem isPreconnected_iff_subset_of_disjoint {s : Set α} : · intro u v hu hv hs huv specialize h u v hu hv hs contrapose! huv - rw [← nonempty_iff_ne_empty] simp [not_subset] at huv rcases huv with ⟨⟨x, hxs, hxu⟩, ⟨y, hys, hyv⟩⟩ have hxv : x ∈ v := or_iff_not_imp_left.mp (hs hxs) hxu @@ -1079,7 +1078,6 @@ theorem isPreconnected_iff_subset_of_disjoint_closed : rw [isPreconnected_closed_iff] at h specialize h u v hu hv hs contrapose! huv - rw [← nonempty_iff_ne_empty] simp [not_subset] at huv rcases huv with ⟨⟨x, hxs, hxu⟩, ⟨y, hys, hyv⟩⟩ have hxv : x ∈ v := or_iff_not_imp_left.mp (hs hxs) hxu @@ -1317,7 +1315,6 @@ theorem isPreconnected_of_forall_constant {s : Set α} unfold IsPreconnected by_contra' rcases this with ⟨u, v, u_op, v_op, hsuv, ⟨x, x_in_s, x_in_u⟩, ⟨y, y_in_s, y_in_v⟩, H⟩ - rw [not_nonempty_iff_eq_empty] at H have hy : y ∉ u := fun y_in_u => eq_empty_iff_forall_not_mem.mp H y ⟨y_in_s, ⟨y_in_u, y_in_v⟩⟩ have : ContinuousOn u.boolIndicator s := by apply (continuousOn_boolIndicator_iff_clopen _ _).mpr ⟨_, _⟩ diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index b0e95f0b727c8..cd883b1ed3ee9 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -245,7 +245,7 @@ theorem continuous_extend : Continuous γ.extend := γ.continuous.Icc_extend' #align path.continuous_extend Path.continuous_extend -theorem _root_.Filter.Tendsto.path_extend {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] +theorem _root_.Filter.Tendsto.path_extend {l r : Y → X} {y : Y} {l₁ : Filter ℝ} {l₂ : Filter X} {γ : ∀ y, Path (l y) (r y)} (hγ : Tendsto (↿γ) (𝓝 y ×ˢ l₁.map (projIcc 0 1 zero_le_one)) l₂) : Tendsto (↿fun x => (γ x).extend) (𝓝 y ×ˢ l₁) l₂ := @@ -259,7 +259,7 @@ theorem _root_.ContinuousAt.path_extend {g : Y → ℝ} {l r : Y → X} (γ : #align continuous_at.path_extend ContinuousAt.path_extend @[simp] -theorem extend_extends {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) {t : ℝ} +theorem extend_extends {a b : X} (γ : Path a b) {t : ℝ} (ht : t ∈ (Icc 0 1 : Set ℝ)) : γ.extend t = γ ⟨t, ht⟩ := IccExtend_of_mem _ γ ht #align path.extend_extends Path.extend_extends @@ -271,29 +271,28 @@ theorem extend_one : γ.extend 1 = y := by simp #align path.extend_one Path.extend_one @[simp] -theorem extend_extends' {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) - (t : (Icc 0 1 : Set ℝ)) : γ.extend t = γ t := +theorem extend_extends' {a b : X} (γ : Path a b) (t : (Icc 0 1 : Set ℝ)) : γ.extend t = γ t := IccExtend_val _ γ t #align path.extend_extends' Path.extend_extends' @[simp] -theorem extend_range {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) : +theorem extend_range {a b : X} (γ : Path a b) : range γ.extend = range γ := IccExtend_range _ γ #align path.extend_range Path.extend_range -theorem extend_of_le_zero {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) {t : ℝ} +theorem extend_of_le_zero {a b : X} (γ : Path a b) {t : ℝ} (ht : t ≤ 0) : γ.extend t = a := (IccExtend_of_le_left _ _ ht).trans γ.source #align path.extend_of_le_zero Path.extend_of_le_zero -theorem extend_of_one_le {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) {t : ℝ} +theorem extend_of_one_le {a b : X} (γ : Path a b) {t : ℝ} (ht : 1 ≤ t) : γ.extend t = b := (IccExtend_of_right_le _ _ ht).trans γ.target #align path.extend_of_one_le Path.extend_of_one_le @[simp] -theorem refl_extend {X : Type*} [TopologicalSpace X] {a : X} : (Path.refl a).extend = fun _ => a := +theorem refl_extend {a : X} : (Path.refl a).extend = fun _ => a := rfl #align path.refl_extend Path.refl_extend @@ -349,14 +348,14 @@ theorem trans_symm (γ : Path x y) (γ' : Path y z) : (γ.trans γ').symm = γ'. #align path.trans_symm Path.trans_symm @[simp] -theorem refl_trans_refl {X : Type*} [TopologicalSpace X] {a : X} : +theorem refl_trans_refl {a : X} : (Path.refl a).trans (Path.refl a) = Path.refl a := by ext simp only [Path.trans, ite_self, one_div, Path.refl_extend] rfl #align path.refl_trans_refl Path.refl_trans_refl -theorem trans_range {X : Type*} [TopologicalSpace X] {a b c : X} (γ₁ : Path a b) (γ₂ : Path b c) : +theorem trans_range {a b c : X} (γ₁ : Path a b) (γ₂ : Path b c) : range (γ₁.trans γ₂) = range γ₁ ∪ range γ₂ := by rw [Path.trans] apply eq_of_subset_of_subset @@ -394,33 +393,32 @@ theorem trans_range {X : Type*} [TopologicalSpace X] {a b c : X} (γ₁ : Path a #align path.trans_range Path.trans_range /-- Image of a path from `x` to `y` by a map which is continuous on the path. -/ -def map' (γ : Path x y) {Y : Type*} [TopologicalSpace Y] {f : X → Y} - (h : ContinuousOn f (range γ)) : Path (f x) (f y) where +def map' (γ : Path x y) {f : X → Y} (h : ContinuousOn f (range γ)) : Path (f x) (f y) where toFun := f ∘ γ continuous_toFun := h.comp_continuous γ.continuous (fun x ↦ mem_range_self x) source' := by simp target' := by simp /-- Image of a path from `x` to `y` by a continuous map -/ -def map (γ : Path x y) {Y : Type*} [TopologicalSpace Y] {f : X → Y} (h : Continuous f) : +def map (γ : Path x y) {f : X → Y} (h : Continuous f) : Path (f x) (f y) := γ.map' h.continuousOn #align path.map Path.map @[simp] -theorem map_coe (γ : Path x y) {Y : Type*} [TopologicalSpace Y] {f : X → Y} (h : Continuous f) : +theorem map_coe (γ : Path x y) {f : X → Y} (h : Continuous f) : (γ.map h : I → Y) = f ∘ γ := by ext t rfl #align path.map_coe Path.map_coe @[simp] -theorem map_symm (γ : Path x y) {Y : Type*} [TopologicalSpace Y] {f : X → Y} (h : Continuous f) : +theorem map_symm (γ : Path x y) {f : X → Y} (h : Continuous f) : (γ.map h).symm = γ.symm.map h := rfl #align path.map_symm Path.map_symm @[simp] -theorem map_trans (γ : Path x y) (γ' : Path y z) {Y : Type*} [TopologicalSpace Y] {f : X → Y} +theorem map_trans (γ : Path x y) (γ' : Path y z) {f : X → Y} (h : Continuous f) : (γ.trans γ').map h = (γ.map h).trans (γ'.map h) := by ext t rw [trans_apply, map_coe, Function.comp_apply, trans_apply] @@ -434,7 +432,7 @@ theorem map_id (γ : Path x y) : γ.map continuous_id = γ := by #align path.map_id Path.map_id @[simp] -theorem map_map (γ : Path x y) {Y : Type*} [TopologicalSpace Y] {Z : Type*} [TopologicalSpace Z] +theorem map_map (γ : Path x y) {Z : Type*} [TopologicalSpace Z] {f : X → Y} (hf : Continuous f) {g : Y → Z} (hg : Continuous g) : (γ.map hf).map hg = γ.map (hg.comp hf) := by ext @@ -450,13 +448,13 @@ def cast (γ : Path x y) {x' y'} (hx : x' = x) (hy : y' = y) : Path x' y' where #align path.cast Path.cast @[simp] -theorem symm_cast {X : Type*} [TopologicalSpace X] {a₁ a₂ b₁ b₂ : X} (γ : Path a₂ b₂) - (ha : a₁ = a₂) (hb : b₁ = b₂) : (γ.cast ha hb).symm = γ.symm.cast hb ha := +theorem symm_cast {a₁ a₂ b₁ b₂ : X} (γ : Path a₂ b₂) (ha : a₁ = a₂) (hb : b₁ = b₂) : + (γ.cast ha hb).symm = γ.symm.cast hb ha := rfl #align path.symm_cast Path.symm_cast @[simp] -theorem trans_cast {X : Type*} [TopologicalSpace X] {a₁ a₂ b₁ b₂ c₁ c₂ : X} (γ : Path a₂ b₂) +theorem trans_cast {a₁ a₂ b₁ b₂ c₁ c₂ : X} (γ : Path a₂ b₂) (γ' : Path b₂ c₂) (ha : a₁ = a₂) (hb : b₁ = b₂) (hc : c₁ = c₂) : (γ.cast ha hb).trans (γ'.cast hb hc) = (γ.trans γ').cast ha hc := rfl @@ -468,7 +466,7 @@ theorem cast_coe (γ : Path x y) {x' y'} (hx : x' = x) (hy : y' = y) : (γ.cast #align path.cast_coe Path.cast_coe @[continuity] -theorem symm_continuous_family {X ι : Type*} [TopologicalSpace X] [TopologicalSpace ι] +theorem symm_continuous_family {ι : Type*} [TopologicalSpace ι] {a b : ι → X} (γ : ∀ t : ι, Path (a t) (b t)) (h : Continuous ↿γ) : Continuous ↿fun t => (γ t).symm := h.comp (continuous_id.prod_map continuous_symm) @@ -480,15 +478,15 @@ theorem continuous_symm : Continuous (symm : Path x y → Path y x) := #align path.continuous_symm Path.continuous_symm @[continuity] -theorem continuous_uncurry_extend_of_continuous_family {X ι : Type*} [TopologicalSpace X] - [TopologicalSpace ι] {a b : ι → X} (γ : ∀ t : ι, Path (a t) (b t)) (h : Continuous ↿γ) : +theorem continuous_uncurry_extend_of_continuous_family {ι : Type*} [TopologicalSpace ι] + {a b : ι → X} (γ : ∀ t : ι, Path (a t) (b t)) (h : Continuous ↿γ) : Continuous ↿fun t => (γ t).extend := by refine' h.comp (continuous_id.prod_map continuous_projIcc) exact zero_le_one #align path.continuous_uncurry_extend_of_continuous_family Path.continuous_uncurry_extend_of_continuous_family @[continuity] -theorem trans_continuous_family {X ι : Type*} [TopologicalSpace X] [TopologicalSpace ι] +theorem trans_continuous_family {ι : Type*} [TopologicalSpace ι] {a b c : ι → X} (γ₁ : ∀ t : ι, Path (a t) (b t)) (h₁ : Continuous ↿γ₁) (γ₂ : ∀ t : ι, Path (b t) (c t)) (h₂ : Continuous ↿γ₂) : Continuous ↿fun t => (γ₁ t).trans (γ₂ t) := by @@ -640,7 +638,7 @@ def truncateOfLE {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) {t (γ.truncate t₀ t₁).cast (by rw [min_eq_left h]) rfl #align path.truncate_of_le Path.truncateOfLE -theorem truncate_range {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) {t₀ t₁ : ℝ} : +theorem truncate_range {a b : X} (γ : Path a b) {t₀ t₁ : ℝ} : range (γ.truncate t₀ t₁) ⊆ range γ := by rw [← γ.extend_range] simp only [range_subset_iff, SetCoe.exists, SetCoe.forall] @@ -651,7 +649,7 @@ theorem truncate_range {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b /-- For a path `γ`, `γ.truncate` gives a "continuous family of paths", by which we mean the uncurried function which maps `(t₀, t₁, s)` to `γ.truncate t₀ t₁ s` is continuous. -/ @[continuity] -theorem truncate_continuous_family {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) : +theorem truncate_continuous_family {a b : X} (γ : Path a b) : Continuous (fun x => γ.truncate x.1 x.2.1 x.2.2 : ℝ × ℝ × I → X) := γ.continuous_extend.comp (((continuous_subtype_val.comp (continuous_snd.comp continuous_snd)).max continuous_fst).min @@ -659,14 +657,14 @@ theorem truncate_continuous_family {X : Type*} [TopologicalSpace X] {a b : X} ( #align path.truncate_continuous_family Path.truncate_continuous_family @[continuity] -theorem truncate_const_continuous_family {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) +theorem truncate_const_continuous_family {a b : X} (γ : Path a b) (t : ℝ) : Continuous ↿(γ.truncate t) := by have key : Continuous (fun x => (t, x) : ℝ × I → ℝ × ℝ × I) := by continuity exact γ.truncate_continuous_family.comp key #align path.truncate_const_continuous_family Path.truncate_const_continuous_family @[simp] -theorem truncate_self {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) (t : ℝ) : +theorem truncate_self {a b : X} (γ : Path a b) (t : ℝ) : γ.truncate t t = (Path.refl <| γ.extend t).cast (by rw [min_self]) rfl := by ext x rw [cast_coe] @@ -675,19 +673,19 @@ theorem truncate_self {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) #align path.truncate_self Path.truncate_self @[simp 1001] -- porting note: increase `simp` priority so left-hand side doesn't simplify -theorem truncate_zero_zero {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) : +theorem truncate_zero_zero {a b : X} (γ : Path a b) : γ.truncate 0 0 = (Path.refl a).cast (by rw [min_self, γ.extend_zero]) γ.extend_zero := by convert γ.truncate_self 0 #align path.truncate_zero_zero Path.truncate_zero_zero @[simp 1001] -- porting note: increase `simp` priority so left-hand side doesn't simplify -theorem truncate_one_one {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) : +theorem truncate_one_one {a b : X} (γ : Path a b) : γ.truncate 1 1 = (Path.refl b).cast (by rw [min_self, γ.extend_one]) γ.extend_one := by convert γ.truncate_self 1 #align path.truncate_one_one Path.truncate_one_one @[simp] -theorem truncate_zero_one {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) : +theorem truncate_zero_one {a b : X} (γ : Path a b) : γ.truncate 0 1 = γ.cast (by simp [zero_le_one, extend_zero]) (by simp) := by ext x rw [cast_coe] @@ -875,6 +873,16 @@ theorem JoinedIn.trans (hxy : JoinedIn F x y) (hyz : JoinedIn F y z) : JoinedIn exact hxy.trans hyz #align joined_in.trans JoinedIn.trans +theorem Specializes.joinedIn (h : x ⤳ y) (hx : x ∈ F) (hy : y ∈ F) : JoinedIn F x y := by + refine ⟨⟨⟨Set.piecewise {1} (const I y) (const I x), ?_⟩, by simp, by simp⟩, fun t ↦ ?_⟩ + · exact isClosed_singleton.continuous_piecewise_of_specializes continuous_const continuous_const + fun _ ↦ h + · simp only [Path.coe_mk_mk, piecewise] + split_ifs <;> assumption + +theorem Inseparable.joinedIn (h : Inseparable x y) (hx : x ∈ F) (hy : y ∈ F) : JoinedIn F x y := + h.specializes.joinedIn hx hy + /-! ### Path component -/ @@ -961,7 +969,8 @@ theorem isPathConnected_iff : fun ⟨⟨b, b_in⟩, h⟩ => ⟨b, b_in, fun x_in => h _ b_in _ x_in⟩⟩ #align is_path_connected_iff isPathConnected_iff -theorem IsPathConnected.image' {Y : Type*} [TopologicalSpace Y] (hF : IsPathConnected F) +/-- If `f` is continuous on `F` and `F` is path-connected, so is `f(F)`. -/ +theorem IsPathConnected.image' (hF : IsPathConnected F) {f : X → Y} (hf : ContinuousOn f F) : IsPathConnected (f '' F) := by rcases hF with ⟨x, x_in, hx⟩ use f x, mem_image_of_mem f x_in @@ -969,10 +978,39 @@ theorem IsPathConnected.image' {Y : Type*} [TopologicalSpace Y] (hF : IsPathConn refine ⟨(hx y_in).somePath.map' ?_, fun t ↦ ⟨_, (hx y_in).somePath_mem t, rfl⟩⟩ exact hf.mono (range_subset_iff.2 (hx y_in).somePath_mem) -theorem IsPathConnected.image {Y : Type*} [TopologicalSpace Y] (hF : IsPathConnected F) {f : X → Y} +/-- If `f` is continuous and `F` is path-connected, so is `f(F)`. -/ +theorem IsPathConnected.image (hF : IsPathConnected F) {f : X → Y} (hf : Continuous f) : IsPathConnected (f '' F) := hF.image' hf.continuousOn #align is_path_connected.image IsPathConnected.image +/-- If `f : X → Y` is a `Inducing`, `F(f)` is path-connected iff `F` is. -/ +nonrec theorem Inducing.isPathConnected_iff {f : X → Y} (hf : Inducing f) : + IsPathConnected F ↔ IsPathConnected (f '' F) := by + refine ⟨fun hF ↦ hF.image hf.continuous, fun hF ↦ ?_⟩ + simp? [isPathConnected_iff] at hF ⊢ says + simp only [isPathConnected_iff, nonempty_image_iff, mem_image, forall_exists_index, + and_imp, forall_apply_eq_imp_iff₂] at hF ⊢ + refine ⟨hF.1, fun x hx y hy ↦ ?_⟩ + rcases hF.2 x hx y hy with ⟨γ, hγ⟩ + choose γ' hγ' hγγ' using hγ + have key₁ : Inseparable x (γ' 0) := by rw [← hf.inseparable_iff, hγγ' 0, γ.source] + have key₂ : Inseparable (γ' 1) y := by rw [← hf.inseparable_iff, hγγ' 1, γ.target] + refine key₁.joinedIn hx (hγ' 0) |>.trans ⟨⟨⟨γ', ?_⟩, rfl, rfl⟩, hγ'⟩ |>.trans + (key₂.joinedIn (hγ' 1) hy) + simpa [hf.continuous_iff] using γ.continuous.congr fun t ↦ (hγγ' t).symm + +/-- If `h : X → Y` is a homeomorphism, `h(s)` is path-connected iff `s` is. -/ +@[simp] +theorem Homeomorph.isPathConnected_image {s : Set X} (h : X ≃ₜ Y) : + IsPathConnected (h '' s) ↔ IsPathConnected s := + h.inducing.isPathConnected_iff.symm + +/-- If `h : X → Y` is a homeomorphism, `h⁻¹(s)` is path-connected iff `s` is. -/ +@[simp] +theorem Homeomorph.isPathConnected_preimage {s : Set Y} (h : X ≃ₜ Y) : + IsPathConnected (h ⁻¹' s) ↔ IsPathConnected s := by + rw [← Homeomorph.image_symm]; exact h.symm.isPathConnected_image + theorem IsPathConnected.mem_pathComponent (h : IsPathConnected F) (x_in : x ∈ F) (y_in : y ∈ F) : y ∈ pathComponent x := (h.joinedIn x x_in y y_in).joined @@ -1006,7 +1044,7 @@ theorem IsPathConnected.preimage_coe {U W : Set X} (hW : IsPathConnected W) (hWU exact ⟨(hx hyW).joined_subtype.somePath.map (continuous_inclusion hWU), by simp⟩ #align is_path_connected.preimage_coe IsPathConnected.preimage_coe -theorem IsPathConnected.exists_path_through_family {X : Type*} [TopologicalSpace X] {n : ℕ} +theorem IsPathConnected.exists_path_through_family {n : ℕ} {s : Set X} (h : IsPathConnected s) (p : Fin (n + 1) → X) (hp : ∀ i, p i ∈ s) : ∃ γ : Path (p 0) (p n), range γ ⊆ s ∧ ∀ i, p i ∈ range γ := by let p' : ℕ → X := fun k => if h : k < n + 1 then p ⟨k, h⟩ else p ⟨0, n.zero_lt_succ⟩ @@ -1062,7 +1100,7 @@ theorem IsPathConnected.exists_path_through_family {X : Type*} [TopologicalSpace rw [Nat.mod_eq_of_lt hi] #align is_path_connected.exists_path_through_family IsPathConnected.exists_path_through_family -theorem IsPathConnected.exists_path_through_family' {X : Type*} [TopologicalSpace X] {n : ℕ} +theorem IsPathConnected.exists_path_through_family' {n : ℕ} {s : Set X} (h : IsPathConnected s) (p : Fin (n + 1) → X) (hp : ∀ i, p i ∈ s) : ∃ (γ : Path (p 0) (p n)) (t : Fin (n + 1) → I), (∀ t, γ t ∈ s) ∧ ∀ i, γ (t i) = p i := by rcases h.exists_path_through_family p hp with ⟨γ, hγ⟩ diff --git a/Mathlib/Topology/Connected/TotallyDisconnected.lean b/Mathlib/Topology/Connected/TotallyDisconnected.lean index 75f5680ce3d1f..3baeaf5a06e8d 100644 --- a/Mathlib/Topology/Connected/TotallyDisconnected.lean +++ b/Mathlib/Topology/Connected/TotallyDisconnected.lean @@ -42,6 +42,7 @@ theorem isTotallyDisconnected_singleton {x} : IsTotallyDisconnected ({x} : Set #align is_totally_disconnected_singleton isTotallyDisconnected_singleton /-- A space is totally disconnected if all of its connected components are singletons. -/ +@[mk_iff] class TotallyDisconnectedSpace (α : Type u) [TopologicalSpace α] : Prop where /-- The universal set `Set.univ` in a totally disconnected space is totally disconnected. -/ isTotallyDisconnected_univ : IsTotallyDisconnected (univ : Set α) @@ -161,10 +162,25 @@ theorem Embedding.isTotallyDisconnected [TopologicalSpace β] {f : α → β} (h isTotallyDisconnected_of_image hf.continuous.continuousOn hf.inj h #align embedding.is_totally_disconnected Embedding.isTotallyDisconnected +lemma Embedding.isTotallyDisconnected_image [TopologicalSpace β] {f : α → β} (hf : Embedding f) + {s : Set α} : IsTotallyDisconnected (f '' s) ↔ IsTotallyDisconnected s := by + refine ⟨hf.isTotallyDisconnected, fun hs u hus hu ↦ ?_⟩ + obtain ⟨v, hvs, rfl⟩ : ∃ v, v ⊆ s ∧ f '' v = u := + ⟨f ⁻¹' u ∩ s, inter_subset_right _ _, by rwa [image_preimage_inter, inter_eq_left]⟩ + rw [hf.toInducing.isPreconnected_image] at hu + exact (hs v hvs hu).image _ + +lemma Embedding.isTotallyDisconnected_range [TopologicalSpace β] {f : α → β} (hf : Embedding f) : + IsTotallyDisconnected (range f) ↔ TotallyDisconnectedSpace α := by + rw [TotallyDisconnectedSpace_iff, ← image_univ, hf.isTotallyDisconnected_image] + +lemma totallyDisconnectedSpace_subtype_iff {s : Set α} : + TotallyDisconnectedSpace s ↔ IsTotallyDisconnected s := by + rw [← embedding_subtype_val.isTotallyDisconnected_range, Subtype.range_val] + instance Subtype.totallyDisconnectedSpace {α : Type*} {p : α → Prop} [TopologicalSpace α] [TotallyDisconnectedSpace α] : TotallyDisconnectedSpace (Subtype p) := - ⟨embedding_subtype_val.isTotallyDisconnected - (isTotallyDisconnected_of_totallyDisconnectedSpace _)⟩ + totallyDisconnectedSpace_subtype_iff.2 (isTotallyDisconnected_of_totallyDisconnectedSpace _) #align subtype.totally_disconnected_space Subtype.totallyDisconnectedSpace end TotallyDisconnected diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index ffba5130508e1..1a06a69100f57 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -491,6 +491,10 @@ theorem continuous_swap : Continuous (Prod.swap : α × β → β × α) := continuous_snd.prod_mk continuous_fst #align continuous_swap continuous_swap +lemma isClosedMap_swap : IsClosedMap (Prod.swap : α × β → β × α) := fun s hs ↦ by + rw [image_swap_eq_preimage_swap] + exact hs.preimage continuous_swap + theorem continuous_uncurry_left {f : α → β → γ} (a : α) (h : Continuous (uncurry f)) : Continuous (f a) := h.comp (Continuous.Prod.mk _) @@ -528,6 +532,10 @@ theorem mem_nhds_prod_iff {a : α} {b : β} {s : Set (α × β)} : s ∈ 𝓝 (a, b) ↔ ∃ u ∈ 𝓝 a, ∃ v ∈ 𝓝 b, u ×ˢ v ⊆ s := by rw [nhds_prod_eq, mem_prod_iff] #align mem_nhds_prod_iff mem_nhds_prod_iff +theorem mem_nhdsWithin_prod_iff {a : α} {b : β} {s : Set (α × β)} {ta : Set α} {tb : Set β} : + s ∈ 𝓝[ta ×ˢ tb] (a, b) ↔ ∃ u ∈ 𝓝[ta] a, ∃ v ∈ 𝓝[tb] b, u ×ˢ v ⊆ s := + by rw [nhdsWithin_prod_eq, mem_prod_iff] + -- porting note: moved up theorem Filter.HasBasis.prod_nhds {ιa ιb : Type*} {pa : ιa → Prop} {pb : ιb → Prop} {sa : ιa → Set α} {sb : ιb → Set β} {a : α} {b : β} (ha : (𝓝 a).HasBasis pa sa) @@ -1025,7 +1033,7 @@ lemma IsClosedMap.restrictPreimage {f : α → β} (hcl : IsClosedMap f) (T : Se IsClosedMap (T.restrictPreimage f) := by rw [isClosedMap_iff_clusterPt] at hcl ⊢ intro A ⟨y, hyT⟩ hy - rw [restrictPreimage, MapClusterPt, ← inducing_subtype_val.mapClusterPt_iff, MapClusterPt, + rw [Set.restrictPreimage, MapClusterPt, ← inducing_subtype_val.mapClusterPt_iff, MapClusterPt, map_map, MapsTo.restrict_commutes, ← map_map, ← MapClusterPt, map_principal] at hy rcases hcl _ y hy with ⟨x, hxy, hx⟩ have hxT : f x ∈ T := hxy ▸ hyT diff --git a/Mathlib/Topology/ContinuousFunction/Algebra.lean b/Mathlib/Topology/ContinuousFunction/Algebra.lean index 4124c29f69b1a..36f5d6251a163 100644 --- a/Mathlib/Topology/ContinuousFunction/Algebra.lean +++ b/Mathlib/Topology/ContinuousFunction/Algebra.lean @@ -55,7 +55,8 @@ variable {α : Type*} {β : Type*} {γ : Type*} variable [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] --- ### "mul" and "add" +/-! ### `mul` and `add` -/ + @[to_additive] instance instMul [Mul β] [ContinuousMul β] : Mul C(α, β) := ⟨fun f g => ⟨f * g, continuous_mul.comp (f.continuous.prod_mk g.continuous : _)⟩⟩ @@ -81,7 +82,8 @@ theorem mul_comp [Mul γ] [ContinuousMul γ] (f₁ f₂ : C(β, γ)) (g : C(α, #align continuous_map.mul_comp ContinuousMap.mul_comp #align continuous_map.add_comp ContinuousMap.add_comp --- ### "one" +/-! ### `one` -/ + @[to_additive] instance [One β] : One C(α, β) := ⟨const α 1⟩ @@ -104,7 +106,8 @@ theorem one_comp [One γ] (g : C(α, β)) : (1 : C(β, γ)).comp g = 1 := #align continuous_map.one_comp ContinuousMap.one_comp #align continuous_map.zero_comp ContinuousMap.zero_comp --- ### "nat_cast" +/-! ### `Nat.cast` -/ + instance [NatCast β] : NatCast C(α, β) := ⟨fun n => ContinuousMap.const _ n⟩ @@ -118,7 +121,8 @@ theorem nat_cast_apply [NatCast β] (n : ℕ) (x : α) : (n : C(α, β)) x = n : rfl #align continuous_map.nat_cast_apply ContinuousMap.nat_cast_apply --- ### "int_cast" +/-! ### `Int.cast` -/ + instance [IntCast β] : IntCast C(α, β) := ⟨fun n => ContinuousMap.const _ n⟩ @@ -132,7 +136,8 @@ theorem int_cast_apply [IntCast β] (n : ℤ) (x : α) : (n : C(α, β)) x = n : rfl #align continuous_map.int_cast_apply ContinuousMap.int_cast_apply --- ### "nsmul" and "pow" +/-! ### `nsmul` and `pow` -/ + instance instNSMul [AddMonoid β] [ContinuousAdd β] : SMul ℕ C(α, β) := ⟨fun n f => ⟨n • ⇑f, f.continuous.nsmul n⟩⟩ #align continuous_map.has_nsmul ContinuousMap.instNSMul @@ -169,7 +174,8 @@ theorem pow_comp [Monoid γ] [ContinuousMul γ] (f : C(β, γ)) (n : ℕ) (g : C -- don't make `nsmul_comp` simp as the linter complains it's redundant WRT `smul_comp` attribute [simp] pow_comp --- ### "inv" and "neg" +/-! ### `inv` and `neg` -/ + @[to_additive] instance [Group β] [TopologicalGroup β] : Inv C(α, β) where inv f := ⟨f⁻¹, f.continuous.inv⟩ @@ -192,7 +198,8 @@ theorem inv_comp [Group γ] [TopologicalGroup γ] (f : C(β, γ)) (g : C(α, β) #align continuous_map.inv_comp ContinuousMap.inv_comp #align continuous_map.neg_comp ContinuousMap.neg_comp --- ### "div" and "sub" +/-! ### `div` and `sub` -/ + @[to_additive] instance [Div β] [ContinuousDiv β] : Div C(α, β) where div f g := ⟨f / g, f.continuous.div' g.continuous⟩ @@ -216,7 +223,8 @@ theorem div_comp [Div γ] [ContinuousDiv γ] (f g : C(β, γ)) (h : C(α, β)) : #align continuous_map.div_comp ContinuousMap.div_comp #align continuous_map.sub_comp ContinuousMap.sub_comp --- ### "zpow" and "zsmul" +/-! ### `zpow` and `zsmul` -/ + instance instZSMul [AddGroup β] [TopologicalAddGroup β] : SMul ℤ C(α, β) where smul z f := ⟨z • ⇑f, f.continuous.zsmul z⟩ #align continuous_map.has_zsmul ContinuousMap.instZSMul diff --git a/Mathlib/Topology/ContinuousFunction/Basic.lean b/Mathlib/Topology/ContinuousFunction/Basic.lean index 83285588e1dfd..c44491fce0e8a 100644 --- a/Mathlib/Topology/ContinuousFunction/Basic.lean +++ b/Mathlib/Topology/ContinuousFunction/Basic.lean @@ -325,12 +325,37 @@ def prodSwap : C(α × β, β × α) := .prodMk .snd .fst end Prod +section Sigma + +variable {I A : Type*} {X : I → Type*} [TopologicalSpace A] [∀ i, TopologicalSpace (X i)] + /-- `Sigma.mk i` as a bundled continuous map. -/ @[simps apply] -def sigmaMk {ι : Type*} {Y : ι → Type*} [∀ i, TopologicalSpace (Y i)] (i : ι) : - C(Y i, Σ i, Y i) where +def sigmaMk (i : I) : C(X i, Σ i, X i) where toFun := Sigma.mk i +/-- +To give a continuous map out of a disjoint union, it suffices to give a continuous map out of +each term. This is `Sigma.uncurry` for continuous maps. +-/ +@[simps] +def sigma (f : ∀ i, C(X i, A)) : C((Σ i, X i), A) where + toFun ig := f ig.fst ig.snd + +variable (A X) in +/-- +Giving a continuous map out of a disjoint union is the same as giving a continuous map out of +each term. This is a version of `Equiv.piCurry` for continuous maps. +-/ +@[simps] +def sigmaEquiv : (∀ i, C(X i, A)) ≃ C((Σ i, X i), A) where + toFun := sigma + invFun f i := f.comp (sigmaMk i) + left_inv := by intro; ext; simp + right_inv := by intro; ext; simp + +end Sigma + section Pi variable {I A : Type*} {X Y : I → Type*} [TopologicalSpace A] [∀ i, TopologicalSpace (X i)] @@ -351,6 +376,18 @@ theorem pi_eval (f : ∀ i, C(A, X i)) (a : A) : (pi f) a = fun i : I => (f i) a def eval (i : I) : C(∀ j, X j, X i) where toFun := Function.eval i +variable (A X) in +/-- +Giving a continuous map out of a disjoint union is the same as giving a continuous map out of +each term +-/ +@[simps] +def piEquiv : (∀ i, C(A, X i)) ≃ C(A, ∀ i, X i) where + toFun := pi + invFun f i := (eval i).comp f + left_inv := by intro; ext; simp [pi] + right_inv := by intro; ext; simp [pi] + /-- Combine a collection of bundled continuous maps `C(X i, Y i)` into a bundled continuous map `C(∀ i, X i, ∀ i, Y i)`. -/ @[simps!] @@ -469,6 +506,77 @@ end Gluing end ContinuousMap +section Lift + +variable {X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] + {f : C(X, Y)} + +/-- `Setoid.quotientKerEquivOfRightInverse` as a homeomorphism. -/ +@[simps!] +def Function.RightInverse.homeomorph {f' : C(Y, X)} (hf : Function.RightInverse f' f) : + Quotient (Setoid.ker f) ≃ₜ Y where + toEquiv := Setoid.quotientKerEquivOfRightInverse _ _ hf + continuous_toFun := quotientMap_quot_mk.continuous_iff.mpr f.continuous + continuous_invFun := continuous_quotient_mk'.comp f'.continuous + +namespace QuotientMap + +/-- +The homeomorphism from the quotient of a quotient map to its codomain. This is +`Setoid.quotientKerEquivOfSurjective` as a homeomorphism. +-/ +@[simps!] +noncomputable def homeomorph (hf : QuotientMap f) : Quotient (Setoid.ker f) ≃ₜ Y where + toEquiv := Setoid.quotientKerEquivOfSurjective _ hf.surjective + continuous_toFun := quotientMap_quot_mk.continuous_iff.mpr hf.continuous + continuous_invFun := by + rw [hf.continuous_iff] + convert continuous_quotient_mk' + ext + simp only [Equiv.invFun_as_coe, Function.comp_apply, + (Setoid.quotientKerEquivOfSurjective f hf.surjective).symm_apply_eq] + rfl + +variable (hf : QuotientMap f) (g : C(X, Z)) (h : Function.FactorsThrough g f) + +/-- Descend a continuous map, which is constant on the fibres, along a quotient map. -/ +@[simps] +noncomputable def lift : C(Y, Z) where + toFun := ((fun i ↦ Quotient.liftOn' i g (fun _ _ (hab : f _ = f _) ↦ h hab)) : + Quotient (Setoid.ker f) → Z) ∘ hf.homeomorph.symm + continuous_toFun := Continuous.comp (continuous_quot_lift _ g.2) (Homeomorph.continuous _) + +/-- +The obvious triangle induced by `QuotientMap.lift` commutes: +``` + g + X --→ Z + | ↗ +f | / hf.lift g h + v / + Y +``` +-/ +@[simp] +theorem lift_comp : (hf.lift g h).comp f = g := by + ext + simpa using h (Function.rightInverse_surjInv _ _) + +/-- `QuotientMap.lift` as an equivalence. -/ +@[simps] +noncomputable def liftEquiv : { g : C(X, Z) // Function.FactorsThrough g f} ≃ C(Y, Z) where + toFun g := hf.lift g g.prop + invFun g := ⟨g.comp f, fun _ _ h ↦ by simp only [ContinuousMap.comp_apply]; rw [h]⟩ + left_inv := by intro; simp + right_inv := by + intro g + ext a + simpa using congrArg g (Function.rightInverse_surjInv hf.surjective a) + +end QuotientMap + +end Lift + namespace Homeomorph variable {α β γ : Type*} [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 4d3ed90b09d3d..970d7bf9c4f68 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -609,6 +609,16 @@ theorem continuous_prod_of_discrete_right [DiscreteTopology β] {f : α × β Continuous f ↔ ∀ b, Continuous (f ⟨·, b⟩) := by simp_rw [continuous_iff_continuousOn_univ]; exact continuousOn_prod_of_discrete_right +theorem isOpenMap_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} : + IsOpenMap f ↔ ∀ a, IsOpenMap (f ⟨a, ·⟩) := by + simp_rw [isOpenMap_iff_nhds_le, Prod.forall, nhds_prod_eq, nhds_discrete, pure_prod, map_map] + rfl + +theorem isOpenMap_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} : + IsOpenMap f ↔ ∀ b, IsOpenMap (f ⟨·, b⟩) := by + simp_rw [isOpenMap_iff_nhds_le, Prod.forall, forall_swap (α := α) (β := β), nhds_prod_eq, + nhds_discrete, prod_pure, map_map]; rfl + theorem continuousWithinAt_pi {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] {f : α → ∀ i, π i} {s : Set α} {x : α} : ContinuousWithinAt f s x ↔ ∀ i, ContinuousWithinAt (fun y => f y i) s x := diff --git a/Mathlib/Topology/Covering.lean b/Mathlib/Topology/Covering.lean index f7157fe094bc4..b5280cf15b8c6 100644 --- a/Mathlib/Topology/Covering.lean +++ b/Mathlib/Topology/Covering.lean @@ -156,25 +156,55 @@ theorem mk (F : X → Type*) [∀ x, TopologicalSpace (F x)] [∀ x, DiscreteTop (IsCoveringMapOn.mk f Set.univ F (fun x _ => e x) fun x _ => h x) #align is_covering_map.mk IsCoveringMap.mk -variable {f} +variable {f} (hf : IsCoveringMap f) -protected theorem continuous (hf : IsCoveringMap f) : Continuous f := +protected theorem continuous : Continuous f := continuous_iff_continuousOn_univ.mpr hf.isCoveringMapOn.continuousOn #align is_covering_map.continuous IsCoveringMap.continuous -protected theorem isLocallyHomeomorph (hf : IsCoveringMap f) : IsLocallyHomeomorph f := +protected theorem isLocallyHomeomorph : IsLocallyHomeomorph f := isLocallyHomeomorph_iff_isLocallyHomeomorphOn_univ.mpr hf.isCoveringMapOn.isLocallyHomeomorphOn #align is_covering_map.is_locally_homeomorph IsCoveringMap.isLocallyHomeomorph -protected theorem isOpenMap (hf : IsCoveringMap f) : IsOpenMap f := +protected theorem isOpenMap : IsOpenMap f := hf.isLocallyHomeomorph.isOpenMap #align is_covering_map.is_open_map IsCoveringMap.isOpenMap -protected theorem quotientMap (hf : IsCoveringMap f) (hf' : Function.Surjective f) : - QuotientMap f := +protected theorem quotientMap (hf' : Function.Surjective f) : QuotientMap f := hf.isOpenMap.to_quotientMap hf.continuous hf' #align is_covering_map.quotient_map IsCoveringMap.quotientMap +protected theorem isSeparatedMap : IsSeparatedMap f := + fun e₁ e₂ he hne ↦ by + obtain ⟨_, t, he₁⟩ := hf (f e₁) + have he₂ := he₁; simp_rw [he] at he₂; rw [← t.mem_source] at he₁ he₂ + refine ⟨t.source ∩ (Prod.snd ∘ t) ⁻¹' {(t e₁).2}, t.source ∩ (Prod.snd ∘ t) ⁻¹' {(t e₂).2}, + ?_, ?_, ⟨he₁, rfl⟩, ⟨he₂, rfl⟩, Set.disjoint_left.mpr fun x h₁ h₂ ↦ hne (t.injOn he₁ he₂ ?_)⟩ + iterate 2 + exact t.continuous_toFun.preimage_open_of_open t.open_source + (continuous_snd.isOpen_preimage _ <| isOpen_discrete _) + refine Prod.ext ?_ (h₁.2.symm.trans h₂.2) + rwa [t.proj_toFun e₁ he₁, t.proj_toFun e₂ he₂] + +variable {A} [TopologicalSpace A] {s : Set A} (hs : IsPreconnected s) {g g₁ g₂ : A → E} + +theorem eq_of_comp_eq [PreconnectedSpace A] (h₁ : Continuous g₁) (h₂ : Continuous g₂) + (he : f ∘ g₁ = f ∘ g₂) (a : A) (ha : g₁ a = g₂ a) : g₁ = g₂ := + hf.isSeparatedMap.eq_of_comp_eq hf.isLocallyHomeomorph.isLocallyInjective h₁ h₂ he a ha + +theorem eqOn_of_comp_eqOn (h₁ : ContinuousOn g₁ s) (h₂ : ContinuousOn g₂ s) + (he : s.EqOn (f ∘ g₁) (f ∘ g₂)) {a : A} (has : a ∈ s) (ha : g₁ a = g₂ a) : s.EqOn g₁ g₂ := + hf.isSeparatedMap.eqOn_of_comp_eqOn hf.isLocallyHomeomorph.isLocallyInjective hs h₁ h₂ he has ha + +theorem const_of_comp [PreconnectedSpace A] (cont : Continuous g) + (he : ∀ a a', f (g a) = f (g a')) (a a') : g a = g a' := + hf.isSeparatedMap.const_of_comp hf.isLocallyHomeomorph.isLocallyInjective cont he a a' + +theorem constOn_of_comp (cont : ContinuousOn g s) + (he : ∀ a ∈ s, ∀ a' ∈ s, f (g a) = f (g a')) + {a a'} (ha : a ∈ s) (ha' : a' ∈ s) : g a = g a' := + hf.isSeparatedMap.constOn_of_comp hf.isLocallyHomeomorph.isLocallyInjective hs cont he ha ha' + end IsCoveringMap variable {f} diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index 07e58c1ace315..2e59f522668d0 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -201,8 +201,7 @@ theorem continuousAt_extend [T3Space γ] {b : β} {f : α → γ} (di : DenseInd have V₁_in : V₁ ∈ 𝓝 b := by filter_upwards [hf] rintro x ⟨c, hc⟩ - unfold_let φ - rwa [di.extend_eq_of_tendsto hc] + rwa [← di.extend_eq_of_tendsto hc] at hc obtain ⟨V₂, V₂_in, V₂_op, hV₂⟩ : ∃ V₂ ∈ 𝓝 b, IsOpen V₂ ∧ ∀ x ∈ i ⁻¹' V₂, f x ∈ V' := by simpa [and_assoc] using ((nhds_basis_opens' b).comap i).tendsto_left_iff.mp (mem_of_mem_nhds V₁_in : b ∈ V₁) V' V'_in diff --git a/Mathlib/Topology/EMetricSpace/Basic.lean b/Mathlib/Topology/EMetricSpace/Basic.lean index 83dd6fa597326..e61e05c1d1142 100644 --- a/Mathlib/Topology/EMetricSpace/Basic.lean +++ b/Mathlib/Topology/EMetricSpace/Basic.lean @@ -8,7 +8,7 @@ import Mathlib.Topology.UniformSpace.Pi import Mathlib.Topology.UniformSpace.UniformConvergence import Mathlib.Topology.UniformSpace.UniformEmbedding -#align_import topology.metric_space.emetric_space from "leanprover-community/mathlib"@"195fcd60ff2bfe392543bceb0ec2adcdb472db4c" +#align_import topology.metric_space.emetric_space from "leanprover-community/mathlib"@"c8f305514e0d47dfaa710f5a52f0d21b588e6328" /-! # Extended metric spaces @@ -29,7 +29,9 @@ to `EMetricSpace` at the end. -/ -open Set Filter Classical Uniformity Topology BigOperators NNReal ENNReal +open Set Filter Classical + +open scoped Uniformity Topology BigOperators Filter NNReal ENNReal Pointwise universe u v w @@ -925,6 +927,12 @@ theorem diam_singleton : diam ({x} : Set α) = 0 := diam_subsingleton subsingleton_singleton #align emetric.diam_singleton EMetric.diam_singleton +@[to_additive (attr := simp)] +theorem diam_one [One α] : diam (1 : Set α) = 0 := + diam_singleton +#align emetric.diam_one EMetric.diam_one +#align emetric.diam_zero EMetric.diam_zero + theorem diam_iUnion_mem_option {ι : Type*} (o : Option ι) (s : ι → Set α) : diam (⋃ i ∈ o, s i) = ⨆ i ∈ o, diam (s i) := by cases o <;> simp #align emetric.diam_Union_mem_option EMetric.diam_iUnion_mem_option diff --git a/Mathlib/Topology/EMetricSpace/Paracompact.lean b/Mathlib/Topology/EMetricSpace/Paracompact.lean index 1dfa424744a91..b4494c9f093de 100644 --- a/Mathlib/Topology/EMetricSpace/Paracompact.lean +++ b/Mathlib/Topology/EMetricSpace/Paracompact.lean @@ -45,7 +45,7 @@ instance (priority := 100) instParacompactSpace [PseudoEMetricSpace α] : Paraco have hpow_le : ∀ {m n : ℕ}, m ≤ n → (2⁻¹ : ℝ≥0∞) ^ n ≤ 2⁻¹ ^ m := @fun m n h => pow_le_pow_of_le_one' (ENNReal.inv_le_one.2 ENNReal.one_lt_two.le) h have h2pow : ∀ n : ℕ, 2 * (2⁻¹ : ℝ≥0∞) ^ (n + 1) = 2⁻¹ ^ n := fun n => by - simp [pow_succ, ← mul_assoc, ENNReal.mul_inv_cancel] + simp [pow_succ, ← mul_assoc, ENNReal.mul_inv_cancel two_ne_zero two_ne_top] -- Consider an open covering `S : Set (Set α)` refine' ⟨fun ι s ho hcov => _⟩ simp only [iUnion_eq_univ_iff] at hcov diff --git a/Mathlib/Topology/ExtremallyDisconnected.lean b/Mathlib/Topology/ExtremallyDisconnected.lean index 7450bded8f911..973964c2da964 100644 --- a/Mathlib/Topology/ExtremallyDisconnected.lean +++ b/Mathlib/Topology/ExtremallyDisconnected.lean @@ -62,7 +62,7 @@ instance [ExtremallyDisconnected X] [T2Space X] : TotallySeparatedSpace X := rw [Set.mem_compl_iff, mem_closure_iff] push_neg refine' ⟨V, ⟨hUV.2.1, hUV.2.2.2.1, _⟩⟩ - rw [Set.not_nonempty_iff_eq_empty, ← Set.disjoint_iff_inter_eq_empty, disjoint_comm] + rw [← Set.disjoint_iff_inter_eq_empty, disjoint_comm] exact hUV.2.2.2.2 } end TotallySeparated diff --git a/Mathlib/Topology/Filter.lean b/Mathlib/Topology/Filter.lean index 5b8cee12a195c..57879a1f75f48 100644 --- a/Mathlib/Topology/Filter.lean +++ b/Mathlib/Topology/Filter.lean @@ -95,7 +95,7 @@ protected theorem HasBasis.nhds {l : Filter α} {p : ι → Prop} {s : ι → Se protected theorem tendsto_pure_self (l : Filter X) : Tendsto (pure : X → Filter X) l (𝓝 l) := by rw [Filter.tendsto_nhds] - refine fun s hs ↦ Eventually.mono hs fun x ↦ id + exact fun s hs ↦ Eventually.mono hs fun x ↦ id /-- Neighborhoods of a countably generated filter is a countably generated filter. -/ instance {l : Filter α} [IsCountablyGenerated l] : IsCountablyGenerated (𝓝 l) := diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index 4fc76bca113b2..d1e8f408fcec9 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -18,7 +18,7 @@ In this file we define * `Inseparable`: a relation saying that two points in a topological space have the same neighbourhoods; equivalently, they can't be separated by an open set; -* `InseparableSetoid X`: same relation, as a `setoid`; +* `InseparableSetoid X`: same relation, as a `Setoid`; * `SeparationQuotient X`: the quotient of `X` by its `InseparableSetoid`. @@ -39,7 +39,7 @@ topological space, separation setoid open Set Filter Function Topology List variable {X Y Z α ι : Type*} {π : ι → Type*} [TopologicalSpace X] [TopologicalSpace Y] - [TopologicalSpace Z] [∀ i, TopologicalSpace (π i)] {x y z : X} {s : Set X} {f : X → Y} + [TopologicalSpace Z] [∀ i, TopologicalSpace (π i)] {x y z : X} {s : Set X} {f g : X → Y} /-! ### `Specializes` relation @@ -228,6 +228,20 @@ theorem not_specializes_iff_exists_closed : ¬x ⤳ y ↔ ∃ S : Set X, IsClose rfl #align not_specializes_iff_exists_closed not_specializes_iff_exists_closed +theorem IsOpen.continuous_piecewise_of_specializes [DecidablePred (· ∈ s)] (hs : IsOpen s) + (hf : Continuous f) (hg : Continuous g) (hspec : ∀ x, f x ⤳ g x) : + Continuous (s.piecewise f g) := by + have : ∀ U, IsOpen U → g ⁻¹' U ⊆ f ⁻¹' U := fun U hU x hx ↦ (hspec x).mem_open hU hx + rw [continuous_def] + intro U hU + rw [piecewise_preimage, ite_eq_of_subset_right _ (this U hU)] + exact hU.preimage hf |>.inter hs |>.union (hU.preimage hg) + +theorem IsClosed.continuous_piecewise_of_specializes [DecidablePred (· ∈ s)] (hs : IsClosed s) + (hf : Continuous f) (hg : Continuous g) (hspec : ∀ x, g x ⤳ f x) : + Continuous (s.piecewise f g) := by + simpa only [piecewise_compl] using hs.isOpen_compl.continuous_piecewise_of_specializes hg hf hspec + variable (X) /-- Specialization forms a preorder on the topological space. -/ @@ -634,3 +648,8 @@ theorem continuous_lift₂ {f : X → Y → Z} {hf : ∀ a b c d, (a ~ᵢ c) → #align separation_quotient.continuous_lift₂ SeparationQuotient.continuous_lift₂ end SeparationQuotient + +theorem continuous_congr_of_inseparable (h : ∀ x, f x ~ᵢ g x) : + Continuous f ↔ Continuous g := by + simp_rw [SeparationQuotient.inducing_mk.continuous_iff (β := Y)] + exact continuous_congr fun x ↦ SeparationQuotient.mk_eq_mk.mpr (h x) diff --git a/Mathlib/Topology/Instances/Complex.lean b/Mathlib/Topology/Instances/Complex.lean index 24af3915fb251..c6a62c3d70c12 100644 --- a/Mathlib/Topology/Instances/Complex.lean +++ b/Mathlib/Topology/Instances/Complex.lean @@ -28,10 +28,8 @@ theorem Complex.subfield_eq_of_closed {K : Subfield ℂ} (hc : IsClosed (K : Set have := (Subalgebra.isSimpleOrder_of_finrank finrank_real_complex).eq_bot_or_eq_top (Subfield.toIntermediateField K this).toSubalgebra - simp_rw [← SetLike.coe_set_eq] at this ⊢ - convert this using 2 - simp only [RingHom.coe_fieldRange, Algebra.coe_bot, coe_algebraMap] - rfl + simp_rw [← SetLike.coe_set_eq, IntermediateField.coe_toSubalgebra] at this ⊢ + exact this suffices range (ofReal' : ℝ → ℂ) ⊆ closure (Set.range ((ofReal' : ℝ → ℂ) ∘ ((↑) : ℚ → ℝ))) by refine' subset_trans this _ rw [← IsClosed.closure_eq hc] diff --git a/Mathlib/Topology/Instances/Discrete.lean b/Mathlib/Topology/Instances/Discrete.lean index 6ab0819eb4d67..e58a8686f03da 100644 --- a/Mathlib/Topology/Instances/Discrete.lean +++ b/Mathlib/Topology/Instances/Discrete.lean @@ -5,7 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Order.SuccPred.Basic import Mathlib.Topology.Order.Basic -import Mathlib.Topology.MetricSpace.MetrizableUniformity +import Mathlib.Topology.Metrizable.Uniformity #align_import topology.instances.discrete from "leanprover-community/mathlib"@"bcfa726826abd57587355b4b5b7e78ad6527b7e4" diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 79e2967655249..ec1da50389cc0 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -9,6 +9,7 @@ import Mathlib.Topology.Algebra.InfiniteSum.Real import Mathlib.Topology.Algebra.Order.LiminfLimsup import Mathlib.Topology.Algebra.Order.T5 import Mathlib.Topology.MetricSpace.Lipschitz +import Mathlib.Topology.Metrizable.Basic #align_import topology.instances.ennreal from "leanprover-community/mathlib"@"ec4b2eeb50364487f80421c0b4c41328a611f30d" @@ -47,6 +48,9 @@ instance : T4Space ℝ≥0∞ := inferInstance instance : SecondCountableTopology ℝ≥0∞ := orderIsoUnitIntervalBirational.toHomeomorph.embedding.secondCountableTopology +instance : MetrizableSpace ENNReal := + orderIsoUnitIntervalBirational.toHomeomorph.embedding.metrizableSpace + theorem embedding_coe : Embedding ((↑) : ℝ≥0 → ℝ≥0∞) := coe_strictMono.embedding_of_ordConnected <| by rw [range_coe']; exact ordConnected_Iio #align ennreal.embedding_coe ENNReal.embedding_coe @@ -302,7 +306,7 @@ protected theorem tendsto_atTop_zero [Nonempty β] [SemilatticeSup β] {f : β theorem tendsto_sub : ∀ {a b : ℝ≥0∞}, (a ≠ ∞ ∨ b ≠ ∞) → Tendsto (fun p : ℝ≥0∞ × ℝ≥0∞ => p.1 - p.2) (𝓝 (a, b)) (𝓝 (a - b)) - | ⊤, ⊤, h => by simp only at h + | ⊤, ⊤, h => by simp only [ne_eq, not_true_eq_false, or_self] at h | ⊤, (b : ℝ≥0), _ => by rw [top_sub_coe, tendsto_nhds_top_iff_nnreal] refine fun x => ((lt_mem_nhds <| @coe_lt_top (b + 1 + x)).prod_nhds @@ -339,11 +343,11 @@ protected theorem tendsto_mul (ha : a ≠ 0 ∨ b ≠ ⊤) (hb : b ≠ 0 ∨ a refine' this.mono fun c hc => _ exact (ENNReal.div_mul_cancel hε.ne' coe_ne_top).symm.trans_lt (mul_lt_mul hc.1 hc.2) induction a using recTopCoe with - | top => simp only [ne_eq, or_false] at hb; simp [ht b hb, top_mul hb] + | top => simp only [ne_eq, or_false, not_true_eq_false] at hb; simp [ht b hb, top_mul hb] | coe a => induction b using recTopCoe with | top => - simp only [ne_eq, or_false] at ha + simp only [ne_eq, or_false, not_true_eq_false] at ha simpa [(· ∘ ·), mul_comm, mul_top ha] using (ht a ha).comp (continuous_swap.tendsto (some a, ⊤)) | coe b => @@ -1452,7 +1456,7 @@ theorem continuous_of_le_add_edist {f : α → ℝ≥0∞} (C : ℝ≥0∞) (hC #align continuous_of_le_add_edist continuous_of_le_add_edist theorem continuous_edist : Continuous fun p : α × α => edist p.1 p.2 := by - apply continuous_of_le_add_edist 2 (by norm_num) + apply continuous_of_le_add_edist 2 (by decide) rintro ⟨x, y⟩ ⟨x', y'⟩ calc edist x y ≤ edist x x' + edist x' y' + edist y' y := edist_triangle4 _ _ _ _ diff --git a/Mathlib/Topology/Instances/Int.lean b/Mathlib/Topology/Instances/Int.lean index 4c6aee914ee13..931a4217703d9 100644 --- a/Mathlib/Topology/Instances/Int.lean +++ b/Mathlib/Topology/Instances/Int.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Data.Int.Interval -import Mathlib.Topology.MetricSpace.Basic +import Mathlib.Topology.MetricSpace.PseudoMetric +import Mathlib.Topology.MetricSpace.Bounded import Mathlib.Order.Filter.Archimedean #align_import topology.instances.int from "leanprover-community/mathlib"@"70fd9563a21e7b963887c9360bd29b2393e6225a" @@ -85,4 +86,3 @@ theorem cofinite_eq : (cofinite : Filter ℤ) = atBot ⊔ atTop := by #align int.cofinite_eq Int.cofinite_eq end Int - diff --git a/Mathlib/Topology/Instances/Rat.lean b/Mathlib/Topology/Instances/Rat.lean index 86e8673608f02..d01ea7c5332f2 100644 --- a/Mathlib/Topology/Instances/Rat.lean +++ b/Mathlib/Topology/Instances/Rat.lean @@ -3,9 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Topology.MetricSpace.Basic import Mathlib.Topology.Algebra.Order.Archimedean -import Mathlib.Topology.Instances.Int import Mathlib.Topology.Instances.Nat import Mathlib.Topology.Instances.Real diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index cd202fe21fa44..2a6dd6d4b6d8c 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -3,16 +3,9 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Topology.MetricSpace.Basic -import Mathlib.Topology.Algebra.UniformGroup import Mathlib.Topology.Algebra.UniformMulAction -import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Algebra.Star import Mathlib.Topology.Algebra.Order.Field -import Mathlib.Topology.Algebra.Order.Archimedean -import Mathlib.RingTheory.Subring.Basic -import Mathlib.GroupTheory.Archimedean -import Mathlib.Algebra.Order.Group.Bounds import Mathlib.Algebra.Periodic import Mathlib.Topology.Instances.Int @@ -86,6 +79,9 @@ theorem Real.cocompact_eq : cocompact ℝ = atBot ⊔ atTop := by rw [← cobounded_eq_cocompact, cobounded_eq] #align real.cocompact_eq Real.cocompact_eq +theorem Real.atBot_le_cocompact : atBot ≤ cocompact ℝ := by simp +theorem Real.atTop_le_cocompact : atTop ≤ cocompact ℝ := by simp + /- TODO(Mario): Prove that these are uniform isomorphisms instead of uniform embeddings lemma uniform_embedding_add_rat {r : ℚ} : uniform_embedding (fun p : ℚ => p + r) := _ diff --git a/Mathlib/Topology/IsLocallyHomeomorph.lean b/Mathlib/Topology/IsLocallyHomeomorph.lean index 419d5dc42ed47..719db40c2d10b 100644 --- a/Mathlib/Topology/IsLocallyHomeomorph.lean +++ b/Mathlib/Topology/IsLocallyHomeomorph.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ import Mathlib.Topology.LocalHomeomorph +import Mathlib.Topology.SeparatedMap #align_import topology.is_locally_homeomorph from "leanprover-community/mathlib"@"e97cf15cd1aec9bd5c193b2ffac5a6dc9118912b" @@ -125,7 +126,7 @@ def IsLocallyHomeomorph := ∀ x : X, ∃ e : LocalHomeomorph X Y, x ∈ e.source ∧ f = e #align is_locally_homeomorph IsLocallyHomeomorph -theorem isLocallyHomeomorph_homeomorph (f : X ≃ₜ Y) : IsLocallyHomeomorph f := +theorem Homeomorph.isLocallyHomeomorph (f : X ≃ₜ Y) : IsLocallyHomeomorph f := fun _ ↦ ⟨f.toLocalHomeomorph, trivial, rfl⟩ variable {f s} @@ -163,6 +164,9 @@ theorem mk (h : ∀ x : X, ∃ e : LocalHomeomorph X Y, x ∈ e.source ∧ ∀ y variable {g f} +lemma isLocallyInjective (hf : IsLocallyHomeomorph f) : IsLocallyInjective f := + fun x ↦ by obtain ⟨f, hx, rfl⟩ := hf x; exact ⟨f.source, f.open_source, hx, f.injOn⟩ + theorem of_comp (hgf : IsLocallyHomeomorph (g ∘ f)) (hg : IsLocallyHomeomorph g) (cont : Continuous f) : IsLocallyHomeomorph f := isLocallyHomeomorph_iff_isLocallyHomeomorphOn_univ.mpr <| @@ -196,7 +200,7 @@ theorem openEmbedding_of_comp (hf : IsLocallyHomeomorph g) (hgf : OpenEmbedding (hgf.isLocallyHomeomorph.of_comp hf cont).openEmbedding_of_injective hgf.inj.of_comp open TopologicalSpace in -/-- Ranges of continuous local sections of a local homeomorphism form a basis of the total space. -/ +/-- Ranges of continuous local sections of a local homeomorphism form a basis of the source space.-/ theorem isTopologicalBasis (hf : IsLocallyHomeomorph f) : IsTopologicalBasis {U : Set X | ∃ V : Set Y, IsOpen V ∧ ∃ s : C(V,X), f ∘ s = (↑) ∧ Set.range s = U} := by refine isTopologicalBasis_of_open_of_nhds ?_ fun x U hx hU ↦ ?_ diff --git a/Mathlib/Topology/LocalHomeomorph.lean b/Mathlib/Topology/LocalHomeomorph.lean index 29aa9ee369d7c..ed34a45d65e58 100644 --- a/Mathlib/Topology/LocalHomeomorph.lean +++ b/Mathlib/Topology/LocalHomeomorph.lean @@ -22,14 +22,14 @@ instead of `e.toFun x` and `e.invFun x`. ## Main definitions -`Homeomorph.toLocalHomeomorph` : associating a local homeomorphism to a homeomorphism, with - `source = target = Set.univ`; -`LocalHomeomorph.symm` : the inverse of a local homeomorphism -`LocalHomeomorph.trans` : the composition of two local homeomorphisms -`LocalHomeomorph.refl` : the identity local homeomorphism -`LocalHomeomorph.ofSet` : the identity on a set `s` -`LocalHomeomorph.EqOnSource` : equivalence relation describing the "right" notion of equality - for local homeomorphisms +* `Homeomorph.toLocalHomeomorph`: associating a local homeomorphism to a homeomorphism, with + `source = target = Set.univ`; +* `LocalHomeomorph.symm`: the inverse of a local homeomorphism +* `LocalHomeomorph.trans`: the composition of two local homeomorphisms +* `LocalHomeomorph.refl`: the identity local homeomorphism +* `LocalHomeomorph.ofSet`: the identity on a set `s` +* `LocalHomeomorph.EqOnSource`: equivalence relation describing the "right" notion of equality + for local homeomorphisms ## Implementation notes @@ -323,7 +323,7 @@ theorem symm_image_target_eq_source (e : LocalHomeomorph α β) : e.symm '' e.ta /-- Two local homeomorphisms are equal when they have equal `toFun`, `invFun` and `source`. It is not sufficient to have equal `toFun` and `source`, as this only determines `invFun` on the target. This would only be true for a weaker notion of equality, arguably the right one, -called `eq_on_source`. -/ +called `EqOnSource`. -/ @[ext] protected theorem ext (e' : LocalHomeomorph α β) (h : ∀ x, e x = e' x) (hinv : ∀ x, e.symm x = e'.symm x) (hs : e.source = e'.source) : e = e' := @@ -940,24 +940,24 @@ theorem eqOnSource_iff (e e' : LocalHomeomorph α β) : Iff.rfl #align local_homeomorph.eq_on_source_iff LocalHomeomorph.eqOnSource_iff -/-- `EqOnSource` is an equivalence relation -/ +/-- `EqOnSource` is an equivalence relation. -/ instance eqOnSourceSetoid : Setoid (LocalHomeomorph α β) := { LocalEquiv.eqOnSourceSetoid.comap toLocalEquiv with r := EqOnSource } theorem eqOnSource_refl : e ≈ e := Setoid.refl _ #align local_homeomorph.eq_on_source_refl LocalHomeomorph.eqOnSource_refl -/-- If two local homeomorphisms are equivalent, so are their inverses -/ +/-- If two local homeomorphisms are equivalent, so are their inverses. -/ theorem EqOnSource.symm' {e e' : LocalHomeomorph α β} (h : e ≈ e') : e.symm ≈ e'.symm := LocalEquiv.EqOnSource.symm' h #align local_homeomorph.eq_on_source.symm' LocalHomeomorph.EqOnSource.symm' -/-- Two equivalent local homeomorphisms have the same source -/ +/-- Two equivalent local homeomorphisms have the same source. -/ theorem EqOnSource.source_eq {e e' : LocalHomeomorph α β} (h : e ≈ e') : e.source = e'.source := h.1 #align local_homeomorph.eq_on_source.source_eq LocalHomeomorph.EqOnSource.source_eq -/-- Two equivalent local homeomorphisms have the same target -/ +/-- Two equivalent local homeomorphisms have the same target. -/ theorem EqOnSource.target_eq {e e' : LocalHomeomorph α β} (h : e ≈ e') : e.target = e'.target := h.symm'.1 #align local_homeomorph.eq_on_source.target_eq LocalHomeomorph.EqOnSource.target_eq @@ -985,6 +985,7 @@ theorem EqOnSource.restr {e e' : LocalHomeomorph α β} (he : e ≈ e') (s : Set LocalEquiv.EqOnSource.restr he _ #align local_homeomorph.eq_on_source.restr LocalHomeomorph.EqOnSource.restr +/- Two equivalent local homeomorphisms are equal when the source and target are `univ`. -/ theorem Set.EqOn.restr_eqOn_source {e e' : LocalHomeomorph α β} (h : EqOn e e' (e.source ∩ e'.source)) : e.restr e'.source ≈ e'.restr e.source := by constructor @@ -1005,10 +1006,10 @@ theorem trans_symm_self : e.symm.trans e ≈ LocalHomeomorph.ofSet e.target e.op e.symm.trans_self_symm #align local_homeomorph.trans_symm_self LocalHomeomorph.trans_symm_self -theorem eq_of_eq_on_source_univ {e e' : LocalHomeomorph α β} (h : e ≈ e') (s : e.source = univ) +theorem eq_of_eqOnSource_univ {e e' : LocalHomeomorph α β} (h : e ≈ e') (s : e.source = univ) (t : e.target = univ) : e = e' := - toLocalEquiv_injective <| LocalEquiv.eq_of_eq_on_source_univ _ _ h s t -#align local_homeomorph.eq_of_eq_on_source_univ LocalHomeomorph.eq_of_eq_on_source_univ + toLocalEquiv_injective <| LocalEquiv.eq_of_eqOnSource_univ _ _ h s t +#align local_homeomorph.eq_of_eq_on_source_univ LocalHomeomorph.eq_of_eqOnSource_univ section Prod diff --git a/Mathlib/Topology/MetricSpace/Antilipschitz.lean b/Mathlib/Topology/MetricSpace/Antilipschitz.lean index 95e859d0c914a..c8acafc12eb0e 100644 --- a/Mathlib/Topology/MetricSpace/Antilipschitz.lean +++ b/Mathlib/Topology/MetricSpace/Antilipschitz.lean @@ -6,7 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Topology.MetricSpace.Lipschitz import Mathlib.Topology.UniformSpace.CompleteSeparated -#align_import topology.metric_space.antilipschitz from "leanprover-community/mathlib"@"97f079b7e89566de3a1143f887713667328c38ba" +#align_import topology.metric_space.antilipschitz from "leanprover-community/mathlib"@"c8f305514e0d47dfaa710f5a52f0d21b588e6328" /-! # Antilipschitz functions @@ -218,7 +218,8 @@ namespace AntilipschitzWith open Metric -variable [PseudoMetricSpace α] [PseudoMetricSpace β] {K : ℝ≥0} {f : α → β} +variable [PseudoMetricSpace α] [PseudoMetricSpace β] [PseudoMetricSpace γ] +variable {K : ℝ≥0} {f : α → β} theorem isBounded_preimage (hf : AntilipschitzWith K f) {s : Set β} (hs : IsBounded s) : IsBounded (f ⁻¹' s) := @@ -243,6 +244,25 @@ protected theorem properSpace {α : Type*} [MetricSpace α] {K : ℝ≥0} {f : exact (hf.image_preimage _).symm #align antilipschitz_with.proper_space AntilipschitzWith.properSpace +theorem isBounded_of_image2_left (f : α → β → γ) {K₁ : ℝ≥0} + (hf : ∀ b, AntilipschitzWith K₁ fun a => f a b) {s : Set α} {t : Set β} + (hst : IsBounded (Set.image2 f s t)) : IsBounded s ∨ IsBounded t := by + contrapose! hst + obtain ⟨b, hb⟩ : t.Nonempty := nonempty_of_not_isBounded hst.2 + have : ¬IsBounded (Set.image2 f s {b}) := by + intro h + apply hst.1 + rw [Set.image2_singleton_right] at h + replace h := (hf b).isBounded_preimage h + refine' h.subset (subset_preimage_image _ _) + exact mt (IsBounded.subset · (image2_subset subset_rfl (singleton_subset_iff.mpr hb))) this +#align antilipschitz_with.bounded_of_image2_left AntilipschitzWith.isBounded_of_image2_left + +theorem isBounded_of_image2_right {f : α → β → γ} {K₂ : ℝ≥0} (hf : ∀ a, AntilipschitzWith K₂ (f a)) + {s : Set α} {t : Set β} (hst : IsBounded (Set.image2 f s t)) : IsBounded s ∨ IsBounded t := + Or.symm <| isBounded_of_image2_left (flip f) hf <| image2_swap f s t ▸ hst +#align antilipschitz_with.bounded_of_image2_right AntilipschitzWith.isBounded_of_image2_right + end AntilipschitzWith theorem LipschitzWith.to_rightInverse [PseudoEMetricSpace α] [PseudoEMetricSpace β] {K : ℝ≥0} diff --git a/Mathlib/Topology/MetricSpace/Basic.lean b/Mathlib/Topology/MetricSpace/Basic.lean index ef10875e3837d..e85dec3035d87 100644 --- a/Mathlib/Topology/MetricSpace/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Basic.lean @@ -3,2891 +3,38 @@ Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel -/ -import Mathlib.Tactic.Positivity -import Mathlib.Topology.Algebra.Order.Compact -import Mathlib.Topology.EMetricSpace.Basic -import Mathlib.Topology.Bornology.Constructions +import Mathlib.Topology.MetricSpace.PseudoMetric -#align_import topology.metric_space.basic from "leanprover-community/mathlib"@"8047de4d911cdef39c2d646165eea972f7f9f539" +#align_import topology.metric_space.basic from "leanprover-community/mathlib"@"c8f305514e0d47dfaa710f5a52f0d21b588e6328" /-! # Metric spaces -This file defines metric spaces. Many definitions and theorems expected -on metric spaces are already introduced on uniform spaces and topological spaces. -For example: open and closed sets, compactness, completeness, continuity and uniform continuity +This file defines metric spaces and shows some of their basic properties. -## Main definitions - -* `Dist α`: Endows a space `α` with a function `dist a b`. -* `PseudoMetricSpace α`: A space endowed with a distance function, which can - be zero even if the two elements are non-equal. -* `Metric.ball x ε`: The set of all points `y` with `dist y x < ε`. -* `Metric.Bounded s`: Whether a subset of a `PseudoMetricSpace` is bounded. -* `MetricSpace α`: A `PseudoMetricSpace` with the guarantee `dist x y = 0 → x = y`. - -Additional useful definitions: - -* `nndist a b`: `dist` as a function to the non-negative reals. -* `Metric.closedBall x ε`: The set of all points `y` with `dist y x ≤ ε`. -* `Metric.sphere x ε`: The set of all points `y` with `dist y x = ε`. -* `ProperSpace α`: A `PseudoMetricSpace` where all closed balls are compact. -* `Metric.diam s` : The `iSup` of the distances of members of `s`. - Defined in terms of `EMetric.diam`, for better handling of the case when it should be infinite. +Many definitions and theorems expected on metric spaces are already introduced on uniform spaces and +topological spaces. This includes open and closed sets, compactness, completeness, continuity +and uniform continuity. TODO (anyone): Add "Main results" section. ## Implementation notes - -Since a lot of elementary properties don't require `eq_of_dist_eq_zero` we start setting up the -theory of `PseudoMetricSpace`, where we don't require `dist x y = 0 → x = y` and we specialize -to `MetricSpace` at the end. +A lot of elementary properties don't require `eq_of_dist_eq_zero`, hence are stated and proven +for `PseudoMetricSpace`s in `PseudoMetric.lean`. ## Tags metric, pseudo_metric, dist -/ - -open Set Filter TopologicalSpace Bornology -open scoped BigOperators ENNReal NNReal Uniformity Topology +open Set Filter Bornology +open scoped NNReal Uniformity universe u v w variable {α : Type u} {β : Type v} {X ι : Type*} - -/-- Construct a uniform structure from a distance function and metric space axioms -/ -def UniformSpace.ofDist (dist : α → α → ℝ) (dist_self : ∀ x : α, dist x x = 0) - (dist_comm : ∀ x y : α, dist x y = dist y x) - (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) : UniformSpace α := - .ofFun dist dist_self dist_comm dist_triangle fun ε ε0 => - ⟨ε / 2, half_pos ε0, fun _x hx _y hy => add_halves ε ▸ add_lt_add hx hy⟩ -#align uniform_space_of_dist UniformSpace.ofDist - --- porting note: dropped the `dist_self` argument -/-- Construct a bornology from a distance function and metric space axioms. -/ -@[reducible] -def Bornology.ofDist {α : Type*} (dist : α → α → ℝ) (dist_comm : ∀ x y, dist x y = dist y x) - (dist_triangle : ∀ x y z, dist x z ≤ dist x y + dist y z) : Bornology α := - Bornology.ofBounded { s : Set α | ∃ C, ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → dist x y ≤ C } - ⟨0, fun x hx y => hx.elim⟩ (fun s ⟨c, hc⟩ t h => ⟨c, fun x hx y hy => hc (h hx) (h hy)⟩) - (fun s hs t ht => by - rcases s.eq_empty_or_nonempty with rfl | ⟨x, hx⟩ - · rwa [empty_union] - rcases t.eq_empty_or_nonempty with rfl | ⟨y, hy⟩ - · rwa [union_empty] - rsuffices ⟨C, hC⟩ : ∃ C, ∀ z ∈ s ∪ t, dist x z ≤ C - · refine ⟨C + C, fun a ha b hb => (dist_triangle a x b).trans ?_⟩ - simpa only [dist_comm] using add_le_add (hC _ ha) (hC _ hb) - rcases hs with ⟨Cs, hs⟩; rcases ht with ⟨Ct, ht⟩ - refine ⟨max Cs (dist x y + Ct), fun z hz => hz.elim - (fun hz => (hs hx hz).trans (le_max_left _ _)) - (fun hz => (dist_triangle x y z).trans <| - (add_le_add le_rfl (ht hy hz)).trans (le_max_right _ _))⟩) - fun z => ⟨dist z z, forall_eq.2 <| forall_eq.2 le_rfl⟩ -#align bornology.of_dist Bornology.ofDistₓ - -/-- The distance function (given an ambient metric space on `α`), which returns - a nonnegative real number `dist x y` given `x y : α`. -/ -@[ext] -class Dist (α : Type*) where - dist : α → α → ℝ -#align has_dist Dist - -export Dist (dist) - --- the uniform structure and the emetric space structure are embedded in the metric space structure --- to avoid instance diamond issues. See Note [forgetful inheritance]. -/-- This is an internal lemma used inside the default of `PseudoMetricSpace.edist`. -/ -private theorem dist_nonneg' {α} {x y : α} (dist : α → α → ℝ) - (dist_self : ∀ x : α, dist x x = 0) (dist_comm : ∀ x y : α, dist x y = dist y x) - (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) : 0 ≤ dist x y := - have : 0 ≤ 2 * dist x y := - calc 0 = dist x x := (dist_self _).symm - _ ≤ dist x y + dist y x := dist_triangle _ _ _ - _ = 2 * dist x y := by rw [two_mul, dist_comm] - nonneg_of_mul_nonneg_right this two_pos - -#noalign pseudo_metric_space.edist_dist_tac -- porting note: todo: restore - -/-- Pseudo metric and Metric spaces - -A pseudo metric space is endowed with a distance for which the requirement `d(x,y)=0 → x = y` might -not hold. A metric space is a pseudo metric space such that `d(x,y)=0 → x = y`. -Each pseudo metric space induces a canonical `UniformSpace` and hence a canonical -`TopologicalSpace` This is enforced in the type class definition, by extending the `UniformSpace` -structure. When instantiating a `PseudoMetricSpace` structure, the uniformity fields are not -necessary, they will be filled in by default. In the same way, each (pseudo) metric space induces a -(pseudo) emetric space structure. It is included in the structure, but filled in by default. --/ -class PseudoMetricSpace (α : Type u) extends Dist α : Type u where - dist_self : ∀ x : α, dist x x = 0 - dist_comm : ∀ x y : α, dist x y = dist y x - dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z - edist : α → α → ℝ≥0∞ := fun x y => ENNReal.some ⟨dist x y, dist_nonneg' _ ‹_› ‹_› ‹_›⟩ - edist_dist : ∀ x y : α, edist x y = ENNReal.ofReal (dist x y) -- porting note: todo: add := by _ - toUniformSpace : UniformSpace α := .ofDist dist dist_self dist_comm dist_triangle - uniformity_dist : 𝓤 α = ⨅ ε > 0, 𝓟 { p : α × α | dist p.1 p.2 < ε } := by intros; rfl - toBornology : Bornology α := Bornology.ofDist dist dist_comm dist_triangle - cobounded_sets : (Bornology.cobounded α).sets = - { s | ∃ C : ℝ, ∀ x ∈ sᶜ, ∀ y ∈ sᶜ, dist x y ≤ C } := by intros; rfl -#align pseudo_metric_space PseudoMetricSpace - -/-- Two pseudo metric space structures with the same distance function coincide. -/ -@[ext] -theorem PseudoMetricSpace.ext {α : Type*} {m m' : PseudoMetricSpace α} - (h : m.toDist = m'.toDist) : m = m' := by - cases' m with d _ _ _ ed hed U hU B hB - cases' m' with d' _ _ _ ed' hed' U' hU' B' hB' - obtain rfl : d = d' := h - congr - · ext x y : 2 - rw [hed, hed'] - · exact UniformSpace.ext (hU.trans hU'.symm) - · ext : 2 - rw [← Filter.mem_sets, ← Filter.mem_sets, hB, hB'] -#align pseudo_metric_space.ext PseudoMetricSpace.ext - variable [PseudoMetricSpace α] -attribute [instance] PseudoMetricSpace.toUniformSpace PseudoMetricSpace.toBornology - --- see Note [lower instance priority] -instance (priority := 200) PseudoMetricSpace.toEDist : EDist α := - ⟨PseudoMetricSpace.edist⟩ -#align pseudo_metric_space.to_has_edist PseudoMetricSpace.toEDist - -/-- Construct a pseudo-metric space structure whose underlying topological space structure -(definitionally) agrees which a pre-existing topology which is compatible with a given distance -function. -/ -def PseudoMetricSpace.ofDistTopology {α : Type u} [TopologicalSpace α] (dist : α → α → ℝ) - (dist_self : ∀ x : α, dist x x = 0) (dist_comm : ∀ x y : α, dist x y = dist y x) - (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) - (H : ∀ s : Set α, IsOpen s ↔ ∀ x ∈ s, ∃ ε > 0, ∀ y, dist x y < ε → y ∈ s) : - PseudoMetricSpace α := - { dist := dist - dist_self := dist_self - dist_comm := dist_comm - dist_triangle := dist_triangle - edist_dist := fun x y => by exact ENNReal.coe_nnreal_eq _ - toUniformSpace := - { toCore := (UniformSpace.ofDist dist dist_self dist_comm dist_triangle).toCore - isOpen_uniformity := fun s => (H s).trans <| forall₂_congr fun x _ => - ((UniformSpace.hasBasis_ofFun (exists_gt (0 : ℝ)) dist _ _ _ _).comap - (Prod.mk x)).mem_iff.symm.trans mem_comap_prod_mk } - uniformity_dist := rfl - toBornology := Bornology.ofDist dist dist_comm dist_triangle - cobounded_sets := rfl } -#align pseudo_metric_space.of_dist_topology PseudoMetricSpace.ofDistTopology - -@[simp] -theorem dist_self (x : α) : dist x x = 0 := - PseudoMetricSpace.dist_self x -#align dist_self dist_self - -theorem dist_comm (x y : α) : dist x y = dist y x := - PseudoMetricSpace.dist_comm x y -#align dist_comm dist_comm - -theorem edist_dist (x y : α) : edist x y = ENNReal.ofReal (dist x y) := - PseudoMetricSpace.edist_dist x y -#align edist_dist edist_dist - -theorem dist_triangle (x y z : α) : dist x z ≤ dist x y + dist y z := - PseudoMetricSpace.dist_triangle x y z -#align dist_triangle dist_triangle - -theorem dist_triangle_left (x y z : α) : dist x y ≤ dist z x + dist z y := by - rw [dist_comm z]; apply dist_triangle -#align dist_triangle_left dist_triangle_left - -theorem dist_triangle_right (x y z : α) : dist x y ≤ dist x z + dist y z := by - rw [dist_comm y]; apply dist_triangle -#align dist_triangle_right dist_triangle_right - -theorem dist_triangle4 (x y z w : α) : dist x w ≤ dist x y + dist y z + dist z w := - calc - dist x w ≤ dist x z + dist z w := dist_triangle x z w - _ ≤ dist x y + dist y z + dist z w := add_le_add_right (dist_triangle x y z) _ -#align dist_triangle4 dist_triangle4 - -theorem dist_triangle4_left (x₁ y₁ x₂ y₂ : α) : - dist x₂ y₂ ≤ dist x₁ y₁ + (dist x₁ x₂ + dist y₁ y₂) := by - rw [add_left_comm, dist_comm x₁, ← add_assoc] - apply dist_triangle4 -#align dist_triangle4_left dist_triangle4_left - -theorem dist_triangle4_right (x₁ y₁ x₂ y₂ : α) : - dist x₁ y₁ ≤ dist x₁ x₂ + dist y₁ y₂ + dist x₂ y₂ := by - rw [add_right_comm, dist_comm y₁] - apply dist_triangle4 -#align dist_triangle4_right dist_triangle4_right - -/-- The triangle (polygon) inequality for sequences of points; `Finset.Ico` version. -/ -theorem dist_le_Ico_sum_dist (f : ℕ → α) {m n} (h : m ≤ n) : - dist (f m) (f n) ≤ ∑ i in Finset.Ico m n, dist (f i) (f (i + 1)) := by - induction n, h using Nat.le_induction with - | base => rw [Finset.Ico_self, Finset.sum_empty, dist_self] - | succ n hle ihn => - calc - dist (f m) (f (n + 1)) ≤ dist (f m) (f n) + dist (f n) (f (n + 1)) := dist_triangle _ _ _ - _ ≤ (∑ i in Finset.Ico m n, _) + _ := add_le_add ihn le_rfl - _ = ∑ i in Finset.Ico m (n + 1), _ := by - { rw [Nat.Ico_succ_right_eq_insert_Ico hle, Finset.sum_insert, add_comm]; simp } -#align dist_le_Ico_sum_dist dist_le_Ico_sum_dist - -/-- The triangle (polygon) inequality for sequences of points; `Finset.range` version. -/ -theorem dist_le_range_sum_dist (f : ℕ → α) (n : ℕ) : - dist (f 0) (f n) ≤ ∑ i in Finset.range n, dist (f i) (f (i + 1)) := - Nat.Ico_zero_eq_range ▸ dist_le_Ico_sum_dist f (Nat.zero_le n) -#align dist_le_range_sum_dist dist_le_range_sum_dist - -/-- A version of `dist_le_Ico_sum_dist` with each intermediate distance replaced -with an upper estimate. -/ -theorem dist_le_Ico_sum_of_dist_le {f : ℕ → α} {m n} (hmn : m ≤ n) {d : ℕ → ℝ} - (hd : ∀ {k}, m ≤ k → k < n → dist (f k) (f (k + 1)) ≤ d k) : - dist (f m) (f n) ≤ ∑ i in Finset.Ico m n, d i := - le_trans (dist_le_Ico_sum_dist f hmn) <| - Finset.sum_le_sum fun _k hk => hd (Finset.mem_Ico.1 hk).1 (Finset.mem_Ico.1 hk).2 -#align dist_le_Ico_sum_of_dist_le dist_le_Ico_sum_of_dist_le - -/-- A version of `dist_le_range_sum_dist` with each intermediate distance replaced -with an upper estimate. -/ -theorem dist_le_range_sum_of_dist_le {f : ℕ → α} (n : ℕ) {d : ℕ → ℝ} - (hd : ∀ {k}, k < n → dist (f k) (f (k + 1)) ≤ d k) : - dist (f 0) (f n) ≤ ∑ i in Finset.range n, d i := - Nat.Ico_zero_eq_range ▸ dist_le_Ico_sum_of_dist_le (zero_le n) fun _ => hd -#align dist_le_range_sum_of_dist_le dist_le_range_sum_of_dist_le - -theorem swap_dist : Function.swap (@dist α _) = dist := by funext x y; exact dist_comm _ _ -#align swap_dist swap_dist - -theorem abs_dist_sub_le (x y z : α) : |dist x z - dist y z| ≤ dist x y := - abs_sub_le_iff.2 - ⟨sub_le_iff_le_add.2 (dist_triangle _ _ _), sub_le_iff_le_add.2 (dist_triangle_left _ _ _)⟩ -#align abs_dist_sub_le abs_dist_sub_le - -theorem dist_nonneg {x y : α} : 0 ≤ dist x y := - dist_nonneg' dist dist_self dist_comm dist_triangle -#align dist_nonneg dist_nonneg - -namespace Mathlib.Meta.Positivity - -open Lean Meta Qq Function - -/-- Extension for the `positivity` tactic: distances are nonnegative. -/ -@[positivity Dist.dist _ _] -def evalDist : PositivityExt where eval {_ _} _zα _pα e := do - let .app (.app _ a) b ← whnfR e | throwError "not dist" - let p ← mkAppOptM ``dist_nonneg #[none, none, a, b] - pure (.nonnegative p) - -end Mathlib.Meta.Positivity - -example {x y : α} : 0 ≤ dist x y := by positivity - -@[simp] theorem abs_dist {a b : α} : |dist a b| = dist a b := abs_of_nonneg dist_nonneg -#align abs_dist abs_dist - -/-- A version of `Dist` that takes value in `ℝ≥0`. -/ -class NNDist (α : Type*) where - nndist : α → α → ℝ≥0 -#align has_nndist NNDist - -export NNDist (nndist) - --- see Note [lower instance priority] -/-- Distance as a nonnegative real number. -/ -instance (priority := 100) PseudoMetricSpace.toNNDist : NNDist α := - ⟨fun a b => ⟨dist a b, dist_nonneg⟩⟩ -#align pseudo_metric_space.to_has_nndist PseudoMetricSpace.toNNDist - -/-- Express `dist` in terms of `nndist`-/ -theorem dist_nndist (x y : α) : dist x y = nndist x y := rfl -#align dist_nndist dist_nndist - -@[simp, norm_cast] -theorem coe_nndist (x y : α) : ↑(nndist x y) = dist x y := rfl -#align coe_nndist coe_nndist - -/-- Express `edist` in terms of `nndist`-/ -theorem edist_nndist (x y : α) : edist x y = nndist x y := by - rw [edist_dist, dist_nndist, ENNReal.ofReal_coe_nnreal] -#align edist_nndist edist_nndist - -/-- Express `nndist` in terms of `edist`-/ -theorem nndist_edist (x y : α) : nndist x y = (edist x y).toNNReal := by - simp [edist_nndist] -#align nndist_edist nndist_edist - -@[simp, norm_cast] -theorem coe_nnreal_ennreal_nndist (x y : α) : ↑(nndist x y) = edist x y := - (edist_nndist x y).symm -#align coe_nnreal_ennreal_nndist coe_nnreal_ennreal_nndist - -@[simp, norm_cast] -theorem edist_lt_coe {x y : α} {c : ℝ≥0} : edist x y < c ↔ nndist x y < c := by - rw [edist_nndist, ENNReal.coe_lt_coe] -#align edist_lt_coe edist_lt_coe - -@[simp, norm_cast] -theorem edist_le_coe {x y : α} {c : ℝ≥0} : edist x y ≤ c ↔ nndist x y ≤ c := by - rw [edist_nndist, ENNReal.coe_le_coe] -#align edist_le_coe edist_le_coe - -/-- In a pseudometric space, the extended distance is always finite-/ -theorem edist_lt_top {α : Type*} [PseudoMetricSpace α] (x y : α) : edist x y < ⊤ := - (edist_dist x y).symm ▸ ENNReal.ofReal_lt_top -#align edist_lt_top edist_lt_top - -/-- In a pseudometric space, the extended distance is always finite-/ -theorem edist_ne_top (x y : α) : edist x y ≠ ⊤ := - (edist_lt_top x y).ne -#align edist_ne_top edist_ne_top - -/-- `nndist x x` vanishes-/ -@[simp] -theorem nndist_self (a : α) : nndist a a = 0 := - (NNReal.coe_eq_zero _).1 (dist_self a) -#align nndist_self nndist_self - --- porting note: `dist_nndist` and `coe_nndist` moved up - -@[simp, norm_cast] -theorem dist_lt_coe {x y : α} {c : ℝ≥0} : dist x y < c ↔ nndist x y < c := - Iff.rfl -#align dist_lt_coe dist_lt_coe - -@[simp, norm_cast] -theorem dist_le_coe {x y : α} {c : ℝ≥0} : dist x y ≤ c ↔ nndist x y ≤ c := - Iff.rfl -#align dist_le_coe dist_le_coe - -@[simp] -theorem edist_lt_ofReal {x y : α} {r : ℝ} : edist x y < ENNReal.ofReal r ↔ dist x y < r := by - rw [edist_dist, ENNReal.ofReal_lt_ofReal_iff_of_nonneg dist_nonneg] -#align edist_lt_of_real edist_lt_ofReal - -@[simp] -theorem edist_le_ofReal {x y : α} {r : ℝ} (hr : 0 ≤ r) : - edist x y ≤ ENNReal.ofReal r ↔ dist x y ≤ r := by - rw [edist_dist, ENNReal.ofReal_le_ofReal_iff hr] -#align edist_le_of_real edist_le_ofReal - -/-- Express `nndist` in terms of `dist`-/ -theorem nndist_dist (x y : α) : nndist x y = Real.toNNReal (dist x y) := by - rw [dist_nndist, Real.toNNReal_coe] -#align nndist_dist nndist_dist - -theorem nndist_comm (x y : α) : nndist x y = nndist y x := NNReal.eq <| dist_comm x y -#align nndist_comm nndist_comm - -/-- Triangle inequality for the nonnegative distance-/ -theorem nndist_triangle (x y z : α) : nndist x z ≤ nndist x y + nndist y z := - dist_triangle _ _ _ -#align nndist_triangle nndist_triangle - -theorem nndist_triangle_left (x y z : α) : nndist x y ≤ nndist z x + nndist z y := - dist_triangle_left _ _ _ -#align nndist_triangle_left nndist_triangle_left - -theorem nndist_triangle_right (x y z : α) : nndist x y ≤ nndist x z + nndist y z := - dist_triangle_right _ _ _ -#align nndist_triangle_right nndist_triangle_right - -/-- Express `dist` in terms of `edist`-/ -theorem dist_edist (x y : α) : dist x y = (edist x y).toReal := by - rw [edist_dist, ENNReal.toReal_ofReal dist_nonneg] -#align dist_edist dist_edist - -namespace Metric - --- instantiate pseudometric space as a topology -variable {x y z : α} {δ ε ε₁ ε₂ : ℝ} {s : Set α} - -/-- `ball x ε` is the set of all points `y` with `dist y x < ε` -/ -def ball (x : α) (ε : ℝ) : Set α := - { y | dist y x < ε } -#align metric.ball Metric.ball - -@[simp] -theorem mem_ball : y ∈ ball x ε ↔ dist y x < ε := - Iff.rfl -#align metric.mem_ball Metric.mem_ball - -theorem mem_ball' : y ∈ ball x ε ↔ dist x y < ε := by rw [dist_comm, mem_ball] -#align metric.mem_ball' Metric.mem_ball' - -theorem pos_of_mem_ball (hy : y ∈ ball x ε) : 0 < ε := - dist_nonneg.trans_lt hy -#align metric.pos_of_mem_ball Metric.pos_of_mem_ball - -theorem mem_ball_self (h : 0 < ε) : x ∈ ball x ε := by - rwa [mem_ball, dist_self] -#align metric.mem_ball_self Metric.mem_ball_self - -@[simp] -theorem nonempty_ball : (ball x ε).Nonempty ↔ 0 < ε := - ⟨fun ⟨_x, hx⟩ => pos_of_mem_ball hx, fun h => ⟨x, mem_ball_self h⟩⟩ -#align metric.nonempty_ball Metric.nonempty_ball - -@[simp] -theorem ball_eq_empty : ball x ε = ∅ ↔ ε ≤ 0 := by - rw [← not_nonempty_iff_eq_empty, nonempty_ball, not_lt] -#align metric.ball_eq_empty Metric.ball_eq_empty - -@[simp] -theorem ball_zero : ball x 0 = ∅ := by rw [ball_eq_empty] -#align metric.ball_zero Metric.ball_zero - -/-- If a point belongs to an open ball, then there is a strictly smaller radius whose ball also -contains it. - -See also `exists_lt_subset_ball`. -/ -theorem exists_lt_mem_ball_of_mem_ball (h : x ∈ ball y ε) : ∃ ε' < ε, x ∈ ball y ε' := by - simp only [mem_ball] at h ⊢ - exact ⟨(dist x y + ε) / 2, by linarith, by linarith⟩ -#align metric.exists_lt_mem_ball_of_mem_ball Metric.exists_lt_mem_ball_of_mem_ball - -theorem ball_eq_ball (ε : ℝ) (x : α) : - UniformSpace.ball x { p | dist p.2 p.1 < ε } = Metric.ball x ε := - rfl -#align metric.ball_eq_ball Metric.ball_eq_ball - -theorem ball_eq_ball' (ε : ℝ) (x : α) : - UniformSpace.ball x { p | dist p.1 p.2 < ε } = Metric.ball x ε := by - ext - simp [dist_comm, UniformSpace.ball] -#align metric.ball_eq_ball' Metric.ball_eq_ball' - -@[simp] -theorem iUnion_ball_nat (x : α) : ⋃ n : ℕ, ball x n = univ := - iUnion_eq_univ_iff.2 fun y => exists_nat_gt (dist y x) -#align metric.Union_ball_nat Metric.iUnion_ball_nat - -@[simp] -theorem iUnion_ball_nat_succ (x : α) : ⋃ n : ℕ, ball x (n + 1) = univ := - iUnion_eq_univ_iff.2 fun y => (exists_nat_gt (dist y x)).imp fun _ h => h.trans (lt_add_one _) -#align metric.Union_ball_nat_succ Metric.iUnion_ball_nat_succ - -/-- `closedBall x ε` is the set of all points `y` with `dist y x ≤ ε` -/ -def closedBall (x : α) (ε : ℝ) := - { y | dist y x ≤ ε } -#align metric.closed_ball Metric.closedBall - -@[simp] theorem mem_closedBall : y ∈ closedBall x ε ↔ dist y x ≤ ε := Iff.rfl -#align metric.mem_closed_ball Metric.mem_closedBall - -theorem mem_closedBall' : y ∈ closedBall x ε ↔ dist x y ≤ ε := by rw [dist_comm, mem_closedBall] -#align metric.mem_closed_ball' Metric.mem_closedBall' - -/-- `sphere x ε` is the set of all points `y` with `dist y x = ε` -/ -def sphere (x : α) (ε : ℝ) := { y | dist y x = ε } -#align metric.sphere Metric.sphere - -@[simp] theorem mem_sphere : y ∈ sphere x ε ↔ dist y x = ε := Iff.rfl -#align metric.mem_sphere Metric.mem_sphere - -theorem mem_sphere' : y ∈ sphere x ε ↔ dist x y = ε := by rw [dist_comm, mem_sphere] -#align metric.mem_sphere' Metric.mem_sphere' - -theorem ne_of_mem_sphere (h : y ∈ sphere x ε) (hε : ε ≠ 0) : y ≠ x := - ne_of_mem_of_not_mem h <| by simpa using hε.symm -#align metric.ne_of_mem_sphere Metric.ne_of_mem_sphere - -theorem nonneg_of_mem_sphere (hy : y ∈ sphere x ε) : 0 ≤ ε := - dist_nonneg.trans_eq hy -#align metric.nonneg_of_mem_sphere Metric.nonneg_of_mem_sphere - -@[simp] -theorem sphere_eq_empty_of_neg (hε : ε < 0) : sphere x ε = ∅ := - Set.eq_empty_iff_forall_not_mem.mpr fun _y hy => (nonneg_of_mem_sphere hy).not_lt hε -#align metric.sphere_eq_empty_of_neg Metric.sphere_eq_empty_of_neg - -theorem sphere_eq_empty_of_subsingleton [Subsingleton α] (hε : ε ≠ 0) : sphere x ε = ∅ := - Set.eq_empty_iff_forall_not_mem.mpr fun _ h => ne_of_mem_sphere h hε (Subsingleton.elim _ _) -#align metric.sphere_eq_empty_of_subsingleton Metric.sphere_eq_empty_of_subsingleton - -theorem sphere_isEmpty_of_subsingleton [Subsingleton α] (hε : ε ≠ 0) : IsEmpty (sphere x ε) := by - rw [sphere_eq_empty_of_subsingleton hε]; infer_instance -#align metric.sphere_is_empty_of_subsingleton Metric.sphere_isEmpty_of_subsingleton - -theorem mem_closedBall_self (h : 0 ≤ ε) : x ∈ closedBall x ε := by - rwa [mem_closedBall, dist_self] -#align metric.mem_closed_ball_self Metric.mem_closedBall_self - -@[simp] -theorem nonempty_closedBall : (closedBall x ε).Nonempty ↔ 0 ≤ ε := - ⟨fun ⟨_x, hx⟩ => dist_nonneg.trans hx, fun h => ⟨x, mem_closedBall_self h⟩⟩ -#align metric.nonempty_closed_ball Metric.nonempty_closedBall - -@[simp] -theorem closedBall_eq_empty : closedBall x ε = ∅ ↔ ε < 0 := by - rw [← not_nonempty_iff_eq_empty, nonempty_closedBall, not_le] -#align metric.closed_ball_eq_empty Metric.closedBall_eq_empty - -/-- Closed balls and spheres coincide when the radius is non-positive -/ -theorem closedBall_eq_sphere_of_nonpos (hε : ε ≤ 0) : closedBall x ε = sphere x ε := - Set.ext fun _ => (hε.trans dist_nonneg).le_iff_eq -#align metric.closed_ball_eq_sphere_of_nonpos Metric.closedBall_eq_sphere_of_nonpos - -theorem ball_subset_closedBall : ball x ε ⊆ closedBall x ε := fun _y hy => - mem_closedBall.2 (le_of_lt hy) -#align metric.ball_subset_closed_ball Metric.ball_subset_closedBall - -theorem sphere_subset_closedBall : sphere x ε ⊆ closedBall x ε := fun _ => le_of_eq -#align metric.sphere_subset_closed_ball Metric.sphere_subset_closedBall - -theorem closedBall_disjoint_ball (h : δ + ε ≤ dist x y) : Disjoint (closedBall x δ) (ball y ε) := - Set.disjoint_left.mpr fun _a ha1 ha2 => - (h.trans <| dist_triangle_left _ _ _).not_lt <| add_lt_add_of_le_of_lt ha1 ha2 -#align metric.closed_ball_disjoint_ball Metric.closedBall_disjoint_ball - -theorem ball_disjoint_closedBall (h : δ + ε ≤ dist x y) : Disjoint (ball x δ) (closedBall y ε) := - (closedBall_disjoint_ball <| by rwa [add_comm, dist_comm]).symm -#align metric.ball_disjoint_closed_ball Metric.ball_disjoint_closedBall - -theorem ball_disjoint_ball (h : δ + ε ≤ dist x y) : Disjoint (ball x δ) (ball y ε) := - (closedBall_disjoint_ball h).mono_left ball_subset_closedBall -#align metric.ball_disjoint_ball Metric.ball_disjoint_ball - -theorem closedBall_disjoint_closedBall (h : δ + ε < dist x y) : - Disjoint (closedBall x δ) (closedBall y ε) := - Set.disjoint_left.mpr fun _a ha1 ha2 => - h.not_le <| (dist_triangle_left _ _ _).trans <| add_le_add ha1 ha2 -#align metric.closed_ball_disjoint_closed_ball Metric.closedBall_disjoint_closedBall - -theorem sphere_disjoint_ball : Disjoint (sphere x ε) (ball x ε) := - Set.disjoint_left.mpr fun _y hy₁ hy₂ => absurd hy₁ <| ne_of_lt hy₂ -#align metric.sphere_disjoint_ball Metric.sphere_disjoint_ball - -@[simp] -theorem ball_union_sphere : ball x ε ∪ sphere x ε = closedBall x ε := - Set.ext fun _y => (@le_iff_lt_or_eq ℝ _ _ _).symm -#align metric.ball_union_sphere Metric.ball_union_sphere - -@[simp] -theorem sphere_union_ball : sphere x ε ∪ ball x ε = closedBall x ε := by - rw [union_comm, ball_union_sphere] -#align metric.sphere_union_ball Metric.sphere_union_ball - -@[simp] -theorem closedBall_diff_sphere : closedBall x ε \ sphere x ε = ball x ε := by - rw [← ball_union_sphere, Set.union_diff_cancel_right sphere_disjoint_ball.symm.le_bot] -#align metric.closed_ball_diff_sphere Metric.closedBall_diff_sphere - -@[simp] -theorem closedBall_diff_ball : closedBall x ε \ ball x ε = sphere x ε := by - rw [← ball_union_sphere, Set.union_diff_cancel_left sphere_disjoint_ball.symm.le_bot] -#align metric.closed_ball_diff_ball Metric.closedBall_diff_ball - -theorem mem_ball_comm : x ∈ ball y ε ↔ y ∈ ball x ε := by rw [mem_ball', mem_ball] -#align metric.mem_ball_comm Metric.mem_ball_comm - -theorem mem_closedBall_comm : x ∈ closedBall y ε ↔ y ∈ closedBall x ε := by - rw [mem_closedBall', mem_closedBall] -#align metric.mem_closed_ball_comm Metric.mem_closedBall_comm - -theorem mem_sphere_comm : x ∈ sphere y ε ↔ y ∈ sphere x ε := by rw [mem_sphere', mem_sphere] -#align metric.mem_sphere_comm Metric.mem_sphere_comm - -theorem ball_subset_ball (h : ε₁ ≤ ε₂) : ball x ε₁ ⊆ ball x ε₂ := fun _y yx => - lt_of_lt_of_le (mem_ball.1 yx) h -#align metric.ball_subset_ball Metric.ball_subset_ball - -theorem closedBall_eq_bInter_ball : closedBall x ε = ⋂ δ > ε, ball x δ := by - ext y; rw [mem_closedBall, ← forall_lt_iff_le', mem_iInter₂]; rfl -#align metric.closed_ball_eq_bInter_ball Metric.closedBall_eq_bInter_ball - -theorem ball_subset_ball' (h : ε₁ + dist x y ≤ ε₂) : ball x ε₁ ⊆ ball y ε₂ := fun z hz => - calc - dist z y ≤ dist z x + dist x y := dist_triangle _ _ _ - _ < ε₁ + dist x y := add_lt_add_right (mem_ball.1 hz) _ - _ ≤ ε₂ := h -#align metric.ball_subset_ball' Metric.ball_subset_ball' - -theorem closedBall_subset_closedBall (h : ε₁ ≤ ε₂) : closedBall x ε₁ ⊆ closedBall x ε₂ := - fun _y (yx : _ ≤ ε₁) => le_trans yx h -#align metric.closed_ball_subset_closed_ball Metric.closedBall_subset_closedBall - -theorem closedBall_subset_closedBall' (h : ε₁ + dist x y ≤ ε₂) : - closedBall x ε₁ ⊆ closedBall y ε₂ := fun z hz => - calc - dist z y ≤ dist z x + dist x y := dist_triangle _ _ _ - _ ≤ ε₁ + dist x y := add_le_add_right (mem_closedBall.1 hz) _ - _ ≤ ε₂ := h -#align metric.closed_ball_subset_closed_ball' Metric.closedBall_subset_closedBall' - -theorem closedBall_subset_ball (h : ε₁ < ε₂) : closedBall x ε₁ ⊆ ball x ε₂ := - fun y (yh : dist y x ≤ ε₁) => lt_of_le_of_lt yh h -#align metric.closed_ball_subset_ball Metric.closedBall_subset_ball - -theorem closedBall_subset_ball' (h : ε₁ + dist x y < ε₂) : - closedBall x ε₁ ⊆ ball y ε₂ := fun z hz => - calc - dist z y ≤ dist z x + dist x y := dist_triangle _ _ _ - _ ≤ ε₁ + dist x y := add_le_add_right (mem_closedBall.1 hz) _ - _ < ε₂ := h -#align metric.closed_ball_subset_ball' Metric.closedBall_subset_ball' - -theorem dist_le_add_of_nonempty_closedBall_inter_closedBall - (h : (closedBall x ε₁ ∩ closedBall y ε₂).Nonempty) : dist x y ≤ ε₁ + ε₂ := - let ⟨z, hz⟩ := h - calc - dist x y ≤ dist z x + dist z y := dist_triangle_left _ _ _ - _ ≤ ε₁ + ε₂ := add_le_add hz.1 hz.2 -#align metric.dist_le_add_of_nonempty_closed_ball_inter_closed_ball Metric.dist_le_add_of_nonempty_closedBall_inter_closedBall - -theorem dist_lt_add_of_nonempty_closedBall_inter_ball (h : (closedBall x ε₁ ∩ ball y ε₂).Nonempty) : - dist x y < ε₁ + ε₂ := - let ⟨z, hz⟩ := h - calc - dist x y ≤ dist z x + dist z y := dist_triangle_left _ _ _ - _ < ε₁ + ε₂ := add_lt_add_of_le_of_lt hz.1 hz.2 -#align metric.dist_lt_add_of_nonempty_closed_ball_inter_ball Metric.dist_lt_add_of_nonempty_closedBall_inter_ball - -theorem dist_lt_add_of_nonempty_ball_inter_closedBall (h : (ball x ε₁ ∩ closedBall y ε₂).Nonempty) : - dist x y < ε₁ + ε₂ := by - rw [inter_comm] at h - rw [add_comm, dist_comm] - exact dist_lt_add_of_nonempty_closedBall_inter_ball h -#align metric.dist_lt_add_of_nonempty_ball_inter_closed_ball Metric.dist_lt_add_of_nonempty_ball_inter_closedBall - -theorem dist_lt_add_of_nonempty_ball_inter_ball (h : (ball x ε₁ ∩ ball y ε₂).Nonempty) : - dist x y < ε₁ + ε₂ := - dist_lt_add_of_nonempty_closedBall_inter_ball <| - h.mono (inter_subset_inter ball_subset_closedBall Subset.rfl) -#align metric.dist_lt_add_of_nonempty_ball_inter_ball Metric.dist_lt_add_of_nonempty_ball_inter_ball - -@[simp] -theorem iUnion_closedBall_nat (x : α) : ⋃ n : ℕ, closedBall x n = univ := - iUnion_eq_univ_iff.2 fun y => exists_nat_ge (dist y x) -#align metric.Union_closed_ball_nat Metric.iUnion_closedBall_nat - -theorem iUnion_inter_closedBall_nat (s : Set α) (x : α) : ⋃ n : ℕ, s ∩ closedBall x n = s := by - rw [← inter_iUnion, iUnion_closedBall_nat, inter_univ] -#align metric.Union_inter_closed_ball_nat Metric.iUnion_inter_closedBall_nat - -theorem ball_subset (h : dist x y ≤ ε₂ - ε₁) : ball x ε₁ ⊆ ball y ε₂ := fun z zx => by - rw [← add_sub_cancel'_right ε₁ ε₂] - exact lt_of_le_of_lt (dist_triangle z x y) (add_lt_add_of_lt_of_le zx h) -#align metric.ball_subset Metric.ball_subset - -theorem ball_half_subset (y) (h : y ∈ ball x (ε / 2)) : ball y (ε / 2) ⊆ ball x ε := - ball_subset <| by rw [sub_self_div_two]; exact le_of_lt h -#align metric.ball_half_subset Metric.ball_half_subset - -theorem exists_ball_subset_ball (h : y ∈ ball x ε) : ∃ ε' > 0, ball y ε' ⊆ ball x ε := - ⟨_, sub_pos.2 h, ball_subset <| by rw [sub_sub_self]⟩ -#align metric.exists_ball_subset_ball Metric.exists_ball_subset_ball - -/-- If a property holds for all points in closed balls of arbitrarily large radii, then it holds for -all points. -/ -theorem forall_of_forall_mem_closedBall (p : α → Prop) (x : α) - (H : ∃ᶠ R : ℝ in atTop, ∀ y ∈ closedBall x R, p y) (y : α) : p y := by - obtain ⟨R, hR, h⟩ : ∃ R ≥ dist y x, ∀ z : α, z ∈ closedBall x R → p z := - frequently_iff.1 H (Ici_mem_atTop (dist y x)) - exact h _ hR -#align metric.forall_of_forall_mem_closed_ball Metric.forall_of_forall_mem_closedBall - -/-- If a property holds for all points in balls of arbitrarily large radii, then it holds for all -points. -/ -theorem forall_of_forall_mem_ball (p : α → Prop) (x : α) - (H : ∃ᶠ R : ℝ in atTop, ∀ y ∈ ball x R, p y) (y : α) : p y := by - obtain ⟨R, hR, h⟩ : ∃ R > dist y x, ∀ z : α, z ∈ ball x R → p z := - frequently_iff.1 H (Ioi_mem_atTop (dist y x)) - exact h _ hR -#align metric.forall_of_forall_mem_ball Metric.forall_of_forall_mem_ball - -theorem isBounded_iff {s : Set α} : - IsBounded s ↔ ∃ C : ℝ, ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → dist x y ≤ C := by - rw [isBounded_def, ← Filter.mem_sets, @PseudoMetricSpace.cobounded_sets α, mem_setOf_eq, - compl_compl] -#align metric.is_bounded_iff Metric.isBounded_iff - -theorem isBounded_iff_eventually {s : Set α} : - IsBounded s ↔ ∀ᶠ C in atTop, ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → dist x y ≤ C := - isBounded_iff.trans - ⟨fun ⟨C, h⟩ => eventually_atTop.2 ⟨C, fun _C' hC' _x hx _y hy => (h hx hy).trans hC'⟩, - Eventually.exists⟩ -#align metric.is_bounded_iff_eventually Metric.isBounded_iff_eventually - -theorem isBounded_iff_exists_ge {s : Set α} (c : ℝ) : - IsBounded s ↔ ∃ C, c ≤ C ∧ ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → dist x y ≤ C := - ⟨fun h => ((eventually_ge_atTop c).and (isBounded_iff_eventually.1 h)).exists, fun h => - isBounded_iff.2 <| h.imp fun _ => And.right⟩ -#align metric.is_bounded_iff_exists_ge Metric.isBounded_iff_exists_ge - -theorem isBounded_iff_nndist {s : Set α} : - IsBounded s ↔ ∃ C : ℝ≥0, ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → nndist x y ≤ C := by - simp only [isBounded_iff_exists_ge 0, NNReal.exists, ← NNReal.coe_le_coe, ← dist_nndist, - NNReal.coe_mk, exists_prop] -#align metric.is_bounded_iff_nndist Metric.isBounded_iff_nndist - -theorem toUniformSpace_eq : - ‹PseudoMetricSpace α›.toUniformSpace = .ofDist dist dist_self dist_comm dist_triangle := - UniformSpace.ext PseudoMetricSpace.uniformity_dist -#align metric.to_uniform_space_eq Metric.toUniformSpace_eq - -theorem uniformity_basis_dist : - (𝓤 α).HasBasis (fun ε : ℝ => 0 < ε) fun ε => { p : α × α | dist p.1 p.2 < ε } := by - rw [toUniformSpace_eq] - exact UniformSpace.hasBasis_ofFun (exists_gt _) _ _ _ _ _ -#align metric.uniformity_basis_dist Metric.uniformity_basis_dist - -/-- Given `f : β → ℝ`, if `f` sends `{i | p i}` to a set of positive numbers -accumulating to zero, then `f i`-neighborhoods of the diagonal form a basis of `𝓤 α`. - -For specific bases see `uniformity_basis_dist`, `uniformity_basis_dist_inv_nat_succ`, -and `uniformity_basis_dist_inv_nat_pos`. -/ -protected theorem mk_uniformity_basis {β : Type*} {p : β → Prop} {f : β → ℝ} - (hf₀ : ∀ i, p i → 0 < f i) (hf : ∀ ⦃ε⦄, 0 < ε → ∃ i, p i ∧ f i ≤ ε) : - (𝓤 α).HasBasis p fun i => { p : α × α | dist p.1 p.2 < f i } := by - refine' ⟨fun s => uniformity_basis_dist.mem_iff.trans _⟩ - constructor - · rintro ⟨ε, ε₀, hε⟩ - rcases hf ε₀ with ⟨i, hi, H⟩ - exact ⟨i, hi, fun x (hx : _ < _) => hε <| lt_of_lt_of_le hx H⟩ - · exact fun ⟨i, hi, H⟩ => ⟨f i, hf₀ i hi, H⟩ -#align metric.mk_uniformity_basis Metric.mk_uniformity_basis - -theorem uniformity_basis_dist_rat : - (𝓤 α).HasBasis (fun r : ℚ => 0 < r) fun r => { p : α × α | dist p.1 p.2 < r } := - Metric.mk_uniformity_basis (fun _ => Rat.cast_pos.2) fun _ε hε => - let ⟨r, hr0, hrε⟩ := exists_rat_btwn hε - ⟨r, Rat.cast_pos.1 hr0, hrε.le⟩ -#align metric.uniformity_basis_dist_rat Metric.uniformity_basis_dist_rat - -theorem uniformity_basis_dist_inv_nat_succ : - (𝓤 α).HasBasis (fun _ => True) fun n : ℕ => { p : α × α | dist p.1 p.2 < 1 / (↑n + 1) } := - Metric.mk_uniformity_basis (fun n _ => div_pos zero_lt_one <| Nat.cast_add_one_pos n) fun _ε ε0 => - (exists_nat_one_div_lt ε0).imp fun _n hn => ⟨trivial, le_of_lt hn⟩ -#align metric.uniformity_basis_dist_inv_nat_succ Metric.uniformity_basis_dist_inv_nat_succ - -theorem uniformity_basis_dist_inv_nat_pos : - (𝓤 α).HasBasis (fun n : ℕ => 0 < n) fun n : ℕ => { p : α × α | dist p.1 p.2 < 1 / ↑n } := - Metric.mk_uniformity_basis (fun n hn => div_pos zero_lt_one <| Nat.cast_pos.2 hn) fun ε ε0 => - let ⟨n, hn⟩ := exists_nat_one_div_lt ε0 - ⟨n + 1, Nat.succ_pos n, by exact_mod_cast hn.le⟩ -#align metric.uniformity_basis_dist_inv_nat_pos Metric.uniformity_basis_dist_inv_nat_pos - -theorem uniformity_basis_dist_pow {r : ℝ} (h0 : 0 < r) (h1 : r < 1) : - (𝓤 α).HasBasis (fun _ : ℕ => True) fun n : ℕ => { p : α × α | dist p.1 p.2 < r ^ n } := - Metric.mk_uniformity_basis (fun _ _ => pow_pos h0 _) fun _ε ε0 => - let ⟨n, hn⟩ := exists_pow_lt_of_lt_one ε0 h1 - ⟨n, trivial, hn.le⟩ -#align metric.uniformity_basis_dist_pow Metric.uniformity_basis_dist_pow - -theorem uniformity_basis_dist_lt {R : ℝ} (hR : 0 < R) : - (𝓤 α).HasBasis (fun r : ℝ => 0 < r ∧ r < R) fun r => { p : α × α | dist p.1 p.2 < r } := - Metric.mk_uniformity_basis (fun _ => And.left) fun r hr => - ⟨min r (R / 2), ⟨lt_min hr (half_pos hR), min_lt_iff.2 <| Or.inr (half_lt_self hR)⟩, - min_le_left _ _⟩ -#align metric.uniformity_basis_dist_lt Metric.uniformity_basis_dist_lt - -/-- Given `f : β → ℝ`, if `f` sends `{i | p i}` to a set of positive numbers -accumulating to zero, then closed neighborhoods of the diagonal of sizes `{f i | p i}` -form a basis of `𝓤 α`. - -Currently we have only one specific basis `uniformity_basis_dist_le` based on this constructor. -More can be easily added if needed in the future. -/ -protected theorem mk_uniformity_basis_le {β : Type*} {p : β → Prop} {f : β → ℝ} - (hf₀ : ∀ x, p x → 0 < f x) (hf : ∀ ε, 0 < ε → ∃ x, p x ∧ f x ≤ ε) : - (𝓤 α).HasBasis p fun x => { p : α × α | dist p.1 p.2 ≤ f x } := by - refine' ⟨fun s => uniformity_basis_dist.mem_iff.trans _⟩ - constructor - · rintro ⟨ε, ε₀, hε⟩ - rcases exists_between ε₀ with ⟨ε', hε'⟩ - rcases hf ε' hε'.1 with ⟨i, hi, H⟩ - exact ⟨i, hi, fun x (hx : _ ≤ _) => hε <| lt_of_le_of_lt (le_trans hx H) hε'.2⟩ - · exact fun ⟨i, hi, H⟩ => ⟨f i, hf₀ i hi, fun x (hx : _ < _) => H (mem_setOf.2 hx.le)⟩ -#align metric.mk_uniformity_basis_le Metric.mk_uniformity_basis_le - -/-- Constant size closed neighborhoods of the diagonal form a basis -of the uniformity filter. -/ -theorem uniformity_basis_dist_le : - (𝓤 α).HasBasis ((0 : ℝ) < ·) fun ε => { p : α × α | dist p.1 p.2 ≤ ε } := - Metric.mk_uniformity_basis_le (fun _ => id) fun ε ε₀ => ⟨ε, ε₀, le_refl ε⟩ -#align metric.uniformity_basis_dist_le Metric.uniformity_basis_dist_le - -theorem uniformity_basis_dist_le_pow {r : ℝ} (h0 : 0 < r) (h1 : r < 1) : - (𝓤 α).HasBasis (fun _ : ℕ => True) fun n : ℕ => { p : α × α | dist p.1 p.2 ≤ r ^ n } := - Metric.mk_uniformity_basis_le (fun _ _ => pow_pos h0 _) fun _ε ε0 => - let ⟨n, hn⟩ := exists_pow_lt_of_lt_one ε0 h1 - ⟨n, trivial, hn.le⟩ -#align metric.uniformity_basis_dist_le_pow Metric.uniformity_basis_dist_le_pow - -theorem mem_uniformity_dist {s : Set (α × α)} : - s ∈ 𝓤 α ↔ ∃ ε > 0, ∀ {a b : α}, dist a b < ε → (a, b) ∈ s := - uniformity_basis_dist.mem_uniformity_iff -#align metric.mem_uniformity_dist Metric.mem_uniformity_dist - -/-- A constant size neighborhood of the diagonal is an entourage. -/ -theorem dist_mem_uniformity {ε : ℝ} (ε0 : 0 < ε) : { p : α × α | dist p.1 p.2 < ε } ∈ 𝓤 α := - mem_uniformity_dist.2 ⟨ε, ε0, id⟩ -#align metric.dist_mem_uniformity Metric.dist_mem_uniformity - -theorem uniformContinuous_iff [PseudoMetricSpace β] {f : α → β} : - UniformContinuous f ↔ ∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, dist a b < δ → dist (f a) (f b) < ε := - uniformity_basis_dist.uniformContinuous_iff uniformity_basis_dist -#align metric.uniform_continuous_iff Metric.uniformContinuous_iff - -theorem uniformContinuousOn_iff [PseudoMetricSpace β] {f : α → β} {s : Set α} : - UniformContinuousOn f s ↔ - ∀ ε > 0, ∃ δ > 0, ∀ x ∈ s, ∀ y ∈ s, dist x y < δ → dist (f x) (f y) < ε := - Metric.uniformity_basis_dist.uniformContinuousOn_iff Metric.uniformity_basis_dist -#align metric.uniform_continuous_on_iff Metric.uniformContinuousOn_iff - -theorem uniformContinuousOn_iff_le [PseudoMetricSpace β] {f : α → β} {s : Set α} : - UniformContinuousOn f s ↔ - ∀ ε > 0, ∃ δ > 0, ∀ x ∈ s, ∀ y ∈ s, dist x y ≤ δ → dist (f x) (f y) ≤ ε := - Metric.uniformity_basis_dist_le.uniformContinuousOn_iff Metric.uniformity_basis_dist_le -#align metric.uniform_continuous_on_iff_le Metric.uniformContinuousOn_iff_le - -nonrec theorem uniformInducing_iff [PseudoMetricSpace β] {f : α → β} : - UniformInducing f ↔ UniformContinuous f ∧ - ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := - uniformInducing_iff'.trans <| Iff.rfl.and <| - ((uniformity_basis_dist.comap _).le_basis_iff uniformity_basis_dist).trans <| by - simp only [subset_def, Prod.forall, gt_iff_lt, preimage_setOf_eq, Prod_map, mem_setOf] - -nonrec theorem uniformEmbedding_iff [PseudoMetricSpace β] {f : α → β} : - UniformEmbedding f ↔ Function.Injective f ∧ UniformContinuous f ∧ - ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := by - rw [uniformEmbedding_iff, and_comm, uniformInducing_iff] -#align metric.uniform_embedding_iff Metric.uniformEmbedding_iff - -/-- If a map between pseudometric spaces is a uniform embedding then the distance between `f x` -and `f y` is controlled in terms of the distance between `x` and `y`. -/ -theorem controlled_of_uniformEmbedding [PseudoMetricSpace β] {f : α → β} (h : UniformEmbedding f) : - (∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, dist a b < δ → dist (f a) (f b) < ε) ∧ - ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := - ⟨uniformContinuous_iff.1 h.uniformContinuous, (uniformEmbedding_iff.1 h).2.2⟩ -#align metric.controlled_of_uniform_embedding Metric.controlled_of_uniformEmbedding - -theorem totallyBounded_iff {s : Set α} : - TotallyBounded s ↔ ∀ ε > 0, ∃ t : Set α, t.Finite ∧ s ⊆ ⋃ y ∈ t, ball y ε := - uniformity_basis_dist.totallyBounded_iff -#align metric.totally_bounded_iff Metric.totallyBounded_iff - -/-- A pseudometric space is totally bounded if one can reconstruct up to any ε>0 any element of the -space from finitely many data. -/ -theorem totallyBounded_of_finite_discretization {s : Set α} - (H : ∀ ε > (0 : ℝ), - ∃ (β : Type u) (_ : Fintype β) (F : s → β), ∀ x y, F x = F y → dist (x : α) y < ε) : - TotallyBounded s := by - cases' s.eq_empty_or_nonempty with hs hs - · rw [hs] - exact totallyBounded_empty - rcases hs with ⟨x0, hx0⟩ - haveI : Inhabited s := ⟨⟨x0, hx0⟩⟩ - refine' totallyBounded_iff.2 fun ε ε0 => _ - rcases H ε ε0 with ⟨β, fβ, F, hF⟩ - let Finv := Function.invFun F - refine' ⟨range (Subtype.val ∘ Finv), finite_range _, fun x xs => _⟩ - let x' := Finv (F ⟨x, xs⟩) - have : F x' = F ⟨x, xs⟩ := Function.invFun_eq ⟨⟨x, xs⟩, rfl⟩ - simp only [Set.mem_iUnion, Set.mem_range] - exact ⟨_, ⟨F ⟨x, xs⟩, rfl⟩, hF _ _ this.symm⟩ -#align metric.totally_bounded_of_finite_discretization Metric.totallyBounded_of_finite_discretization - -theorem finite_approx_of_totallyBounded {s : Set α} (hs : TotallyBounded s) : - ∀ ε > 0, ∃ t, t ⊆ s ∧ Set.Finite t ∧ s ⊆ ⋃ y ∈ t, ball y ε := by - intro ε ε_pos - rw [totallyBounded_iff_subset] at hs - exact hs _ (dist_mem_uniformity ε_pos) -#align metric.finite_approx_of_totally_bounded Metric.finite_approx_of_totallyBounded - -/-- Expressing uniform convergence using `dist` -/ -theorem tendstoUniformlyOnFilter_iff {F : ι → β → α} {f : β → α} {p : Filter ι} {p' : Filter β} : - TendstoUniformlyOnFilter F f p p' ↔ - ∀ ε > 0, ∀ᶠ n : ι × β in p ×ˢ p', dist (f n.snd) (F n.fst n.snd) < ε := by - refine' ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu => _⟩ - rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ - refine' (H ε εpos).mono fun n hn => hε hn -#align metric.tendsto_uniformly_on_filter_iff Metric.tendstoUniformlyOnFilter_iff - -/-- Expressing locally uniform convergence on a set using `dist`. -/ -theorem tendstoLocallyUniformlyOn_iff [TopologicalSpace β] {F : ι → β → α} {f : β → α} - {p : Filter ι} {s : Set β} : - TendstoLocallyUniformlyOn F f p s ↔ - ∀ ε > 0, ∀ x ∈ s, ∃ t ∈ 𝓝[s] x, ∀ᶠ n in p, ∀ y ∈ t, dist (f y) (F n y) < ε := by - refine' ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu x hx => _⟩ - rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ - rcases H ε εpos x hx with ⟨t, ht, Ht⟩ - exact ⟨t, ht, Ht.mono fun n hs x hx => hε (hs x hx)⟩ -#align metric.tendsto_locally_uniformly_on_iff Metric.tendstoLocallyUniformlyOn_iff - -/-- Expressing uniform convergence on a set using `dist`. -/ -theorem tendstoUniformlyOn_iff {F : ι → β → α} {f : β → α} {p : Filter ι} {s : Set β} : - TendstoUniformlyOn F f p s ↔ ∀ ε > 0, ∀ᶠ n in p, ∀ x ∈ s, dist (f x) (F n x) < ε := by - refine' ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu => _⟩ - rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ - exact (H ε εpos).mono fun n hs x hx => hε (hs x hx) -#align metric.tendsto_uniformly_on_iff Metric.tendstoUniformlyOn_iff - -/-- Expressing locally uniform convergence using `dist`. -/ -theorem tendstoLocallyUniformly_iff [TopologicalSpace β] {F : ι → β → α} {f : β → α} - {p : Filter ι} : - TendstoLocallyUniformly F f p ↔ - ∀ ε > 0, ∀ x : β, ∃ t ∈ 𝓝 x, ∀ᶠ n in p, ∀ y ∈ t, dist (f y) (F n y) < ε := by - simp only [← tendstoLocallyUniformlyOn_univ, tendstoLocallyUniformlyOn_iff, nhdsWithin_univ, - mem_univ, forall_const, exists_prop] -#align metric.tendsto_locally_uniformly_iff Metric.tendstoLocallyUniformly_iff - -/-- Expressing uniform convergence using `dist`. -/ -theorem tendstoUniformly_iff {F : ι → β → α} {f : β → α} {p : Filter ι} : - TendstoUniformly F f p ↔ ∀ ε > 0, ∀ᶠ n in p, ∀ x, dist (f x) (F n x) < ε := by - rw [← tendstoUniformlyOn_univ, tendstoUniformlyOn_iff] - simp -#align metric.tendsto_uniformly_iff Metric.tendstoUniformly_iff - -protected theorem cauchy_iff {f : Filter α} : - Cauchy f ↔ NeBot f ∧ ∀ ε > 0, ∃ t ∈ f, ∀ x ∈ t, ∀ y ∈ t, dist x y < ε := - uniformity_basis_dist.cauchy_iff -#align metric.cauchy_iff Metric.cauchy_iff - -theorem nhds_basis_ball : (𝓝 x).HasBasis (0 < ·) (ball x) := - nhds_basis_uniformity uniformity_basis_dist -#align metric.nhds_basis_ball Metric.nhds_basis_ball - -theorem mem_nhds_iff : s ∈ 𝓝 x ↔ ∃ ε > 0, ball x ε ⊆ s := - nhds_basis_ball.mem_iff -#align metric.mem_nhds_iff Metric.mem_nhds_iff - -theorem eventually_nhds_iff {p : α → Prop} : - (∀ᶠ y in 𝓝 x, p y) ↔ ∃ ε > 0, ∀ ⦃y⦄, dist y x < ε → p y := - mem_nhds_iff -#align metric.eventually_nhds_iff Metric.eventually_nhds_iff - -theorem eventually_nhds_iff_ball {p : α → Prop} : - (∀ᶠ y in 𝓝 x, p y) ↔ ∃ ε > 0, ∀ y ∈ ball x ε, p y := - mem_nhds_iff -#align metric.eventually_nhds_iff_ball Metric.eventually_nhds_iff_ball - -/-- A version of `Filter.eventually_prod_iff` where the first filter consists of neighborhoods -in a pseudo-metric space.-/ -theorem eventually_nhds_prod_iff {f : Filter ι} {x₀ : α} {p : α × ι → Prop} : - (∀ᶠ x in 𝓝 x₀ ×ˢ f, p x) ↔ ∃ ε > (0 : ℝ), ∃ pa : ι → Prop, (∀ᶠ i in f, pa i) ∧ - ∀ {x}, dist x x₀ < ε → ∀ {i}, pa i → p (x, i) := by - refine (nhds_basis_ball.prod f.basis_sets).eventually_iff.trans ?_ - simp only [Prod.exists, forall_prod_set, id, mem_ball, and_assoc, exists_and_left, and_imp] - rfl -#align metric.eventually_nhds_prod_iff Metric.eventually_nhds_prod_iff - -/-- A version of `Filter.eventually_prod_iff` where the second filter consists of neighborhoods -in a pseudo-metric space.-/ -theorem eventually_prod_nhds_iff {f : Filter ι} {x₀ : α} {p : ι × α → Prop} : - (∀ᶠ x in f ×ˢ 𝓝 x₀, p x) ↔ ∃ pa : ι → Prop, (∀ᶠ i in f, pa i) ∧ - ∃ ε > 0, ∀ {i}, pa i → ∀ {x}, dist x x₀ < ε → p (i, x) := by - rw [eventually_swap_iff, Metric.eventually_nhds_prod_iff] - constructor <;> - · rintro ⟨a1, a2, a3, a4, a5⟩ - exact ⟨a3, a4, a1, a2, fun b1 b2 b3 => a5 b3 b1⟩ -#align metric.eventually_prod_nhds_iff Metric.eventually_prod_nhds_iff - -theorem nhds_basis_closedBall : (𝓝 x).HasBasis (fun ε : ℝ => 0 < ε) (closedBall x) := - nhds_basis_uniformity uniformity_basis_dist_le -#align metric.nhds_basis_closed_ball Metric.nhds_basis_closedBall - -theorem nhds_basis_ball_inv_nat_succ : - (𝓝 x).HasBasis (fun _ => True) fun n : ℕ => ball x (1 / (↑n + 1)) := - nhds_basis_uniformity uniformity_basis_dist_inv_nat_succ -#align metric.nhds_basis_ball_inv_nat_succ Metric.nhds_basis_ball_inv_nat_succ - -theorem nhds_basis_ball_inv_nat_pos : - (𝓝 x).HasBasis (fun n => 0 < n) fun n : ℕ => ball x (1 / ↑n) := - nhds_basis_uniformity uniformity_basis_dist_inv_nat_pos -#align metric.nhds_basis_ball_inv_nat_pos Metric.nhds_basis_ball_inv_nat_pos - -theorem nhds_basis_ball_pow {r : ℝ} (h0 : 0 < r) (h1 : r < 1) : - (𝓝 x).HasBasis (fun _ => True) fun n : ℕ => ball x (r ^ n) := - nhds_basis_uniformity (uniformity_basis_dist_pow h0 h1) -#align metric.nhds_basis_ball_pow Metric.nhds_basis_ball_pow - -theorem nhds_basis_closedBall_pow {r : ℝ} (h0 : 0 < r) (h1 : r < 1) : - (𝓝 x).HasBasis (fun _ => True) fun n : ℕ => closedBall x (r ^ n) := - nhds_basis_uniformity (uniformity_basis_dist_le_pow h0 h1) -#align metric.nhds_basis_closed_ball_pow Metric.nhds_basis_closedBall_pow - -theorem isOpen_iff : IsOpen s ↔ ∀ x ∈ s, ∃ ε > 0, ball x ε ⊆ s := by - simp only [isOpen_iff_mem_nhds, mem_nhds_iff] -#align metric.is_open_iff Metric.isOpen_iff - -theorem isOpen_ball : IsOpen (ball x ε) := - isOpen_iff.2 fun _ => exists_ball_subset_ball -#align metric.is_open_ball Metric.isOpen_ball - -theorem ball_mem_nhds (x : α) {ε : ℝ} (ε0 : 0 < ε) : ball x ε ∈ 𝓝 x := - isOpen_ball.mem_nhds (mem_ball_self ε0) -#align metric.ball_mem_nhds Metric.ball_mem_nhds - -theorem closedBall_mem_nhds (x : α) {ε : ℝ} (ε0 : 0 < ε) : closedBall x ε ∈ 𝓝 x := - mem_of_superset (ball_mem_nhds x ε0) ball_subset_closedBall -#align metric.closed_ball_mem_nhds Metric.closedBall_mem_nhds - -theorem closedBall_mem_nhds_of_mem {x c : α} {ε : ℝ} (h : x ∈ ball c ε) : closedBall c ε ∈ 𝓝 x := - mem_of_superset (isOpen_ball.mem_nhds h) ball_subset_closedBall -#align metric.closed_ball_mem_nhds_of_mem Metric.closedBall_mem_nhds_of_mem - -theorem nhdsWithin_basis_ball {s : Set α} : - (𝓝[s] x).HasBasis (fun ε : ℝ => 0 < ε) fun ε => ball x ε ∩ s := - nhdsWithin_hasBasis nhds_basis_ball s -#align metric.nhds_within_basis_ball Metric.nhdsWithin_basis_ball - -theorem mem_nhdsWithin_iff {t : Set α} : s ∈ 𝓝[t] x ↔ ∃ ε > 0, ball x ε ∩ t ⊆ s := - nhdsWithin_basis_ball.mem_iff -#align metric.mem_nhds_within_iff Metric.mem_nhdsWithin_iff - -theorem tendsto_nhdsWithin_nhdsWithin [PseudoMetricSpace β] {t : Set β} {f : α → β} {a b} : - Tendsto f (𝓝[s] a) (𝓝[t] b) ↔ - ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → dist x a < δ → f x ∈ t ∧ dist (f x) b < ε := - (nhdsWithin_basis_ball.tendsto_iff nhdsWithin_basis_ball).trans <| by - simp only [inter_comm _ s, inter_comm _ t, mem_inter_iff, and_imp]; rfl -#align metric.tendsto_nhds_within_nhds_within Metric.tendsto_nhdsWithin_nhdsWithin - -theorem tendsto_nhdsWithin_nhds [PseudoMetricSpace β] {f : α → β} {a b} : - Tendsto f (𝓝[s] a) (𝓝 b) ↔ - ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → dist x a < δ → dist (f x) b < ε := by - rw [← nhdsWithin_univ b, tendsto_nhdsWithin_nhdsWithin] - simp only [mem_univ, true_and_iff] -#align metric.tendsto_nhds_within_nhds Metric.tendsto_nhdsWithin_nhds - -theorem tendsto_nhds_nhds [PseudoMetricSpace β] {f : α → β} {a b} : - Tendsto f (𝓝 a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, dist x a < δ → dist (f x) b < ε := - nhds_basis_ball.tendsto_iff nhds_basis_ball -#align metric.tendsto_nhds_nhds Metric.tendsto_nhds_nhds - -theorem continuousAt_iff [PseudoMetricSpace β] {f : α → β} {a : α} : - ContinuousAt f a ↔ ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, dist x a < δ → dist (f x) (f a) < ε := by - rw [ContinuousAt, tendsto_nhds_nhds] -#align metric.continuous_at_iff Metric.continuousAt_iff - -theorem continuousWithinAt_iff [PseudoMetricSpace β] {f : α → β} {a : α} {s : Set α} : - ContinuousWithinAt f s a ↔ - ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → dist x a < δ → dist (f x) (f a) < ε := - by rw [ContinuousWithinAt, tendsto_nhdsWithin_nhds] -#align metric.continuous_within_at_iff Metric.continuousWithinAt_iff - -theorem continuousOn_iff [PseudoMetricSpace β] {f : α → β} {s : Set α} : - ContinuousOn f s ↔ ∀ b ∈ s, ∀ ε > 0, ∃ δ > 0, ∀ a ∈ s, dist a b < δ → dist (f a) (f b) < ε := by - simp [ContinuousOn, continuousWithinAt_iff] -#align metric.continuous_on_iff Metric.continuousOn_iff - -theorem continuous_iff [PseudoMetricSpace β] {f : α → β} : - Continuous f ↔ ∀ b, ∀ ε > 0, ∃ δ > 0, ∀ a, dist a b < δ → dist (f a) (f b) < ε := - continuous_iff_continuousAt.trans <| forall_congr' fun _ => tendsto_nhds_nhds -#align metric.continuous_iff Metric.continuous_iff - -theorem tendsto_nhds {f : Filter β} {u : β → α} {a : α} : - Tendsto u f (𝓝 a) ↔ ∀ ε > 0, ∀ᶠ x in f, dist (u x) a < ε := - nhds_basis_ball.tendsto_right_iff -#align metric.tendsto_nhds Metric.tendsto_nhds - -theorem continuousAt_iff' [TopologicalSpace β] {f : β → α} {b : β} : - ContinuousAt f b ↔ ∀ ε > 0, ∀ᶠ x in 𝓝 b, dist (f x) (f b) < ε := by - rw [ContinuousAt, tendsto_nhds] -#align metric.continuous_at_iff' Metric.continuousAt_iff' - -theorem continuousWithinAt_iff' [TopologicalSpace β] {f : β → α} {b : β} {s : Set β} : - ContinuousWithinAt f s b ↔ ∀ ε > 0, ∀ᶠ x in 𝓝[s] b, dist (f x) (f b) < ε := by - rw [ContinuousWithinAt, tendsto_nhds] -#align metric.continuous_within_at_iff' Metric.continuousWithinAt_iff' - -theorem continuousOn_iff' [TopologicalSpace β] {f : β → α} {s : Set β} : - ContinuousOn f s ↔ ∀ b ∈ s, ∀ ε > 0, ∀ᶠ x in 𝓝[s] b, dist (f x) (f b) < ε := by - simp [ContinuousOn, continuousWithinAt_iff'] -#align metric.continuous_on_iff' Metric.continuousOn_iff' - -theorem continuous_iff' [TopologicalSpace β] {f : β → α} : - Continuous f ↔ ∀ (a), ∀ ε > 0, ∀ᶠ x in 𝓝 a, dist (f x) (f a) < ε := - continuous_iff_continuousAt.trans <| forall_congr' fun _ => tendsto_nhds -#align metric.continuous_iff' Metric.continuous_iff' - -theorem tendsto_atTop [Nonempty β] [SemilatticeSup β] {u : β → α} {a : α} : - Tendsto u atTop (𝓝 a) ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, dist (u n) a < ε := - (atTop_basis.tendsto_iff nhds_basis_ball).trans <| by - simp only [true_and]; rfl -#align metric.tendsto_at_top Metric.tendsto_atTop - -/-- A variant of `tendsto_atTop` that -uses `∃ N, ∀ n > N, ...` rather than `∃ N, ∀ n ≥ N, ...` --/ -theorem tendsto_atTop' [Nonempty β] [SemilatticeSup β] [NoMaxOrder β] {u : β → α} {a : α} : - Tendsto u atTop (𝓝 a) ↔ ∀ ε > 0, ∃ N, ∀ n > N, dist (u n) a < ε := - (atTop_basis_Ioi.tendsto_iff nhds_basis_ball).trans <| by - simp only [true_and, gt_iff_lt, mem_Ioi, mem_ball] -#align metric.tendsto_at_top' Metric.tendsto_atTop' - -theorem isOpen_singleton_iff {α : Type*} [PseudoMetricSpace α] {x : α} : - IsOpen ({x} : Set α) ↔ ∃ ε > 0, ∀ y, dist y x < ε → y = x := by - simp [isOpen_iff, subset_singleton_iff, mem_ball] -#align metric.is_open_singleton_iff Metric.isOpen_singleton_iff - -/-- Given a point `x` in a discrete subset `s` of a pseudometric space, there is an open ball -centered at `x` and intersecting `s` only at `x`. -/ -theorem exists_ball_inter_eq_singleton_of_mem_discrete [DiscreteTopology s] {x : α} (hx : x ∈ s) : - ∃ ε > 0, Metric.ball x ε ∩ s = {x} := - nhds_basis_ball.exists_inter_eq_singleton_of_mem_discrete hx -#align metric.exists_ball_inter_eq_singleton_of_mem_discrete Metric.exists_ball_inter_eq_singleton_of_mem_discrete - -/-- Given a point `x` in a discrete subset `s` of a pseudometric space, there is a closed ball -of positive radius centered at `x` and intersecting `s` only at `x`. -/ -theorem exists_closedBall_inter_eq_singleton_of_discrete [DiscreteTopology s] {x : α} (hx : x ∈ s) : - ∃ ε > 0, Metric.closedBall x ε ∩ s = {x} := - nhds_basis_closedBall.exists_inter_eq_singleton_of_mem_discrete hx -#align metric.exists_closed_ball_inter_eq_singleton_of_discrete Metric.exists_closedBall_inter_eq_singleton_of_discrete - -theorem _root_.Dense.exists_dist_lt {s : Set α} (hs : Dense s) (x : α) {ε : ℝ} (hε : 0 < ε) : - ∃ y ∈ s, dist x y < ε := by - have : (ball x ε).Nonempty := by simp [hε] - simpa only [mem_ball'] using hs.exists_mem_open isOpen_ball this -#align dense.exists_dist_lt Dense.exists_dist_lt - -nonrec theorem _root_.DenseRange.exists_dist_lt {β : Type*} {f : β → α} (hf : DenseRange f) (x : α) - {ε : ℝ} (hε : 0 < ε) : ∃ y, dist x (f y) < ε := - exists_range_iff.1 (hf.exists_dist_lt x hε) -#align dense_range.exists_dist_lt DenseRange.exists_dist_lt - -end Metric - -open Metric - -/-Instantiate a pseudometric space as a pseudoemetric space. Before we can state the instance, -we need to show that the uniform structure coming from the edistance and the -distance coincide. -/ - --- porting note: new -theorem Metric.uniformity_edist_aux {α} (d : α → α → ℝ≥0) : - ⨅ ε > (0 : ℝ), 𝓟 { p : α × α | ↑(d p.1 p.2) < ε } = - ⨅ ε > (0 : ℝ≥0∞), 𝓟 { p : α × α | ↑(d p.1 p.2) < ε } := by - simp only [le_antisymm_iff, le_iInf_iff, le_principal_iff] - refine ⟨fun ε hε => ?_, fun ε hε => ?_⟩ - · rcases ENNReal.lt_iff_exists_nnreal_btwn.1 hε with ⟨ε', ε'0, ε'ε⟩ - refine mem_iInf_of_mem (ε' : ℝ) (mem_iInf_of_mem (ENNReal.coe_pos.1 ε'0) ?_) - exact fun x hx => lt_trans (ENNReal.coe_lt_coe.2 hx) ε'ε - · lift ε to ℝ≥0 using le_of_lt hε - refine mem_iInf_of_mem (ε : ℝ≥0∞) (mem_iInf_of_mem (ENNReal.coe_pos.2 hε) ?_) - exact fun _ => ENNReal.coe_lt_coe.1 - -theorem Metric.uniformity_edist : 𝓤 α = ⨅ ε > 0, 𝓟 { p : α × α | edist p.1 p.2 < ε } := by - simp only [PseudoMetricSpace.uniformity_dist, dist_nndist, edist_nndist, - Metric.uniformity_edist_aux] -#align metric.uniformity_edist Metric.uniformity_edist - --- see Note [lower instance priority] -/-- A pseudometric space induces a pseudoemetric space -/ -instance (priority := 100) PseudoMetricSpace.toPseudoEMetricSpace : PseudoEMetricSpace α := - { ‹PseudoMetricSpace α› with - edist_self := by simp [edist_dist] - edist_comm := fun _ _ => by simp only [edist_dist, dist_comm] - edist_triangle := fun x y z => by - simp only [edist_dist, ← ENNReal.ofReal_add, dist_nonneg] - rw [ENNReal.ofReal_le_ofReal_iff _] - · exact dist_triangle _ _ _ - · simpa using add_le_add (dist_nonneg : 0 ≤ dist x y) dist_nonneg - uniformity_edist := Metric.uniformity_edist } -#align pseudo_metric_space.to_pseudo_emetric_space PseudoMetricSpace.toPseudoEMetricSpace - -/-- Expressing the uniformity in terms of `edist` -/ -@[deprecated _root_.uniformity_basis_edist] -protected theorem Metric.uniformity_basis_edist : - (𝓤 α).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) fun ε => { p | edist p.1 p.2 < ε } := - uniformity_basis_edist -#align pseudo_metric.uniformity_basis_edist Metric.uniformity_basis_edist - -/-- In a pseudometric space, an open ball of infinite radius is the whole space -/ -theorem Metric.eball_top_eq_univ (x : α) : EMetric.ball x ∞ = Set.univ := - Set.eq_univ_iff_forall.mpr fun y => edist_lt_top y x -#align metric.eball_top_eq_univ Metric.eball_top_eq_univ - -/-- Balls defined using the distance or the edistance coincide -/ -@[simp] -theorem Metric.emetric_ball {x : α} {ε : ℝ} : EMetric.ball x (ENNReal.ofReal ε) = ball x ε := by - ext y - simp only [EMetric.mem_ball, mem_ball, edist_dist] - exact ENNReal.ofReal_lt_ofReal_iff_of_nonneg dist_nonneg -#align metric.emetric_ball Metric.emetric_ball - -/-- Balls defined using the distance or the edistance coincide -/ -@[simp] -theorem Metric.emetric_ball_nnreal {x : α} {ε : ℝ≥0} : EMetric.ball x ε = ball x ε := by - rw [← Metric.emetric_ball] - simp -#align metric.emetric_ball_nnreal Metric.emetric_ball_nnreal - -/-- Closed balls defined using the distance or the edistance coincide -/ -theorem Metric.emetric_closedBall {x : α} {ε : ℝ} (h : 0 ≤ ε) : - EMetric.closedBall x (ENNReal.ofReal ε) = closedBall x ε := by - ext y; simp [edist_le_ofReal h] -#align metric.emetric_closed_ball Metric.emetric_closedBall - -/-- Closed balls defined using the distance or the edistance coincide -/ -@[simp] -theorem Metric.emetric_closedBall_nnreal {x : α} {ε : ℝ≥0} : - EMetric.closedBall x ε = closedBall x ε := by - rw [← Metric.emetric_closedBall ε.coe_nonneg, ENNReal.ofReal_coe_nnreal] -#align metric.emetric_closed_ball_nnreal Metric.emetric_closedBall_nnreal - -@[simp] -theorem Metric.emetric_ball_top (x : α) : EMetric.ball x ⊤ = univ := - eq_univ_of_forall fun _ => edist_lt_top _ _ -#align metric.emetric_ball_top Metric.emetric_ball_top - -theorem Metric.inseparable_iff {x y : α} : Inseparable x y ↔ dist x y = 0 := by - rw [EMetric.inseparable_iff, edist_nndist, dist_nndist, ENNReal.coe_eq_zero, NNReal.coe_eq_zero] -#align metric.inseparable_iff Metric.inseparable_iff - -/-- Build a new pseudometric space from an old one where the bundled uniform structure is provably -(but typically non-definitionaly) equal to some given uniform structure. -See Note [forgetful inheritance]. --/ -@[reducible] -def PseudoMetricSpace.replaceUniformity {α} [U : UniformSpace α] (m : PseudoMetricSpace α) - (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : PseudoMetricSpace α := - { m with - toUniformSpace := U - uniformity_dist := H.trans PseudoMetricSpace.uniformity_dist } -#align pseudo_metric_space.replace_uniformity PseudoMetricSpace.replaceUniformity - -theorem PseudoMetricSpace.replaceUniformity_eq {α} [U : UniformSpace α] (m : PseudoMetricSpace α) - (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : m.replaceUniformity H = m := by - ext - rfl -#align pseudo_metric_space.replace_uniformity_eq PseudoMetricSpace.replaceUniformity_eq - --- ensure that the bornology is unchanged when replacing the uniformity. -example {α} [U : UniformSpace α] (m : PseudoMetricSpace α) - (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : - (PseudoMetricSpace.replaceUniformity m H).toBornology = m.toBornology := rfl - -/-- Build a new pseudo metric space from an old one where the bundled topological structure is -provably (but typically non-definitionaly) equal to some given topological structure. -See Note [forgetful inheritance]. --/ -@[reducible] -def PseudoMetricSpace.replaceTopology {γ} [U : TopologicalSpace γ] (m : PseudoMetricSpace γ) - (H : U = m.toUniformSpace.toTopologicalSpace) : PseudoMetricSpace γ := - @PseudoMetricSpace.replaceUniformity γ (m.toUniformSpace.replaceTopology H) m rfl -#align pseudo_metric_space.replace_topology PseudoMetricSpace.replaceTopology - -theorem PseudoMetricSpace.replaceTopology_eq {γ} [U : TopologicalSpace γ] (m : PseudoMetricSpace γ) - (H : U = m.toUniformSpace.toTopologicalSpace) : m.replaceTopology H = m := by - ext - rfl -#align pseudo_metric_space.replace_topology_eq PseudoMetricSpace.replaceTopology_eq - -/-- One gets a pseudometric space from an emetric space if the edistance -is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the -uniformity are defeq in the pseudometric space and the pseudoemetric space. In this definition, the -distance is given separately, to be able to prescribe some expression which is not defeq to the -push-forward of the edistance to reals. -/ -def PseudoEMetricSpace.toPseudoMetricSpaceOfDist {α : Type u} [e : PseudoEMetricSpace α] - (dist : α → α → ℝ) (edist_ne_top : ∀ x y : α, edist x y ≠ ⊤) - (h : ∀ x y, dist x y = ENNReal.toReal (edist x y)) : PseudoMetricSpace α where - dist := dist - dist_self x := by simp [h] - dist_comm x y := by simp [h, edist_comm] - dist_triangle x y z := by - simp only [h] - exact ENNReal.toReal_le_add (edist_triangle _ _ _) (edist_ne_top _ _) (edist_ne_top _ _) - edist := edist - edist_dist _ _ := by simp only [h, ENNReal.ofReal_toReal (edist_ne_top _ _)] - toUniformSpace := e.toUniformSpace - uniformity_dist := e.uniformity_edist.trans <| by - simpa only [ENNReal.coe_toNNReal (edist_ne_top _ _), h] - using (Metric.uniformity_edist_aux fun x y : α => (edist x y).toNNReal).symm -#align pseudo_emetric_space.to_pseudo_metric_space_of_dist PseudoEMetricSpace.toPseudoMetricSpaceOfDist - -/-- One gets a pseudometric space from an emetric space if the edistance -is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the -uniformity are defeq in the pseudometric space and the emetric space. -/ -@[reducible] -def PseudoEMetricSpace.toPseudoMetricSpace {α : Type u} [PseudoEMetricSpace α] - (h : ∀ x y : α, edist x y ≠ ⊤) : PseudoMetricSpace α := - PseudoEMetricSpace.toPseudoMetricSpaceOfDist (fun x y => ENNReal.toReal (edist x y)) h fun _ _ => - rfl -#align pseudo_emetric_space.to_pseudo_metric_space PseudoEMetricSpace.toPseudoMetricSpace - -/-- Build a new pseudometric space from an old one where the bundled bornology structure is provably -(but typically non-definitionaly) equal to some given bornology structure. -See Note [forgetful inheritance]. --/ -@[reducible] -def PseudoMetricSpace.replaceBornology {α} [B : Bornology α] (m : PseudoMetricSpace α) - (H : ∀ s, @IsBounded _ B s ↔ @IsBounded _ PseudoMetricSpace.toBornology s) : - PseudoMetricSpace α := - { m with - toBornology := B - cobounded_sets := Set.ext <| compl_surjective.forall.2 fun s => - (H s).trans <| by rw [isBounded_iff, mem_setOf_eq, compl_compl] } -#align pseudo_metric_space.replace_bornology PseudoMetricSpace.replaceBornology - -theorem PseudoMetricSpace.replaceBornology_eq {α} [m : PseudoMetricSpace α] [B : Bornology α] - (H : ∀ s, @IsBounded _ B s ↔ @IsBounded _ PseudoMetricSpace.toBornology s) : - PseudoMetricSpace.replaceBornology _ H = m := by - ext - rfl -#align pseudo_metric_space.replace_bornology_eq PseudoMetricSpace.replaceBornology_eq - --- ensure that the uniformity is unchanged when replacing the bornology. -example {α} [B : Bornology α] (m : PseudoMetricSpace α) - (H : ∀ s, @IsBounded _ B s ↔ @IsBounded _ PseudoMetricSpace.toBornology s) : - (PseudoMetricSpace.replaceBornology m H).toUniformSpace = m.toUniformSpace := rfl - -/-- A very useful criterion to show that a space is complete is to show that all sequences -which satisfy a bound of the form `dist (u n) (u m) < B N` for all `n m ≥ N` are -converging. This is often applied for `B N = 2^{-N}`, i.e., with a very fast convergence to -`0`, which makes it possible to use arguments of converging series, while this is impossible -to do in general for arbitrary Cauchy sequences. -/ -theorem Metric.complete_of_convergent_controlled_sequences (B : ℕ → Real) (hB : ∀ n, 0 < B n) - (H : ∀ u : ℕ → α, (∀ N n m : ℕ, N ≤ n → N ≤ m → dist (u n) (u m) < B N) → - ∃ x, Tendsto u atTop (𝓝 x)) : - CompleteSpace α := - UniformSpace.complete_of_convergent_controlled_sequences - (fun n => { p : α × α | dist p.1 p.2 < B n }) (fun n => dist_mem_uniformity <| hB n) H -#align metric.complete_of_convergent_controlled_sequences Metric.complete_of_convergent_controlled_sequences - -theorem Metric.complete_of_cauchySeq_tendsto : - (∀ u : ℕ → α, CauchySeq u → ∃ a, Tendsto u atTop (𝓝 a)) → CompleteSpace α := - EMetric.complete_of_cauchySeq_tendsto -#align metric.complete_of_cauchy_seq_tendsto Metric.complete_of_cauchySeq_tendsto - -section Real - -/-- Instantiate the reals as a pseudometric space. -/ -instance Real.pseudoMetricSpace : PseudoMetricSpace ℝ where - dist x y := |x - y| - dist_self := by simp [abs_zero] - dist_comm x y := abs_sub_comm _ _ - dist_triangle x y z := abs_sub_le _ _ _ - edist_dist := fun x y => by exact ENNReal.coe_nnreal_eq _ -#align real.pseudo_metric_space Real.pseudoMetricSpace - -theorem Real.dist_eq (x y : ℝ) : dist x y = |x - y| := rfl -#align real.dist_eq Real.dist_eq - -theorem Real.nndist_eq (x y : ℝ) : nndist x y = Real.nnabs (x - y) := rfl -#align real.nndist_eq Real.nndist_eq - -theorem Real.nndist_eq' (x y : ℝ) : nndist x y = Real.nnabs (y - x) := - nndist_comm _ _ -#align real.nndist_eq' Real.nndist_eq' - -theorem Real.dist_0_eq_abs (x : ℝ) : dist x 0 = |x| := by simp [Real.dist_eq] -#align real.dist_0_eq_abs Real.dist_0_eq_abs - -theorem Real.dist_left_le_of_mem_uIcc {x y z : ℝ} (h : y ∈ uIcc x z) : dist x y ≤ dist x z := by - simpa only [dist_comm x] using abs_sub_left_of_mem_uIcc h -#align real.dist_left_le_of_mem_uIcc Real.dist_left_le_of_mem_uIcc - -theorem Real.dist_right_le_of_mem_uIcc {x y z : ℝ} (h : y ∈ uIcc x z) : dist y z ≤ dist x z := by - simpa only [dist_comm _ z] using abs_sub_right_of_mem_uIcc h -#align real.dist_right_le_of_mem_uIcc Real.dist_right_le_of_mem_uIcc - -theorem Real.dist_le_of_mem_uIcc {x y x' y' : ℝ} (hx : x ∈ uIcc x' y') (hy : y ∈ uIcc x' y') : - dist x y ≤ dist x' y' := - abs_sub_le_of_uIcc_subset_uIcc <| uIcc_subset_uIcc (by rwa [uIcc_comm]) (by rwa [uIcc_comm]) -#align real.dist_le_of_mem_uIcc Real.dist_le_of_mem_uIcc - -theorem Real.dist_le_of_mem_Icc {x y x' y' : ℝ} (hx : x ∈ Icc x' y') (hy : y ∈ Icc x' y') : - dist x y ≤ y' - x' := by - simpa only [Real.dist_eq, abs_of_nonpos (sub_nonpos.2 <| hx.1.trans hx.2), neg_sub] using - Real.dist_le_of_mem_uIcc (Icc_subset_uIcc hx) (Icc_subset_uIcc hy) -#align real.dist_le_of_mem_Icc Real.dist_le_of_mem_Icc - -theorem Real.dist_le_of_mem_Icc_01 {x y : ℝ} (hx : x ∈ Icc (0 : ℝ) 1) (hy : y ∈ Icc (0 : ℝ) 1) : - dist x y ≤ 1 := by simpa only [sub_zero] using Real.dist_le_of_mem_Icc hx hy -#align real.dist_le_of_mem_Icc_01 Real.dist_le_of_mem_Icc_01 - -instance : OrderTopology ℝ := - orderTopology_of_nhds_abs fun x => by - simp only [nhds_basis_ball.eq_biInf, ball, Real.dist_eq, abs_sub_comm] - -theorem Real.ball_eq_Ioo (x r : ℝ) : ball x r = Ioo (x - r) (x + r) := - Set.ext fun y => by - rw [mem_ball, dist_comm, Real.dist_eq, abs_sub_lt_iff, mem_Ioo, ← sub_lt_iff_lt_add', - sub_lt_comm] -#align real.ball_eq_Ioo Real.ball_eq_Ioo - -theorem Real.closedBall_eq_Icc {x r : ℝ} : closedBall x r = Icc (x - r) (x + r) := by - ext y - rw [mem_closedBall, dist_comm, Real.dist_eq, abs_sub_le_iff, mem_Icc, ← sub_le_iff_le_add', - sub_le_comm] -#align real.closed_ball_eq_Icc Real.closedBall_eq_Icc - -theorem Real.Ioo_eq_ball (x y : ℝ) : Ioo x y = ball ((x + y) / 2) ((y - x) / 2) := by - rw [Real.ball_eq_Ioo, ← sub_div, add_comm, ← sub_add, add_sub_cancel', add_self_div_two, - ← add_div, add_assoc, add_sub_cancel'_right, add_self_div_two] -#align real.Ioo_eq_ball Real.Ioo_eq_ball - -theorem Real.Icc_eq_closedBall (x y : ℝ) : Icc x y = closedBall ((x + y) / 2) ((y - x) / 2) := by - rw [Real.closedBall_eq_Icc, ← sub_div, add_comm, ← sub_add, add_sub_cancel', add_self_div_two, ← - add_div, add_assoc, add_sub_cancel'_right, add_self_div_two] -#align real.Icc_eq_closed_ball Real.Icc_eq_closedBall - -section MetricOrdered - -variable [Preorder α] [CompactIccSpace α] - -theorem totallyBounded_Icc (a b : α) : TotallyBounded (Icc a b) := - isCompact_Icc.totallyBounded -#align totally_bounded_Icc totallyBounded_Icc - -theorem totallyBounded_Ico (a b : α) : TotallyBounded (Ico a b) := - totallyBounded_subset Ico_subset_Icc_self (totallyBounded_Icc a b) -#align totally_bounded_Ico totallyBounded_Ico - -theorem totallyBounded_Ioc (a b : α) : TotallyBounded (Ioc a b) := - totallyBounded_subset Ioc_subset_Icc_self (totallyBounded_Icc a b) -#align totally_bounded_Ioc totallyBounded_Ioc - -theorem totallyBounded_Ioo (a b : α) : TotallyBounded (Ioo a b) := - totallyBounded_subset Ioo_subset_Icc_self (totallyBounded_Icc a b) -#align totally_bounded_Ioo totallyBounded_Ioo - -end MetricOrdered - -/-- Special case of the sandwich theorem; see `tendsto_of_tendsto_of_tendsto_of_le_of_le'` for the -general case. -/ -theorem squeeze_zero' {α} {f g : α → ℝ} {t₀ : Filter α} (hf : ∀ᶠ t in t₀, 0 ≤ f t) - (hft : ∀ᶠ t in t₀, f t ≤ g t) (g0 : Tendsto g t₀ (nhds 0)) : Tendsto f t₀ (𝓝 0) := - tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds g0 hf hft -#align squeeze_zero' squeeze_zero' - -/-- Special case of the sandwich theorem; see `tendsto_of_tendsto_of_tendsto_of_le_of_le` -and `tendsto_of_tendsto_of_tendsto_of_le_of_le'` for the general case. -/ -theorem squeeze_zero {α} {f g : α → ℝ} {t₀ : Filter α} (hf : ∀ t, 0 ≤ f t) (hft : ∀ t, f t ≤ g t) - (g0 : Tendsto g t₀ (𝓝 0)) : Tendsto f t₀ (𝓝 0) := - squeeze_zero' (eventually_of_forall hf) (eventually_of_forall hft) g0 -#align squeeze_zero squeeze_zero - -theorem Metric.uniformity_eq_comap_nhds_zero : - 𝓤 α = comap (fun p : α × α => dist p.1 p.2) (𝓝 (0 : ℝ)) := by - ext s - simp only [mem_uniformity_dist, (nhds_basis_ball.comap _).mem_iff] - simp [subset_def, Real.dist_0_eq_abs] -#align metric.uniformity_eq_comap_nhds_zero Metric.uniformity_eq_comap_nhds_zero - -theorem cauchySeq_iff_tendsto_dist_atTop_0 [Nonempty β] [SemilatticeSup β] {u : β → α} : - CauchySeq u ↔ Tendsto (fun n : β × β => dist (u n.1) (u n.2)) atTop (𝓝 0) := by - rw [cauchySeq_iff_tendsto, Metric.uniformity_eq_comap_nhds_zero, tendsto_comap_iff]; rfl -#align cauchy_seq_iff_tendsto_dist_at_top_0 cauchySeq_iff_tendsto_dist_atTop_0 - -theorem tendsto_uniformity_iff_dist_tendsto_zero {f : ι → α × α} {p : Filter ι} : - Tendsto f p (𝓤 α) ↔ Tendsto (fun x => dist (f x).1 (f x).2) p (𝓝 0) := by - rw [Metric.uniformity_eq_comap_nhds_zero, tendsto_comap_iff]; rfl -#align tendsto_uniformity_iff_dist_tendsto_zero tendsto_uniformity_iff_dist_tendsto_zero - -theorem Filter.Tendsto.congr_dist {f₁ f₂ : ι → α} {p : Filter ι} {a : α} - (h₁ : Tendsto f₁ p (𝓝 a)) (h : Tendsto (fun x => dist (f₁ x) (f₂ x)) p (𝓝 0)) : - Tendsto f₂ p (𝓝 a) := - h₁.congr_uniformity <| tendsto_uniformity_iff_dist_tendsto_zero.2 h -#align filter.tendsto.congr_dist Filter.Tendsto.congr_dist - -alias tendsto_of_tendsto_of_dist := Filter.Tendsto.congr_dist -#align tendsto_of_tendsto_of_dist tendsto_of_tendsto_of_dist - -theorem tendsto_iff_of_dist {f₁ f₂ : ι → α} {p : Filter ι} {a : α} - (h : Tendsto (fun x => dist (f₁ x) (f₂ x)) p (𝓝 0)) : Tendsto f₁ p (𝓝 a) ↔ Tendsto f₂ p (𝓝 a) := - Uniform.tendsto_congr <| tendsto_uniformity_iff_dist_tendsto_zero.2 h -#align tendsto_iff_of_dist tendsto_iff_of_dist - -/-- If `u` is a neighborhood of `x`, then for small enough `r`, the closed ball -`Metric.closedBall x r` is contained in `u`. -/ -theorem eventually_closedBall_subset {x : α} {u : Set α} (hu : u ∈ 𝓝 x) : - ∀ᶠ r in 𝓝 (0 : ℝ), closedBall x r ⊆ u := by - obtain ⟨ε, εpos, hε⟩ : ∃ ε, 0 < ε ∧ closedBall x ε ⊆ u := nhds_basis_closedBall.mem_iff.1 hu - have : Iic ε ∈ 𝓝 (0 : ℝ) := Iic_mem_nhds εpos - filter_upwards [this] with _ hr using Subset.trans (closedBall_subset_closedBall hr) hε -#align eventually_closed_ball_subset eventually_closedBall_subset - -end Real - -section CauchySeq - -variable [Nonempty β] [SemilatticeSup β] - -/-- In a pseudometric space, Cauchy sequences are characterized by the fact that, eventually, -the distance between its elements is arbitrarily small -/ --- porting note: @[nolint ge_or_gt] doesn't exist -theorem Metric.cauchySeq_iff {u : β → α} : - CauchySeq u ↔ ∀ ε > 0, ∃ N, ∀ m ≥ N, ∀ n ≥ N, dist (u m) (u n) < ε := - uniformity_basis_dist.cauchySeq_iff -#align metric.cauchy_seq_iff Metric.cauchySeq_iff - -/-- A variation around the pseudometric characterization of Cauchy sequences -/ -theorem Metric.cauchySeq_iff' {u : β → α} : - CauchySeq u ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, dist (u n) (u N) < ε := - uniformity_basis_dist.cauchySeq_iff' -#align metric.cauchy_seq_iff' Metric.cauchySeq_iff' - --- see Note [nolint_ge] -/-- In a pseudometric space, uniform Cauchy sequences are characterized by the fact that, -eventually, the distance between all its elements is uniformly, arbitrarily small -/ --- porting note: no attr @[nolint ge_or_gt] -theorem Metric.uniformCauchySeqOn_iff {γ : Type*} {F : β → γ → α} {s : Set γ} : - UniformCauchySeqOn F atTop s ↔ ∀ ε > (0 : ℝ), - ∃ N : β, ∀ m ≥ N, ∀ n ≥ N, ∀ x ∈ s, dist (F m x) (F n x) < ε := by - constructor - · intro h ε hε - let u := { a : α × α | dist a.fst a.snd < ε } - have hu : u ∈ 𝓤 α := Metric.mem_uniformity_dist.mpr ⟨ε, hε, by simp⟩ - rw [← @Filter.eventually_atTop_prod_self' _ _ _ fun m => - ∀ x ∈ s, dist (F m.fst x) (F m.snd x) < ε] - specialize h u hu - rw [prod_atTop_atTop_eq] at h - exact h.mono fun n h x hx => h x hx - · intro h u hu - rcases Metric.mem_uniformity_dist.mp hu with ⟨ε, hε, hab⟩ - rcases h ε hε with ⟨N, hN⟩ - rw [prod_atTop_atTop_eq, eventually_atTop] - use (N, N) - intro b hb x hx - rcases hb with ⟨hbl, hbr⟩ - exact hab (hN b.fst hbl.ge b.snd hbr.ge x hx) -#align metric.uniform_cauchy_seq_on_iff Metric.uniformCauchySeqOn_iff - -/-- If the distance between `s n` and `s m`, `n ≤ m` is bounded above by `b n` -and `b` converges to zero, then `s` is a Cauchy sequence. -/ -theorem cauchySeq_of_le_tendsto_0' {s : β → α} (b : β → ℝ) - (h : ∀ n m : β, n ≤ m → dist (s n) (s m) ≤ b n) (h₀ : Tendsto b atTop (𝓝 0)) : CauchySeq s := - Metric.cauchySeq_iff'.2 fun ε ε0 => (h₀.eventually (gt_mem_nhds ε0)).exists.imp fun N hN n hn => - calc dist (s n) (s N) = dist (s N) (s n) := dist_comm _ _ - _ ≤ b N := h _ _ hn - _ < ε := hN -#align cauchy_seq_of_le_tendsto_0' cauchySeq_of_le_tendsto_0' - -/-- If the distance between `s n` and `s m`, `n, m ≥ N` is bounded above by `b N` -and `b` converges to zero, then `s` is a Cauchy sequence. -/ -theorem cauchySeq_of_le_tendsto_0 {s : β → α} (b : β → ℝ) - (h : ∀ n m N : β, N ≤ n → N ≤ m → dist (s n) (s m) ≤ b N) (h₀ : Tendsto b atTop (𝓝 0)) : - CauchySeq s := - cauchySeq_of_le_tendsto_0' b (fun _n _m hnm => h _ _ _ le_rfl hnm) h₀ -#align cauchy_seq_of_le_tendsto_0 cauchySeq_of_le_tendsto_0 - -/-- A Cauchy sequence on the natural numbers is bounded. -/ -theorem cauchySeq_bdd {u : ℕ → α} (hu : CauchySeq u) : ∃ R > 0, ∀ m n, dist (u m) (u n) < R := by - rcases Metric.cauchySeq_iff'.1 hu 1 zero_lt_one with ⟨N, hN⟩ - suffices : ∃ R > 0, ∀ n, dist (u n) (u N) < R - · rcases this with ⟨R, R0, H⟩ - exact ⟨_, add_pos R0 R0, fun m n => - lt_of_le_of_lt (dist_triangle_right _ _ _) (add_lt_add (H m) (H n))⟩ - let R := Finset.sup (Finset.range N) fun n => nndist (u n) (u N) - refine' ⟨↑R + 1, add_pos_of_nonneg_of_pos R.2 zero_lt_one, fun n => _⟩ - cases' le_or_lt N n with h h - · exact lt_of_lt_of_le (hN _ h) (le_add_of_nonneg_left R.2) - · have : _ ≤ R := Finset.le_sup (Finset.mem_range.2 h) - exact lt_of_le_of_lt this (lt_add_of_pos_right _ zero_lt_one) -#align cauchy_seq_bdd cauchySeq_bdd - -/-- Yet another metric characterization of Cauchy sequences on integers. This one is often the -most efficient. -/ -theorem cauchySeq_iff_le_tendsto_0 {s : ℕ → α} : - CauchySeq s ↔ - ∃ b : ℕ → ℝ, - (∀ n, 0 ≤ b n) ∧ - (∀ n m N : ℕ, N ≤ n → N ≤ m → dist (s n) (s m) ≤ b N) ∧ Tendsto b atTop (𝓝 0) := - ⟨fun hs => by - /- `s` is a Cauchy sequence. The sequence `b` will be constructed by taking - the supremum of the distances between `s n` and `s m` for `n m ≥ N`. - First, we prove that all these distances are bounded, as otherwise the Sup - would not make sense. -/ - let S N := (fun p : ℕ × ℕ => dist (s p.1) (s p.2)) '' { p | p.1 ≥ N ∧ p.2 ≥ N } - have hS : ∀ N, ∃ x, ∀ y ∈ S N, y ≤ x := by - rcases cauchySeq_bdd hs with ⟨R, -, hR⟩ - refine' fun N => ⟨R, _⟩ - rintro _ ⟨⟨m, n⟩, _, rfl⟩ - exact le_of_lt (hR m n) - -- Prove that it bounds the distances of points in the Cauchy sequence - have ub : ∀ m n N, N ≤ m → N ≤ n → dist (s m) (s n) ≤ sSup (S N) := fun m n N hm hn => - le_csSup (hS N) ⟨⟨_, _⟩, ⟨hm, hn⟩, rfl⟩ - have S0m : ∀ n, (0 : ℝ) ∈ S n := fun n => ⟨⟨n, n⟩, ⟨le_rfl, le_rfl⟩, dist_self _⟩ - have S0 := fun n => le_csSup (hS n) (S0m n) - -- Prove that it tends to `0`, by using the Cauchy property of `s` - refine' ⟨fun N => sSup (S N), S0, ub, Metric.tendsto_atTop.2 fun ε ε0 => _⟩ - refine' (Metric.cauchySeq_iff.1 hs (ε / 2) (half_pos ε0)).imp fun N hN n hn => _ - rw [Real.dist_0_eq_abs, abs_of_nonneg (S0 n)] - refine' lt_of_le_of_lt (csSup_le ⟨_, S0m _⟩ _) (half_lt_self ε0) - rintro _ ⟨⟨m', n'⟩, ⟨hm', hn'⟩, rfl⟩ - exact le_of_lt (hN _ (le_trans hn hm') _ (le_trans hn hn')), - fun ⟨b, _, b_bound, b_lim⟩ => cauchySeq_of_le_tendsto_0 b b_bound b_lim⟩ -#align cauchy_seq_iff_le_tendsto_0 cauchySeq_iff_le_tendsto_0 - -end CauchySeq - -/-- Pseudometric space structure pulled back by a function. -/ -@[reducible] -def PseudoMetricSpace.induced {α β} (f : α → β) (m : PseudoMetricSpace β) : - PseudoMetricSpace α where - dist x y := dist (f x) (f y) - dist_self x := dist_self _ - dist_comm x y := dist_comm _ _ - dist_triangle x y z := dist_triangle _ _ _ - edist x y := edist (f x) (f y) - edist_dist x y := edist_dist _ _ - toUniformSpace := UniformSpace.comap f m.toUniformSpace - uniformity_dist := (uniformity_basis_dist.comap _).eq_biInf - toBornology := Bornology.induced f - cobounded_sets := Set.ext fun s => mem_comap_iff_compl.trans <| by - simp only [← isBounded_def, isBounded_iff, ball_image_iff, mem_setOf] -#align pseudo_metric_space.induced PseudoMetricSpace.induced - -/-- Pull back a pseudometric space structure by an inducing map. This is a version of -`PseudoMetricSpace.induced` useful in case if the domain already has a `TopologicalSpace` -structure. -/ -def Inducing.comapPseudoMetricSpace {α β} [TopologicalSpace α] [m : PseudoMetricSpace β] {f : α → β} - (hf : Inducing f) : PseudoMetricSpace α := - .replaceTopology (.induced f m) hf.induced -#align inducing.comap_pseudo_metric_space Inducing.comapPseudoMetricSpace - -/-- Pull back a pseudometric space structure by a uniform inducing map. This is a version of -`PseudoMetricSpace.induced` useful in case if the domain already has a `UniformSpace` -structure. -/ -def UniformInducing.comapPseudoMetricSpace {α β} [UniformSpace α] [m : PseudoMetricSpace β] - (f : α → β) (h : UniformInducing f) : PseudoMetricSpace α := - .replaceUniformity (.induced f m) h.comap_uniformity.symm -#align uniform_inducing.comap_pseudo_metric_space UniformInducing.comapPseudoMetricSpace - -instance Subtype.pseudoMetricSpace {p : α → Prop} : PseudoMetricSpace (Subtype p) := - PseudoMetricSpace.induced Subtype.val ‹_› -#align subtype.pseudo_metric_space Subtype.pseudoMetricSpace - -theorem Subtype.dist_eq {p : α → Prop} (x y : Subtype p) : dist x y = dist (x : α) y := - rfl -#align subtype.dist_eq Subtype.dist_eq - -theorem Subtype.nndist_eq {p : α → Prop} (x y : Subtype p) : nndist x y = nndist (x : α) y := - rfl -#align subtype.nndist_eq Subtype.nndist_eq - -namespace MulOpposite - -@[to_additive] -instance instPseudoMetricSpaceMulOpposite : PseudoMetricSpace αᵐᵒᵖ := - PseudoMetricSpace.induced MulOpposite.unop ‹_› - -@[to_additive (attr := simp)] -theorem dist_unop (x y : αᵐᵒᵖ) : dist (unop x) (unop y) = dist x y := rfl -#align mul_opposite.dist_unop MulOpposite.dist_unop -#align add_opposite.dist_unop AddOpposite.dist_unop - -@[to_additive (attr := simp)] -theorem dist_op (x y : α) : dist (op x) (op y) = dist x y := rfl -#align mul_opposite.dist_op MulOpposite.dist_op -#align add_opposite.dist_op AddOpposite.dist_op - -@[to_additive (attr := simp)] -theorem nndist_unop (x y : αᵐᵒᵖ) : nndist (unop x) (unop y) = nndist x y := rfl -#align mul_opposite.nndist_unop MulOpposite.nndist_unop -#align add_opposite.nndist_unop AddOpposite.nndist_unop - -@[to_additive (attr := simp)] -theorem nndist_op (x y : α) : nndist (op x) (op y) = nndist x y := rfl -#align mul_opposite.nndist_op MulOpposite.nndist_op -#align add_opposite.nndist_op AddOpposite.nndist_op - -end MulOpposite - -section NNReal - -instance : PseudoMetricSpace ℝ≥0 := Subtype.pseudoMetricSpace - -theorem NNReal.dist_eq (a b : ℝ≥0) : dist a b = |(a : ℝ) - b| := rfl -#align nnreal.dist_eq NNReal.dist_eq - -theorem NNReal.nndist_eq (a b : ℝ≥0) : nndist a b = max (a - b) (b - a) := - eq_of_forall_ge_iff fun _ => by - simp only [← NNReal.coe_le_coe, coe_nndist, dist_eq, max_le_iff, abs_sub_le_iff, - tsub_le_iff_right, NNReal.coe_add] -#align nnreal.nndist_eq NNReal.nndist_eq - -@[simp] -theorem NNReal.nndist_zero_eq_val (z : ℝ≥0) : nndist 0 z = z := by - simp only [NNReal.nndist_eq, max_eq_right, tsub_zero, zero_tsub, zero_le'] -#align nnreal.nndist_zero_eq_val NNReal.nndist_zero_eq_val - -@[simp] -theorem NNReal.nndist_zero_eq_val' (z : ℝ≥0) : nndist z 0 = z := by - rw [nndist_comm] - exact NNReal.nndist_zero_eq_val z -#align nnreal.nndist_zero_eq_val' NNReal.nndist_zero_eq_val' - -theorem NNReal.le_add_nndist (a b : ℝ≥0) : a ≤ b + nndist a b := by - suffices (a : ℝ) ≤ (b : ℝ) + dist a b by - rwa [← NNReal.coe_le_coe, NNReal.coe_add, coe_nndist] - rw [← sub_le_iff_le_add'] - exact le_of_abs_le (dist_eq a b).ge -#align nnreal.le_add_nndist NNReal.le_add_nndist - -end NNReal - -section ULift - -variable [PseudoMetricSpace β] - -instance : PseudoMetricSpace (ULift β) := - PseudoMetricSpace.induced ULift.down ‹_› - -theorem ULift.dist_eq (x y : ULift β) : dist x y = dist x.down y.down := rfl -#align ulift.dist_eq ULift.dist_eq - -theorem ULift.nndist_eq (x y : ULift β) : nndist x y = nndist x.down y.down := rfl -#align ulift.nndist_eq ULift.nndist_eq - -@[simp] -theorem ULift.dist_up_up (x y : β) : dist (ULift.up x) (ULift.up y) = dist x y := rfl -#align ulift.dist_up_up ULift.dist_up_up - -@[simp] -theorem ULift.nndist_up_up (x y : β) : nndist (ULift.up x) (ULift.up y) = nndist x y := rfl -#align ulift.nndist_up_up ULift.nndist_up_up - -end ULift - -section Prod - -variable [PseudoMetricSpace β] - --- porting note: added `let`, otherwise `simp` failed -instance Prod.pseudoMetricSpaceMax : PseudoMetricSpace (α × β) := - let i := PseudoEMetricSpace.toPseudoMetricSpaceOfDist - (fun x y : α × β => dist x.1 y.1 ⊔ dist x.2 y.2) - (fun x y => (max_lt (edist_lt_top _ _) (edist_lt_top _ _)).ne) fun x y => by - simp only [sup_eq_max, dist_edist, ← ENNReal.toReal_max (edist_ne_top _ _) (edist_ne_top _ _), - Prod.edist_eq] - i.replaceBornology fun s => by - simp only [← isBounded_image_fst_and_snd, isBounded_iff_eventually, ball_image_iff, ← - eventually_and, ← forall_and, ← max_le_iff] - rfl -#align prod.pseudo_metric_space_max Prod.pseudoMetricSpaceMax - -theorem Prod.dist_eq {x y : α × β} : dist x y = max (dist x.1 y.1) (dist x.2 y.2) := rfl -#align prod.dist_eq Prod.dist_eq - -@[simp] -theorem dist_prod_same_left {x : α} {y₁ y₂ : β} : dist (x, y₁) (x, y₂) = dist y₁ y₂ := by - simp [Prod.dist_eq, dist_nonneg] -#align dist_prod_same_left dist_prod_same_left - -@[simp] -theorem dist_prod_same_right {x₁ x₂ : α} {y : β} : dist (x₁, y) (x₂, y) = dist x₁ x₂ := by - simp [Prod.dist_eq, dist_nonneg] -#align dist_prod_same_right dist_prod_same_right - -theorem ball_prod_same (x : α) (y : β) (r : ℝ) : ball x r ×ˢ ball y r = ball (x, y) r := - ext fun z => by simp [Prod.dist_eq] -#align ball_prod_same ball_prod_same - -theorem closedBall_prod_same (x : α) (y : β) (r : ℝ) : - closedBall x r ×ˢ closedBall y r = closedBall (x, y) r := - ext fun z => by simp [Prod.dist_eq] -#align closed_ball_prod_same closedBall_prod_same - -theorem sphere_prod (x : α × β) (r : ℝ) : - sphere x r = sphere x.1 r ×ˢ closedBall x.2 r ∪ closedBall x.1 r ×ˢ sphere x.2 r := by - obtain hr | rfl | hr := lt_trichotomy r 0 - · simp [hr] - · cases x - simp_rw [← closedBall_eq_sphere_of_nonpos le_rfl, union_self, closedBall_prod_same] - · ext ⟨x', y'⟩ - simp_rw [Set.mem_union, Set.mem_prod, Metric.mem_closedBall, Metric.mem_sphere, Prod.dist_eq, - max_eq_iff] - refine' or_congr (and_congr_right _) (and_comm.trans (and_congr_left _)) - all_goals rintro rfl; rfl -#align sphere_prod sphere_prod - -end Prod - --- porting note: 3 new lemmas -theorem dist_dist_dist_le_left (x y z : α) : dist (dist x z) (dist y z) ≤ dist x y := - abs_dist_sub_le .. - -theorem dist_dist_dist_le_right (x y z : α) : dist (dist x y) (dist x z) ≤ dist y z := by - simpa only [dist_comm x] using dist_dist_dist_le_left y z x - -theorem dist_dist_dist_le (x y x' y' : α) : dist (dist x y) (dist x' y') ≤ dist x x' + dist y y' := - (dist_triangle _ _ _).trans <| - add_le_add (dist_dist_dist_le_left _ _ _) (dist_dist_dist_le_right _ _ _) - -theorem uniformContinuous_dist : UniformContinuous fun p : α × α => dist p.1 p.2 := - Metric.uniformContinuous_iff.2 fun ε ε0 => - ⟨ε / 2, half_pos ε0, fun {a b} h => - calc dist (dist a.1 a.2) (dist b.1 b.2) ≤ dist a.1 b.1 + dist a.2 b.2 := - dist_dist_dist_le _ _ _ _ - _ ≤ dist a b + dist a b := add_le_add (le_max_left _ _) (le_max_right _ _) - _ < ε / 2 + ε / 2 := add_lt_add h h - _ = ε := add_halves ε⟩ -#align uniform_continuous_dist uniformContinuous_dist - -protected theorem UniformContinuous.dist [UniformSpace β] {f g : β → α} (hf : UniformContinuous f) - (hg : UniformContinuous g) : UniformContinuous fun b => dist (f b) (g b) := - uniformContinuous_dist.comp (hf.prod_mk hg) -#align uniform_continuous.dist UniformContinuous.dist - -@[continuity] -theorem continuous_dist : Continuous fun p : α × α => dist p.1 p.2 := - uniformContinuous_dist.continuous -#align continuous_dist continuous_dist - -@[continuity] -protected theorem Continuous.dist [TopologicalSpace β] {f g : β → α} (hf : Continuous f) - (hg : Continuous g) : Continuous fun b => dist (f b) (g b) := - continuous_dist.comp (hf.prod_mk hg : _) -#align continuous.dist Continuous.dist - -protected theorem Filter.Tendsto.dist {f g : β → α} {x : Filter β} {a b : α} - (hf : Tendsto f x (𝓝 a)) (hg : Tendsto g x (𝓝 b)) : - Tendsto (fun x => dist (f x) (g x)) x (𝓝 (dist a b)) := - (continuous_dist.tendsto (a, b)).comp (hf.prod_mk_nhds hg) -#align filter.tendsto.dist Filter.Tendsto.dist - -theorem nhds_comap_dist (a : α) : ((𝓝 (0 : ℝ)).comap (dist · a)) = 𝓝 a := by - simp only [@nhds_eq_comap_uniformity α, Metric.uniformity_eq_comap_nhds_zero, comap_comap, - (· ∘ ·), dist_comm] -#align nhds_comap_dist nhds_comap_dist - -theorem tendsto_iff_dist_tendsto_zero {f : β → α} {x : Filter β} {a : α} : - Tendsto f x (𝓝 a) ↔ Tendsto (fun b => dist (f b) a) x (𝓝 0) := by - rw [← nhds_comap_dist a, tendsto_comap_iff]; rfl -#align tendsto_iff_dist_tendsto_zero tendsto_iff_dist_tendsto_zero - -theorem continuous_iff_continuous_dist [TopologicalSpace β] {f : β → α} : - Continuous f ↔ Continuous fun x : β × β => dist (f x.1) (f x.2) := - ⟨fun h => h.fst'.dist h.snd', fun h => - continuous_iff_continuousAt.2 fun _ => tendsto_iff_dist_tendsto_zero.2 <| - (h.comp (continuous_id.prod_mk continuous_const)).tendsto' _ _ <| dist_self _⟩ -#align continuous_iff_continuous_dist continuous_iff_continuous_dist - -theorem uniformContinuous_nndist : UniformContinuous fun p : α × α => nndist p.1 p.2 := - uniformContinuous_dist.subtype_mk _ -#align uniform_continuous_nndist uniformContinuous_nndist - -protected theorem UniformContinuous.nndist [UniformSpace β] {f g : β → α} (hf : UniformContinuous f) - (hg : UniformContinuous g) : UniformContinuous fun b => nndist (f b) (g b) := - uniformContinuous_nndist.comp (hf.prod_mk hg) -#align uniform_continuous.nndist UniformContinuous.nndist - -theorem continuous_nndist : Continuous fun p : α × α => nndist p.1 p.2 := - uniformContinuous_nndist.continuous -#align continuous_nndist continuous_nndist - -protected theorem Continuous.nndist [TopologicalSpace β] {f g : β → α} (hf : Continuous f) - (hg : Continuous g) : Continuous fun b => nndist (f b) (g b) := - continuous_nndist.comp (hf.prod_mk hg : _) -#align continuous.nndist Continuous.nndist - -protected theorem Filter.Tendsto.nndist {f g : β → α} {x : Filter β} {a b : α} - (hf : Tendsto f x (𝓝 a)) (hg : Tendsto g x (𝓝 b)) : - Tendsto (fun x => nndist (f x) (g x)) x (𝓝 (nndist a b)) := - (continuous_nndist.tendsto (a, b)).comp (hf.prod_mk_nhds hg) -#align filter.tendsto.nndist Filter.Tendsto.nndist - -namespace Metric - -variable {x y z : α} {ε ε₁ ε₂ : ℝ} {s : Set α} - -theorem isClosed_ball : IsClosed (closedBall x ε) := - isClosed_le (continuous_id.dist continuous_const) continuous_const -#align metric.is_closed_ball Metric.isClosed_ball - -theorem isClosed_sphere : IsClosed (sphere x ε) := - isClosed_eq (continuous_id.dist continuous_const) continuous_const -#align metric.is_closed_sphere Metric.isClosed_sphere - -@[simp] -theorem closure_closedBall : closure (closedBall x ε) = closedBall x ε := - isClosed_ball.closure_eq -#align metric.closure_closed_ball Metric.closure_closedBall - -@[simp] -theorem closure_sphere : closure (sphere x ε) = sphere x ε := - isClosed_sphere.closure_eq -#align metric.closure_sphere Metric.closure_sphere - -theorem closure_ball_subset_closedBall : closure (ball x ε) ⊆ closedBall x ε := - closure_minimal ball_subset_closedBall isClosed_ball -#align metric.closure_ball_subset_closed_ball Metric.closure_ball_subset_closedBall - -theorem frontier_ball_subset_sphere : frontier (ball x ε) ⊆ sphere x ε := - frontier_lt_subset_eq (continuous_id.dist continuous_const) continuous_const -#align metric.frontier_ball_subset_sphere Metric.frontier_ball_subset_sphere - -theorem frontier_closedBall_subset_sphere : frontier (closedBall x ε) ⊆ sphere x ε := - frontier_le_subset_eq (continuous_id.dist continuous_const) continuous_const -#align metric.frontier_closed_ball_subset_sphere Metric.frontier_closedBall_subset_sphere - -theorem ball_subset_interior_closedBall : ball x ε ⊆ interior (closedBall x ε) := - interior_maximal ball_subset_closedBall isOpen_ball -#align metric.ball_subset_interior_closed_ball Metric.ball_subset_interior_closedBall - -/-- ε-characterization of the closure in pseudometric spaces-/ -theorem mem_closure_iff {s : Set α} {a : α} : a ∈ closure s ↔ ∀ ε > 0, ∃ b ∈ s, dist a b < ε := - (mem_closure_iff_nhds_basis nhds_basis_ball).trans <| by simp only [mem_ball, dist_comm] -#align metric.mem_closure_iff Metric.mem_closure_iff - -theorem mem_closure_range_iff {e : β → α} {a : α} : - a ∈ closure (range e) ↔ ∀ ε > 0, ∃ k : β, dist a (e k) < ε := by - simp only [mem_closure_iff, exists_range_iff] -#align metric.mem_closure_range_iff Metric.mem_closure_range_iff - -theorem mem_closure_range_iff_nat {e : β → α} {a : α} : - a ∈ closure (range e) ↔ ∀ n : ℕ, ∃ k : β, dist a (e k) < 1 / ((n : ℝ) + 1) := - (mem_closure_iff_nhds_basis nhds_basis_ball_inv_nat_succ).trans <| by - simp only [mem_ball, dist_comm, exists_range_iff, forall_const] -#align metric.mem_closure_range_iff_nat Metric.mem_closure_range_iff_nat - -theorem mem_of_closed' {s : Set α} (hs : IsClosed s) {a : α} : - a ∈ s ↔ ∀ ε > 0, ∃ b ∈ s, dist a b < ε := by - simpa only [hs.closure_eq] using @mem_closure_iff _ _ s a -#align metric.mem_of_closed' Metric.mem_of_closed' - -theorem closedBall_zero' (x : α) : closedBall x 0 = closure {x} := - Subset.antisymm - (fun _y hy => - mem_closure_iff.2 fun _ε ε0 => ⟨x, mem_singleton x, (mem_closedBall.1 hy).trans_lt ε0⟩) - (closure_minimal (singleton_subset_iff.2 (dist_self x).le) isClosed_ball) -#align metric.closed_ball_zero' Metric.closedBall_zero' - -lemma eventually_isCompact_closedBall [LocallyCompactSpace α] (x : α) : - ∀ᶠ r in 𝓝 (0 : ℝ), IsCompact (closedBall x r) := by - rcases local_compact_nhds (x := x) (n := univ) univ_mem with ⟨s, hs, -, s_compact⟩ - filter_upwards [eventually_closedBall_subset hs] with r hr - exact IsCompact.of_isClosed_subset s_compact isClosed_ball hr - -lemma exists_isCompact_closedBall [LocallyCompactSpace α] (x : α) : - ∃ r, 0 < r ∧ IsCompact (closedBall x r) := by - have : ∀ᶠ r in 𝓝[>] 0, IsCompact (closedBall x r) := - eventually_nhdsWithin_of_eventually_nhds (eventually_isCompact_closedBall x) - simpa only [and_comm] using (this.and self_mem_nhdsWithin).exists - -theorem dense_iff {s : Set α} : Dense s ↔ ∀ x, ∀ r > 0, (ball x r ∩ s).Nonempty := - forall_congr' fun x => by - simp only [mem_closure_iff, Set.Nonempty, exists_prop, mem_inter_iff, mem_ball', and_comm] -#align metric.dense_iff Metric.dense_iff - -theorem denseRange_iff {f : β → α} : DenseRange f ↔ ∀ x, ∀ r > 0, ∃ y, dist x (f y) < r := - forall_congr' fun x => by simp only [mem_closure_iff, exists_range_iff] -#align metric.dense_range_iff Metric.denseRange_iff - --- porting note: `TopologicalSpace.IsSeparable.separableSpace` moved to `EMetricSpace` - -/-- The preimage of a separable set by an inducing map is separable. -/ -protected theorem _root_.Inducing.isSeparable_preimage {f : β → α} [TopologicalSpace β] - (hf : Inducing f) {s : Set α} (hs : IsSeparable s) : IsSeparable (f ⁻¹' s) := by - have : SeparableSpace s := hs.separableSpace - have : SecondCountableTopology s := UniformSpace.secondCountable_of_separable _ - have : Inducing ((mapsTo_preimage f s).restrict _ _ _) := - (hf.comp inducing_subtype_val).codRestrict _ - have := this.secondCountableTopology - exact isSeparable_of_separableSpace_subtype _ -#align inducing.is_separable_preimage Inducing.isSeparable_preimage - -protected theorem _root_.Embedding.isSeparable_preimage {f : β → α} [TopologicalSpace β] - (hf : Embedding f) {s : Set α} (hs : IsSeparable s) : IsSeparable (f ⁻¹' s) := - hf.toInducing.isSeparable_preimage hs -#align embedding.is_separable_preimage Embedding.isSeparable_preimage - -/-- If a map is continuous on a separable set `s`, then the image of `s` is also separable. -/ -theorem _root_.ContinuousOn.isSeparable_image [TopologicalSpace β] {f : α → β} {s : Set α} - (hf : ContinuousOn f s) (hs : IsSeparable s) : IsSeparable (f '' s) := by - rw [image_eq_range, ← image_univ] - exact (isSeparable_univ_iff.2 hs.separableSpace).image hf.restrict -#align continuous_on.is_separable_image ContinuousOn.isSeparable_image - -end Metric - -section Pi - -open Finset - -variable {π : β → Type*} [Fintype β] [∀ b, PseudoMetricSpace (π b)] - -/-- A finite product of pseudometric spaces is a pseudometric space, with the sup distance. -/ -instance pseudoMetricSpacePi : PseudoMetricSpace (∀ b, π b) := by - /- we construct the instance from the pseudoemetric space instance to avoid checking again that - the uniformity is the same as the product uniformity, but we register nevertheless a nice - formula for the distance -/ - let i := PseudoEMetricSpace.toPseudoMetricSpaceOfDist - (fun f g : ∀ b, π b => ((sup univ fun b => nndist (f b) (g b) : ℝ≥0) : ℝ)) - (fun f g => ((Finset.sup_lt_iff bot_lt_top).2 fun b _ => edist_lt_top _ _).ne) - (fun f g => by - simp only [edist_pi_def, edist_nndist, ← ENNReal.coe_finset_sup, ENNReal.coe_toReal]) - refine i.replaceBornology fun s => ?_ - simp only [← isBounded_def, isBounded_iff_eventually, ← forall_isBounded_image_eval_iff, - ball_image_iff, ← Filter.eventually_all, Function.eval_apply, @dist_nndist (π _)] - refine' eventually_congr ((eventually_ge_atTop 0).mono fun C hC => _) - lift C to ℝ≥0 using hC - refine' ⟨fun H x hx y hy => NNReal.coe_le_coe.2 <| Finset.sup_le fun b _ => H b x hx y hy, - fun H b x hx y hy => NNReal.coe_le_coe.2 _⟩ - simpa only using Finset.sup_le_iff.1 (NNReal.coe_le_coe.1 <| H hx hy) b (Finset.mem_univ b) -#align pseudo_metric_space_pi pseudoMetricSpacePi - -theorem nndist_pi_def (f g : ∀ b, π b) : nndist f g = sup univ fun b => nndist (f b) (g b) := - NNReal.eq rfl -#align nndist_pi_def nndist_pi_def - -theorem dist_pi_def (f g : ∀ b, π b) : dist f g = (sup univ fun b => nndist (f b) (g b) : ℝ≥0) := - rfl -#align dist_pi_def dist_pi_def - -theorem nndist_pi_le_iff {f g : ∀ b, π b} {r : ℝ≥0} : - nndist f g ≤ r ↔ ∀ b, nndist (f b) (g b) ≤ r := by simp [nndist_pi_def] -#align nndist_pi_le_iff nndist_pi_le_iff - -theorem nndist_pi_lt_iff {f g : ∀ b, π b} {r : ℝ≥0} (hr : 0 < r) : - nndist f g < r ↔ ∀ b, nndist (f b) (g b) < r := by - simp [nndist_pi_def, Finset.sup_lt_iff (show ⊥ < r from hr)] -#align nndist_pi_lt_iff nndist_pi_lt_iff - -theorem nndist_pi_eq_iff {f g : ∀ b, π b} {r : ℝ≥0} (hr : 0 < r) : - nndist f g = r ↔ (∃ i, nndist (f i) (g i) = r) ∧ ∀ b, nndist (f b) (g b) ≤ r := by - rw [eq_iff_le_not_lt, nndist_pi_lt_iff hr, nndist_pi_le_iff, not_forall, and_comm] - simp_rw [not_lt, and_congr_left_iff, le_antisymm_iff] - intro h - refine' exists_congr fun b => _ - apply (and_iff_right <| h _).symm -#align nndist_pi_eq_iff nndist_pi_eq_iff - -theorem dist_pi_lt_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 < r) : - dist f g < r ↔ ∀ b, dist (f b) (g b) < r := by - lift r to ℝ≥0 using hr.le - exact nndist_pi_lt_iff hr -#align dist_pi_lt_iff dist_pi_lt_iff - -theorem dist_pi_le_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 ≤ r) : - dist f g ≤ r ↔ ∀ b, dist (f b) (g b) ≤ r := by - lift r to ℝ≥0 using hr - exact nndist_pi_le_iff -#align dist_pi_le_iff dist_pi_le_iff - -theorem dist_pi_eq_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 < r) : - dist f g = r ↔ (∃ i, dist (f i) (g i) = r) ∧ ∀ b, dist (f b) (g b) ≤ r := by - lift r to ℝ≥0 using hr.le - simp_rw [← coe_nndist, NNReal.coe_eq, nndist_pi_eq_iff hr, NNReal.coe_le_coe] -#align dist_pi_eq_iff dist_pi_eq_iff - -theorem dist_pi_le_iff' [Nonempty β] {f g : ∀ b, π b} {r : ℝ} : - dist f g ≤ r ↔ ∀ b, dist (f b) (g b) ≤ r := by - by_cases hr : 0 ≤ r - · exact dist_pi_le_iff hr - · exact iff_of_false (fun h => hr <| dist_nonneg.trans h) fun h => - hr <| dist_nonneg.trans <| h <| Classical.arbitrary _ -#align dist_pi_le_iff' dist_pi_le_iff' - -theorem dist_pi_const_le (a b : α) : (dist (fun _ : β => a) fun _ => b) ≤ dist a b := - (dist_pi_le_iff dist_nonneg).2 fun _ => le_rfl -#align dist_pi_const_le dist_pi_const_le - -theorem nndist_pi_const_le (a b : α) : (nndist (fun _ : β => a) fun _ => b) ≤ nndist a b := - nndist_pi_le_iff.2 fun _ => le_rfl -#align nndist_pi_const_le nndist_pi_const_le - -@[simp] -theorem dist_pi_const [Nonempty β] (a b : α) : (dist (fun _ : β => a) fun _ => b) = dist a b := by - simpa only [dist_edist] using congr_arg ENNReal.toReal (edist_pi_const a b) -#align dist_pi_const dist_pi_const - -@[simp] -theorem nndist_pi_const [Nonempty β] (a b : α) : - (nndist (fun _ : β => a) fun _ => b) = nndist a b := - NNReal.eq <| dist_pi_const a b -#align nndist_pi_const nndist_pi_const - -theorem nndist_le_pi_nndist (f g : ∀ b, π b) (b : β) : nndist (f b) (g b) ≤ nndist f g := by - rw [← ENNReal.coe_le_coe, ← edist_nndist, ← edist_nndist] - exact edist_le_pi_edist f g b -#align nndist_le_pi_nndist nndist_le_pi_nndist - -theorem dist_le_pi_dist (f g : ∀ b, π b) (b : β) : dist (f b) (g b) ≤ dist f g := by - simp only [dist_nndist, NNReal.coe_le_coe, nndist_le_pi_nndist f g b] -#align dist_le_pi_dist dist_le_pi_dist - -/-- An open ball in a product space is a product of open balls. See also `ball_pi'` -for a version assuming `Nonempty β` instead of `0 < r`. -/ -theorem ball_pi (x : ∀ b, π b) {r : ℝ} (hr : 0 < r) : - ball x r = Set.pi univ fun b => ball (x b) r := by - ext p - simp [dist_pi_lt_iff hr] -#align ball_pi ball_pi - -/-- An open ball in a product space is a product of open balls. See also `ball_pi` -for a version assuming `0 < r` instead of `Nonempty β`. -/ -theorem ball_pi' [Nonempty β] (x : ∀ b, π b) (r : ℝ) : - ball x r = Set.pi univ fun b => ball (x b) r := - (lt_or_le 0 r).elim (ball_pi x) fun hr => by simp [ball_eq_empty.2 hr] -#align ball_pi' ball_pi' - -/-- A closed ball in a product space is a product of closed balls. See also `closedBall_pi'` -for a version assuming `Nonempty β` instead of `0 ≤ r`. -/ -theorem closedBall_pi (x : ∀ b, π b) {r : ℝ} (hr : 0 ≤ r) : - closedBall x r = Set.pi univ fun b => closedBall (x b) r := by - ext p - simp [dist_pi_le_iff hr] -#align closed_ball_pi closedBall_pi - -/-- A closed ball in a product space is a product of closed balls. See also `closedBall_pi` -for a version assuming `0 ≤ r` instead of `Nonempty β`. -/ -theorem closedBall_pi' [Nonempty β] (x : ∀ b, π b) (r : ℝ) : - closedBall x r = Set.pi univ fun b => closedBall (x b) r := - (le_or_lt 0 r).elim (closedBall_pi x) fun hr => by simp [closedBall_eq_empty.2 hr] -#align closed_ball_pi' closedBall_pi' - -/-- A sphere in a product space is a union of spheres on each component restricted to the closed -ball. -/ -theorem sphere_pi (x : ∀ b, π b) {r : ℝ} (h : 0 < r ∨ Nonempty β) : - sphere x r = (⋃ i : β, Function.eval i ⁻¹' sphere (x i) r) ∩ closedBall x r := by - obtain hr | rfl | hr := lt_trichotomy r 0 - · simp [hr] - · rw [closedBall_eq_sphere_of_nonpos le_rfl, eq_comm, Set.inter_eq_right] - letI := h.resolve_left (lt_irrefl _) - inhabit β - refine' subset_iUnion_of_subset default _ - intro x hx - replace hx := hx.le - rw [dist_pi_le_iff le_rfl] at hx - exact le_antisymm (hx default) dist_nonneg - · ext - simp [dist_pi_eq_iff hr, dist_pi_le_iff hr.le] -#align sphere_pi sphere_pi - -@[simp] -theorem Fin.nndist_insertNth_insertNth {n : ℕ} {α : Fin (n + 1) → Type*} - [∀ i, PseudoMetricSpace (α i)] (i : Fin (n + 1)) (x y : α i) (f g : ∀ j, α (i.succAbove j)) : - nndist (i.insertNth x f) (i.insertNth y g) = max (nndist x y) (nndist f g) := - eq_of_forall_ge_iff fun c => by simp [nndist_pi_le_iff, i.forall_iff_succAbove] -#align fin.nndist_insert_nth_insert_nth Fin.nndist_insertNth_insertNth - -@[simp] -theorem Fin.dist_insertNth_insertNth {n : ℕ} {α : Fin (n + 1) → Type*} - [∀ i, PseudoMetricSpace (α i)] (i : Fin (n + 1)) (x y : α i) (f g : ∀ j, α (i.succAbove j)) : - dist (i.insertNth x f) (i.insertNth y g) = max (dist x y) (dist f g) := by - simp only [dist_nndist, Fin.nndist_insertNth_insertNth, NNReal.coe_max] -#align fin.dist_insert_nth_insert_nth Fin.dist_insertNth_insertNth - -theorem Real.dist_le_of_mem_pi_Icc {x y x' y' : β → ℝ} (hx : x ∈ Icc x' y') (hy : y ∈ Icc x' y') : - dist x y ≤ dist x' y' := by - refine' (dist_pi_le_iff dist_nonneg).2 fun b => - (Real.dist_le_of_mem_uIcc _ _).trans (dist_le_pi_dist x' y' b) <;> refine' Icc_subset_uIcc _ - exacts [⟨hx.1 _, hx.2 _⟩, ⟨hy.1 _, hy.2 _⟩] -#align real.dist_le_of_mem_pi_Icc Real.dist_le_of_mem_pi_Icc - -end Pi - -section Compact - -/-- Any compact set in a pseudometric space can be covered by finitely many balls of a given -positive radius -/ -theorem finite_cover_balls_of_compact {α : Type u} [PseudoMetricSpace α] {s : Set α} - (hs : IsCompact s) {e : ℝ} (he : 0 < e) : - ∃ t, t ⊆ s ∧ Set.Finite t ∧ s ⊆ ⋃ x ∈ t, ball x e := - let ⟨t, hts, ht⟩ := hs.elim_nhds_subcover _ (fun x _ => ball_mem_nhds x he) - ⟨t, hts, t.finite_toSet, ht⟩ -#align finite_cover_balls_of_compact finite_cover_balls_of_compact - -alias IsCompact.finite_cover_balls := finite_cover_balls_of_compact -#align is_compact.finite_cover_balls IsCompact.finite_cover_balls - -end Compact - -section ProperSpace - -open Metric - -/-- A pseudometric space is proper if all closed balls are compact. -/ -class ProperSpace (α : Type u) [PseudoMetricSpace α] : Prop where - isCompact_closedBall : ∀ x : α, ∀ r, IsCompact (closedBall x r) -#align proper_space ProperSpace - -export ProperSpace (isCompact_closedBall) - -/-- In a proper pseudometric space, all spheres are compact. -/ -theorem isCompact_sphere {α : Type*} [PseudoMetricSpace α] [ProperSpace α] (x : α) (r : ℝ) : - IsCompact (sphere x r) := - (isCompact_closedBall x r).of_isClosed_subset isClosed_sphere sphere_subset_closedBall -#align is_compact_sphere isCompact_sphere - -/-- In a proper pseudometric space, any sphere is a `CompactSpace` when considered as a subtype. -/ -instance Metric.sphere.compactSpace {α : Type*} [PseudoMetricSpace α] [ProperSpace α] - (x : α) (r : ℝ) : CompactSpace (sphere x r) := - isCompact_iff_compactSpace.mp (isCompact_sphere _ _) - --- see Note [lower instance priority] -/-- A proper pseudo metric space is sigma compact, and therefore second countable. -/ -instance (priority := 100) secondCountable_of_proper [ProperSpace α] : - SecondCountableTopology α := by - -- We already have `sigmaCompactSpace_of_locallyCompact_secondCountable`, so we don't - -- add an instance for `SigmaCompactSpace`. - suffices SigmaCompactSpace α by exact EMetric.secondCountable_of_sigmaCompact α - rcases em (Nonempty α) with (⟨⟨x⟩⟩ | hn) - · exact ⟨⟨fun n => closedBall x n, fun n => isCompact_closedBall _ _, iUnion_closedBall_nat _⟩⟩ - · exact ⟨⟨fun _ => ∅, fun _ => isCompact_empty, iUnion_eq_univ_iff.2 fun x => (hn ⟨x⟩).elim⟩⟩ -#align second_countable_of_proper secondCountable_of_proper - -/-- If all closed balls of large enough radius are compact, then the space is proper. Especially -useful when the lower bound for the radius is 0. -/ -theorem properSpace_of_compact_closedBall_of_le (R : ℝ) - (h : ∀ x : α, ∀ r, R ≤ r → IsCompact (closedBall x r)) : ProperSpace α := - ⟨fun x r => IsCompact.of_isClosed_subset (h x (max r R) (le_max_right _ _)) isClosed_ball - (closedBall_subset_closedBall <| le_max_left _ _)⟩ -#align proper_space_of_compact_closed_ball_of_le properSpace_of_compact_closedBall_of_le - --- A compact pseudometric space is proper --- see Note [lower instance priority] -instance (priority := 100) proper_of_compact [CompactSpace α] : ProperSpace α := - ⟨fun _ _ => isClosed_ball.isCompact⟩ -#align proper_of_compact proper_of_compact - --- see Note [lower instance priority] -/-- A proper space is locally compact -/ -instance (priority := 100) locally_compact_of_proper [ProperSpace α] : LocallyCompactSpace α := - locallyCompactSpace_of_hasBasis (fun _ => nhds_basis_closedBall) fun _ _ _ => - isCompact_closedBall _ _ -#align locally_compact_of_proper locally_compact_of_proper - --- see Note [lower instance priority] -/-- A proper space is complete -/ -instance (priority := 100) complete_of_proper [ProperSpace α] : CompleteSpace α := - ⟨fun {f} hf => by - /- We want to show that the Cauchy filter `f` is converging. It suffices to find a closed - ball (therefore compact by properness) where it is nontrivial. -/ - obtain ⟨t, t_fset, ht⟩ : ∃ t ∈ f, ∀ x ∈ t, ∀ y ∈ t, dist x y < 1 := - (Metric.cauchy_iff.1 hf).2 1 zero_lt_one - rcases hf.1.nonempty_of_mem t_fset with ⟨x, xt⟩ - have : closedBall x 1 ∈ f := mem_of_superset t_fset fun y yt => (ht y yt x xt).le - rcases (isCompact_iff_totallyBounded_isComplete.1 (isCompact_closedBall x 1)).2 f hf - (le_principal_iff.2 this) with - ⟨y, -, hy⟩ - exact ⟨y, hy⟩⟩ -#align complete_of_proper complete_of_proper - -/-- A binary product of proper spaces is proper. -/ -instance prod_properSpace {α : Type*} {β : Type*} [PseudoMetricSpace α] [PseudoMetricSpace β] - [ProperSpace α] [ProperSpace β] : ProperSpace (α × β) where - isCompact_closedBall := by - rintro ⟨x, y⟩ r - rw [← closedBall_prod_same x y] - exact (isCompact_closedBall x r).prod (isCompact_closedBall y r) -#align prod_proper_space prod_properSpace - -/-- A finite product of proper spaces is proper. -/ -instance pi_properSpace {π : β → Type*} [Fintype β] [∀ b, PseudoMetricSpace (π b)] - [h : ∀ b, ProperSpace (π b)] : ProperSpace (∀ b, π b) := by - refine' properSpace_of_compact_closedBall_of_le 0 fun x r hr => _ - rw [closedBall_pi _ hr] - exact isCompact_univ_pi fun _ => isCompact_closedBall _ _ -#align pi_proper_space pi_properSpace - -variable [ProperSpace α] {x : α} {r : ℝ} {s : Set α} - -/-- If a nonempty ball in a proper space includes a closed set `s`, then there exists a nonempty -ball with the same center and a strictly smaller radius that includes `s`. -/ -theorem exists_pos_lt_subset_ball (hr : 0 < r) (hs : IsClosed s) (h : s ⊆ ball x r) : - ∃ r' ∈ Ioo 0 r, s ⊆ ball x r' := by - rcases eq_empty_or_nonempty s with (rfl | hne) - · exact ⟨r / 2, ⟨half_pos hr, half_lt_self hr⟩, empty_subset _⟩ - have : IsCompact s := - (isCompact_closedBall x r).of_isClosed_subset hs (h.trans ball_subset_closedBall) - obtain ⟨y, hys, hy⟩ : ∃ y ∈ s, s ⊆ closedBall x (dist y x) := - this.exists_forall_ge hne (continuous_id.dist continuous_const).continuousOn - have hyr : dist y x < r := h hys - rcases exists_between hyr with ⟨r', hyr', hrr'⟩ - exact ⟨r', ⟨dist_nonneg.trans_lt hyr', hrr'⟩, hy.trans <| closedBall_subset_ball hyr'⟩ -#align exists_pos_lt_subset_ball exists_pos_lt_subset_ball - -/-- If a ball in a proper space includes a closed set `s`, then there exists a ball with the same -center and a strictly smaller radius that includes `s`. -/ -theorem exists_lt_subset_ball (hs : IsClosed s) (h : s ⊆ ball x r) : ∃ r' < r, s ⊆ ball x r' := by - cases' le_or_lt r 0 with hr hr - · rw [ball_eq_empty.2 hr, subset_empty_iff] at h - subst s - exact (exists_lt r).imp fun r' hr' => ⟨hr', empty_subset _⟩ - · exact (exists_pos_lt_subset_ball hr hs h).imp fun r' hr' => ⟨hr'.1.2, hr'.2⟩ -#align exists_lt_subset_ball exists_lt_subset_ball - -end ProperSpace - -theorem IsCompact.isSeparable {s : Set α} (hs : IsCompact s) : IsSeparable s := - haveI : CompactSpace s := isCompact_iff_compactSpace.mp hs - isSeparable_of_separableSpace_subtype s -#align is_compact.is_separable IsCompact.isSeparable - -namespace Metric - -section SecondCountable - -open TopologicalSpace - -/-- A pseudometric space is second countable if, for every `ε > 0`, there is a countable set which -is `ε`-dense. -/ -theorem secondCountable_of_almost_dense_set - (H : ∀ ε > (0 : ℝ), ∃ s : Set α, s.Countable ∧ ∀ x, ∃ y ∈ s, dist x y ≤ ε) : - SecondCountableTopology α := by - refine' EMetric.secondCountable_of_almost_dense_set fun ε ε0 => _ - rcases ENNReal.lt_iff_exists_nnreal_btwn.1 ε0 with ⟨ε', ε'0, ε'ε⟩ - choose s hsc y hys hyx using H ε' (by exact_mod_cast ε'0) - refine' ⟨s, hsc, iUnion₂_eq_univ_iff.2 fun x => ⟨y x, hys _, le_trans _ ε'ε.le⟩⟩ - exact_mod_cast hyx x -#align metric.second_countable_of_almost_dense_set Metric.secondCountable_of_almost_dense_set - -end SecondCountable - -end Metric - -theorem lebesgue_number_lemma_of_metric {s : Set α} {ι : Sort*} {c : ι → Set α} (hs : IsCompact s) - (hc₁ : ∀ i, IsOpen (c i)) (hc₂ : s ⊆ ⋃ i, c i) : ∃ δ > 0, ∀ x ∈ s, ∃ i, ball x δ ⊆ c i := - let ⟨_n, en, hn⟩ := lebesgue_number_lemma hs hc₁ hc₂ - let ⟨δ, δ0, hδ⟩ := mem_uniformity_dist.1 en - ⟨δ, δ0, fun x hx => let ⟨i, hi⟩ := hn x hx; ⟨i, fun _y hy => hi (hδ (mem_ball'.mp hy))⟩⟩ -#align lebesgue_number_lemma_of_metric lebesgue_number_lemma_of_metric - -theorem lebesgue_number_lemma_of_metric_sUnion {s : Set α} {c : Set (Set α)} (hs : IsCompact s) - (hc₁ : ∀ t ∈ c, IsOpen t) (hc₂ : s ⊆ ⋃₀ c) : ∃ δ > 0, ∀ x ∈ s, ∃ t ∈ c, ball x δ ⊆ t := by - rw [sUnion_eq_iUnion] at hc₂; simpa using lebesgue_number_lemma_of_metric hs (by simpa) hc₂ -#align lebesgue_number_lemma_of_metric_sUnion lebesgue_number_lemma_of_metric_sUnion - -namespace Metric - -#align metric.bounded Bornology.IsBounded - -section Bounded - -variable {x : α} {s t : Set α} {r : ℝ} - -#noalign metric.bounded_iff_is_bounded - -#align metric.bounded_empty Bornology.isBounded_empty -#align metric.bounded_iff_mem_bounded Bornology.isBounded_iff_forall_mem -#align metric.bounded.mono Bornology.IsBounded.subset - -/-- Closed balls are bounded -/ -theorem isBounded_closedBall : IsBounded (closedBall x r) := - isBounded_iff.2 ⟨r + r, fun y hy z hz => - calc dist y z ≤ dist y x + dist z x := dist_triangle_right _ _ _ - _ ≤ r + r := add_le_add hy hz⟩ -#align metric.bounded_closed_ball Metric.isBounded_closedBall - -/-- Open balls are bounded -/ -theorem isBounded_ball : IsBounded (ball x r) := - isBounded_closedBall.subset ball_subset_closedBall -#align metric.bounded_ball Metric.isBounded_ball - -/-- Spheres are bounded -/ -theorem isBounded_sphere : IsBounded (sphere x r) := - isBounded_closedBall.subset sphere_subset_closedBall -#align metric.bounded_sphere Metric.isBounded_sphere - -/-- Given a point, a bounded subset is included in some ball around this point -/ -theorem isBounded_iff_subset_closedBall (c : α) : IsBounded s ↔ ∃ r, s ⊆ closedBall c r := - ⟨fun h ↦ (isBounded_iff.1 (h.insert c)).imp fun _r hr _x hx ↦ hr (.inr hx) (mem_insert _ _), - fun ⟨_r, hr⟩ ↦ isBounded_closedBall.subset hr⟩ -#align metric.bounded_iff_subset_ball Metric.isBounded_iff_subset_closedBall - -theorem _root_.Bornology.IsBounded.subset_closedBall (h : IsBounded s) (c : α) : - ∃ r, s ⊆ closedBall c r := - (isBounded_iff_subset_closedBall c).1 h -#align metric.bounded.subset_ball Bornology.IsBounded.subset_closedBall - -theorem _root_.Bornology.IsBounded.subset_ball_lt (h : IsBounded s) (a : ℝ) (c : α) : - ∃ r, a < r ∧ s ⊆ ball c r := - let ⟨r, hr⟩ := h.subset_closedBall c - ⟨max r a + 1, (le_max_right _ _).trans_lt (lt_add_one _), hr.trans <| closedBall_subset_ball <| - (le_max_left _ _).trans_lt (lt_add_one _)⟩ - -theorem _root_.Bornology.IsBounded.subset_ball (h : IsBounded s) (c : α) : ∃ r, s ⊆ ball c r := - (h.subset_ball_lt 0 c).imp fun _ ↦ And.right - -theorem isBounded_iff_subset_ball (c : α) : IsBounded s ↔ ∃ r, s ⊆ ball c r := - ⟨(IsBounded.subset_ball · c), fun ⟨_r, hr⟩ ↦ isBounded_ball.subset hr⟩ - -theorem _root_.Bornology.IsBounded.subset_closedBall_lt (h : IsBounded s) (a : ℝ) (c : α) : - ∃ r, a < r ∧ s ⊆ closedBall c r := - let ⟨r, har, hr⟩ := h.subset_ball_lt a c - ⟨r, har, hr.trans ball_subset_closedBall⟩ -#align metric.bounded.subset_ball_lt Bornology.IsBounded.subset_closedBall_lt - -theorem isBounded_closure_of_isBounded (h : IsBounded s) : IsBounded (closure s) := - let ⟨C, h⟩ := isBounded_iff.1 h - isBounded_iff.2 ⟨C, fun _a ha _b hb => (isClosed_le' C).closure_subset <| - map_mem_closure₂ continuous_dist ha hb h⟩ -#align metric.bounded_closure_of_bounded Metric.isBounded_closure_of_isBounded - -protected theorem _root_.Bornology.IsBounded.closure (h : IsBounded s) : IsBounded (closure s) := - isBounded_closure_of_isBounded h -#align metric.bounded.closure Bornology.IsBounded.closure - -@[simp] -theorem isBounded_closure_iff : IsBounded (closure s) ↔ IsBounded s := - ⟨fun h => h.subset subset_closure, fun h => h.closure⟩ -#align metric.bounded_closure_iff Metric.isBounded_closure_iff - -#align metric.bounded_union Bornology.isBounded_union -#align metric.bounded.union Bornology.IsBounded.union -#align metric.bounded_bUnion Bornology.isBounded_biUnion -#align metric.bounded.prod Bornology.IsBounded.prod - -theorem hasBasis_cobounded_compl_closedBall (c : α) : - (cobounded α).HasBasis (fun _ ↦ True) (fun r ↦ (closedBall c r)ᶜ) := - ⟨compl_surjective.forall.2 fun _ ↦ (isBounded_iff_subset_closedBall c).trans <| by simp⟩ - -theorem hasBasis_cobounded_compl_ball (c : α) : - (cobounded α).HasBasis (fun _ ↦ True) (fun r ↦ (ball c r)ᶜ) := - ⟨compl_surjective.forall.2 fun _ ↦ (isBounded_iff_subset_ball c).trans <| by simp⟩ - -@[simp] -theorem comap_dist_right_atTop (c : α) : comap (dist · c) atTop = cobounded α := - (atTop_basis.comap _).eq_of_same_basis <| by - simpa only [compl_def, mem_ball, not_lt] using hasBasis_cobounded_compl_ball c - -@[simp] -theorem comap_dist_left_atTop (c : α) : comap (dist c) atTop = cobounded α := by - simpa only [dist_comm _ c] using comap_dist_right_atTop c - -@[simp] -theorem tendsto_dist_right_atTop_iff (c : α) {f : β → α} {l : Filter β} : - Tendsto (fun x ↦ dist (f x) c) l atTop ↔ Tendsto f l (cobounded α) := by - rw [← comap_dist_right_atTop c, tendsto_comap_iff]; rfl - -@[simp] -theorem tendsto_dist_left_atTop_iff (c : α) {f : β → α} {l : Filter β} : - Tendsto (fun x ↦ dist c (f x)) l atTop ↔ Tendsto f l (cobounded α) := by - simp only [dist_comm c, tendsto_dist_right_atTop_iff] - -theorem tendsto_dist_right_cobounded_atTop (c : α) : Tendsto (dist · c) (cobounded α) atTop := - tendsto_iff_comap.2 (comap_dist_right_atTop c).ge - -theorem tendsto_dist_left_cobounded_atTop (c : α) : Tendsto (dist c) (cobounded α) atTop := - tendsto_iff_comap.2 (comap_dist_left_atTop c).ge - -/-- A totally bounded set is bounded -/ -theorem _root_.TotallyBounded.isBounded {s : Set α} (h : TotallyBounded s) : IsBounded s := - -- We cover the totally bounded set by finitely many balls of radius 1, - -- and then argue that a finite union of bounded sets is bounded - let ⟨_t, fint, subs⟩ := (totallyBounded_iff.mp h) 1 zero_lt_one - ((isBounded_biUnion fint).2 fun _ _ => isBounded_ball).subset subs -#align totally_bounded.bounded TotallyBounded.isBounded - -/-- A compact set is bounded -/ -theorem _root_.IsCompact.isBounded {s : Set α} (h : IsCompact s) : IsBounded s := - -- A compact set is totally bounded, thus bounded - h.totallyBounded.isBounded -#align is_compact.bounded IsCompact.isBounded - -#align metric.bounded_of_finite Set.Finite.isBounded -#align set.finite.bounded Set.Finite.isBounded -#align metric.bounded_singleton Bornology.isBounded_singleton - -theorem cobounded_le_cocompact : cobounded α ≤ cocompact α := - hasBasis_cocompact.ge_iff.2 fun _s hs ↦ hs.isBounded -#align comap_dist_right_at_top_le_cocompact Metric.cobounded_le_cocompactₓ -#align comap_dist_left_at_top_le_cocompact Metric.cobounded_le_cocompactₓ - -/-- Characterization of the boundedness of the range of a function -/ -theorem isBounded_range_iff {f : β → α} : IsBounded (range f) ↔ ∃ C, ∀ x y, dist (f x) (f y) ≤ C := - isBounded_iff.trans <| by simp only [forall_range_iff] -#align metric.bounded_range_iff Metric.isBounded_range_iff - -theorem isBounded_image_iff {f : β → α} {s : Set β} : - IsBounded (f '' s) ↔ ∃ C, ∀ x ∈ s, ∀ y ∈ s, dist (f x) (f y) ≤ C := - isBounded_iff.trans <| by simp only [ball_image_iff] - -theorem isBounded_range_of_tendsto_cofinite_uniformity {f : β → α} - (hf : Tendsto (Prod.map f f) (.cofinite ×ˢ .cofinite) (𝓤 α)) : IsBounded (range f) := by - rcases (hasBasis_cofinite.prod_self.tendsto_iff uniformity_basis_dist).1 hf 1 zero_lt_one with - ⟨s, hsf, hs1⟩ - rw [← image_union_image_compl_eq_range] - refine (hsf.image f).isBounded.union (isBounded_image_iff.2 ⟨1, fun x hx y hy ↦ ?_⟩) - exact le_of_lt (hs1 (x, y) ⟨hx, hy⟩) -#align metric.bounded_range_of_tendsto_cofinite_uniformity Metric.isBounded_range_of_tendsto_cofinite_uniformity - -theorem isBounded_range_of_cauchy_map_cofinite {f : β → α} (hf : Cauchy (map f cofinite)) : - IsBounded (range f) := - isBounded_range_of_tendsto_cofinite_uniformity <| (cauchy_map_iff.1 hf).2 -#align metric.bounded_range_of_cauchy_map_cofinite Metric.isBounded_range_of_cauchy_map_cofinite - -theorem _root_.CauchySeq.isBounded_range {f : ℕ → α} (hf : CauchySeq f) : IsBounded (range f) := - isBounded_range_of_cauchy_map_cofinite <| by rwa [Nat.cofinite_eq_atTop] -#align cauchy_seq.bounded_range CauchySeq.isBounded_range - -theorem isBounded_range_of_tendsto_cofinite {f : β → α} {a : α} (hf : Tendsto f cofinite (𝓝 a)) : - IsBounded (range f) := - isBounded_range_of_tendsto_cofinite_uniformity <| - (hf.prod_map hf).mono_right <| nhds_prod_eq.symm.trans_le (nhds_le_uniformity a) -#align metric.bounded_range_of_tendsto_cofinite Metric.isBounded_range_of_tendsto_cofinite - -/-- In a compact space, all sets are bounded -/ -theorem isBounded_of_compactSpace [CompactSpace α] : IsBounded s := - isCompact_univ.isBounded.subset (subset_univ _) -#align metric.bounded_of_compact_space Metric.isBounded_of_compactSpace - -theorem isBounded_range_of_tendsto (u : ℕ → α) {x : α} (hu : Tendsto u atTop (𝓝 x)) : - IsBounded (range u) := - hu.cauchySeq.isBounded_range -#align metric.bounded_range_of_tendsto Metric.isBounded_range_of_tendsto - -theorem disjoint_nhds_cobounded (x : α) : Disjoint (𝓝 x) (cobounded α) := - disjoint_of_disjoint_of_mem disjoint_compl_right (ball_mem_nhds _ one_pos) isBounded_ball - -theorem disjoint_cobounded_nhds (x : α) : Disjoint (cobounded α) (𝓝 x) := - (disjoint_nhds_cobounded x).symm - -theorem disjoint_nhdsSet_cobounded {s : Set α} (hs : IsCompact s) : Disjoint (𝓝ˢ s) (cobounded α) := - hs.disjoint_nhdsSet_left.2 fun _ _ ↦ disjoint_nhds_cobounded _ - -theorem disjoint_cobounded_nhdsSet {s : Set α} (hs : IsCompact s) : Disjoint (cobounded α) (𝓝ˢ s) := - (disjoint_nhdsSet_cobounded hs).symm - -/-- If a function is continuous within a set `s` at every point of a compact set `k`, then it is -bounded on some open neighborhood of `k` in `s`. -/ -theorem exists_isOpen_isBounded_image_inter_of_isCompact_of_forall_continuousWithinAt - [TopologicalSpace β] {k s : Set β} {f : β → α} (hk : IsCompact k) - (hf : ∀ x ∈ k, ContinuousWithinAt f s x) : - ∃ t, k ⊆ t ∧ IsOpen t ∧ IsBounded (f '' (t ∩ s)) := by - have : Disjoint (𝓝ˢ k ⊓ 𝓟 s) (comap f (cobounded α)) := by - rw [disjoint_assoc, inf_comm, hk.disjoint_nhdsSet_left] - exact fun x hx ↦ disjoint_left_comm.2 <| - tendsto_comap.disjoint (disjoint_cobounded_nhds _) (hf x hx) - rcases ((((hasBasis_nhdsSet _).inf_principal _)).disjoint_iff ((basis_sets _).comap _)).1 this - with ⟨U, ⟨hUo, hkU⟩, t, ht, hd⟩ - refine ⟨U, hkU, hUo, (isBounded_compl_iff.2 ht).subset ?_⟩ - rwa [image_subset_iff, preimage_compl, subset_compl_iff_disjoint_right] -#align metric.exists_is_open_bounded_image_inter_of_is_compact_of_forall_continuous_within_at Metric.exists_isOpen_isBounded_image_inter_of_isCompact_of_forall_continuousWithinAt - -/-- If a function is continuous at every point of a compact set `k`, then it is bounded on -some open neighborhood of `k`. -/ -theorem exists_isOpen_isBounded_image_of_isCompact_of_forall_continuousAt [TopologicalSpace β] - {k : Set β} {f : β → α} (hk : IsCompact k) (hf : ∀ x ∈ k, ContinuousAt f x) : - ∃ t, k ⊆ t ∧ IsOpen t ∧ IsBounded (f '' t) := by - simp_rw [← continuousWithinAt_univ] at hf - simpa only [inter_univ] using - exists_isOpen_isBounded_image_inter_of_isCompact_of_forall_continuousWithinAt hk hf -#align metric.exists_is_open_bounded_image_of_is_compact_of_forall_continuous_at Metric.exists_isOpen_isBounded_image_of_isCompact_of_forall_continuousAt - -/-- If a function is continuous on a set `s` containing a compact set `k`, then it is bounded on -some open neighborhood of `k` in `s`. -/ -theorem exists_isOpen_isBounded_image_inter_of_isCompact_of_continuousOn [TopologicalSpace β] - {k s : Set β} {f : β → α} (hk : IsCompact k) (hks : k ⊆ s) (hf : ContinuousOn f s) : - ∃ t, k ⊆ t ∧ IsOpen t ∧ IsBounded (f '' (t ∩ s)) := - exists_isOpen_isBounded_image_inter_of_isCompact_of_forall_continuousWithinAt hk fun x hx => - hf x (hks hx) -#align metric.exists_is_open_bounded_image_inter_of_is_compact_of_continuous_on Metric.exists_isOpen_isBounded_image_inter_of_isCompact_of_continuousOn - -/-- If a function is continuous on a neighborhood of a compact set `k`, then it is bounded on -some open neighborhood of `k`. -/ -theorem exists_isOpen_isBounded_image_of_isCompact_of_continuousOn [TopologicalSpace β] - {k s : Set β} {f : β → α} (hk : IsCompact k) (hs : IsOpen s) (hks : k ⊆ s) - (hf : ContinuousOn f s) : ∃ t, k ⊆ t ∧ IsOpen t ∧ IsBounded (f '' t) := - exists_isOpen_isBounded_image_of_isCompact_of_forall_continuousAt hk fun _x hx => - hf.continuousAt (hs.mem_nhds (hks hx)) -#align metric.exists_is_open_bounded_image_of_is_compact_of_continuous_on Metric.exists_isOpen_isBounded_image_of_isCompact_of_continuousOn - -/-- The **Heine–Borel theorem**: In a proper space, a closed bounded set is compact. -/ -theorem isCompact_of_isClosed_isBounded [ProperSpace α] (hc : IsClosed s) (hb : IsBounded s) : - IsCompact s := by - rcases eq_empty_or_nonempty s with (rfl | ⟨x, -⟩) - · exact isCompact_empty - · rcases hb.subset_closedBall x with ⟨r, hr⟩ - exact (isCompact_closedBall x r).of_isClosed_subset hc hr -#align metric.is_compact_of_is_closed_bounded Metric.isCompact_of_isClosed_isBounded - -/-- The **Heine–Borel theorem**: In a proper space, the closure of a bounded set is compact. -/ -theorem _root_.Bornology.IsBounded.isCompact_closure [ProperSpace α] (h : IsBounded s) : - IsCompact (closure s) := - isCompact_of_isClosed_isBounded isClosed_closure h.closure -#align metric.bounded.is_compact_closure Bornology.IsBounded.isCompact_closure - --- porting note: todo: assume `[MetricSpace α]` instead of `[PseudoMetricSpace α] [T2Space α]` -/-- The **Heine–Borel theorem**: -In a proper Hausdorff space, a set is compact if and only if it is closed and bounded. -/ -theorem isCompact_iff_isClosed_bounded [T2Space α] [ProperSpace α] : - IsCompact s ↔ IsClosed s ∧ IsBounded s := - ⟨fun h => ⟨h.isClosed, h.isBounded⟩, fun h => isCompact_of_isClosed_isBounded h.1 h.2⟩ -#align metric.is_compact_iff_is_closed_bounded Metric.isCompact_iff_isClosed_bounded - -theorem compactSpace_iff_isBounded_univ [ProperSpace α] : - CompactSpace α ↔ IsBounded (univ : Set α) := - ⟨@isBounded_of_compactSpace α _ _, fun hb => ⟨isCompact_of_isClosed_isBounded isClosed_univ hb⟩⟩ -#align metric.compact_space_iff_bounded_univ Metric.compactSpace_iff_isBounded_univ - -section ConditionallyCompleteLinearOrder - -variable [Preorder α] [CompactIccSpace α] - -theorem isBounded_Icc (a b : α) : IsBounded (Icc a b) := - (totallyBounded_Icc a b).isBounded -#align metric.bounded_Icc Metric.isBounded_Icc - -theorem isBounded_Ico (a b : α) : IsBounded (Ico a b) := - (totallyBounded_Ico a b).isBounded -#align metric.bounded_Ico Metric.isBounded_Ico - -theorem isBounded_Ioc (a b : α) : IsBounded (Ioc a b) := - (totallyBounded_Ioc a b).isBounded -#align metric.bounded_Ioc Metric.isBounded_Ioc - -theorem isBounded_Ioo (a b : α) : IsBounded (Ioo a b) := - (totallyBounded_Ioo a b).isBounded -#align metric.bounded_Ioo Metric.isBounded_Ioo - -/-- In a pseudo metric space with a conditionally complete linear order such that the order and the - metric structure give the same topology, any order-bounded set is metric-bounded. -/ -theorem isBounded_of_bddAbove_of_bddBelow {s : Set α} (h₁ : BddAbove s) (h₂ : BddBelow s) : - IsBounded s := - let ⟨u, hu⟩ := h₁ - let ⟨l, hl⟩ := h₂ - (isBounded_Icc l u).subset (fun _x hx => mem_Icc.mpr ⟨hl hx, hu hx⟩) -#align metric.bounded_of_bdd_above_of_bdd_below Metric.isBounded_of_bddAbove_of_bddBelow - -end ConditionallyCompleteLinearOrder - -end Bounded - -section Diam - -variable {s : Set α} {x y z : α} - -/-- The diameter of a set in a metric space. To get controllable behavior even when the diameter -should be infinite, we express it in terms of the emetric.diameter -/ -noncomputable def diam (s : Set α) : ℝ := - ENNReal.toReal (EMetric.diam s) -#align metric.diam Metric.diam - -/-- The diameter of a set is always nonnegative -/ -theorem diam_nonneg : 0 ≤ diam s := - ENNReal.toReal_nonneg -#align metric.diam_nonneg Metric.diam_nonneg - -theorem diam_subsingleton (hs : s.Subsingleton) : diam s = 0 := by - simp only [diam, EMetric.diam_subsingleton hs, ENNReal.zero_toReal] -#align metric.diam_subsingleton Metric.diam_subsingleton - -/-- The empty set has zero diameter -/ -@[simp] -theorem diam_empty : diam (∅ : Set α) = 0 := - diam_subsingleton subsingleton_empty -#align metric.diam_empty Metric.diam_empty - -/-- A singleton has zero diameter -/ -@[simp] -theorem diam_singleton : diam ({x} : Set α) = 0 := - diam_subsingleton subsingleton_singleton -#align metric.diam_singleton Metric.diam_singleton - --- Does not work as a simp-lemma, since {x, y} reduces to (insert y {x}) -theorem diam_pair : diam ({x, y} : Set α) = dist x y := by - simp only [diam, EMetric.diam_pair, dist_edist] -#align metric.diam_pair Metric.diam_pair - --- Does not work as a simp-lemma, since {x, y, z} reduces to (insert z (insert y {x})) -theorem diam_triple : - Metric.diam ({x, y, z} : Set α) = max (max (dist x y) (dist x z)) (dist y z) := by - simp only [Metric.diam, EMetric.diam_triple, dist_edist] - rw [ENNReal.toReal_max, ENNReal.toReal_max] <;> apply_rules [ne_of_lt, edist_lt_top, max_lt] -#align metric.diam_triple Metric.diam_triple - -/-- If the distance between any two points in a set is bounded by some constant `C`, -then `ENNReal.ofReal C` bounds the emetric diameter of this set. -/ -theorem ediam_le_of_forall_dist_le {C : ℝ} (h : ∀ x ∈ s, ∀ y ∈ s, dist x y ≤ C) : - EMetric.diam s ≤ ENNReal.ofReal C := - EMetric.diam_le fun x hx y hy => (edist_dist x y).symm ▸ ENNReal.ofReal_le_ofReal (h x hx y hy) -#align metric.ediam_le_of_forall_dist_le Metric.ediam_le_of_forall_dist_le - -/-- If the distance between any two points in a set is bounded by some non-negative constant, -this constant bounds the diameter. -/ -theorem diam_le_of_forall_dist_le {C : ℝ} (h₀ : 0 ≤ C) (h : ∀ x ∈ s, ∀ y ∈ s, dist x y ≤ C) : - diam s ≤ C := - ENNReal.toReal_le_of_le_ofReal h₀ (ediam_le_of_forall_dist_le h) -#align metric.diam_le_of_forall_dist_le Metric.diam_le_of_forall_dist_le - -/-- If the distance between any two points in a nonempty set is bounded by some constant, -this constant bounds the diameter. -/ -theorem diam_le_of_forall_dist_le_of_nonempty (hs : s.Nonempty) {C : ℝ} - (h : ∀ x ∈ s, ∀ y ∈ s, dist x y ≤ C) : diam s ≤ C := - have h₀ : 0 ≤ C := - let ⟨x, hx⟩ := hs - le_trans dist_nonneg (h x hx x hx) - diam_le_of_forall_dist_le h₀ h -#align metric.diam_le_of_forall_dist_le_of_nonempty Metric.diam_le_of_forall_dist_le_of_nonempty - -/-- The distance between two points in a set is controlled by the diameter of the set. -/ -theorem dist_le_diam_of_mem' (h : EMetric.diam s ≠ ⊤) (hx : x ∈ s) (hy : y ∈ s) : - dist x y ≤ diam s := by - rw [diam, dist_edist] - rw [ENNReal.toReal_le_toReal (edist_ne_top _ _) h] - exact EMetric.edist_le_diam_of_mem hx hy -#align metric.dist_le_diam_of_mem' Metric.dist_le_diam_of_mem' - -/-- Characterize the boundedness of a set in terms of the finiteness of its emetric.diameter. -/ -theorem isBounded_iff_ediam_ne_top : IsBounded s ↔ EMetric.diam s ≠ ⊤ := - isBounded_iff.trans <| Iff.intro - (fun ⟨_C, hC⟩ => ne_top_of_le_ne_top ENNReal.ofReal_ne_top <| ediam_le_of_forall_dist_le hC) - fun h => ⟨diam s, fun _x hx _y hy => dist_le_diam_of_mem' h hx hy⟩ -#align metric.bounded_iff_ediam_ne_top Metric.isBounded_iff_ediam_ne_top - -alias ⟨_root_.Bornology.IsBounded.ediam_ne_top, _⟩ := isBounded_iff_ediam_ne_top -#align metric.bounded.ediam_ne_top Bornology.IsBounded.ediam_ne_top - -theorem ediam_eq_top_iff_unbounded : EMetric.diam s = ⊤ ↔ ¬IsBounded s := - isBounded_iff_ediam_ne_top.not_left.symm - -theorem ediam_univ_eq_top_iff_noncompact [ProperSpace α] : - EMetric.diam (univ : Set α) = ∞ ↔ NoncompactSpace α := by - rw [← not_compactSpace_iff, compactSpace_iff_isBounded_univ, isBounded_iff_ediam_ne_top, - Classical.not_not] -#align metric.ediam_univ_eq_top_iff_noncompact Metric.ediam_univ_eq_top_iff_noncompact - -@[simp] -theorem ediam_univ_of_noncompact [ProperSpace α] [NoncompactSpace α] : - EMetric.diam (univ : Set α) = ∞ := - ediam_univ_eq_top_iff_noncompact.mpr ‹_› -#align metric.ediam_univ_of_noncompact Metric.ediam_univ_of_noncompact - -@[simp] -theorem diam_univ_of_noncompact [ProperSpace α] [NoncompactSpace α] : diam (univ : Set α) = 0 := by - simp [diam] -#align metric.diam_univ_of_noncompact Metric.diam_univ_of_noncompact - -/-- The distance between two points in a set is controlled by the diameter of the set. -/ -theorem dist_le_diam_of_mem (h : IsBounded s) (hx : x ∈ s) (hy : y ∈ s) : dist x y ≤ diam s := - dist_le_diam_of_mem' h.ediam_ne_top hx hy -#align metric.dist_le_diam_of_mem Metric.dist_le_diam_of_mem - -theorem ediam_of_unbounded (h : ¬IsBounded s) : EMetric.diam s = ∞ := ediam_eq_top_iff_unbounded.2 h -#align metric.ediam_of_unbounded Metric.ediam_of_unbounded - -/-- An unbounded set has zero diameter. If you would prefer to get the value ∞, use `EMetric.diam`. -This lemma makes it possible to avoid side conditions in some situations -/ -theorem diam_eq_zero_of_unbounded (h : ¬IsBounded s) : diam s = 0 := by - rw [diam, ediam_of_unbounded h, ENNReal.top_toReal] -#align metric.diam_eq_zero_of_unbounded Metric.diam_eq_zero_of_unbounded - -/-- If `s ⊆ t`, then the diameter of `s` is bounded by that of `t`, provided `t` is bounded. -/ -theorem diam_mono {s t : Set α} (h : s ⊆ t) (ht : IsBounded t) : diam s ≤ diam t := - ENNReal.toReal_mono ht.ediam_ne_top <| EMetric.diam_mono h -#align metric.diam_mono Metric.diam_mono - -/-- The diameter of a union is controlled by the sum of the diameters, and the distance between -any two points in each of the sets. This lemma is true without any side condition, since it is -obviously true if `s ∪ t` is unbounded. -/ -theorem diam_union {t : Set α} (xs : x ∈ s) (yt : y ∈ t) : - diam (s ∪ t) ≤ diam s + dist x y + diam t := by - simp only [diam, dist_edist] - refine (ENNReal.toReal_le_add' (EMetric.diam_union xs yt) ?_ ?_).trans - (add_le_add_right ENNReal.toReal_add_le _) - · simp only [ENNReal.add_eq_top, edist_ne_top, or_false] - exact fun h ↦ top_unique <| h ▸ EMetric.diam_mono (subset_union_left _ _) - · exact fun h ↦ top_unique <| h ▸ EMetric.diam_mono (subset_union_right _ _) -#align metric.diam_union Metric.diam_union - -/-- If two sets intersect, the diameter of the union is bounded by the sum of the diameters. -/ -theorem diam_union' {t : Set α} (h : (s ∩ t).Nonempty) : diam (s ∪ t) ≤ diam s + diam t := by - rcases h with ⟨x, ⟨xs, xt⟩⟩ - simpa using diam_union xs xt -#align metric.diam_union' Metric.diam_union' - -theorem diam_le_of_subset_closedBall {r : ℝ} (hr : 0 ≤ r) (h : s ⊆ closedBall x r) : - diam s ≤ 2 * r := - diam_le_of_forall_dist_le (mul_nonneg zero_le_two hr) fun a ha b hb => - calc - dist a b ≤ dist a x + dist b x := dist_triangle_right _ _ _ - _ ≤ r + r := (add_le_add (h ha) (h hb)) - _ = 2 * r := by simp [mul_two, mul_comm] -#align metric.diam_le_of_subset_closed_ball Metric.diam_le_of_subset_closedBall - -/-- The diameter of a closed ball of radius `r` is at most `2 r`. -/ -theorem diam_closedBall {r : ℝ} (h : 0 ≤ r) : diam (closedBall x r) ≤ 2 * r := - diam_le_of_subset_closedBall h Subset.rfl -#align metric.diam_closed_ball Metric.diam_closedBall - -/-- The diameter of a ball of radius `r` is at most `2 r`. -/ -theorem diam_ball {r : ℝ} (h : 0 ≤ r) : diam (ball x r) ≤ 2 * r := - diam_le_of_subset_closedBall h ball_subset_closedBall -#align metric.diam_ball Metric.diam_ball - -/-- If a family of complete sets with diameter tending to `0` is such that each finite intersection -is nonempty, then the total intersection is also nonempty. -/ -theorem _root_.IsComplete.nonempty_iInter_of_nonempty_biInter {s : ℕ → Set α} - (h0 : IsComplete (s 0)) (hs : ∀ n, IsClosed (s n)) (h's : ∀ n, IsBounded (s n)) - (h : ∀ N, (⋂ n ≤ N, s n).Nonempty) (h' : Tendsto (fun n => diam (s n)) atTop (𝓝 0)) : - (⋂ n, s n).Nonempty := by - let u N := (h N).some - have I : ∀ n N, n ≤ N → u N ∈ s n := by - intro n N hn - apply mem_of_subset_of_mem _ (h N).choose_spec - intro x hx - simp only [mem_iInter] at hx - exact hx n hn - have : CauchySeq u := by - apply cauchySeq_of_le_tendsto_0 _ _ h' - intro m n N hm hn - exact dist_le_diam_of_mem (h's N) (I _ _ hm) (I _ _ hn) - obtain ⟨x, -, xlim⟩ : ∃ x ∈ s 0, Tendsto (fun n : ℕ => u n) atTop (𝓝 x) := - cauchySeq_tendsto_of_isComplete h0 (fun n => I 0 n (zero_le _)) this - refine' ⟨x, mem_iInter.2 fun n => _⟩ - apply (hs n).mem_of_tendsto xlim - filter_upwards [Ici_mem_atTop n] with p hp - exact I n p hp -#align is_complete.nonempty_Inter_of_nonempty_bInter IsComplete.nonempty_iInter_of_nonempty_biInter - -/-- In a complete space, if a family of closed sets with diameter tending to `0` is such that each -finite intersection is nonempty, then the total intersection is also nonempty. -/ -theorem nonempty_iInter_of_nonempty_biInter [CompleteSpace α] {s : ℕ → Set α} - (hs : ∀ n, IsClosed (s n)) (h's : ∀ n, IsBounded (s n)) (h : ∀ N, (⋂ n ≤ N, s n).Nonempty) - (h' : Tendsto (fun n => diam (s n)) atTop (𝓝 0)) : (⋂ n, s n).Nonempty := - (hs 0).isComplete.nonempty_iInter_of_nonempty_biInter hs h's h h' -#align metric.nonempty_Inter_of_nonempty_bInter Metric.nonempty_iInter_of_nonempty_biInter - -end Diam - -theorem exists_isLocalMin_mem_ball [ProperSpace α] [TopologicalSpace β] - [ConditionallyCompleteLinearOrder β] [OrderTopology β] {f : α → β} {a z : α} {r : ℝ} - (hf : ContinuousOn f (closedBall a r)) (hz : z ∈ closedBall a r) - (hf1 : ∀ z' ∈ sphere a r, f z < f z') : ∃ z ∈ ball a r, IsLocalMin f z := by - simp_rw [← closedBall_diff_ball] at hf1 - exact (isCompact_closedBall a r).exists_isLocalMin_mem_open ball_subset_closedBall hf hz hf1 - isOpen_ball -#align metric.exists_local_min_mem_ball Metric.exists_isLocalMin_mem_ball - -end Metric - -namespace Mathlib.Meta.Positivity - -open Lean Meta Qq Function - -/-- Extension for the `positivity` tactic: the diameter of a set is always nonnegative. -/ -@[positivity Metric.diam _] -def evalDiam : PositivityExt where eval {_ _} _zα _pα e := do - let .app _ s ← whnfR e | throwError "not Metric.diam" - let p ← mkAppOptM ``Metric.diam_nonneg #[none, none, s] - pure (.nonnegative p) - -end Mathlib.Meta.Positivity - -theorem Metric.cobounded_eq_cocompact [ProperSpace α] : cobounded α = cocompact α := by - nontriviality α; inhabit α - exact cobounded_le_cocompact.antisymm <| (hasBasis_cobounded_compl_closedBall default).ge_iff.2 - fun _ _ ↦ (isCompact_closedBall _ _).compl_mem_cocompact -#align comap_dist_right_at_top_eq_cocompact Metric.cobounded_eq_cocompact - -theorem tendsto_dist_right_cocompact_atTop [ProperSpace α] (x : α) : - Tendsto (dist · x) (cocompact α) atTop := - (tendsto_dist_right_cobounded_atTop x).mono_left cobounded_eq_cocompact.ge -#align tendsto_dist_right_cocompact_at_top tendsto_dist_right_cocompact_atTop - -theorem tendsto_dist_left_cocompact_atTop [ProperSpace α] (x : α) : - Tendsto (dist x) (cocompact α) atTop := - (tendsto_dist_left_cobounded_atTop x).mono_left cobounded_eq_cocompact.ge -#align tendsto_dist_left_cocompact_at_top tendsto_dist_left_cocompact_atTop - -theorem comap_dist_left_atTop_eq_cocompact [ProperSpace α] (x : α) : - comap (dist x) atTop = cocompact α := by simp [cobounded_eq_cocompact] -#align comap_dist_left_at_top_eq_cocompact comap_dist_left_atTop_eq_cocompact - -theorem tendsto_cocompact_of_tendsto_dist_comp_atTop {f : β → α} {l : Filter β} (x : α) - (h : Tendsto (fun y => dist (f y) x) l atTop) : Tendsto f l (cocompact α) := - ((tendsto_dist_right_atTop_iff _).1 h).mono_right cobounded_le_cocompact -#align tendsto_cocompact_of_tendsto_dist_comp_at_top tendsto_cocompact_of_tendsto_dist_comp_atTop - /-- We now define `MetricSpace`, extending `PseudoMetricSpace`. -/ class MetricSpace (α : Type u) extends PseudoMetricSpace α : Type u where eq_of_dist_eq_zero : ∀ {x y : α}, dist x y = 0 → x = y @@ -3031,12 +178,6 @@ theorem uniformEmbedding_bot_of_pairwise_le_dist {β : Type*} {ε : ℝ} (hε : uniformEmbedding_of_spaced_out (dist_mem_uniformity hε) <| by simpa using hf #align metric.uniform_embedding_bot_of_pairwise_le_dist Metric.uniformEmbedding_bot_of_pairwise_le_dist -theorem finite_isBounded_inter_isClosed [ProperSpace α] {K s : Set α} [DiscreteTopology s] - (hK : IsBounded K) (hs : IsClosed s) : Set.Finite (K ∩ s) := by - refine Set.Finite.subset (IsCompact.finite ?_ ?_) (Set.inter_subset_inter_left s subset_closure) - · exact hK.isCompact_closure.inter_right hs - · exact DiscreteTopology.of_subset inferInstance (Set.inter_subset_right _ s) - end Metric /-- Build a new metric space from an old one where the bundled uniform structure is provably diff --git a/Mathlib/Topology/MetricSpace/Bounded.lean b/Mathlib/Topology/MetricSpace/Bounded.lean new file mode 100644 index 0000000000000..281f94be3db5e --- /dev/null +++ b/Mathlib/Topology/MetricSpace/Bounded.lean @@ -0,0 +1,584 @@ +/- +Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel +-/ +import Mathlib.Topology.MetricSpace.PseudoMetric +import Mathlib.Topology.MetricSpace.Cauchy + +/-! +## Boundedness in (pseudo)-metric spaces + +This file contains one definition, and various results on boundedness in pseudo-metric spaces. +* `Metric.diam s` : The `iSup` of the distances of members of `s`. + Defined in terms of `EMetric.diam`, for better handling of the case when it should be infinite. + +* `isBounded_iff_subset_closedBall`: a non-empty set is bounded if and only if + it is is included in some closed ball +* describing the cobounded filter, relating to the cocompact filter +* `IsCompact.isBounded`: compact sets are bounded +* `TotallyBounded.isBounded`: totally bounded sets are bounded +* `isCompact_iff_isClosed_bounded`, the **Heine–Borel theorem**: + in a proper space, a set is compact if and only if it is closed and bounded. +* `cobounded_eq_cocompact`: in a proper space, cobounded and compact sets are the same +diameter of a subset, and its relation to boundedness + +## Tags + +metric, pseudo_metric, bounded, diameter, Heine-Borel theorem +-/ + +open Set Filter Bornology +open scoped ENNReal Uniformity Topology Pointwise + +universe u v w + +variable {α : Type u} {β : Type v} {X ι : Type*} +variable [PseudoMetricSpace α] + +namespace Metric + +#align metric.bounded Bornology.IsBounded + +section Bounded + +variable {x : α} {s t : Set α} {r : ℝ} + +#noalign metric.bounded_iff_is_bounded + +#align metric.bounded_empty Bornology.isBounded_empty +#align metric.bounded_iff_mem_bounded Bornology.isBounded_iff_forall_mem +#align metric.bounded.mono Bornology.IsBounded.subset + +/-- Closed balls are bounded -/ +theorem isBounded_closedBall : IsBounded (closedBall x r) := + isBounded_iff.2 ⟨r + r, fun y hy z hz => + calc dist y z ≤ dist y x + dist z x := dist_triangle_right _ _ _ + _ ≤ r + r := add_le_add hy hz⟩ +#align metric.bounded_closed_ball Metric.isBounded_closedBall + +/-- Open balls are bounded -/ +theorem isBounded_ball : IsBounded (ball x r) := + isBounded_closedBall.subset ball_subset_closedBall +#align metric.bounded_ball Metric.isBounded_ball + +/-- Spheres are bounded -/ +theorem isBounded_sphere : IsBounded (sphere x r) := + isBounded_closedBall.subset sphere_subset_closedBall +#align metric.bounded_sphere Metric.isBounded_sphere + +/-- Given a point, a bounded subset is included in some ball around this point -/ +theorem isBounded_iff_subset_closedBall (c : α) : IsBounded s ↔ ∃ r, s ⊆ closedBall c r := + ⟨fun h ↦ (isBounded_iff.1 (h.insert c)).imp fun _r hr _x hx ↦ hr (.inr hx) (mem_insert _ _), + fun ⟨_r, hr⟩ ↦ isBounded_closedBall.subset hr⟩ +#align metric.bounded_iff_subset_ball Metric.isBounded_iff_subset_closedBall + +theorem _root_.Bornology.IsBounded.subset_closedBall (h : IsBounded s) (c : α) : + ∃ r, s ⊆ closedBall c r := + (isBounded_iff_subset_closedBall c).1 h +#align metric.bounded.subset_ball Bornology.IsBounded.subset_closedBall + +theorem _root_.Bornology.IsBounded.subset_ball_lt (h : IsBounded s) (a : ℝ) (c : α) : + ∃ r, a < r ∧ s ⊆ ball c r := + let ⟨r, hr⟩ := h.subset_closedBall c + ⟨max r a + 1, (le_max_right _ _).trans_lt (lt_add_one _), hr.trans <| closedBall_subset_ball <| + (le_max_left _ _).trans_lt (lt_add_one _)⟩ + +theorem _root_.Bornology.IsBounded.subset_ball (h : IsBounded s) (c : α) : ∃ r, s ⊆ ball c r := + (h.subset_ball_lt 0 c).imp fun _ ↦ And.right + +theorem isBounded_iff_subset_ball (c : α) : IsBounded s ↔ ∃ r, s ⊆ ball c r := + ⟨(IsBounded.subset_ball · c), fun ⟨_r, hr⟩ ↦ isBounded_ball.subset hr⟩ + +theorem _root_.Bornology.IsBounded.subset_closedBall_lt (h : IsBounded s) (a : ℝ) (c : α) : + ∃ r, a < r ∧ s ⊆ closedBall c r := + let ⟨r, har, hr⟩ := h.subset_ball_lt a c + ⟨r, har, hr.trans ball_subset_closedBall⟩ +#align metric.bounded.subset_ball_lt Bornology.IsBounded.subset_closedBall_lt + +theorem isBounded_closure_of_isBounded (h : IsBounded s) : IsBounded (closure s) := + let ⟨C, h⟩ := isBounded_iff.1 h + isBounded_iff.2 ⟨C, fun _a ha _b hb => (isClosed_le' C).closure_subset <| + map_mem_closure₂ continuous_dist ha hb h⟩ +#align metric.bounded_closure_of_bounded Metric.isBounded_closure_of_isBounded + +protected theorem _root_.Bornology.IsBounded.closure (h : IsBounded s) : IsBounded (closure s) := + isBounded_closure_of_isBounded h +#align metric.bounded.closure Bornology.IsBounded.closure + +@[simp] +theorem isBounded_closure_iff : IsBounded (closure s) ↔ IsBounded s := + ⟨fun h => h.subset subset_closure, fun h => h.closure⟩ +#align metric.bounded_closure_iff Metric.isBounded_closure_iff + +#align metric.bounded_union Bornology.isBounded_union +#align metric.bounded.union Bornology.IsBounded.union +#align metric.bounded_bUnion Bornology.isBounded_biUnion +#align metric.bounded.prod Bornology.IsBounded.prod + +theorem hasBasis_cobounded_compl_closedBall (c : α) : + (cobounded α).HasBasis (fun _ ↦ True) (fun r ↦ (closedBall c r)ᶜ) := + ⟨compl_surjective.forall.2 fun _ ↦ (isBounded_iff_subset_closedBall c).trans <| by simp⟩ + +theorem hasBasis_cobounded_compl_ball (c : α) : + (cobounded α).HasBasis (fun _ ↦ True) (fun r ↦ (ball c r)ᶜ) := + ⟨compl_surjective.forall.2 fun _ ↦ (isBounded_iff_subset_ball c).trans <| by simp⟩ + +@[simp] +theorem comap_dist_right_atTop (c : α) : comap (dist · c) atTop = cobounded α := + (atTop_basis.comap _).eq_of_same_basis <| by + simpa only [compl_def, mem_ball, not_lt] using hasBasis_cobounded_compl_ball c + +@[simp] +theorem comap_dist_left_atTop (c : α) : comap (dist c) atTop = cobounded α := by + simpa only [dist_comm _ c] using comap_dist_right_atTop c + +@[simp] +theorem tendsto_dist_right_atTop_iff (c : α) {f : β → α} {l : Filter β} : + Tendsto (fun x ↦ dist (f x) c) l atTop ↔ Tendsto f l (cobounded α) := by + rw [← comap_dist_right_atTop c, tendsto_comap_iff]; rfl + +@[simp] +theorem tendsto_dist_left_atTop_iff (c : α) {f : β → α} {l : Filter β} : + Tendsto (fun x ↦ dist c (f x)) l atTop ↔ Tendsto f l (cobounded α) := by + simp only [dist_comm c, tendsto_dist_right_atTop_iff] + +theorem tendsto_dist_right_cobounded_atTop (c : α) : Tendsto (dist · c) (cobounded α) atTop := + tendsto_iff_comap.2 (comap_dist_right_atTop c).ge + +theorem tendsto_dist_left_cobounded_atTop (c : α) : Tendsto (dist c) (cobounded α) atTop := + tendsto_iff_comap.2 (comap_dist_left_atTop c).ge + +/-- A totally bounded set is bounded -/ +theorem _root_.TotallyBounded.isBounded {s : Set α} (h : TotallyBounded s) : IsBounded s := + -- We cover the totally bounded set by finitely many balls of radius 1, + -- and then argue that a finite union of bounded sets is bounded + let ⟨_t, fint, subs⟩ := (totallyBounded_iff.mp h) 1 zero_lt_one + ((isBounded_biUnion fint).2 fun _ _ => isBounded_ball).subset subs +#align totally_bounded.bounded TotallyBounded.isBounded + +/-- A compact set is bounded -/ +theorem _root_.IsCompact.isBounded {s : Set α} (h : IsCompact s) : IsBounded s := + -- A compact set is totally bounded, thus bounded + h.totallyBounded.isBounded +#align is_compact.bounded IsCompact.isBounded + +#align metric.bounded_of_finite Set.Finite.isBounded +#align set.finite.bounded Set.Finite.isBounded +#align metric.bounded_singleton Bornology.isBounded_singleton + +theorem cobounded_le_cocompact : cobounded α ≤ cocompact α := + hasBasis_cocompact.ge_iff.2 fun _s hs ↦ hs.isBounded +#align comap_dist_right_at_top_le_cocompact Metric.cobounded_le_cocompactₓ +#align comap_dist_left_at_top_le_cocompact Metric.cobounded_le_cocompactₓ + +/-- Characterization of the boundedness of the range of a function -/ +theorem isBounded_range_iff {f : β → α} : IsBounded (range f) ↔ ∃ C, ∀ x y, dist (f x) (f y) ≤ C := + isBounded_iff.trans <| by simp only [forall_range_iff] +#align metric.bounded_range_iff Metric.isBounded_range_iff + +theorem isBounded_image_iff {f : β → α} {s : Set β} : + IsBounded (f '' s) ↔ ∃ C, ∀ x ∈ s, ∀ y ∈ s, dist (f x) (f y) ≤ C := + isBounded_iff.trans <| by simp only [ball_image_iff] + +theorem isBounded_range_of_tendsto_cofinite_uniformity {f : β → α} + (hf : Tendsto (Prod.map f f) (.cofinite ×ˢ .cofinite) (𝓤 α)) : IsBounded (range f) := by + rcases (hasBasis_cofinite.prod_self.tendsto_iff uniformity_basis_dist).1 hf 1 zero_lt_one with + ⟨s, hsf, hs1⟩ + rw [← image_union_image_compl_eq_range] + refine (hsf.image f).isBounded.union (isBounded_image_iff.2 ⟨1, fun x hx y hy ↦ ?_⟩) + exact le_of_lt (hs1 (x, y) ⟨hx, hy⟩) +#align metric.bounded_range_of_tendsto_cofinite_uniformity Metric.isBounded_range_of_tendsto_cofinite_uniformity + +theorem isBounded_range_of_cauchy_map_cofinite {f : β → α} (hf : Cauchy (map f cofinite)) : + IsBounded (range f) := + isBounded_range_of_tendsto_cofinite_uniformity <| (cauchy_map_iff.1 hf).2 +#align metric.bounded_range_of_cauchy_map_cofinite Metric.isBounded_range_of_cauchy_map_cofinite + +theorem _root_.CauchySeq.isBounded_range {f : ℕ → α} (hf : CauchySeq f) : IsBounded (range f) := + isBounded_range_of_cauchy_map_cofinite <| by rwa [Nat.cofinite_eq_atTop] +#align cauchy_seq.bounded_range CauchySeq.isBounded_range + +theorem isBounded_range_of_tendsto_cofinite {f : β → α} {a : α} (hf : Tendsto f cofinite (𝓝 a)) : + IsBounded (range f) := + isBounded_range_of_tendsto_cofinite_uniformity <| + (hf.prod_map hf).mono_right <| nhds_prod_eq.symm.trans_le (nhds_le_uniformity a) +#align metric.bounded_range_of_tendsto_cofinite Metric.isBounded_range_of_tendsto_cofinite + +/-- In a compact space, all sets are bounded -/ +theorem isBounded_of_compactSpace [CompactSpace α] : IsBounded s := + isCompact_univ.isBounded.subset (subset_univ _) +#align metric.bounded_of_compact_space Metric.isBounded_of_compactSpace + +theorem isBounded_range_of_tendsto (u : ℕ → α) {x : α} (hu : Tendsto u atTop (𝓝 x)) : + IsBounded (range u) := + hu.cauchySeq.isBounded_range +#align metric.bounded_range_of_tendsto Metric.isBounded_range_of_tendsto + +theorem disjoint_nhds_cobounded (x : α) : Disjoint (𝓝 x) (cobounded α) := + disjoint_of_disjoint_of_mem disjoint_compl_right (ball_mem_nhds _ one_pos) isBounded_ball + +theorem disjoint_cobounded_nhds (x : α) : Disjoint (cobounded α) (𝓝 x) := + (disjoint_nhds_cobounded x).symm + +theorem disjoint_nhdsSet_cobounded {s : Set α} (hs : IsCompact s) : Disjoint (𝓝ˢ s) (cobounded α) := + hs.disjoint_nhdsSet_left.2 fun _ _ ↦ disjoint_nhds_cobounded _ + +theorem disjoint_cobounded_nhdsSet {s : Set α} (hs : IsCompact s) : Disjoint (cobounded α) (𝓝ˢ s) := + (disjoint_nhdsSet_cobounded hs).symm + +/-- If a function is continuous within a set `s` at every point of a compact set `k`, then it is +bounded on some open neighborhood of `k` in `s`. -/ +theorem exists_isOpen_isBounded_image_inter_of_isCompact_of_forall_continuousWithinAt + [TopologicalSpace β] {k s : Set β} {f : β → α} (hk : IsCompact k) + (hf : ∀ x ∈ k, ContinuousWithinAt f s x) : + ∃ t, k ⊆ t ∧ IsOpen t ∧ IsBounded (f '' (t ∩ s)) := by + have : Disjoint (𝓝ˢ k ⊓ 𝓟 s) (comap f (cobounded α)) := by + rw [disjoint_assoc, inf_comm, hk.disjoint_nhdsSet_left] + exact fun x hx ↦ disjoint_left_comm.2 <| + tendsto_comap.disjoint (disjoint_cobounded_nhds _) (hf x hx) + rcases ((((hasBasis_nhdsSet _).inf_principal _)).disjoint_iff ((basis_sets _).comap _)).1 this + with ⟨U, ⟨hUo, hkU⟩, t, ht, hd⟩ + refine ⟨U, hkU, hUo, (isBounded_compl_iff.2 ht).subset ?_⟩ + rwa [image_subset_iff, preimage_compl, subset_compl_iff_disjoint_right] +#align metric.exists_is_open_bounded_image_inter_of_is_compact_of_forall_continuous_within_at Metric.exists_isOpen_isBounded_image_inter_of_isCompact_of_forall_continuousWithinAt + +/-- If a function is continuous at every point of a compact set `k`, then it is bounded on +some open neighborhood of `k`. -/ +theorem exists_isOpen_isBounded_image_of_isCompact_of_forall_continuousAt [TopologicalSpace β] + {k : Set β} {f : β → α} (hk : IsCompact k) (hf : ∀ x ∈ k, ContinuousAt f x) : + ∃ t, k ⊆ t ∧ IsOpen t ∧ IsBounded (f '' t) := by + simp_rw [← continuousWithinAt_univ] at hf + simpa only [inter_univ] using + exists_isOpen_isBounded_image_inter_of_isCompact_of_forall_continuousWithinAt hk hf +#align metric.exists_is_open_bounded_image_of_is_compact_of_forall_continuous_at Metric.exists_isOpen_isBounded_image_of_isCompact_of_forall_continuousAt + +/-- If a function is continuous on a set `s` containing a compact set `k`, then it is bounded on +some open neighborhood of `k` in `s`. -/ +theorem exists_isOpen_isBounded_image_inter_of_isCompact_of_continuousOn [TopologicalSpace β] + {k s : Set β} {f : β → α} (hk : IsCompact k) (hks : k ⊆ s) (hf : ContinuousOn f s) : + ∃ t, k ⊆ t ∧ IsOpen t ∧ IsBounded (f '' (t ∩ s)) := + exists_isOpen_isBounded_image_inter_of_isCompact_of_forall_continuousWithinAt hk fun x hx => + hf x (hks hx) +#align metric.exists_is_open_bounded_image_inter_of_is_compact_of_continuous_on Metric.exists_isOpen_isBounded_image_inter_of_isCompact_of_continuousOn + +/-- If a function is continuous on a neighborhood of a compact set `k`, then it is bounded on +some open neighborhood of `k`. -/ +theorem exists_isOpen_isBounded_image_of_isCompact_of_continuousOn [TopologicalSpace β] + {k s : Set β} {f : β → α} (hk : IsCompact k) (hs : IsOpen s) (hks : k ⊆ s) + (hf : ContinuousOn f s) : ∃ t, k ⊆ t ∧ IsOpen t ∧ IsBounded (f '' t) := + exists_isOpen_isBounded_image_of_isCompact_of_forall_continuousAt hk fun _x hx => + hf.continuousAt (hs.mem_nhds (hks hx)) +#align metric.exists_is_open_bounded_image_of_is_compact_of_continuous_on Metric.exists_isOpen_isBounded_image_of_isCompact_of_continuousOn + +/-- The **Heine–Borel theorem**: In a proper space, a closed bounded set is compact. -/ +theorem isCompact_of_isClosed_isBounded [ProperSpace α] (hc : IsClosed s) (hb : IsBounded s) : + IsCompact s := by + rcases eq_empty_or_nonempty s with (rfl | ⟨x, -⟩) + · exact isCompact_empty + · rcases hb.subset_closedBall x with ⟨r, hr⟩ + exact (isCompact_closedBall x r).of_isClosed_subset hc hr +#align metric.is_compact_of_is_closed_bounded Metric.isCompact_of_isClosed_isBounded + +/-- The **Heine–Borel theorem**: In a proper space, the closure of a bounded set is compact. -/ +theorem _root_.Bornology.IsBounded.isCompact_closure [ProperSpace α] (h : IsBounded s) : + IsCompact (closure s) := + isCompact_of_isClosed_isBounded isClosed_closure h.closure +#align metric.bounded.is_compact_closure Bornology.IsBounded.isCompact_closure + +-- porting note: todo: assume `[MetricSpace α]` instead of `[PseudoMetricSpace α] [T2Space α]` +/-- The **Heine–Borel theorem**: +In a proper Hausdorff space, a set is compact if and only if it is closed and bounded. -/ +theorem isCompact_iff_isClosed_bounded [T2Space α] [ProperSpace α] : + IsCompact s ↔ IsClosed s ∧ IsBounded s := + ⟨fun h => ⟨h.isClosed, h.isBounded⟩, fun h => isCompact_of_isClosed_isBounded h.1 h.2⟩ +#align metric.is_compact_iff_is_closed_bounded Metric.isCompact_iff_isClosed_bounded + +theorem compactSpace_iff_isBounded_univ [ProperSpace α] : + CompactSpace α ↔ IsBounded (univ : Set α) := + ⟨@isBounded_of_compactSpace α _ _, fun hb => ⟨isCompact_of_isClosed_isBounded isClosed_univ hb⟩⟩ +#align metric.compact_space_iff_bounded_univ Metric.compactSpace_iff_isBounded_univ + +section ConditionallyCompleteLinearOrder + +variable [Preorder α] [CompactIccSpace α] + +theorem isBounded_Icc (a b : α) : IsBounded (Icc a b) := + (totallyBounded_Icc a b).isBounded +#align metric.bounded_Icc Metric.isBounded_Icc + +theorem isBounded_Ico (a b : α) : IsBounded (Ico a b) := + (totallyBounded_Ico a b).isBounded +#align metric.bounded_Ico Metric.isBounded_Ico + +theorem isBounded_Ioc (a b : α) : IsBounded (Ioc a b) := + (totallyBounded_Ioc a b).isBounded +#align metric.bounded_Ioc Metric.isBounded_Ioc + +theorem isBounded_Ioo (a b : α) : IsBounded (Ioo a b) := + (totallyBounded_Ioo a b).isBounded +#align metric.bounded_Ioo Metric.isBounded_Ioo + +/-- In a pseudo metric space with a conditionally complete linear order such that the order and the + metric structure give the same topology, any order-bounded set is metric-bounded. -/ +theorem isBounded_of_bddAbove_of_bddBelow {s : Set α} (h₁ : BddAbove s) (h₂ : BddBelow s) : + IsBounded s := + let ⟨u, hu⟩ := h₁ + let ⟨l, hl⟩ := h₂ + (isBounded_Icc l u).subset (fun _x hx => mem_Icc.mpr ⟨hl hx, hu hx⟩) +#align metric.bounded_of_bdd_above_of_bdd_below Metric.isBounded_of_bddAbove_of_bddBelow + +end ConditionallyCompleteLinearOrder + +end Bounded + +section Diam + +variable {s : Set α} {x y z : α} + +/-- The diameter of a set in a metric space. To get controllable behavior even when the diameter +should be infinite, we express it in terms of the emetric.diameter -/ +noncomputable def diam (s : Set α) : ℝ := + ENNReal.toReal (EMetric.diam s) +#align metric.diam Metric.diam + +/-- The diameter of a set is always nonnegative -/ +theorem diam_nonneg : 0 ≤ diam s := + ENNReal.toReal_nonneg +#align metric.diam_nonneg Metric.diam_nonneg + +theorem diam_subsingleton (hs : s.Subsingleton) : diam s = 0 := by + simp only [diam, EMetric.diam_subsingleton hs, ENNReal.zero_toReal] +#align metric.diam_subsingleton Metric.diam_subsingleton + +/-- The empty set has zero diameter -/ +@[simp] +theorem diam_empty : diam (∅ : Set α) = 0 := + diam_subsingleton subsingleton_empty +#align metric.diam_empty Metric.diam_empty + +/-- A singleton has zero diameter -/ +@[simp] +theorem diam_singleton : diam ({x} : Set α) = 0 := + diam_subsingleton subsingleton_singleton +#align metric.diam_singleton Metric.diam_singleton + +@[to_additive (attr := simp)] +theorem diam_one [One α] : diam (1 : Set α) = 0 := + diam_singleton +#align metric.diam_one Metric.diam_one +#align metric.diam_zero Metric.diam_zero + +-- Does not work as a simp-lemma, since {x, y} reduces to (insert y {x}) +theorem diam_pair : diam ({x, y} : Set α) = dist x y := by + simp only [diam, EMetric.diam_pair, dist_edist] +#align metric.diam_pair Metric.diam_pair + +-- Does not work as a simp-lemma, since {x, y, z} reduces to (insert z (insert y {x})) +theorem diam_triple : + Metric.diam ({x, y, z} : Set α) = max (max (dist x y) (dist x z)) (dist y z) := by + simp only [Metric.diam, EMetric.diam_triple, dist_edist] + rw [ENNReal.toReal_max, ENNReal.toReal_max] <;> apply_rules [ne_of_lt, edist_lt_top, max_lt] +#align metric.diam_triple Metric.diam_triple + +/-- If the distance between any two points in a set is bounded by some constant `C`, +then `ENNReal.ofReal C` bounds the emetric diameter of this set. -/ +theorem ediam_le_of_forall_dist_le {C : ℝ} (h : ∀ x ∈ s, ∀ y ∈ s, dist x y ≤ C) : + EMetric.diam s ≤ ENNReal.ofReal C := + EMetric.diam_le fun x hx y hy => (edist_dist x y).symm ▸ ENNReal.ofReal_le_ofReal (h x hx y hy) +#align metric.ediam_le_of_forall_dist_le Metric.ediam_le_of_forall_dist_le + +/-- If the distance between any two points in a set is bounded by some non-negative constant, +this constant bounds the diameter. -/ +theorem diam_le_of_forall_dist_le {C : ℝ} (h₀ : 0 ≤ C) (h : ∀ x ∈ s, ∀ y ∈ s, dist x y ≤ C) : + diam s ≤ C := + ENNReal.toReal_le_of_le_ofReal h₀ (ediam_le_of_forall_dist_le h) +#align metric.diam_le_of_forall_dist_le Metric.diam_le_of_forall_dist_le + +/-- If the distance between any two points in a nonempty set is bounded by some constant, +this constant bounds the diameter. -/ +theorem diam_le_of_forall_dist_le_of_nonempty (hs : s.Nonempty) {C : ℝ} + (h : ∀ x ∈ s, ∀ y ∈ s, dist x y ≤ C) : diam s ≤ C := + have h₀ : 0 ≤ C := + let ⟨x, hx⟩ := hs + le_trans dist_nonneg (h x hx x hx) + diam_le_of_forall_dist_le h₀ h +#align metric.diam_le_of_forall_dist_le_of_nonempty Metric.diam_le_of_forall_dist_le_of_nonempty + +/-- The distance between two points in a set is controlled by the diameter of the set. -/ +theorem dist_le_diam_of_mem' (h : EMetric.diam s ≠ ⊤) (hx : x ∈ s) (hy : y ∈ s) : + dist x y ≤ diam s := by + rw [diam, dist_edist] + rw [ENNReal.toReal_le_toReal (edist_ne_top _ _) h] + exact EMetric.edist_le_diam_of_mem hx hy +#align metric.dist_le_diam_of_mem' Metric.dist_le_diam_of_mem' + +/-- Characterize the boundedness of a set in terms of the finiteness of its emetric.diameter. -/ +theorem isBounded_iff_ediam_ne_top : IsBounded s ↔ EMetric.diam s ≠ ⊤ := + isBounded_iff.trans <| Iff.intro + (fun ⟨_C, hC⟩ => ne_top_of_le_ne_top ENNReal.ofReal_ne_top <| ediam_le_of_forall_dist_le hC) + fun h => ⟨diam s, fun _x hx _y hy => dist_le_diam_of_mem' h hx hy⟩ +#align metric.bounded_iff_ediam_ne_top Metric.isBounded_iff_ediam_ne_top + +alias ⟨_root_.Bornology.IsBounded.ediam_ne_top, _⟩ := isBounded_iff_ediam_ne_top +#align metric.bounded.ediam_ne_top Bornology.IsBounded.ediam_ne_top + +theorem ediam_eq_top_iff_unbounded : EMetric.diam s = ⊤ ↔ ¬IsBounded s := + isBounded_iff_ediam_ne_top.not_left.symm + +theorem ediam_univ_eq_top_iff_noncompact [ProperSpace α] : + EMetric.diam (univ : Set α) = ∞ ↔ NoncompactSpace α := by + rw [← not_compactSpace_iff, compactSpace_iff_isBounded_univ, isBounded_iff_ediam_ne_top, + Classical.not_not] +#align metric.ediam_univ_eq_top_iff_noncompact Metric.ediam_univ_eq_top_iff_noncompact + +@[simp] +theorem ediam_univ_of_noncompact [ProperSpace α] [NoncompactSpace α] : + EMetric.diam (univ : Set α) = ∞ := + ediam_univ_eq_top_iff_noncompact.mpr ‹_› +#align metric.ediam_univ_of_noncompact Metric.ediam_univ_of_noncompact + +@[simp] +theorem diam_univ_of_noncompact [ProperSpace α] [NoncompactSpace α] : diam (univ : Set α) = 0 := by + simp [diam] +#align metric.diam_univ_of_noncompact Metric.diam_univ_of_noncompact + +/-- The distance between two points in a set is controlled by the diameter of the set. -/ +theorem dist_le_diam_of_mem (h : IsBounded s) (hx : x ∈ s) (hy : y ∈ s) : dist x y ≤ diam s := + dist_le_diam_of_mem' h.ediam_ne_top hx hy +#align metric.dist_le_diam_of_mem Metric.dist_le_diam_of_mem + +theorem ediam_of_unbounded (h : ¬IsBounded s) : EMetric.diam s = ∞ := ediam_eq_top_iff_unbounded.2 h +#align metric.ediam_of_unbounded Metric.ediam_of_unbounded + +/-- An unbounded set has zero diameter. If you would prefer to get the value ∞, use `EMetric.diam`. +This lemma makes it possible to avoid side conditions in some situations -/ +theorem diam_eq_zero_of_unbounded (h : ¬IsBounded s) : diam s = 0 := by + rw [diam, ediam_of_unbounded h, ENNReal.top_toReal] +#align metric.diam_eq_zero_of_unbounded Metric.diam_eq_zero_of_unbounded + +/-- If `s ⊆ t`, then the diameter of `s` is bounded by that of `t`, provided `t` is bounded. -/ +theorem diam_mono {s t : Set α} (h : s ⊆ t) (ht : IsBounded t) : diam s ≤ diam t := + ENNReal.toReal_mono ht.ediam_ne_top <| EMetric.diam_mono h +#align metric.diam_mono Metric.diam_mono + +/-- The diameter of a union is controlled by the sum of the diameters, and the distance between +any two points in each of the sets. This lemma is true without any side condition, since it is +obviously true if `s ∪ t` is unbounded. -/ +theorem diam_union {t : Set α} (xs : x ∈ s) (yt : y ∈ t) : + diam (s ∪ t) ≤ diam s + dist x y + diam t := by + simp only [diam, dist_edist] + refine (ENNReal.toReal_le_add' (EMetric.diam_union xs yt) ?_ ?_).trans + (add_le_add_right ENNReal.toReal_add_le _) + · simp only [ENNReal.add_eq_top, edist_ne_top, or_false] + exact fun h ↦ top_unique <| h ▸ EMetric.diam_mono (subset_union_left _ _) + · exact fun h ↦ top_unique <| h ▸ EMetric.diam_mono (subset_union_right _ _) +#align metric.diam_union Metric.diam_union + +/-- If two sets intersect, the diameter of the union is bounded by the sum of the diameters. -/ +theorem diam_union' {t : Set α} (h : (s ∩ t).Nonempty) : diam (s ∪ t) ≤ diam s + diam t := by + rcases h with ⟨x, ⟨xs, xt⟩⟩ + simpa using diam_union xs xt +#align metric.diam_union' Metric.diam_union' + +theorem diam_le_of_subset_closedBall {r : ℝ} (hr : 0 ≤ r) (h : s ⊆ closedBall x r) : + diam s ≤ 2 * r := + diam_le_of_forall_dist_le (mul_nonneg zero_le_two hr) fun a ha b hb => + calc + dist a b ≤ dist a x + dist b x := dist_triangle_right _ _ _ + _ ≤ r + r := (add_le_add (h ha) (h hb)) + _ = 2 * r := by simp [mul_two, mul_comm] +#align metric.diam_le_of_subset_closed_ball Metric.diam_le_of_subset_closedBall + +/-- The diameter of a closed ball of radius `r` is at most `2 r`. -/ +theorem diam_closedBall {r : ℝ} (h : 0 ≤ r) : diam (closedBall x r) ≤ 2 * r := + diam_le_of_subset_closedBall h Subset.rfl +#align metric.diam_closed_ball Metric.diam_closedBall + +/-- The diameter of a ball of radius `r` is at most `2 r`. -/ +theorem diam_ball {r : ℝ} (h : 0 ≤ r) : diam (ball x r) ≤ 2 * r := + diam_le_of_subset_closedBall h ball_subset_closedBall +#align metric.diam_ball Metric.diam_ball + +/-- If a family of complete sets with diameter tending to `0` is such that each finite intersection +is nonempty, then the total intersection is also nonempty. -/ +theorem _root_.IsComplete.nonempty_iInter_of_nonempty_biInter {s : ℕ → Set α} + (h0 : IsComplete (s 0)) (hs : ∀ n, IsClosed (s n)) (h's : ∀ n, IsBounded (s n)) + (h : ∀ N, (⋂ n ≤ N, s n).Nonempty) (h' : Tendsto (fun n => diam (s n)) atTop (𝓝 0)) : + (⋂ n, s n).Nonempty := by + let u N := (h N).some + have I : ∀ n N, n ≤ N → u N ∈ s n := by + intro n N hn + apply mem_of_subset_of_mem _ (h N).choose_spec + intro x hx + simp only [mem_iInter] at hx + exact hx n hn + have : CauchySeq u := by + apply cauchySeq_of_le_tendsto_0 _ _ h' + intro m n N hm hn + exact dist_le_diam_of_mem (h's N) (I _ _ hm) (I _ _ hn) + obtain ⟨x, -, xlim⟩ : ∃ x ∈ s 0, Tendsto (fun n : ℕ => u n) atTop (𝓝 x) := + cauchySeq_tendsto_of_isComplete h0 (fun n => I 0 n (zero_le _)) this + refine' ⟨x, mem_iInter.2 fun n => _⟩ + apply (hs n).mem_of_tendsto xlim + filter_upwards [Ici_mem_atTop n] with p hp + exact I n p hp +#align is_complete.nonempty_Inter_of_nonempty_bInter IsComplete.nonempty_iInter_of_nonempty_biInter + +/-- In a complete space, if a family of closed sets with diameter tending to `0` is such that each +finite intersection is nonempty, then the total intersection is also nonempty. -/ +theorem nonempty_iInter_of_nonempty_biInter [CompleteSpace α] {s : ℕ → Set α} + (hs : ∀ n, IsClosed (s n)) (h's : ∀ n, IsBounded (s n)) (h : ∀ N, (⋂ n ≤ N, s n).Nonempty) + (h' : Tendsto (fun n => diam (s n)) atTop (𝓝 0)) : (⋂ n, s n).Nonempty := + (hs 0).isComplete.nonempty_iInter_of_nonempty_biInter hs h's h h' +#align metric.nonempty_Inter_of_nonempty_bInter Metric.nonempty_iInter_of_nonempty_biInter + +end Diam + +end Metric + +namespace Mathlib.Meta.Positivity + +open Lean Meta Qq Function + +/-- Extension for the `positivity` tactic: the diameter of a set is always nonnegative. -/ +@[positivity Metric.diam _] +def evalDiam : PositivityExt where eval {_ _} _zα _pα e := do + let .app _ s ← whnfR e | throwError "not Metric.diam" + let p ← mkAppOptM ``Metric.diam_nonneg #[none, none, s] + pure (.nonnegative p) + +end Mathlib.Meta.Positivity + +open Metric + +theorem Metric.cobounded_eq_cocompact [ProperSpace α] : cobounded α = cocompact α := by + nontriviality α; inhabit α + exact cobounded_le_cocompact.antisymm <| (hasBasis_cobounded_compl_closedBall default).ge_iff.2 + fun _ _ ↦ (isCompact_closedBall _ _).compl_mem_cocompact +#align comap_dist_right_at_top_eq_cocompact Metric.cobounded_eq_cocompact + +theorem tendsto_dist_right_cocompact_atTop [ProperSpace α] (x : α) : + Tendsto (dist · x) (cocompact α) atTop := + (tendsto_dist_right_cobounded_atTop x).mono_left cobounded_eq_cocompact.ge +#align tendsto_dist_right_cocompact_at_top tendsto_dist_right_cocompact_atTop + +theorem tendsto_dist_left_cocompact_atTop [ProperSpace α] (x : α) : + Tendsto (dist x) (cocompact α) atTop := + (tendsto_dist_left_cobounded_atTop x).mono_left cobounded_eq_cocompact.ge +#align tendsto_dist_left_cocompact_at_top tendsto_dist_left_cocompact_atTop + +theorem comap_dist_left_atTop_eq_cocompact [ProperSpace α] (x : α) : + comap (dist x) atTop = cocompact α := by simp [cobounded_eq_cocompact] +#align comap_dist_left_at_top_eq_cocompact comap_dist_left_atTop_eq_cocompact + +theorem tendsto_cocompact_of_tendsto_dist_comp_atTop {f : β → α} {l : Filter β} (x : α) + (h : Tendsto (fun y => dist (f y) x) l atTop) : Tendsto f l (cocompact α) := + ((tendsto_dist_right_atTop_iff _).1 h).mono_right cobounded_le_cocompact +#align tendsto_cocompact_of_tendsto_dist_comp_at_top tendsto_cocompact_of_tendsto_dist_comp_atTop + +theorem Metric.finite_isBounded_inter_isClosed [ProperSpace α] {K s : Set α} [DiscreteTopology s] + (hK : IsBounded K) (hs : IsClosed s) : Set.Finite (K ∩ s) := by + refine Set.Finite.subset (IsCompact.finite ?_ ?_) (Set.inter_subset_inter_left s subset_closure) + · exact hK.isCompact_closure.inter_right hs + · exact DiscreteTopology.of_subset inferInstance (Set.inter_subset_right _ s) diff --git a/Mathlib/Topology/MetricSpace/Cauchy.lean b/Mathlib/Topology/MetricSpace/Cauchy.lean new file mode 100644 index 0000000000000..f9fdde148de4e --- /dev/null +++ b/Mathlib/Topology/MetricSpace/Cauchy.lean @@ -0,0 +1,160 @@ +/- +Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel +-/ +import Mathlib.Topology.MetricSpace.Basic + +/-! +## Cauchy sequences in (pseudo-)metric spaces + +Various results on Cauchy sequences in (pseudo-)metric spaces, including + +* `Metric.complete_of_cauchySeq_tendsto` A pseudo-metric space is complete iff each Cauchy sequences +converges to some limit point. +* `cauchySeq_bdd`: a Cauchy sequence on the natural numbers is bounded +* various characterisation of Cauchy and uniformly Cauchy sequences + +## Tags + +metric, pseudo_metric, Cauchy sequence +-/ + +open Filter +open scoped Uniformity Topology + +universe u v w + +variable {α : Type u} {β : Type v} {X ι : Type*} +variable [PseudoMetricSpace α] + +/-- A very useful criterion to show that a space is complete is to show that all sequences +which satisfy a bound of the form `dist (u n) (u m) < B N` for all `n m ≥ N` are +converging. This is often applied for `B N = 2^{-N}`, i.e., with a very fast convergence to +`0`, which makes it possible to use arguments of converging series, while this is impossible +to do in general for arbitrary Cauchy sequences. -/ +theorem Metric.complete_of_convergent_controlled_sequences (B : ℕ → Real) (hB : ∀ n, 0 < B n) + (H : ∀ u : ℕ → α, (∀ N n m : ℕ, N ≤ n → N ≤ m → dist (u n) (u m) < B N) → + ∃ x, Tendsto u atTop (𝓝 x)) : + CompleteSpace α := + UniformSpace.complete_of_convergent_controlled_sequences + (fun n => { p : α × α | dist p.1 p.2 < B n }) (fun n => dist_mem_uniformity <| hB n) H +#align metric.complete_of_convergent_controlled_sequences Metric.complete_of_convergent_controlled_sequences + +/-- A pseudo-metric space is complete iff every Cauchy sequence converges. -/ +theorem Metric.complete_of_cauchySeq_tendsto : + (∀ u : ℕ → α, CauchySeq u → ∃ a, Tendsto u atTop (𝓝 a)) → CompleteSpace α := + EMetric.complete_of_cauchySeq_tendsto +#align metric.complete_of_cauchy_seq_tendsto Metric.complete_of_cauchySeq_tendsto + +section CauchySeq + +variable [Nonempty β] [SemilatticeSup β] + +/-- In a pseudometric space, Cauchy sequences are characterized by the fact that, eventually, +the distance between its elements is arbitrarily small -/ +-- porting note: @[nolint ge_or_gt] doesn't exist +theorem Metric.cauchySeq_iff {u : β → α} : + CauchySeq u ↔ ∀ ε > 0, ∃ N, ∀ m ≥ N, ∀ n ≥ N, dist (u m) (u n) < ε := + uniformity_basis_dist.cauchySeq_iff +#align metric.cauchy_seq_iff Metric.cauchySeq_iff + +/-- A variation around the pseudometric characterization of Cauchy sequences -/ +theorem Metric.cauchySeq_iff' {u : β → α} : + CauchySeq u ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, dist (u n) (u N) < ε := + uniformity_basis_dist.cauchySeq_iff' +#align metric.cauchy_seq_iff' Metric.cauchySeq_iff' + +-- see Note [nolint_ge] +/-- In a pseudometric space, uniform Cauchy sequences are characterized by the fact that, +eventually, the distance between all its elements is uniformly, arbitrarily small. -/ +-- porting note: no attr @[nolint ge_or_gt] +theorem Metric.uniformCauchySeqOn_iff {γ : Type*} {F : β → γ → α} {s : Set γ} : + UniformCauchySeqOn F atTop s ↔ ∀ ε > (0 : ℝ), + ∃ N : β, ∀ m ≥ N, ∀ n ≥ N, ∀ x ∈ s, dist (F m x) (F n x) < ε := by + constructor + · intro h ε hε + let u := { a : α × α | dist a.fst a.snd < ε } + have hu : u ∈ 𝓤 α := Metric.mem_uniformity_dist.mpr ⟨ε, hε, by simp⟩ + rw [← @Filter.eventually_atTop_prod_self' _ _ _ fun m => + ∀ x ∈ s, dist (F m.fst x) (F m.snd x) < ε] + specialize h u hu + rw [prod_atTop_atTop_eq] at h + exact h.mono fun n h x hx => h x hx + · intro h u hu + rcases Metric.mem_uniformity_dist.mp hu with ⟨ε, hε, hab⟩ + rcases h ε hε with ⟨N, hN⟩ + rw [prod_atTop_atTop_eq, eventually_atTop] + use (N, N) + intro b hb x hx + rcases hb with ⟨hbl, hbr⟩ + exact hab (hN b.fst hbl.ge b.snd hbr.ge x hx) +#align metric.uniform_cauchy_seq_on_iff Metric.uniformCauchySeqOn_iff + +/-- If the distance between `s n` and `s m`, `n ≤ m` is bounded above by `b n` +and `b` converges to zero, then `s` is a Cauchy sequence. -/ +theorem cauchySeq_of_le_tendsto_0' {s : β → α} (b : β → ℝ) + (h : ∀ n m : β, n ≤ m → dist (s n) (s m) ≤ b n) (h₀ : Tendsto b atTop (𝓝 0)) : CauchySeq s := + Metric.cauchySeq_iff'.2 fun ε ε0 => (h₀.eventually (gt_mem_nhds ε0)).exists.imp fun N hN n hn => + calc dist (s n) (s N) = dist (s N) (s n) := dist_comm _ _ + _ ≤ b N := h _ _ hn + _ < ε := hN +#align cauchy_seq_of_le_tendsto_0' cauchySeq_of_le_tendsto_0' + +/-- If the distance between `s n` and `s m`, `n, m ≥ N` is bounded above by `b N` +and `b` converges to zero, then `s` is a Cauchy sequence. -/ +theorem cauchySeq_of_le_tendsto_0 {s : β → α} (b : β → ℝ) + (h : ∀ n m N : β, N ≤ n → N ≤ m → dist (s n) (s m) ≤ b N) (h₀ : Tendsto b atTop (𝓝 0)) : + CauchySeq s := + cauchySeq_of_le_tendsto_0' b (fun _n _m hnm => h _ _ _ le_rfl hnm) h₀ +#align cauchy_seq_of_le_tendsto_0 cauchySeq_of_le_tendsto_0 + +/-- A Cauchy sequence on the natural numbers is bounded. -/ +theorem cauchySeq_bdd {u : ℕ → α} (hu : CauchySeq u) : ∃ R > 0, ∀ m n, dist (u m) (u n) < R := by + rcases Metric.cauchySeq_iff'.1 hu 1 zero_lt_one with ⟨N, hN⟩ + suffices : ∃ R > 0, ∀ n, dist (u n) (u N) < R + · rcases this with ⟨R, R0, H⟩ + exact ⟨_, add_pos R0 R0, fun m n => + lt_of_le_of_lt (dist_triangle_right _ _ _) (add_lt_add (H m) (H n))⟩ + let R := Finset.sup (Finset.range N) fun n => nndist (u n) (u N) + refine' ⟨↑R + 1, add_pos_of_nonneg_of_pos R.2 zero_lt_one, fun n => _⟩ + cases' le_or_lt N n with h h + · exact lt_of_lt_of_le (hN _ h) (le_add_of_nonneg_left R.2) + · have : _ ≤ R := Finset.le_sup (Finset.mem_range.2 h) + exact lt_of_le_of_lt this (lt_add_of_pos_right _ zero_lt_one) +#align cauchy_seq_bdd cauchySeq_bdd + +/-- Yet another metric characterization of Cauchy sequences on integers. This one is often the +most efficient. -/ +theorem cauchySeq_iff_le_tendsto_0 {s : ℕ → α} : + CauchySeq s ↔ + ∃ b : ℕ → ℝ, + (∀ n, 0 ≤ b n) ∧ + (∀ n m N : ℕ, N ≤ n → N ≤ m → dist (s n) (s m) ≤ b N) ∧ Tendsto b atTop (𝓝 0) := + ⟨fun hs => by + /- `s` is a Cauchy sequence. The sequence `b` will be constructed by taking + the supremum of the distances between `s n` and `s m` for `n m ≥ N`. + First, we prove that all these distances are bounded, as otherwise the Sup + would not make sense. -/ + let S N := (fun p : ℕ × ℕ => dist (s p.1) (s p.2)) '' { p | p.1 ≥ N ∧ p.2 ≥ N } + have hS : ∀ N, ∃ x, ∀ y ∈ S N, y ≤ x := by + rcases cauchySeq_bdd hs with ⟨R, -, hR⟩ + refine' fun N => ⟨R, _⟩ + rintro _ ⟨⟨m, n⟩, _, rfl⟩ + exact le_of_lt (hR m n) + -- Prove that it bounds the distances of points in the Cauchy sequence + have ub : ∀ m n N, N ≤ m → N ≤ n → dist (s m) (s n) ≤ sSup (S N) := fun m n N hm hn => + le_csSup (hS N) ⟨⟨_, _⟩, ⟨hm, hn⟩, rfl⟩ + have S0m : ∀ n, (0 : ℝ) ∈ S n := fun n => ⟨⟨n, n⟩, ⟨le_rfl, le_rfl⟩, dist_self _⟩ + have S0 := fun n => le_csSup (hS n) (S0m n) + -- Prove that it tends to `0`, by using the Cauchy property of `s` + refine' ⟨fun N => sSup (S N), S0, ub, Metric.tendsto_atTop.2 fun ε ε0 => _⟩ + refine' (Metric.cauchySeq_iff.1 hs (ε / 2) (half_pos ε0)).imp fun N hN n hn => _ + rw [Real.dist_0_eq_abs, abs_of_nonneg (S0 n)] + refine' lt_of_le_of_lt (csSup_le ⟨_, S0m _⟩ _) (half_lt_self ε0) + rintro _ ⟨⟨m', n'⟩, ⟨hm', hn'⟩, rfl⟩ + exact le_of_lt (hN _ (le_trans hn hm') _ (le_trans hn hn')), + fun ⟨b, _, b_bound, b_lim⟩ => cauchySeq_of_le_tendsto_0 b b_bound b_lim⟩ +#align cauchy_seq_iff_le_tendsto_0 cauchySeq_iff_le_tendsto_0 + +end CauchySeq diff --git a/Mathlib/Topology/MetricSpace/Equicontinuity.lean b/Mathlib/Topology/MetricSpace/Equicontinuity.lean index 6eb1be3e68419..98f88a49aa12c 100644 --- a/Mathlib/Topology/MetricSpace/Equicontinuity.lean +++ b/Mathlib/Topology/MetricSpace/Equicontinuity.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ -import Mathlib.Topology.MetricSpace.Basic +import Mathlib.Topology.MetricSpace.PseudoMetric import Mathlib.Topology.UniformSpace.Equicontinuity #align_import topology.metric_space.equicontinuity from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" diff --git a/Mathlib/Topology/MetricSpace/Lipschitz.lean b/Mathlib/Topology/MetricSpace/Lipschitz.lean index 701b324b920b8..7d3808063b35d 100644 --- a/Mathlib/Topology/MetricSpace/Lipschitz.lean +++ b/Mathlib/Topology/MetricSpace/Lipschitz.lean @@ -7,8 +7,9 @@ import Mathlib.Data.Set.Intervals.ProjIcc import Mathlib.Topology.Algebra.Order.Field import Mathlib.Topology.Bornology.Hom import Mathlib.Topology.MetricSpace.Basic +import Mathlib.Topology.MetricSpace.Bounded -#align_import topology.metric_space.lipschitz from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" +#align_import topology.metric_space.lipschitz from "leanprover-community/mathlib"@"c8f305514e0d47dfaa710f5a52f0d21b588e6328" /-! # Lipschitz continuous functions @@ -540,6 +541,19 @@ protected theorem prod {g : α → γ} {Kf Kg : ℝ≥0} (hf : LipschitzOnWith K rw [ENNReal.coe_mono.map_max, Prod.edist_eq, ENNReal.max_mul] exact max_le_max (hf hx hy) (hg hx hy) +theorem ediam_image2_le (f : α → β → γ) {K₁ K₂ : ℝ≥0} (s : Set α) (t : Set β) + (hf₁ : ∀ b ∈ t, LipschitzOnWith K₁ (fun a => f a b) s) + (hf₂ : ∀ a ∈ s, LipschitzOnWith K₂ (f a) t) : + EMetric.diam (Set.image2 f s t) ≤ ↑K₁ * EMetric.diam s + ↑K₂ * EMetric.diam t := by + apply EMetric.diam_le + rintro _ ⟨a₁, b₁, ha₁, hb₁, rfl⟩ _ ⟨a₂, b₂, ha₂, hb₂, rfl⟩ + refine' (edist_triangle _ (f a₂ b₁) _).trans _ + exact + add_le_add + ((hf₁ b₁ hb₁ ha₁ ha₂).trans <| ENNReal.mul_left_mono <| EMetric.edist_le_diam_of_mem ha₁ ha₂) + ((hf₂ a₂ ha₂ hb₁ hb₂).trans <| ENNReal.mul_left_mono <| EMetric.edist_le_diam_of_mem hb₁ hb₂) +#align lipschitz_on_with.ediam_image2_le LipschitzOnWith.ediam_image2_le + end EMetric section Metric @@ -591,6 +605,18 @@ protected theorem iff_le_add_mul {f : α → ℝ} {K : ℝ≥0} : ⟨LipschitzOnWith.le_add_mul, LipschitzOnWith.of_le_add_mul K⟩ #align lipschitz_on_with.iff_le_add_mul LipschitzOnWith.iff_le_add_mul +theorem isBounded_image2 (f : α → β → γ) {K₁ K₂ : ℝ≥0} {s : Set α} {t : Set β} + (hs : Bornology.IsBounded s) (ht : Bornology.IsBounded t) + (hf₁ : ∀ b ∈ t, LipschitzOnWith K₁ (fun a => f a b) s) + (hf₂ : ∀ a ∈ s, LipschitzOnWith K₂ (f a) t) : Bornology.IsBounded (Set.image2 f s t) := + Metric.isBounded_iff_ediam_ne_top.2 <| + ne_top_of_le_ne_top + (ENNReal.add_ne_top.mpr + ⟨ENNReal.mul_ne_top ENNReal.coe_ne_top hs.ediam_ne_top, + ENNReal.mul_ne_top ENNReal.coe_ne_top ht.ediam_ne_top⟩) + (ediam_image2_le _ _ _ hf₁ hf₂) +#align lipschitz_on_with.bounded_image2 LipschitzOnWith.isBounded_image2 + end Metric end LipschitzOnWith diff --git a/Mathlib/Topology/MetricSpace/PseudoMetric.lean b/Mathlib/Topology/MetricSpace/PseudoMetric.lean new file mode 100644 index 0000000000000..f5d4fb2f55fab --- /dev/null +++ b/Mathlib/Topology/MetricSpace/PseudoMetric.lean @@ -0,0 +1,2223 @@ +/- +Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel +-/ +import Mathlib.Tactic.Positivity +import Mathlib.Topology.Algebra.Order.Compact +import Mathlib.Topology.EMetricSpace.Basic +import Mathlib.Topology.Bornology.Constructions + +/-! +## Pseudo-metric spaces + +This file defines pseudo-metric spaces: these differ from metric spaces by not imposing the +condition `dist x y = 0 → x = y`. +Many definitions and theorems expected on (pseudo-)metric spaces are already introduced on uniform +spaces and topological spaces. For example: open and closed sets, compactness, completeness, +continuity and uniform continuity. + +## Main definitions + +* `Dist α`: Endows a space `α` with a function `dist a b`. +* `PseudoMetricSpace α`: A space endowed with a distance function, which can + be zero even if the two elements are non-equal. +* `Metric.ball x ε`: The set of all points `y` with `dist y x < ε`. +* `Metric.Bounded s`: Whether a subset of a `PseudoMetricSpace` is bounded. +* `MetricSpace α`: A `PseudoMetricSpace` with the guarantee `dist x y = 0 → x = y`. + +Additional useful definitions: + +* `nndist a b`: `dist` as a function to the non-negative reals. +* `Metric.closedBall x ε`: The set of all points `y` with `dist y x ≤ ε`. +* `Metric.sphere x ε`: The set of all points `y` with `dist y x = ε`. +* `ProperSpace α`: A `PseudoMetricSpace` where all closed balls are compact. + +TODO (anyone): Add "Main results" section. + +## Tags + +pseudo_metric, dist +-/ + +open Set Filter TopologicalSpace Bornology +open scoped BigOperators ENNReal NNReal Uniformity Topology + +universe u v w + +variable {α : Type u} {β : Type v} {X ι : Type*} + +/-- Construct a uniform structure from a distance function and metric space axioms -/ +def UniformSpace.ofDist (dist : α → α → ℝ) (dist_self : ∀ x : α, dist x x = 0) + (dist_comm : ∀ x y : α, dist x y = dist y x) + (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) : UniformSpace α := + .ofFun dist dist_self dist_comm dist_triangle fun ε ε0 => + ⟨ε / 2, half_pos ε0, fun _x hx _y hy => add_halves ε ▸ add_lt_add hx hy⟩ +#align uniform_space_of_dist UniformSpace.ofDist + +-- porting note: dropped the `dist_self` argument +/-- Construct a bornology from a distance function and metric space axioms. -/ +@[reducible] +def Bornology.ofDist {α : Type*} (dist : α → α → ℝ) (dist_comm : ∀ x y, dist x y = dist y x) + (dist_triangle : ∀ x y z, dist x z ≤ dist x y + dist y z) : Bornology α := + Bornology.ofBounded { s : Set α | ∃ C, ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → dist x y ≤ C } + ⟨0, fun x hx y => hx.elim⟩ (fun s ⟨c, hc⟩ t h => ⟨c, fun x hx y hy => hc (h hx) (h hy)⟩) + (fun s hs t ht => by + rcases s.eq_empty_or_nonempty with rfl | ⟨x, hx⟩ + · rwa [empty_union] + rcases t.eq_empty_or_nonempty with rfl | ⟨y, hy⟩ + · rwa [union_empty] + rsuffices ⟨C, hC⟩ : ∃ C, ∀ z ∈ s ∪ t, dist x z ≤ C + · refine ⟨C + C, fun a ha b hb => (dist_triangle a x b).trans ?_⟩ + simpa only [dist_comm] using add_le_add (hC _ ha) (hC _ hb) + rcases hs with ⟨Cs, hs⟩; rcases ht with ⟨Ct, ht⟩ + refine ⟨max Cs (dist x y + Ct), fun z hz => hz.elim + (fun hz => (hs hx hz).trans (le_max_left _ _)) + (fun hz => (dist_triangle x y z).trans <| + (add_le_add le_rfl (ht hy hz)).trans (le_max_right _ _))⟩) + fun z => ⟨dist z z, forall_eq.2 <| forall_eq.2 le_rfl⟩ +#align bornology.of_dist Bornology.ofDistₓ + +/-- The distance function (given an ambient metric space on `α`), which returns + a nonnegative real number `dist x y` given `x y : α`. -/ +@[ext] +class Dist (α : Type*) where + dist : α → α → ℝ +#align has_dist Dist + +export Dist (dist) + +-- the uniform structure and the emetric space structure are embedded in the metric space structure +-- to avoid instance diamond issues. See Note [forgetful inheritance]. +/-- This is an internal lemma used inside the default of `PseudoMetricSpace.edist`. -/ +private theorem dist_nonneg' {α} {x y : α} (dist : α → α → ℝ) + (dist_self : ∀ x : α, dist x x = 0) (dist_comm : ∀ x y : α, dist x y = dist y x) + (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) : 0 ≤ dist x y := + have : 0 ≤ 2 * dist x y := + calc 0 = dist x x := (dist_self _).symm + _ ≤ dist x y + dist y x := dist_triangle _ _ _ + _ = 2 * dist x y := by rw [two_mul, dist_comm] + nonneg_of_mul_nonneg_right this two_pos + +#noalign pseudo_metric_space.edist_dist_tac -- porting note: todo: restore + +/-- Pseudo metric and Metric spaces + +A pseudo metric space is endowed with a distance for which the requirement `d(x,y)=0 → x = y` might +not hold. A metric space is a pseudo metric space such that `d(x,y)=0 → x = y`. +Each pseudo metric space induces a canonical `UniformSpace` and hence a canonical +`TopologicalSpace` This is enforced in the type class definition, by extending the `UniformSpace` +structure. When instantiating a `PseudoMetricSpace` structure, the uniformity fields are not +necessary, they will be filled in by default. In the same way, each (pseudo) metric space induces a +(pseudo) emetric space structure. It is included in the structure, but filled in by default. +-/ +class PseudoMetricSpace (α : Type u) extends Dist α : Type u where + dist_self : ∀ x : α, dist x x = 0 + dist_comm : ∀ x y : α, dist x y = dist y x + dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z + edist : α → α → ℝ≥0∞ := fun x y => ENNReal.some ⟨dist x y, dist_nonneg' _ ‹_› ‹_› ‹_›⟩ + edist_dist : ∀ x y : α, edist x y = ENNReal.ofReal (dist x y) -- porting note: todo: add := by _ + toUniformSpace : UniformSpace α := .ofDist dist dist_self dist_comm dist_triangle + uniformity_dist : 𝓤 α = ⨅ ε > 0, 𝓟 { p : α × α | dist p.1 p.2 < ε } := by intros; rfl + toBornology : Bornology α := Bornology.ofDist dist dist_comm dist_triangle + cobounded_sets : (Bornology.cobounded α).sets = + { s | ∃ C : ℝ, ∀ x ∈ sᶜ, ∀ y ∈ sᶜ, dist x y ≤ C } := by intros; rfl +#align pseudo_metric_space PseudoMetricSpace + +/-- Two pseudo metric space structures with the same distance function coincide. -/ +@[ext] +theorem PseudoMetricSpace.ext {α : Type*} {m m' : PseudoMetricSpace α} + (h : m.toDist = m'.toDist) : m = m' := by + cases' m with d _ _ _ ed hed U hU B hB + cases' m' with d' _ _ _ ed' hed' U' hU' B' hB' + obtain rfl : d = d' := h + congr + · ext x y : 2 + rw [hed, hed'] + · exact UniformSpace.ext (hU.trans hU'.symm) + · ext : 2 + rw [← Filter.mem_sets, ← Filter.mem_sets, hB, hB'] +#align pseudo_metric_space.ext PseudoMetricSpace.ext + +variable [PseudoMetricSpace α] + +attribute [instance] PseudoMetricSpace.toUniformSpace PseudoMetricSpace.toBornology + +-- see Note [lower instance priority] +instance (priority := 200) PseudoMetricSpace.toEDist : EDist α := + ⟨PseudoMetricSpace.edist⟩ +#align pseudo_metric_space.to_has_edist PseudoMetricSpace.toEDist + +/-- Construct a pseudo-metric space structure whose underlying topological space structure +(definitionally) agrees which a pre-existing topology which is compatible with a given distance +function. -/ +def PseudoMetricSpace.ofDistTopology {α : Type u} [TopologicalSpace α] (dist : α → α → ℝ) + (dist_self : ∀ x : α, dist x x = 0) (dist_comm : ∀ x y : α, dist x y = dist y x) + (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) + (H : ∀ s : Set α, IsOpen s ↔ ∀ x ∈ s, ∃ ε > 0, ∀ y, dist x y < ε → y ∈ s) : + PseudoMetricSpace α := + { dist := dist + dist_self := dist_self + dist_comm := dist_comm + dist_triangle := dist_triangle + edist_dist := fun x y => by exact ENNReal.coe_nnreal_eq _ + toUniformSpace := + { toCore := (UniformSpace.ofDist dist dist_self dist_comm dist_triangle).toCore + isOpen_uniformity := fun s => (H s).trans <| forall₂_congr fun x _ => + ((UniformSpace.hasBasis_ofFun (exists_gt (0 : ℝ)) dist _ _ _ _).comap + (Prod.mk x)).mem_iff.symm.trans mem_comap_prod_mk } + uniformity_dist := rfl + toBornology := Bornology.ofDist dist dist_comm dist_triangle + cobounded_sets := rfl } +#align pseudo_metric_space.of_dist_topology PseudoMetricSpace.ofDistTopology + +@[simp] +theorem dist_self (x : α) : dist x x = 0 := + PseudoMetricSpace.dist_self x +#align dist_self dist_self + +theorem dist_comm (x y : α) : dist x y = dist y x := + PseudoMetricSpace.dist_comm x y +#align dist_comm dist_comm + +theorem edist_dist (x y : α) : edist x y = ENNReal.ofReal (dist x y) := + PseudoMetricSpace.edist_dist x y +#align edist_dist edist_dist + +theorem dist_triangle (x y z : α) : dist x z ≤ dist x y + dist y z := + PseudoMetricSpace.dist_triangle x y z +#align dist_triangle dist_triangle + +theorem dist_triangle_left (x y z : α) : dist x y ≤ dist z x + dist z y := by + rw [dist_comm z]; apply dist_triangle +#align dist_triangle_left dist_triangle_left + +theorem dist_triangle_right (x y z : α) : dist x y ≤ dist x z + dist y z := by + rw [dist_comm y]; apply dist_triangle +#align dist_triangle_right dist_triangle_right + +theorem dist_triangle4 (x y z w : α) : dist x w ≤ dist x y + dist y z + dist z w := + calc + dist x w ≤ dist x z + dist z w := dist_triangle x z w + _ ≤ dist x y + dist y z + dist z w := add_le_add_right (dist_triangle x y z) _ +#align dist_triangle4 dist_triangle4 + +theorem dist_triangle4_left (x₁ y₁ x₂ y₂ : α) : + dist x₂ y₂ ≤ dist x₁ y₁ + (dist x₁ x₂ + dist y₁ y₂) := by + rw [add_left_comm, dist_comm x₁, ← add_assoc] + apply dist_triangle4 +#align dist_triangle4_left dist_triangle4_left + +theorem dist_triangle4_right (x₁ y₁ x₂ y₂ : α) : + dist x₁ y₁ ≤ dist x₁ x₂ + dist y₁ y₂ + dist x₂ y₂ := by + rw [add_right_comm, dist_comm y₁] + apply dist_triangle4 +#align dist_triangle4_right dist_triangle4_right + +/-- The triangle (polygon) inequality for sequences of points; `Finset.Ico` version. -/ +theorem dist_le_Ico_sum_dist (f : ℕ → α) {m n} (h : m ≤ n) : + dist (f m) (f n) ≤ ∑ i in Finset.Ico m n, dist (f i) (f (i + 1)) := by + induction n, h using Nat.le_induction with + | base => rw [Finset.Ico_self, Finset.sum_empty, dist_self] + | succ n hle ihn => + calc + dist (f m) (f (n + 1)) ≤ dist (f m) (f n) + dist (f n) (f (n + 1)) := dist_triangle _ _ _ + _ ≤ (∑ i in Finset.Ico m n, _) + _ := add_le_add ihn le_rfl + _ = ∑ i in Finset.Ico m (n + 1), _ := by + { rw [Nat.Ico_succ_right_eq_insert_Ico hle, Finset.sum_insert, add_comm]; simp } +#align dist_le_Ico_sum_dist dist_le_Ico_sum_dist + +/-- The triangle (polygon) inequality for sequences of points; `Finset.range` version. -/ +theorem dist_le_range_sum_dist (f : ℕ → α) (n : ℕ) : + dist (f 0) (f n) ≤ ∑ i in Finset.range n, dist (f i) (f (i + 1)) := + Nat.Ico_zero_eq_range ▸ dist_le_Ico_sum_dist f (Nat.zero_le n) +#align dist_le_range_sum_dist dist_le_range_sum_dist + +/-- A version of `dist_le_Ico_sum_dist` with each intermediate distance replaced +with an upper estimate. -/ +theorem dist_le_Ico_sum_of_dist_le {f : ℕ → α} {m n} (hmn : m ≤ n) {d : ℕ → ℝ} + (hd : ∀ {k}, m ≤ k → k < n → dist (f k) (f (k + 1)) ≤ d k) : + dist (f m) (f n) ≤ ∑ i in Finset.Ico m n, d i := + le_trans (dist_le_Ico_sum_dist f hmn) <| + Finset.sum_le_sum fun _k hk => hd (Finset.mem_Ico.1 hk).1 (Finset.mem_Ico.1 hk).2 +#align dist_le_Ico_sum_of_dist_le dist_le_Ico_sum_of_dist_le + +/-- A version of `dist_le_range_sum_dist` with each intermediate distance replaced +with an upper estimate. -/ +theorem dist_le_range_sum_of_dist_le {f : ℕ → α} (n : ℕ) {d : ℕ → ℝ} + (hd : ∀ {k}, k < n → dist (f k) (f (k + 1)) ≤ d k) : + dist (f 0) (f n) ≤ ∑ i in Finset.range n, d i := + Nat.Ico_zero_eq_range ▸ dist_le_Ico_sum_of_dist_le (zero_le n) fun _ => hd +#align dist_le_range_sum_of_dist_le dist_le_range_sum_of_dist_le + +theorem swap_dist : Function.swap (@dist α _) = dist := by funext x y; exact dist_comm _ _ +#align swap_dist swap_dist + +theorem abs_dist_sub_le (x y z : α) : |dist x z - dist y z| ≤ dist x y := + abs_sub_le_iff.2 + ⟨sub_le_iff_le_add.2 (dist_triangle _ _ _), sub_le_iff_le_add.2 (dist_triangle_left _ _ _)⟩ +#align abs_dist_sub_le abs_dist_sub_le + +theorem dist_nonneg {x y : α} : 0 ≤ dist x y := + dist_nonneg' dist dist_self dist_comm dist_triangle +#align dist_nonneg dist_nonneg + +namespace Mathlib.Meta.Positivity + +open Lean Meta Qq Function + +/-- Extension for the `positivity` tactic: distances are nonnegative. -/ +@[positivity Dist.dist _ _] +def evalDist : PositivityExt where eval {_ _} _zα _pα e := do + let .app (.app _ a) b ← whnfR e | throwError "not dist" + let p ← mkAppOptM ``dist_nonneg #[none, none, a, b] + pure (.nonnegative p) + +end Mathlib.Meta.Positivity + +example {x y : α} : 0 ≤ dist x y := by positivity + +@[simp] theorem abs_dist {a b : α} : |dist a b| = dist a b := abs_of_nonneg dist_nonneg +#align abs_dist abs_dist + +/-- A version of `Dist` that takes value in `ℝ≥0`. -/ +class NNDist (α : Type*) where + nndist : α → α → ℝ≥0 +#align has_nndist NNDist + +export NNDist (nndist) + +-- see Note [lower instance priority] +/-- Distance as a nonnegative real number. -/ +instance (priority := 100) PseudoMetricSpace.toNNDist : NNDist α := + ⟨fun a b => ⟨dist a b, dist_nonneg⟩⟩ +#align pseudo_metric_space.to_has_nndist PseudoMetricSpace.toNNDist + +/-- Express `dist` in terms of `nndist`-/ +theorem dist_nndist (x y : α) : dist x y = nndist x y := rfl +#align dist_nndist dist_nndist + +@[simp, norm_cast] +theorem coe_nndist (x y : α) : ↑(nndist x y) = dist x y := rfl +#align coe_nndist coe_nndist + +/-- Express `edist` in terms of `nndist`-/ +theorem edist_nndist (x y : α) : edist x y = nndist x y := by + rw [edist_dist, dist_nndist, ENNReal.ofReal_coe_nnreal] +#align edist_nndist edist_nndist + +/-- Express `nndist` in terms of `edist`-/ +theorem nndist_edist (x y : α) : nndist x y = (edist x y).toNNReal := by + simp [edist_nndist] +#align nndist_edist nndist_edist + +@[simp, norm_cast] +theorem coe_nnreal_ennreal_nndist (x y : α) : ↑(nndist x y) = edist x y := + (edist_nndist x y).symm +#align coe_nnreal_ennreal_nndist coe_nnreal_ennreal_nndist + +@[simp, norm_cast] +theorem edist_lt_coe {x y : α} {c : ℝ≥0} : edist x y < c ↔ nndist x y < c := by + rw [edist_nndist, ENNReal.coe_lt_coe] +#align edist_lt_coe edist_lt_coe + +@[simp, norm_cast] +theorem edist_le_coe {x y : α} {c : ℝ≥0} : edist x y ≤ c ↔ nndist x y ≤ c := by + rw [edist_nndist, ENNReal.coe_le_coe] +#align edist_le_coe edist_le_coe + +/-- In a pseudometric space, the extended distance is always finite-/ +theorem edist_lt_top {α : Type*} [PseudoMetricSpace α] (x y : α) : edist x y < ⊤ := + (edist_dist x y).symm ▸ ENNReal.ofReal_lt_top +#align edist_lt_top edist_lt_top + +/-- In a pseudometric space, the extended distance is always finite-/ +theorem edist_ne_top (x y : α) : edist x y ≠ ⊤ := + (edist_lt_top x y).ne +#align edist_ne_top edist_ne_top + +/-- `nndist x x` vanishes-/ +@[simp] +theorem nndist_self (a : α) : nndist a a = 0 := + (NNReal.coe_eq_zero _).1 (dist_self a) +#align nndist_self nndist_self + +-- porting note: `dist_nndist` and `coe_nndist` moved up + +@[simp, norm_cast] +theorem dist_lt_coe {x y : α} {c : ℝ≥0} : dist x y < c ↔ nndist x y < c := + Iff.rfl +#align dist_lt_coe dist_lt_coe + +@[simp, norm_cast] +theorem dist_le_coe {x y : α} {c : ℝ≥0} : dist x y ≤ c ↔ nndist x y ≤ c := + Iff.rfl +#align dist_le_coe dist_le_coe + +@[simp] +theorem edist_lt_ofReal {x y : α} {r : ℝ} : edist x y < ENNReal.ofReal r ↔ dist x y < r := by + rw [edist_dist, ENNReal.ofReal_lt_ofReal_iff_of_nonneg dist_nonneg] +#align edist_lt_of_real edist_lt_ofReal + +@[simp] +theorem edist_le_ofReal {x y : α} {r : ℝ} (hr : 0 ≤ r) : + edist x y ≤ ENNReal.ofReal r ↔ dist x y ≤ r := by + rw [edist_dist, ENNReal.ofReal_le_ofReal_iff hr] +#align edist_le_of_real edist_le_ofReal + +/-- Express `nndist` in terms of `dist`-/ +theorem nndist_dist (x y : α) : nndist x y = Real.toNNReal (dist x y) := by + rw [dist_nndist, Real.toNNReal_coe] +#align nndist_dist nndist_dist + +theorem nndist_comm (x y : α) : nndist x y = nndist y x := NNReal.eq <| dist_comm x y +#align nndist_comm nndist_comm + +/-- Triangle inequality for the nonnegative distance-/ +theorem nndist_triangle (x y z : α) : nndist x z ≤ nndist x y + nndist y z := + dist_triangle _ _ _ +#align nndist_triangle nndist_triangle + +theorem nndist_triangle_left (x y z : α) : nndist x y ≤ nndist z x + nndist z y := + dist_triangle_left _ _ _ +#align nndist_triangle_left nndist_triangle_left + +theorem nndist_triangle_right (x y z : α) : nndist x y ≤ nndist x z + nndist y z := + dist_triangle_right _ _ _ +#align nndist_triangle_right nndist_triangle_right + +/-- Express `dist` in terms of `edist`-/ +theorem dist_edist (x y : α) : dist x y = (edist x y).toReal := by + rw [edist_dist, ENNReal.toReal_ofReal dist_nonneg] +#align dist_edist dist_edist + +namespace Metric + +-- instantiate pseudometric space as a topology +variable {x y z : α} {δ ε ε₁ ε₂ : ℝ} {s : Set α} + +/-- `ball x ε` is the set of all points `y` with `dist y x < ε` -/ +def ball (x : α) (ε : ℝ) : Set α := + { y | dist y x < ε } +#align metric.ball Metric.ball + +@[simp] +theorem mem_ball : y ∈ ball x ε ↔ dist y x < ε := + Iff.rfl +#align metric.mem_ball Metric.mem_ball + +theorem mem_ball' : y ∈ ball x ε ↔ dist x y < ε := by rw [dist_comm, mem_ball] +#align metric.mem_ball' Metric.mem_ball' + +theorem pos_of_mem_ball (hy : y ∈ ball x ε) : 0 < ε := + dist_nonneg.trans_lt hy +#align metric.pos_of_mem_ball Metric.pos_of_mem_ball + +theorem mem_ball_self (h : 0 < ε) : x ∈ ball x ε := by + rwa [mem_ball, dist_self] +#align metric.mem_ball_self Metric.mem_ball_self + +@[simp] +theorem nonempty_ball : (ball x ε).Nonempty ↔ 0 < ε := + ⟨fun ⟨_x, hx⟩ => pos_of_mem_ball hx, fun h => ⟨x, mem_ball_self h⟩⟩ +#align metric.nonempty_ball Metric.nonempty_ball + +@[simp] +theorem ball_eq_empty : ball x ε = ∅ ↔ ε ≤ 0 := by + rw [← not_nonempty_iff_eq_empty, nonempty_ball, not_lt] +#align metric.ball_eq_empty Metric.ball_eq_empty + +@[simp] +theorem ball_zero : ball x 0 = ∅ := by rw [ball_eq_empty] +#align metric.ball_zero Metric.ball_zero + +/-- If a point belongs to an open ball, then there is a strictly smaller radius whose ball also +contains it. + +See also `exists_lt_subset_ball`. -/ +theorem exists_lt_mem_ball_of_mem_ball (h : x ∈ ball y ε) : ∃ ε' < ε, x ∈ ball y ε' := by + simp only [mem_ball] at h ⊢ + exact ⟨(dist x y + ε) / 2, by linarith, by linarith⟩ +#align metric.exists_lt_mem_ball_of_mem_ball Metric.exists_lt_mem_ball_of_mem_ball + +theorem ball_eq_ball (ε : ℝ) (x : α) : + UniformSpace.ball x { p | dist p.2 p.1 < ε } = Metric.ball x ε := + rfl +#align metric.ball_eq_ball Metric.ball_eq_ball + +theorem ball_eq_ball' (ε : ℝ) (x : α) : + UniformSpace.ball x { p | dist p.1 p.2 < ε } = Metric.ball x ε := by + ext + simp [dist_comm, UniformSpace.ball] +#align metric.ball_eq_ball' Metric.ball_eq_ball' + +@[simp] +theorem iUnion_ball_nat (x : α) : ⋃ n : ℕ, ball x n = univ := + iUnion_eq_univ_iff.2 fun y => exists_nat_gt (dist y x) +#align metric.Union_ball_nat Metric.iUnion_ball_nat + +@[simp] +theorem iUnion_ball_nat_succ (x : α) : ⋃ n : ℕ, ball x (n + 1) = univ := + iUnion_eq_univ_iff.2 fun y => (exists_nat_gt (dist y x)).imp fun _ h => h.trans (lt_add_one _) +#align metric.Union_ball_nat_succ Metric.iUnion_ball_nat_succ + +/-- `closedBall x ε` is the set of all points `y` with `dist y x ≤ ε` -/ +def closedBall (x : α) (ε : ℝ) := + { y | dist y x ≤ ε } +#align metric.closed_ball Metric.closedBall + +@[simp] theorem mem_closedBall : y ∈ closedBall x ε ↔ dist y x ≤ ε := Iff.rfl +#align metric.mem_closed_ball Metric.mem_closedBall + +theorem mem_closedBall' : y ∈ closedBall x ε ↔ dist x y ≤ ε := by rw [dist_comm, mem_closedBall] +#align metric.mem_closed_ball' Metric.mem_closedBall' + +/-- `sphere x ε` is the set of all points `y` with `dist y x = ε` -/ +def sphere (x : α) (ε : ℝ) := { y | dist y x = ε } +#align metric.sphere Metric.sphere + +@[simp] theorem mem_sphere : y ∈ sphere x ε ↔ dist y x = ε := Iff.rfl +#align metric.mem_sphere Metric.mem_sphere + +theorem mem_sphere' : y ∈ sphere x ε ↔ dist x y = ε := by rw [dist_comm, mem_sphere] +#align metric.mem_sphere' Metric.mem_sphere' + +theorem ne_of_mem_sphere (h : y ∈ sphere x ε) (hε : ε ≠ 0) : y ≠ x := + ne_of_mem_of_not_mem h <| by simpa using hε.symm +#align metric.ne_of_mem_sphere Metric.ne_of_mem_sphere + +theorem nonneg_of_mem_sphere (hy : y ∈ sphere x ε) : 0 ≤ ε := + dist_nonneg.trans_eq hy +#align metric.nonneg_of_mem_sphere Metric.nonneg_of_mem_sphere + +@[simp] +theorem sphere_eq_empty_of_neg (hε : ε < 0) : sphere x ε = ∅ := + Set.eq_empty_iff_forall_not_mem.mpr fun _y hy => (nonneg_of_mem_sphere hy).not_lt hε +#align metric.sphere_eq_empty_of_neg Metric.sphere_eq_empty_of_neg + +theorem sphere_eq_empty_of_subsingleton [Subsingleton α] (hε : ε ≠ 0) : sphere x ε = ∅ := + Set.eq_empty_iff_forall_not_mem.mpr fun _ h => ne_of_mem_sphere h hε (Subsingleton.elim _ _) +#align metric.sphere_eq_empty_of_subsingleton Metric.sphere_eq_empty_of_subsingleton + +theorem sphere_isEmpty_of_subsingleton [Subsingleton α] (hε : ε ≠ 0) : IsEmpty (sphere x ε) := by + rw [sphere_eq_empty_of_subsingleton hε]; infer_instance +#align metric.sphere_is_empty_of_subsingleton Metric.sphere_isEmpty_of_subsingleton + +theorem mem_closedBall_self (h : 0 ≤ ε) : x ∈ closedBall x ε := by + rwa [mem_closedBall, dist_self] +#align metric.mem_closed_ball_self Metric.mem_closedBall_self + +@[simp] +theorem nonempty_closedBall : (closedBall x ε).Nonempty ↔ 0 ≤ ε := + ⟨fun ⟨_x, hx⟩ => dist_nonneg.trans hx, fun h => ⟨x, mem_closedBall_self h⟩⟩ +#align metric.nonempty_closed_ball Metric.nonempty_closedBall + +@[simp] +theorem closedBall_eq_empty : closedBall x ε = ∅ ↔ ε < 0 := by + rw [← not_nonempty_iff_eq_empty, nonempty_closedBall, not_le] +#align metric.closed_ball_eq_empty Metric.closedBall_eq_empty + +/-- Closed balls and spheres coincide when the radius is non-positive -/ +theorem closedBall_eq_sphere_of_nonpos (hε : ε ≤ 0) : closedBall x ε = sphere x ε := + Set.ext fun _ => (hε.trans dist_nonneg).le_iff_eq +#align metric.closed_ball_eq_sphere_of_nonpos Metric.closedBall_eq_sphere_of_nonpos + +theorem ball_subset_closedBall : ball x ε ⊆ closedBall x ε := fun _y hy => + mem_closedBall.2 (le_of_lt hy) +#align metric.ball_subset_closed_ball Metric.ball_subset_closedBall + +theorem sphere_subset_closedBall : sphere x ε ⊆ closedBall x ε := fun _ => le_of_eq +#align metric.sphere_subset_closed_ball Metric.sphere_subset_closedBall + +lemma sphere_subset_ball {r R : ℝ} (h : r < R) : sphere x r ⊆ ball x R := fun _x hx ↦ + (mem_sphere.1 hx).trans_lt h + +theorem closedBall_disjoint_ball (h : δ + ε ≤ dist x y) : Disjoint (closedBall x δ) (ball y ε) := + Set.disjoint_left.mpr fun _a ha1 ha2 => + (h.trans <| dist_triangle_left _ _ _).not_lt <| add_lt_add_of_le_of_lt ha1 ha2 +#align metric.closed_ball_disjoint_ball Metric.closedBall_disjoint_ball + +theorem ball_disjoint_closedBall (h : δ + ε ≤ dist x y) : Disjoint (ball x δ) (closedBall y ε) := + (closedBall_disjoint_ball <| by rwa [add_comm, dist_comm]).symm +#align metric.ball_disjoint_closed_ball Metric.ball_disjoint_closedBall + +theorem ball_disjoint_ball (h : δ + ε ≤ dist x y) : Disjoint (ball x δ) (ball y ε) := + (closedBall_disjoint_ball h).mono_left ball_subset_closedBall +#align metric.ball_disjoint_ball Metric.ball_disjoint_ball + +theorem closedBall_disjoint_closedBall (h : δ + ε < dist x y) : + Disjoint (closedBall x δ) (closedBall y ε) := + Set.disjoint_left.mpr fun _a ha1 ha2 => + h.not_le <| (dist_triangle_left _ _ _).trans <| add_le_add ha1 ha2 +#align metric.closed_ball_disjoint_closed_ball Metric.closedBall_disjoint_closedBall + +theorem sphere_disjoint_ball : Disjoint (sphere x ε) (ball x ε) := + Set.disjoint_left.mpr fun _y hy₁ hy₂ => absurd hy₁ <| ne_of_lt hy₂ +#align metric.sphere_disjoint_ball Metric.sphere_disjoint_ball + +@[simp] +theorem ball_union_sphere : ball x ε ∪ sphere x ε = closedBall x ε := + Set.ext fun _y => (@le_iff_lt_or_eq ℝ _ _ _).symm +#align metric.ball_union_sphere Metric.ball_union_sphere + +@[simp] +theorem sphere_union_ball : sphere x ε ∪ ball x ε = closedBall x ε := by + rw [union_comm, ball_union_sphere] +#align metric.sphere_union_ball Metric.sphere_union_ball + +@[simp] +theorem closedBall_diff_sphere : closedBall x ε \ sphere x ε = ball x ε := by + rw [← ball_union_sphere, Set.union_diff_cancel_right sphere_disjoint_ball.symm.le_bot] +#align metric.closed_ball_diff_sphere Metric.closedBall_diff_sphere + +@[simp] +theorem closedBall_diff_ball : closedBall x ε \ ball x ε = sphere x ε := by + rw [← ball_union_sphere, Set.union_diff_cancel_left sphere_disjoint_ball.symm.le_bot] +#align metric.closed_ball_diff_ball Metric.closedBall_diff_ball + +theorem mem_ball_comm : x ∈ ball y ε ↔ y ∈ ball x ε := by rw [mem_ball', mem_ball] +#align metric.mem_ball_comm Metric.mem_ball_comm + +theorem mem_closedBall_comm : x ∈ closedBall y ε ↔ y ∈ closedBall x ε := by + rw [mem_closedBall', mem_closedBall] +#align metric.mem_closed_ball_comm Metric.mem_closedBall_comm + +theorem mem_sphere_comm : x ∈ sphere y ε ↔ y ∈ sphere x ε := by rw [mem_sphere', mem_sphere] +#align metric.mem_sphere_comm Metric.mem_sphere_comm + +theorem ball_subset_ball (h : ε₁ ≤ ε₂) : ball x ε₁ ⊆ ball x ε₂ := fun _y yx => + lt_of_lt_of_le (mem_ball.1 yx) h +#align metric.ball_subset_ball Metric.ball_subset_ball + +theorem closedBall_eq_bInter_ball : closedBall x ε = ⋂ δ > ε, ball x δ := by + ext y; rw [mem_closedBall, ← forall_lt_iff_le', mem_iInter₂]; rfl +#align metric.closed_ball_eq_bInter_ball Metric.closedBall_eq_bInter_ball + +theorem ball_subset_ball' (h : ε₁ + dist x y ≤ ε₂) : ball x ε₁ ⊆ ball y ε₂ := fun z hz => + calc + dist z y ≤ dist z x + dist x y := dist_triangle _ _ _ + _ < ε₁ + dist x y := add_lt_add_right (mem_ball.1 hz) _ + _ ≤ ε₂ := h +#align metric.ball_subset_ball' Metric.ball_subset_ball' + +theorem closedBall_subset_closedBall (h : ε₁ ≤ ε₂) : closedBall x ε₁ ⊆ closedBall x ε₂ := + fun _y (yx : _ ≤ ε₁) => le_trans yx h +#align metric.closed_ball_subset_closed_ball Metric.closedBall_subset_closedBall + +theorem closedBall_subset_closedBall' (h : ε₁ + dist x y ≤ ε₂) : + closedBall x ε₁ ⊆ closedBall y ε₂ := fun z hz => + calc + dist z y ≤ dist z x + dist x y := dist_triangle _ _ _ + _ ≤ ε₁ + dist x y := add_le_add_right (mem_closedBall.1 hz) _ + _ ≤ ε₂ := h +#align metric.closed_ball_subset_closed_ball' Metric.closedBall_subset_closedBall' + +theorem closedBall_subset_ball (h : ε₁ < ε₂) : closedBall x ε₁ ⊆ ball x ε₂ := + fun y (yh : dist y x ≤ ε₁) => lt_of_le_of_lt yh h +#align metric.closed_ball_subset_ball Metric.closedBall_subset_ball + +theorem closedBall_subset_ball' (h : ε₁ + dist x y < ε₂) : + closedBall x ε₁ ⊆ ball y ε₂ := fun z hz => + calc + dist z y ≤ dist z x + dist x y := dist_triangle _ _ _ + _ ≤ ε₁ + dist x y := add_le_add_right (mem_closedBall.1 hz) _ + _ < ε₂ := h +#align metric.closed_ball_subset_ball' Metric.closedBall_subset_ball' + +theorem dist_le_add_of_nonempty_closedBall_inter_closedBall + (h : (closedBall x ε₁ ∩ closedBall y ε₂).Nonempty) : dist x y ≤ ε₁ + ε₂ := + let ⟨z, hz⟩ := h + calc + dist x y ≤ dist z x + dist z y := dist_triangle_left _ _ _ + _ ≤ ε₁ + ε₂ := add_le_add hz.1 hz.2 +#align metric.dist_le_add_of_nonempty_closed_ball_inter_closed_ball Metric.dist_le_add_of_nonempty_closedBall_inter_closedBall + +theorem dist_lt_add_of_nonempty_closedBall_inter_ball (h : (closedBall x ε₁ ∩ ball y ε₂).Nonempty) : + dist x y < ε₁ + ε₂ := + let ⟨z, hz⟩ := h + calc + dist x y ≤ dist z x + dist z y := dist_triangle_left _ _ _ + _ < ε₁ + ε₂ := add_lt_add_of_le_of_lt hz.1 hz.2 +#align metric.dist_lt_add_of_nonempty_closed_ball_inter_ball Metric.dist_lt_add_of_nonempty_closedBall_inter_ball + +theorem dist_lt_add_of_nonempty_ball_inter_closedBall (h : (ball x ε₁ ∩ closedBall y ε₂).Nonempty) : + dist x y < ε₁ + ε₂ := by + rw [inter_comm] at h + rw [add_comm, dist_comm] + exact dist_lt_add_of_nonempty_closedBall_inter_ball h +#align metric.dist_lt_add_of_nonempty_ball_inter_closed_ball Metric.dist_lt_add_of_nonempty_ball_inter_closedBall + +theorem dist_lt_add_of_nonempty_ball_inter_ball (h : (ball x ε₁ ∩ ball y ε₂).Nonempty) : + dist x y < ε₁ + ε₂ := + dist_lt_add_of_nonempty_closedBall_inter_ball <| + h.mono (inter_subset_inter ball_subset_closedBall Subset.rfl) +#align metric.dist_lt_add_of_nonempty_ball_inter_ball Metric.dist_lt_add_of_nonempty_ball_inter_ball + +@[simp] +theorem iUnion_closedBall_nat (x : α) : ⋃ n : ℕ, closedBall x n = univ := + iUnion_eq_univ_iff.2 fun y => exists_nat_ge (dist y x) +#align metric.Union_closed_ball_nat Metric.iUnion_closedBall_nat + +theorem iUnion_inter_closedBall_nat (s : Set α) (x : α) : ⋃ n : ℕ, s ∩ closedBall x n = s := by + rw [← inter_iUnion, iUnion_closedBall_nat, inter_univ] +#align metric.Union_inter_closed_ball_nat Metric.iUnion_inter_closedBall_nat + +theorem ball_subset (h : dist x y ≤ ε₂ - ε₁) : ball x ε₁ ⊆ ball y ε₂ := fun z zx => by + rw [← add_sub_cancel'_right ε₁ ε₂] + exact lt_of_le_of_lt (dist_triangle z x y) (add_lt_add_of_lt_of_le zx h) +#align metric.ball_subset Metric.ball_subset + +theorem ball_half_subset (y) (h : y ∈ ball x (ε / 2)) : ball y (ε / 2) ⊆ ball x ε := + ball_subset <| by rw [sub_self_div_two]; exact le_of_lt h +#align metric.ball_half_subset Metric.ball_half_subset + +theorem exists_ball_subset_ball (h : y ∈ ball x ε) : ∃ ε' > 0, ball y ε' ⊆ ball x ε := + ⟨_, sub_pos.2 h, ball_subset <| by rw [sub_sub_self]⟩ +#align metric.exists_ball_subset_ball Metric.exists_ball_subset_ball + +/-- If a property holds for all points in closed balls of arbitrarily large radii, then it holds for +all points. -/ +theorem forall_of_forall_mem_closedBall (p : α → Prop) (x : α) + (H : ∃ᶠ R : ℝ in atTop, ∀ y ∈ closedBall x R, p y) (y : α) : p y := by + obtain ⟨R, hR, h⟩ : ∃ R ≥ dist y x, ∀ z : α, z ∈ closedBall x R → p z := + frequently_iff.1 H (Ici_mem_atTop (dist y x)) + exact h _ hR +#align metric.forall_of_forall_mem_closed_ball Metric.forall_of_forall_mem_closedBall + +/-- If a property holds for all points in balls of arbitrarily large radii, then it holds for all +points. -/ +theorem forall_of_forall_mem_ball (p : α → Prop) (x : α) + (H : ∃ᶠ R : ℝ in atTop, ∀ y ∈ ball x R, p y) (y : α) : p y := by + obtain ⟨R, hR, h⟩ : ∃ R > dist y x, ∀ z : α, z ∈ ball x R → p z := + frequently_iff.1 H (Ioi_mem_atTop (dist y x)) + exact h _ hR +#align metric.forall_of_forall_mem_ball Metric.forall_of_forall_mem_ball + +theorem isBounded_iff {s : Set α} : + IsBounded s ↔ ∃ C : ℝ, ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → dist x y ≤ C := by + rw [isBounded_def, ← Filter.mem_sets, @PseudoMetricSpace.cobounded_sets α, mem_setOf_eq, + compl_compl] +#align metric.is_bounded_iff Metric.isBounded_iff + +theorem isBounded_iff_eventually {s : Set α} : + IsBounded s ↔ ∀ᶠ C in atTop, ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → dist x y ≤ C := + isBounded_iff.trans + ⟨fun ⟨C, h⟩ => eventually_atTop.2 ⟨C, fun _C' hC' _x hx _y hy => (h hx hy).trans hC'⟩, + Eventually.exists⟩ +#align metric.is_bounded_iff_eventually Metric.isBounded_iff_eventually + +theorem isBounded_iff_exists_ge {s : Set α} (c : ℝ) : + IsBounded s ↔ ∃ C, c ≤ C ∧ ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → dist x y ≤ C := + ⟨fun h => ((eventually_ge_atTop c).and (isBounded_iff_eventually.1 h)).exists, fun h => + isBounded_iff.2 <| h.imp fun _ => And.right⟩ +#align metric.is_bounded_iff_exists_ge Metric.isBounded_iff_exists_ge + +theorem isBounded_iff_nndist {s : Set α} : + IsBounded s ↔ ∃ C : ℝ≥0, ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → nndist x y ≤ C := by + simp only [isBounded_iff_exists_ge 0, NNReal.exists, ← NNReal.coe_le_coe, ← dist_nndist, + NNReal.coe_mk, exists_prop] +#align metric.is_bounded_iff_nndist Metric.isBounded_iff_nndist + +theorem toUniformSpace_eq : + ‹PseudoMetricSpace α›.toUniformSpace = .ofDist dist dist_self dist_comm dist_triangle := + UniformSpace.ext PseudoMetricSpace.uniformity_dist +#align metric.to_uniform_space_eq Metric.toUniformSpace_eq + +theorem uniformity_basis_dist : + (𝓤 α).HasBasis (fun ε : ℝ => 0 < ε) fun ε => { p : α × α | dist p.1 p.2 < ε } := by + rw [toUniformSpace_eq] + exact UniformSpace.hasBasis_ofFun (exists_gt _) _ _ _ _ _ +#align metric.uniformity_basis_dist Metric.uniformity_basis_dist + +/-- Given `f : β → ℝ`, if `f` sends `{i | p i}` to a set of positive numbers +accumulating to zero, then `f i`-neighborhoods of the diagonal form a basis of `𝓤 α`. + +For specific bases see `uniformity_basis_dist`, `uniformity_basis_dist_inv_nat_succ`, +and `uniformity_basis_dist_inv_nat_pos`. -/ +protected theorem mk_uniformity_basis {β : Type*} {p : β → Prop} {f : β → ℝ} + (hf₀ : ∀ i, p i → 0 < f i) (hf : ∀ ⦃ε⦄, 0 < ε → ∃ i, p i ∧ f i ≤ ε) : + (𝓤 α).HasBasis p fun i => { p : α × α | dist p.1 p.2 < f i } := by + refine' ⟨fun s => uniformity_basis_dist.mem_iff.trans _⟩ + constructor + · rintro ⟨ε, ε₀, hε⟩ + rcases hf ε₀ with ⟨i, hi, H⟩ + exact ⟨i, hi, fun x (hx : _ < _) => hε <| lt_of_lt_of_le hx H⟩ + · exact fun ⟨i, hi, H⟩ => ⟨f i, hf₀ i hi, H⟩ +#align metric.mk_uniformity_basis Metric.mk_uniformity_basis + +theorem uniformity_basis_dist_rat : + (𝓤 α).HasBasis (fun r : ℚ => 0 < r) fun r => { p : α × α | dist p.1 p.2 < r } := + Metric.mk_uniformity_basis (fun _ => Rat.cast_pos.2) fun _ε hε => + let ⟨r, hr0, hrε⟩ := exists_rat_btwn hε + ⟨r, Rat.cast_pos.1 hr0, hrε.le⟩ +#align metric.uniformity_basis_dist_rat Metric.uniformity_basis_dist_rat + +theorem uniformity_basis_dist_inv_nat_succ : + (𝓤 α).HasBasis (fun _ => True) fun n : ℕ => { p : α × α | dist p.1 p.2 < 1 / (↑n + 1) } := + Metric.mk_uniformity_basis (fun n _ => div_pos zero_lt_one <| Nat.cast_add_one_pos n) fun _ε ε0 => + (exists_nat_one_div_lt ε0).imp fun _n hn => ⟨trivial, le_of_lt hn⟩ +#align metric.uniformity_basis_dist_inv_nat_succ Metric.uniformity_basis_dist_inv_nat_succ + +theorem uniformity_basis_dist_inv_nat_pos : + (𝓤 α).HasBasis (fun n : ℕ => 0 < n) fun n : ℕ => { p : α × α | dist p.1 p.2 < 1 / ↑n } := + Metric.mk_uniformity_basis (fun n hn => div_pos zero_lt_one <| Nat.cast_pos.2 hn) fun ε ε0 => + let ⟨n, hn⟩ := exists_nat_one_div_lt ε0 + ⟨n + 1, Nat.succ_pos n, by exact_mod_cast hn.le⟩ +#align metric.uniformity_basis_dist_inv_nat_pos Metric.uniformity_basis_dist_inv_nat_pos + +theorem uniformity_basis_dist_pow {r : ℝ} (h0 : 0 < r) (h1 : r < 1) : + (𝓤 α).HasBasis (fun _ : ℕ => True) fun n : ℕ => { p : α × α | dist p.1 p.2 < r ^ n } := + Metric.mk_uniformity_basis (fun _ _ => pow_pos h0 _) fun _ε ε0 => + let ⟨n, hn⟩ := exists_pow_lt_of_lt_one ε0 h1 + ⟨n, trivial, hn.le⟩ +#align metric.uniformity_basis_dist_pow Metric.uniformity_basis_dist_pow + +theorem uniformity_basis_dist_lt {R : ℝ} (hR : 0 < R) : + (𝓤 α).HasBasis (fun r : ℝ => 0 < r ∧ r < R) fun r => { p : α × α | dist p.1 p.2 < r } := + Metric.mk_uniformity_basis (fun _ => And.left) fun r hr => + ⟨min r (R / 2), ⟨lt_min hr (half_pos hR), min_lt_iff.2 <| Or.inr (half_lt_self hR)⟩, + min_le_left _ _⟩ +#align metric.uniformity_basis_dist_lt Metric.uniformity_basis_dist_lt + +/-- Given `f : β → ℝ`, if `f` sends `{i | p i}` to a set of positive numbers +accumulating to zero, then closed neighborhoods of the diagonal of sizes `{f i | p i}` +form a basis of `𝓤 α`. + +Currently we have only one specific basis `uniformity_basis_dist_le` based on this constructor. +More can be easily added if needed in the future. -/ +protected theorem mk_uniformity_basis_le {β : Type*} {p : β → Prop} {f : β → ℝ} + (hf₀ : ∀ x, p x → 0 < f x) (hf : ∀ ε, 0 < ε → ∃ x, p x ∧ f x ≤ ε) : + (𝓤 α).HasBasis p fun x => { p : α × α | dist p.1 p.2 ≤ f x } := by + refine' ⟨fun s => uniformity_basis_dist.mem_iff.trans _⟩ + constructor + · rintro ⟨ε, ε₀, hε⟩ + rcases exists_between ε₀ with ⟨ε', hε'⟩ + rcases hf ε' hε'.1 with ⟨i, hi, H⟩ + exact ⟨i, hi, fun x (hx : _ ≤ _) => hε <| lt_of_le_of_lt (le_trans hx H) hε'.2⟩ + · exact fun ⟨i, hi, H⟩ => ⟨f i, hf₀ i hi, fun x (hx : _ < _) => H (mem_setOf.2 hx.le)⟩ +#align metric.mk_uniformity_basis_le Metric.mk_uniformity_basis_le + +/-- Constant size closed neighborhoods of the diagonal form a basis +of the uniformity filter. -/ +theorem uniformity_basis_dist_le : + (𝓤 α).HasBasis ((0 : ℝ) < ·) fun ε => { p : α × α | dist p.1 p.2 ≤ ε } := + Metric.mk_uniformity_basis_le (fun _ => id) fun ε ε₀ => ⟨ε, ε₀, le_refl ε⟩ +#align metric.uniformity_basis_dist_le Metric.uniformity_basis_dist_le + +theorem uniformity_basis_dist_le_pow {r : ℝ} (h0 : 0 < r) (h1 : r < 1) : + (𝓤 α).HasBasis (fun _ : ℕ => True) fun n : ℕ => { p : α × α | dist p.1 p.2 ≤ r ^ n } := + Metric.mk_uniformity_basis_le (fun _ _ => pow_pos h0 _) fun _ε ε0 => + let ⟨n, hn⟩ := exists_pow_lt_of_lt_one ε0 h1 + ⟨n, trivial, hn.le⟩ +#align metric.uniformity_basis_dist_le_pow Metric.uniformity_basis_dist_le_pow + +theorem mem_uniformity_dist {s : Set (α × α)} : + s ∈ 𝓤 α ↔ ∃ ε > 0, ∀ {a b : α}, dist a b < ε → (a, b) ∈ s := + uniformity_basis_dist.mem_uniformity_iff +#align metric.mem_uniformity_dist Metric.mem_uniformity_dist + +/-- A constant size neighborhood of the diagonal is an entourage. -/ +theorem dist_mem_uniformity {ε : ℝ} (ε0 : 0 < ε) : { p : α × α | dist p.1 p.2 < ε } ∈ 𝓤 α := + mem_uniformity_dist.2 ⟨ε, ε0, id⟩ +#align metric.dist_mem_uniformity Metric.dist_mem_uniformity + +theorem uniformContinuous_iff [PseudoMetricSpace β] {f : α → β} : + UniformContinuous f ↔ ∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, dist a b < δ → dist (f a) (f b) < ε := + uniformity_basis_dist.uniformContinuous_iff uniformity_basis_dist +#align metric.uniform_continuous_iff Metric.uniformContinuous_iff + +theorem uniformContinuousOn_iff [PseudoMetricSpace β] {f : α → β} {s : Set α} : + UniformContinuousOn f s ↔ + ∀ ε > 0, ∃ δ > 0, ∀ x ∈ s, ∀ y ∈ s, dist x y < δ → dist (f x) (f y) < ε := + Metric.uniformity_basis_dist.uniformContinuousOn_iff Metric.uniformity_basis_dist +#align metric.uniform_continuous_on_iff Metric.uniformContinuousOn_iff + +theorem uniformContinuousOn_iff_le [PseudoMetricSpace β] {f : α → β} {s : Set α} : + UniformContinuousOn f s ↔ + ∀ ε > 0, ∃ δ > 0, ∀ x ∈ s, ∀ y ∈ s, dist x y ≤ δ → dist (f x) (f y) ≤ ε := + Metric.uniformity_basis_dist_le.uniformContinuousOn_iff Metric.uniformity_basis_dist_le +#align metric.uniform_continuous_on_iff_le Metric.uniformContinuousOn_iff_le + +nonrec theorem uniformInducing_iff [PseudoMetricSpace β] {f : α → β} : + UniformInducing f ↔ UniformContinuous f ∧ + ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := + uniformInducing_iff'.trans <| Iff.rfl.and <| + ((uniformity_basis_dist.comap _).le_basis_iff uniformity_basis_dist).trans <| by + simp only [subset_def, Prod.forall, gt_iff_lt, preimage_setOf_eq, Prod_map, mem_setOf] + +nonrec theorem uniformEmbedding_iff [PseudoMetricSpace β] {f : α → β} : + UniformEmbedding f ↔ Function.Injective f ∧ UniformContinuous f ∧ + ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := by + rw [uniformEmbedding_iff, and_comm, uniformInducing_iff] +#align metric.uniform_embedding_iff Metric.uniformEmbedding_iff + +/-- If a map between pseudometric spaces is a uniform embedding then the distance between `f x` +and `f y` is controlled in terms of the distance between `x` and `y`. -/ +theorem controlled_of_uniformEmbedding [PseudoMetricSpace β] {f : α → β} (h : UniformEmbedding f) : + (∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, dist a b < δ → dist (f a) (f b) < ε) ∧ + ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := + ⟨uniformContinuous_iff.1 h.uniformContinuous, (uniformEmbedding_iff.1 h).2.2⟩ +#align metric.controlled_of_uniform_embedding Metric.controlled_of_uniformEmbedding + +theorem totallyBounded_iff {s : Set α} : + TotallyBounded s ↔ ∀ ε > 0, ∃ t : Set α, t.Finite ∧ s ⊆ ⋃ y ∈ t, ball y ε := + uniformity_basis_dist.totallyBounded_iff +#align metric.totally_bounded_iff Metric.totallyBounded_iff + +/-- A pseudometric space is totally bounded if one can reconstruct up to any ε>0 any element of the +space from finitely many data. -/ +theorem totallyBounded_of_finite_discretization {s : Set α} + (H : ∀ ε > (0 : ℝ), + ∃ (β : Type u) (_ : Fintype β) (F : s → β), ∀ x y, F x = F y → dist (x : α) y < ε) : + TotallyBounded s := by + cases' s.eq_empty_or_nonempty with hs hs + · rw [hs] + exact totallyBounded_empty + rcases hs with ⟨x0, hx0⟩ + haveI : Inhabited s := ⟨⟨x0, hx0⟩⟩ + refine' totallyBounded_iff.2 fun ε ε0 => _ + rcases H ε ε0 with ⟨β, fβ, F, hF⟩ + let Finv := Function.invFun F + refine' ⟨range (Subtype.val ∘ Finv), finite_range _, fun x xs => _⟩ + let x' := Finv (F ⟨x, xs⟩) + have : F x' = F ⟨x, xs⟩ := Function.invFun_eq ⟨⟨x, xs⟩, rfl⟩ + simp only [Set.mem_iUnion, Set.mem_range] + exact ⟨_, ⟨F ⟨x, xs⟩, rfl⟩, hF _ _ this.symm⟩ +#align metric.totally_bounded_of_finite_discretization Metric.totallyBounded_of_finite_discretization + +theorem finite_approx_of_totallyBounded {s : Set α} (hs : TotallyBounded s) : + ∀ ε > 0, ∃ t, t ⊆ s ∧ Set.Finite t ∧ s ⊆ ⋃ y ∈ t, ball y ε := by + intro ε ε_pos + rw [totallyBounded_iff_subset] at hs + exact hs _ (dist_mem_uniformity ε_pos) +#align metric.finite_approx_of_totally_bounded Metric.finite_approx_of_totallyBounded + +/-- Expressing uniform convergence using `dist` -/ +theorem tendstoUniformlyOnFilter_iff {F : ι → β → α} {f : β → α} {p : Filter ι} {p' : Filter β} : + TendstoUniformlyOnFilter F f p p' ↔ + ∀ ε > 0, ∀ᶠ n : ι × β in p ×ˢ p', dist (f n.snd) (F n.fst n.snd) < ε := by + refine' ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu => _⟩ + rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ + refine' (H ε εpos).mono fun n hn => hε hn +#align metric.tendsto_uniformly_on_filter_iff Metric.tendstoUniformlyOnFilter_iff + +/-- Expressing locally uniform convergence on a set using `dist`. -/ +theorem tendstoLocallyUniformlyOn_iff [TopologicalSpace β] {F : ι → β → α} {f : β → α} + {p : Filter ι} {s : Set β} : + TendstoLocallyUniformlyOn F f p s ↔ + ∀ ε > 0, ∀ x ∈ s, ∃ t ∈ 𝓝[s] x, ∀ᶠ n in p, ∀ y ∈ t, dist (f y) (F n y) < ε := by + refine' ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu x hx => _⟩ + rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ + rcases H ε εpos x hx with ⟨t, ht, Ht⟩ + exact ⟨t, ht, Ht.mono fun n hs x hx => hε (hs x hx)⟩ +#align metric.tendsto_locally_uniformly_on_iff Metric.tendstoLocallyUniformlyOn_iff + +/-- Expressing uniform convergence on a set using `dist`. -/ +theorem tendstoUniformlyOn_iff {F : ι → β → α} {f : β → α} {p : Filter ι} {s : Set β} : + TendstoUniformlyOn F f p s ↔ ∀ ε > 0, ∀ᶠ n in p, ∀ x ∈ s, dist (f x) (F n x) < ε := by + refine' ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu => _⟩ + rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ + exact (H ε εpos).mono fun n hs x hx => hε (hs x hx) +#align metric.tendsto_uniformly_on_iff Metric.tendstoUniformlyOn_iff + +/-- Expressing locally uniform convergence using `dist`. -/ +theorem tendstoLocallyUniformly_iff [TopologicalSpace β] {F : ι → β → α} {f : β → α} + {p : Filter ι} : + TendstoLocallyUniformly F f p ↔ + ∀ ε > 0, ∀ x : β, ∃ t ∈ 𝓝 x, ∀ᶠ n in p, ∀ y ∈ t, dist (f y) (F n y) < ε := by + simp only [← tendstoLocallyUniformlyOn_univ, tendstoLocallyUniformlyOn_iff, nhdsWithin_univ, + mem_univ, forall_const, exists_prop] +#align metric.tendsto_locally_uniformly_iff Metric.tendstoLocallyUniformly_iff + +/-- Expressing uniform convergence using `dist`. -/ +theorem tendstoUniformly_iff {F : ι → β → α} {f : β → α} {p : Filter ι} : + TendstoUniformly F f p ↔ ∀ ε > 0, ∀ᶠ n in p, ∀ x, dist (f x) (F n x) < ε := by + rw [← tendstoUniformlyOn_univ, tendstoUniformlyOn_iff] + simp +#align metric.tendsto_uniformly_iff Metric.tendstoUniformly_iff + +protected theorem cauchy_iff {f : Filter α} : + Cauchy f ↔ NeBot f ∧ ∀ ε > 0, ∃ t ∈ f, ∀ x ∈ t, ∀ y ∈ t, dist x y < ε := + uniformity_basis_dist.cauchy_iff +#align metric.cauchy_iff Metric.cauchy_iff + +theorem nhds_basis_ball : (𝓝 x).HasBasis (0 < ·) (ball x) := + nhds_basis_uniformity uniformity_basis_dist +#align metric.nhds_basis_ball Metric.nhds_basis_ball + +theorem mem_nhds_iff : s ∈ 𝓝 x ↔ ∃ ε > 0, ball x ε ⊆ s := + nhds_basis_ball.mem_iff +#align metric.mem_nhds_iff Metric.mem_nhds_iff + +theorem eventually_nhds_iff {p : α → Prop} : + (∀ᶠ y in 𝓝 x, p y) ↔ ∃ ε > 0, ∀ ⦃y⦄, dist y x < ε → p y := + mem_nhds_iff +#align metric.eventually_nhds_iff Metric.eventually_nhds_iff + +theorem eventually_nhds_iff_ball {p : α → Prop} : + (∀ᶠ y in 𝓝 x, p y) ↔ ∃ ε > 0, ∀ y ∈ ball x ε, p y := + mem_nhds_iff +#align metric.eventually_nhds_iff_ball Metric.eventually_nhds_iff_ball + +/-- A version of `Filter.eventually_prod_iff` where the first filter consists of neighborhoods +in a pseudo-metric space.-/ +theorem eventually_nhds_prod_iff {f : Filter ι} {x₀ : α} {p : α × ι → Prop} : + (∀ᶠ x in 𝓝 x₀ ×ˢ f, p x) ↔ ∃ ε > (0 : ℝ), ∃ pa : ι → Prop, (∀ᶠ i in f, pa i) ∧ + ∀ {x}, dist x x₀ < ε → ∀ {i}, pa i → p (x, i) := by + refine (nhds_basis_ball.prod f.basis_sets).eventually_iff.trans ?_ + simp only [Prod.exists, forall_prod_set, id, mem_ball, and_assoc, exists_and_left, and_imp] + rfl +#align metric.eventually_nhds_prod_iff Metric.eventually_nhds_prod_iff + +/-- A version of `Filter.eventually_prod_iff` where the second filter consists of neighborhoods +in a pseudo-metric space.-/ +theorem eventually_prod_nhds_iff {f : Filter ι} {x₀ : α} {p : ι × α → Prop} : + (∀ᶠ x in f ×ˢ 𝓝 x₀, p x) ↔ ∃ pa : ι → Prop, (∀ᶠ i in f, pa i) ∧ + ∃ ε > 0, ∀ {i}, pa i → ∀ {x}, dist x x₀ < ε → p (i, x) := by + rw [eventually_swap_iff, Metric.eventually_nhds_prod_iff] + constructor <;> + · rintro ⟨a1, a2, a3, a4, a5⟩ + exact ⟨a3, a4, a1, a2, fun b1 b2 b3 => a5 b3 b1⟩ +#align metric.eventually_prod_nhds_iff Metric.eventually_prod_nhds_iff + +theorem nhds_basis_closedBall : (𝓝 x).HasBasis (fun ε : ℝ => 0 < ε) (closedBall x) := + nhds_basis_uniformity uniformity_basis_dist_le +#align metric.nhds_basis_closed_ball Metric.nhds_basis_closedBall + +theorem nhds_basis_ball_inv_nat_succ : + (𝓝 x).HasBasis (fun _ => True) fun n : ℕ => ball x (1 / (↑n + 1)) := + nhds_basis_uniformity uniformity_basis_dist_inv_nat_succ +#align metric.nhds_basis_ball_inv_nat_succ Metric.nhds_basis_ball_inv_nat_succ + +theorem nhds_basis_ball_inv_nat_pos : + (𝓝 x).HasBasis (fun n => 0 < n) fun n : ℕ => ball x (1 / ↑n) := + nhds_basis_uniformity uniformity_basis_dist_inv_nat_pos +#align metric.nhds_basis_ball_inv_nat_pos Metric.nhds_basis_ball_inv_nat_pos + +theorem nhds_basis_ball_pow {r : ℝ} (h0 : 0 < r) (h1 : r < 1) : + (𝓝 x).HasBasis (fun _ => True) fun n : ℕ => ball x (r ^ n) := + nhds_basis_uniformity (uniformity_basis_dist_pow h0 h1) +#align metric.nhds_basis_ball_pow Metric.nhds_basis_ball_pow + +theorem nhds_basis_closedBall_pow {r : ℝ} (h0 : 0 < r) (h1 : r < 1) : + (𝓝 x).HasBasis (fun _ => True) fun n : ℕ => closedBall x (r ^ n) := + nhds_basis_uniformity (uniformity_basis_dist_le_pow h0 h1) +#align metric.nhds_basis_closed_ball_pow Metric.nhds_basis_closedBall_pow + +theorem isOpen_iff : IsOpen s ↔ ∀ x ∈ s, ∃ ε > 0, ball x ε ⊆ s := by + simp only [isOpen_iff_mem_nhds, mem_nhds_iff] +#align metric.is_open_iff Metric.isOpen_iff + +theorem isOpen_ball : IsOpen (ball x ε) := + isOpen_iff.2 fun _ => exists_ball_subset_ball +#align metric.is_open_ball Metric.isOpen_ball + +theorem ball_mem_nhds (x : α) {ε : ℝ} (ε0 : 0 < ε) : ball x ε ∈ 𝓝 x := + isOpen_ball.mem_nhds (mem_ball_self ε0) +#align metric.ball_mem_nhds Metric.ball_mem_nhds + +theorem closedBall_mem_nhds (x : α) {ε : ℝ} (ε0 : 0 < ε) : closedBall x ε ∈ 𝓝 x := + mem_of_superset (ball_mem_nhds x ε0) ball_subset_closedBall +#align metric.closed_ball_mem_nhds Metric.closedBall_mem_nhds + +theorem closedBall_mem_nhds_of_mem {x c : α} {ε : ℝ} (h : x ∈ ball c ε) : closedBall c ε ∈ 𝓝 x := + mem_of_superset (isOpen_ball.mem_nhds h) ball_subset_closedBall +#align metric.closed_ball_mem_nhds_of_mem Metric.closedBall_mem_nhds_of_mem + +theorem nhdsWithin_basis_ball {s : Set α} : + (𝓝[s] x).HasBasis (fun ε : ℝ => 0 < ε) fun ε => ball x ε ∩ s := + nhdsWithin_hasBasis nhds_basis_ball s +#align metric.nhds_within_basis_ball Metric.nhdsWithin_basis_ball + +theorem mem_nhdsWithin_iff {t : Set α} : s ∈ 𝓝[t] x ↔ ∃ ε > 0, ball x ε ∩ t ⊆ s := + nhdsWithin_basis_ball.mem_iff +#align metric.mem_nhds_within_iff Metric.mem_nhdsWithin_iff + +theorem tendsto_nhdsWithin_nhdsWithin [PseudoMetricSpace β] {t : Set β} {f : α → β} {a b} : + Tendsto f (𝓝[s] a) (𝓝[t] b) ↔ + ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → dist x a < δ → f x ∈ t ∧ dist (f x) b < ε := + (nhdsWithin_basis_ball.tendsto_iff nhdsWithin_basis_ball).trans <| by + simp only [inter_comm _ s, inter_comm _ t, mem_inter_iff, and_imp]; rfl +#align metric.tendsto_nhds_within_nhds_within Metric.tendsto_nhdsWithin_nhdsWithin + +theorem tendsto_nhdsWithin_nhds [PseudoMetricSpace β] {f : α → β} {a b} : + Tendsto f (𝓝[s] a) (𝓝 b) ↔ + ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → dist x a < δ → dist (f x) b < ε := by + rw [← nhdsWithin_univ b, tendsto_nhdsWithin_nhdsWithin] + simp only [mem_univ, true_and_iff] +#align metric.tendsto_nhds_within_nhds Metric.tendsto_nhdsWithin_nhds + +theorem tendsto_nhds_nhds [PseudoMetricSpace β] {f : α → β} {a b} : + Tendsto f (𝓝 a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, dist x a < δ → dist (f x) b < ε := + nhds_basis_ball.tendsto_iff nhds_basis_ball +#align metric.tendsto_nhds_nhds Metric.tendsto_nhds_nhds + +theorem continuousAt_iff [PseudoMetricSpace β] {f : α → β} {a : α} : + ContinuousAt f a ↔ ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, dist x a < δ → dist (f x) (f a) < ε := by + rw [ContinuousAt, tendsto_nhds_nhds] +#align metric.continuous_at_iff Metric.continuousAt_iff + +theorem continuousWithinAt_iff [PseudoMetricSpace β] {f : α → β} {a : α} {s : Set α} : + ContinuousWithinAt f s a ↔ + ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → dist x a < δ → dist (f x) (f a) < ε := + by rw [ContinuousWithinAt, tendsto_nhdsWithin_nhds] +#align metric.continuous_within_at_iff Metric.continuousWithinAt_iff + +theorem continuousOn_iff [PseudoMetricSpace β] {f : α → β} {s : Set α} : + ContinuousOn f s ↔ ∀ b ∈ s, ∀ ε > 0, ∃ δ > 0, ∀ a ∈ s, dist a b < δ → dist (f a) (f b) < ε := by + simp [ContinuousOn, continuousWithinAt_iff] +#align metric.continuous_on_iff Metric.continuousOn_iff + +theorem continuous_iff [PseudoMetricSpace β] {f : α → β} : + Continuous f ↔ ∀ b, ∀ ε > 0, ∃ δ > 0, ∀ a, dist a b < δ → dist (f a) (f b) < ε := + continuous_iff_continuousAt.trans <| forall_congr' fun _ => tendsto_nhds_nhds +#align metric.continuous_iff Metric.continuous_iff + +theorem tendsto_nhds {f : Filter β} {u : β → α} {a : α} : + Tendsto u f (𝓝 a) ↔ ∀ ε > 0, ∀ᶠ x in f, dist (u x) a < ε := + nhds_basis_ball.tendsto_right_iff +#align metric.tendsto_nhds Metric.tendsto_nhds + +theorem continuousAt_iff' [TopologicalSpace β] {f : β → α} {b : β} : + ContinuousAt f b ↔ ∀ ε > 0, ∀ᶠ x in 𝓝 b, dist (f x) (f b) < ε := by + rw [ContinuousAt, tendsto_nhds] +#align metric.continuous_at_iff' Metric.continuousAt_iff' + +theorem continuousWithinAt_iff' [TopologicalSpace β] {f : β → α} {b : β} {s : Set β} : + ContinuousWithinAt f s b ↔ ∀ ε > 0, ∀ᶠ x in 𝓝[s] b, dist (f x) (f b) < ε := by + rw [ContinuousWithinAt, tendsto_nhds] +#align metric.continuous_within_at_iff' Metric.continuousWithinAt_iff' + +theorem continuousOn_iff' [TopologicalSpace β] {f : β → α} {s : Set β} : + ContinuousOn f s ↔ ∀ b ∈ s, ∀ ε > 0, ∀ᶠ x in 𝓝[s] b, dist (f x) (f b) < ε := by + simp [ContinuousOn, continuousWithinAt_iff'] +#align metric.continuous_on_iff' Metric.continuousOn_iff' + +theorem continuous_iff' [TopologicalSpace β] {f : β → α} : + Continuous f ↔ ∀ (a), ∀ ε > 0, ∀ᶠ x in 𝓝 a, dist (f x) (f a) < ε := + continuous_iff_continuousAt.trans <| forall_congr' fun _ => tendsto_nhds +#align metric.continuous_iff' Metric.continuous_iff' + +theorem tendsto_atTop [Nonempty β] [SemilatticeSup β] {u : β → α} {a : α} : + Tendsto u atTop (𝓝 a) ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, dist (u n) a < ε := + (atTop_basis.tendsto_iff nhds_basis_ball).trans <| by + simp only [true_and]; rfl +#align metric.tendsto_at_top Metric.tendsto_atTop + +/-- A variant of `tendsto_atTop` that +uses `∃ N, ∀ n > N, ...` rather than `∃ N, ∀ n ≥ N, ...` +-/ +theorem tendsto_atTop' [Nonempty β] [SemilatticeSup β] [NoMaxOrder β] {u : β → α} {a : α} : + Tendsto u atTop (𝓝 a) ↔ ∀ ε > 0, ∃ N, ∀ n > N, dist (u n) a < ε := + (atTop_basis_Ioi.tendsto_iff nhds_basis_ball).trans <| by + simp only [true_and, gt_iff_lt, mem_Ioi, mem_ball] +#align metric.tendsto_at_top' Metric.tendsto_atTop' + +theorem isOpen_singleton_iff {α : Type*} [PseudoMetricSpace α] {x : α} : + IsOpen ({x} : Set α) ↔ ∃ ε > 0, ∀ y, dist y x < ε → y = x := by + simp [isOpen_iff, subset_singleton_iff, mem_ball] +#align metric.is_open_singleton_iff Metric.isOpen_singleton_iff + +/-- Given a point `x` in a discrete subset `s` of a pseudometric space, there is an open ball +centered at `x` and intersecting `s` only at `x`. -/ +theorem exists_ball_inter_eq_singleton_of_mem_discrete [DiscreteTopology s] {x : α} (hx : x ∈ s) : + ∃ ε > 0, Metric.ball x ε ∩ s = {x} := + nhds_basis_ball.exists_inter_eq_singleton_of_mem_discrete hx +#align metric.exists_ball_inter_eq_singleton_of_mem_discrete Metric.exists_ball_inter_eq_singleton_of_mem_discrete + +/-- Given a point `x` in a discrete subset `s` of a pseudometric space, there is a closed ball +of positive radius centered at `x` and intersecting `s` only at `x`. -/ +theorem exists_closedBall_inter_eq_singleton_of_discrete [DiscreteTopology s] {x : α} (hx : x ∈ s) : + ∃ ε > 0, Metric.closedBall x ε ∩ s = {x} := + nhds_basis_closedBall.exists_inter_eq_singleton_of_mem_discrete hx +#align metric.exists_closed_ball_inter_eq_singleton_of_discrete Metric.exists_closedBall_inter_eq_singleton_of_discrete + +theorem _root_.Dense.exists_dist_lt {s : Set α} (hs : Dense s) (x : α) {ε : ℝ} (hε : 0 < ε) : + ∃ y ∈ s, dist x y < ε := by + have : (ball x ε).Nonempty := by simp [hε] + simpa only [mem_ball'] using hs.exists_mem_open isOpen_ball this +#align dense.exists_dist_lt Dense.exists_dist_lt + +nonrec theorem _root_.DenseRange.exists_dist_lt {β : Type*} {f : β → α} (hf : DenseRange f) (x : α) + {ε : ℝ} (hε : 0 < ε) : ∃ y, dist x (f y) < ε := + exists_range_iff.1 (hf.exists_dist_lt x hε) +#align dense_range.exists_dist_lt DenseRange.exists_dist_lt + +end Metric + +open Metric + +/-Instantiate a pseudometric space as a pseudoemetric space. Before we can state the instance, +we need to show that the uniform structure coming from the edistance and the +distance coincide. -/ + +-- porting note: new +theorem Metric.uniformity_edist_aux {α} (d : α → α → ℝ≥0) : + ⨅ ε > (0 : ℝ), 𝓟 { p : α × α | ↑(d p.1 p.2) < ε } = + ⨅ ε > (0 : ℝ≥0∞), 𝓟 { p : α × α | ↑(d p.1 p.2) < ε } := by + simp only [le_antisymm_iff, le_iInf_iff, le_principal_iff] + refine ⟨fun ε hε => ?_, fun ε hε => ?_⟩ + · rcases ENNReal.lt_iff_exists_nnreal_btwn.1 hε with ⟨ε', ε'0, ε'ε⟩ + refine mem_iInf_of_mem (ε' : ℝ) (mem_iInf_of_mem (ENNReal.coe_pos.1 ε'0) ?_) + exact fun x hx => lt_trans (ENNReal.coe_lt_coe.2 hx) ε'ε + · lift ε to ℝ≥0 using le_of_lt hε + refine mem_iInf_of_mem (ε : ℝ≥0∞) (mem_iInf_of_mem (ENNReal.coe_pos.2 hε) ?_) + exact fun _ => ENNReal.coe_lt_coe.1 + +theorem Metric.uniformity_edist : 𝓤 α = ⨅ ε > 0, 𝓟 { p : α × α | edist p.1 p.2 < ε } := by + simp only [PseudoMetricSpace.uniformity_dist, dist_nndist, edist_nndist, + Metric.uniformity_edist_aux] +#align metric.uniformity_edist Metric.uniformity_edist + +-- see Note [lower instance priority] +/-- A pseudometric space induces a pseudoemetric space -/ +instance (priority := 100) PseudoMetricSpace.toPseudoEMetricSpace : PseudoEMetricSpace α := + { ‹PseudoMetricSpace α› with + edist_self := by simp [edist_dist] + edist_comm := fun _ _ => by simp only [edist_dist, dist_comm] + edist_triangle := fun x y z => by + simp only [edist_dist, ← ENNReal.ofReal_add, dist_nonneg] + rw [ENNReal.ofReal_le_ofReal_iff _] + · exact dist_triangle _ _ _ + · simpa using add_le_add (dist_nonneg : 0 ≤ dist x y) dist_nonneg + uniformity_edist := Metric.uniformity_edist } +#align pseudo_metric_space.to_pseudo_emetric_space PseudoMetricSpace.toPseudoEMetricSpace + +/-- Expressing the uniformity in terms of `edist` -/ +@[deprecated _root_.uniformity_basis_edist] +protected theorem Metric.uniformity_basis_edist : + (𝓤 α).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) fun ε => { p | edist p.1 p.2 < ε } := + uniformity_basis_edist +#align pseudo_metric.uniformity_basis_edist Metric.uniformity_basis_edist + +/-- In a pseudometric space, an open ball of infinite radius is the whole space -/ +theorem Metric.eball_top_eq_univ (x : α) : EMetric.ball x ∞ = Set.univ := + Set.eq_univ_iff_forall.mpr fun y => edist_lt_top y x +#align metric.eball_top_eq_univ Metric.eball_top_eq_univ + +/-- Balls defined using the distance or the edistance coincide -/ +@[simp] +theorem Metric.emetric_ball {x : α} {ε : ℝ} : EMetric.ball x (ENNReal.ofReal ε) = ball x ε := by + ext y + simp only [EMetric.mem_ball, mem_ball, edist_dist] + exact ENNReal.ofReal_lt_ofReal_iff_of_nonneg dist_nonneg +#align metric.emetric_ball Metric.emetric_ball + +/-- Balls defined using the distance or the edistance coincide -/ +@[simp] +theorem Metric.emetric_ball_nnreal {x : α} {ε : ℝ≥0} : EMetric.ball x ε = ball x ε := by + rw [← Metric.emetric_ball] + simp +#align metric.emetric_ball_nnreal Metric.emetric_ball_nnreal + +/-- Closed balls defined using the distance or the edistance coincide -/ +theorem Metric.emetric_closedBall {x : α} {ε : ℝ} (h : 0 ≤ ε) : + EMetric.closedBall x (ENNReal.ofReal ε) = closedBall x ε := by + ext y; simp [edist_le_ofReal h] +#align metric.emetric_closed_ball Metric.emetric_closedBall + +/-- Closed balls defined using the distance or the edistance coincide -/ +@[simp] +theorem Metric.emetric_closedBall_nnreal {x : α} {ε : ℝ≥0} : + EMetric.closedBall x ε = closedBall x ε := by + rw [← Metric.emetric_closedBall ε.coe_nonneg, ENNReal.ofReal_coe_nnreal] +#align metric.emetric_closed_ball_nnreal Metric.emetric_closedBall_nnreal + +@[simp] +theorem Metric.emetric_ball_top (x : α) : EMetric.ball x ⊤ = univ := + eq_univ_of_forall fun _ => edist_lt_top _ _ +#align metric.emetric_ball_top Metric.emetric_ball_top + +theorem Metric.inseparable_iff {x y : α} : Inseparable x y ↔ dist x y = 0 := by + rw [EMetric.inseparable_iff, edist_nndist, dist_nndist, ENNReal.coe_eq_zero, NNReal.coe_eq_zero] +#align metric.inseparable_iff Metric.inseparable_iff + +/-- Build a new pseudometric space from an old one where the bundled uniform structure is provably +(but typically non-definitionaly) equal to some given uniform structure. +See Note [forgetful inheritance]. +-/ +@[reducible] +def PseudoMetricSpace.replaceUniformity {α} [U : UniformSpace α] (m : PseudoMetricSpace α) + (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : PseudoMetricSpace α := + { m with + toUniformSpace := U + uniformity_dist := H.trans PseudoMetricSpace.uniformity_dist } +#align pseudo_metric_space.replace_uniformity PseudoMetricSpace.replaceUniformity + +theorem PseudoMetricSpace.replaceUniformity_eq {α} [U : UniformSpace α] (m : PseudoMetricSpace α) + (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : m.replaceUniformity H = m := by + ext + rfl +#align pseudo_metric_space.replace_uniformity_eq PseudoMetricSpace.replaceUniformity_eq + +-- ensure that the bornology is unchanged when replacing the uniformity. +example {α} [U : UniformSpace α] (m : PseudoMetricSpace α) + (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : + (PseudoMetricSpace.replaceUniformity m H).toBornology = m.toBornology := rfl + +/-- Build a new pseudo metric space from an old one where the bundled topological structure is +provably (but typically non-definitionaly) equal to some given topological structure. +See Note [forgetful inheritance]. +-/ +@[reducible] +def PseudoMetricSpace.replaceTopology {γ} [U : TopologicalSpace γ] (m : PseudoMetricSpace γ) + (H : U = m.toUniformSpace.toTopologicalSpace) : PseudoMetricSpace γ := + @PseudoMetricSpace.replaceUniformity γ (m.toUniformSpace.replaceTopology H) m rfl +#align pseudo_metric_space.replace_topology PseudoMetricSpace.replaceTopology + +theorem PseudoMetricSpace.replaceTopology_eq {γ} [U : TopologicalSpace γ] (m : PseudoMetricSpace γ) + (H : U = m.toUniformSpace.toTopologicalSpace) : m.replaceTopology H = m := by + ext + rfl +#align pseudo_metric_space.replace_topology_eq PseudoMetricSpace.replaceTopology_eq + +/-- One gets a pseudometric space from an emetric space if the edistance +is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the +uniformity are defeq in the pseudometric space and the pseudoemetric space. In this definition, the +distance is given separately, to be able to prescribe some expression which is not defeq to the +push-forward of the edistance to reals. -/ +def PseudoEMetricSpace.toPseudoMetricSpaceOfDist {α : Type u} [e : PseudoEMetricSpace α] + (dist : α → α → ℝ) (edist_ne_top : ∀ x y : α, edist x y ≠ ⊤) + (h : ∀ x y, dist x y = ENNReal.toReal (edist x y)) : PseudoMetricSpace α where + dist := dist + dist_self x := by simp [h] + dist_comm x y := by simp [h, edist_comm] + dist_triangle x y z := by + simp only [h] + exact ENNReal.toReal_le_add (edist_triangle _ _ _) (edist_ne_top _ _) (edist_ne_top _ _) + edist := edist + edist_dist _ _ := by simp only [h, ENNReal.ofReal_toReal (edist_ne_top _ _)] + toUniformSpace := e.toUniformSpace + uniformity_dist := e.uniformity_edist.trans <| by + simpa only [ENNReal.coe_toNNReal (edist_ne_top _ _), h] + using (Metric.uniformity_edist_aux fun x y : α => (edist x y).toNNReal).symm +#align pseudo_emetric_space.to_pseudo_metric_space_of_dist PseudoEMetricSpace.toPseudoMetricSpaceOfDist + +/-- One gets a pseudometric space from an emetric space if the edistance +is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the +uniformity are defeq in the pseudometric space and the emetric space. -/ +@[reducible] +def PseudoEMetricSpace.toPseudoMetricSpace {α : Type u} [PseudoEMetricSpace α] + (h : ∀ x y : α, edist x y ≠ ⊤) : PseudoMetricSpace α := + PseudoEMetricSpace.toPseudoMetricSpaceOfDist (fun x y => ENNReal.toReal (edist x y)) h fun _ _ => + rfl +#align pseudo_emetric_space.to_pseudo_metric_space PseudoEMetricSpace.toPseudoMetricSpace + +/-- Build a new pseudometric space from an old one where the bundled bornology structure is provably +(but typically non-definitionaly) equal to some given bornology structure. +See Note [forgetful inheritance]. +-/ +@[reducible] +def PseudoMetricSpace.replaceBornology {α} [B : Bornology α] (m : PseudoMetricSpace α) + (H : ∀ s, @IsBounded _ B s ↔ @IsBounded _ PseudoMetricSpace.toBornology s) : + PseudoMetricSpace α := + { m with + toBornology := B + cobounded_sets := Set.ext <| compl_surjective.forall.2 fun s => + (H s).trans <| by rw [isBounded_iff, mem_setOf_eq, compl_compl] } +#align pseudo_metric_space.replace_bornology PseudoMetricSpace.replaceBornology + +theorem PseudoMetricSpace.replaceBornology_eq {α} [m : PseudoMetricSpace α] [B : Bornology α] + (H : ∀ s, @IsBounded _ B s ↔ @IsBounded _ PseudoMetricSpace.toBornology s) : + PseudoMetricSpace.replaceBornology _ H = m := by + ext + rfl +#align pseudo_metric_space.replace_bornology_eq PseudoMetricSpace.replaceBornology_eq + +-- ensure that the uniformity is unchanged when replacing the bornology. +example {α} [B : Bornology α] (m : PseudoMetricSpace α) + (H : ∀ s, @IsBounded _ B s ↔ @IsBounded _ PseudoMetricSpace.toBornology s) : + (PseudoMetricSpace.replaceBornology m H).toUniformSpace = m.toUniformSpace := rfl + +section Real + +/-- Instantiate the reals as a pseudometric space. -/ +instance Real.pseudoMetricSpace : PseudoMetricSpace ℝ where + dist x y := |x - y| + dist_self := by simp [abs_zero] + dist_comm x y := abs_sub_comm _ _ + dist_triangle x y z := abs_sub_le _ _ _ + edist_dist := fun x y => by exact ENNReal.coe_nnreal_eq _ +#align real.pseudo_metric_space Real.pseudoMetricSpace + +theorem Real.dist_eq (x y : ℝ) : dist x y = |x - y| := rfl +#align real.dist_eq Real.dist_eq + +theorem Real.nndist_eq (x y : ℝ) : nndist x y = Real.nnabs (x - y) := rfl +#align real.nndist_eq Real.nndist_eq + +theorem Real.nndist_eq' (x y : ℝ) : nndist x y = Real.nnabs (y - x) := + nndist_comm _ _ +#align real.nndist_eq' Real.nndist_eq' + +theorem Real.dist_0_eq_abs (x : ℝ) : dist x 0 = |x| := by simp [Real.dist_eq] +#align real.dist_0_eq_abs Real.dist_0_eq_abs + +theorem Real.dist_left_le_of_mem_uIcc {x y z : ℝ} (h : y ∈ uIcc x z) : dist x y ≤ dist x z := by + simpa only [dist_comm x] using abs_sub_left_of_mem_uIcc h +#align real.dist_left_le_of_mem_uIcc Real.dist_left_le_of_mem_uIcc + +theorem Real.dist_right_le_of_mem_uIcc {x y z : ℝ} (h : y ∈ uIcc x z) : dist y z ≤ dist x z := by + simpa only [dist_comm _ z] using abs_sub_right_of_mem_uIcc h +#align real.dist_right_le_of_mem_uIcc Real.dist_right_le_of_mem_uIcc + +theorem Real.dist_le_of_mem_uIcc {x y x' y' : ℝ} (hx : x ∈ uIcc x' y') (hy : y ∈ uIcc x' y') : + dist x y ≤ dist x' y' := + abs_sub_le_of_uIcc_subset_uIcc <| uIcc_subset_uIcc (by rwa [uIcc_comm]) (by rwa [uIcc_comm]) +#align real.dist_le_of_mem_uIcc Real.dist_le_of_mem_uIcc + +theorem Real.dist_le_of_mem_Icc {x y x' y' : ℝ} (hx : x ∈ Icc x' y') (hy : y ∈ Icc x' y') : + dist x y ≤ y' - x' := by + simpa only [Real.dist_eq, abs_of_nonpos (sub_nonpos.2 <| hx.1.trans hx.2), neg_sub] using + Real.dist_le_of_mem_uIcc (Icc_subset_uIcc hx) (Icc_subset_uIcc hy) +#align real.dist_le_of_mem_Icc Real.dist_le_of_mem_Icc + +theorem Real.dist_le_of_mem_Icc_01 {x y : ℝ} (hx : x ∈ Icc (0 : ℝ) 1) (hy : y ∈ Icc (0 : ℝ) 1) : + dist x y ≤ 1 := by simpa only [sub_zero] using Real.dist_le_of_mem_Icc hx hy +#align real.dist_le_of_mem_Icc_01 Real.dist_le_of_mem_Icc_01 + +instance : OrderTopology ℝ := + orderTopology_of_nhds_abs fun x => by + simp only [nhds_basis_ball.eq_biInf, ball, Real.dist_eq, abs_sub_comm] + +theorem Real.ball_eq_Ioo (x r : ℝ) : ball x r = Ioo (x - r) (x + r) := + Set.ext fun y => by + rw [mem_ball, dist_comm, Real.dist_eq, abs_sub_lt_iff, mem_Ioo, ← sub_lt_iff_lt_add', + sub_lt_comm] +#align real.ball_eq_Ioo Real.ball_eq_Ioo + +theorem Real.closedBall_eq_Icc {x r : ℝ} : closedBall x r = Icc (x - r) (x + r) := by + ext y + rw [mem_closedBall, dist_comm, Real.dist_eq, abs_sub_le_iff, mem_Icc, ← sub_le_iff_le_add', + sub_le_comm] +#align real.closed_ball_eq_Icc Real.closedBall_eq_Icc + +theorem Real.Ioo_eq_ball (x y : ℝ) : Ioo x y = ball ((x + y) / 2) ((y - x) / 2) := by + rw [Real.ball_eq_Ioo, ← sub_div, add_comm, ← sub_add, add_sub_cancel', add_self_div_two, + ← add_div, add_assoc, add_sub_cancel'_right, add_self_div_two] +#align real.Ioo_eq_ball Real.Ioo_eq_ball + +theorem Real.Icc_eq_closedBall (x y : ℝ) : Icc x y = closedBall ((x + y) / 2) ((y - x) / 2) := by + rw [Real.closedBall_eq_Icc, ← sub_div, add_comm, ← sub_add, add_sub_cancel', add_self_div_two, ← + add_div, add_assoc, add_sub_cancel'_right, add_self_div_two] +#align real.Icc_eq_closed_ball Real.Icc_eq_closedBall + +section MetricOrdered + +variable [Preorder α] [CompactIccSpace α] + +theorem totallyBounded_Icc (a b : α) : TotallyBounded (Icc a b) := + isCompact_Icc.totallyBounded +#align totally_bounded_Icc totallyBounded_Icc + +theorem totallyBounded_Ico (a b : α) : TotallyBounded (Ico a b) := + totallyBounded_subset Ico_subset_Icc_self (totallyBounded_Icc a b) +#align totally_bounded_Ico totallyBounded_Ico + +theorem totallyBounded_Ioc (a b : α) : TotallyBounded (Ioc a b) := + totallyBounded_subset Ioc_subset_Icc_self (totallyBounded_Icc a b) +#align totally_bounded_Ioc totallyBounded_Ioc + +theorem totallyBounded_Ioo (a b : α) : TotallyBounded (Ioo a b) := + totallyBounded_subset Ioo_subset_Icc_self (totallyBounded_Icc a b) +#align totally_bounded_Ioo totallyBounded_Ioo + +end MetricOrdered + +/-- Special case of the sandwich theorem; see `tendsto_of_tendsto_of_tendsto_of_le_of_le'` for the +general case. -/ +theorem squeeze_zero' {α} {f g : α → ℝ} {t₀ : Filter α} (hf : ∀ᶠ t in t₀, 0 ≤ f t) + (hft : ∀ᶠ t in t₀, f t ≤ g t) (g0 : Tendsto g t₀ (nhds 0)) : Tendsto f t₀ (𝓝 0) := + tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds g0 hf hft +#align squeeze_zero' squeeze_zero' + +/-- Special case of the sandwich theorem; see `tendsto_of_tendsto_of_tendsto_of_le_of_le` +and `tendsto_of_tendsto_of_tendsto_of_le_of_le'` for the general case. -/ +theorem squeeze_zero {α} {f g : α → ℝ} {t₀ : Filter α} (hf : ∀ t, 0 ≤ f t) (hft : ∀ t, f t ≤ g t) + (g0 : Tendsto g t₀ (𝓝 0)) : Tendsto f t₀ (𝓝 0) := + squeeze_zero' (eventually_of_forall hf) (eventually_of_forall hft) g0 +#align squeeze_zero squeeze_zero + +theorem Metric.uniformity_eq_comap_nhds_zero : + 𝓤 α = comap (fun p : α × α => dist p.1 p.2) (𝓝 (0 : ℝ)) := by + ext s + simp only [mem_uniformity_dist, (nhds_basis_ball.comap _).mem_iff] + simp [subset_def, Real.dist_0_eq_abs] +#align metric.uniformity_eq_comap_nhds_zero Metric.uniformity_eq_comap_nhds_zero + +theorem cauchySeq_iff_tendsto_dist_atTop_0 [Nonempty β] [SemilatticeSup β] {u : β → α} : + CauchySeq u ↔ Tendsto (fun n : β × β => dist (u n.1) (u n.2)) atTop (𝓝 0) := by + rw [cauchySeq_iff_tendsto, Metric.uniformity_eq_comap_nhds_zero, tendsto_comap_iff]; rfl +#align cauchy_seq_iff_tendsto_dist_at_top_0 cauchySeq_iff_tendsto_dist_atTop_0 + +theorem tendsto_uniformity_iff_dist_tendsto_zero {f : ι → α × α} {p : Filter ι} : + Tendsto f p (𝓤 α) ↔ Tendsto (fun x => dist (f x).1 (f x).2) p (𝓝 0) := by + rw [Metric.uniformity_eq_comap_nhds_zero, tendsto_comap_iff]; rfl +#align tendsto_uniformity_iff_dist_tendsto_zero tendsto_uniformity_iff_dist_tendsto_zero + +theorem Filter.Tendsto.congr_dist {f₁ f₂ : ι → α} {p : Filter ι} {a : α} + (h₁ : Tendsto f₁ p (𝓝 a)) (h : Tendsto (fun x => dist (f₁ x) (f₂ x)) p (𝓝 0)) : + Tendsto f₂ p (𝓝 a) := + h₁.congr_uniformity <| tendsto_uniformity_iff_dist_tendsto_zero.2 h +#align filter.tendsto.congr_dist Filter.Tendsto.congr_dist + +alias tendsto_of_tendsto_of_dist := Filter.Tendsto.congr_dist +#align tendsto_of_tendsto_of_dist tendsto_of_tendsto_of_dist + +theorem tendsto_iff_of_dist {f₁ f₂ : ι → α} {p : Filter ι} {a : α} + (h : Tendsto (fun x => dist (f₁ x) (f₂ x)) p (𝓝 0)) : Tendsto f₁ p (𝓝 a) ↔ Tendsto f₂ p (𝓝 a) := + Uniform.tendsto_congr <| tendsto_uniformity_iff_dist_tendsto_zero.2 h +#align tendsto_iff_of_dist tendsto_iff_of_dist + +/-- If `u` is a neighborhood of `x`, then for small enough `r`, the closed ball +`Metric.closedBall x r` is contained in `u`. -/ +theorem eventually_closedBall_subset {x : α} {u : Set α} (hu : u ∈ 𝓝 x) : + ∀ᶠ r in 𝓝 (0 : ℝ), closedBall x r ⊆ u := by + obtain ⟨ε, εpos, hε⟩ : ∃ ε, 0 < ε ∧ closedBall x ε ⊆ u := nhds_basis_closedBall.mem_iff.1 hu + have : Iic ε ∈ 𝓝 (0 : ℝ) := Iic_mem_nhds εpos + filter_upwards [this] with _ hr using Subset.trans (closedBall_subset_closedBall hr) hε +#align eventually_closed_ball_subset eventually_closedBall_subset + +end Real + +/-- Pseudometric space structure pulled back by a function. -/ +@[reducible] +def PseudoMetricSpace.induced {α β} (f : α → β) (m : PseudoMetricSpace β) : + PseudoMetricSpace α where + dist x y := dist (f x) (f y) + dist_self x := dist_self _ + dist_comm x y := dist_comm _ _ + dist_triangle x y z := dist_triangle _ _ _ + edist x y := edist (f x) (f y) + edist_dist x y := edist_dist _ _ + toUniformSpace := UniformSpace.comap f m.toUniformSpace + uniformity_dist := (uniformity_basis_dist.comap _).eq_biInf + toBornology := Bornology.induced f + cobounded_sets := Set.ext fun s => mem_comap_iff_compl.trans <| by + simp only [← isBounded_def, isBounded_iff, ball_image_iff, mem_setOf] +#align pseudo_metric_space.induced PseudoMetricSpace.induced + +/-- Pull back a pseudometric space structure by an inducing map. This is a version of +`PseudoMetricSpace.induced` useful in case if the domain already has a `TopologicalSpace` +structure. -/ +def Inducing.comapPseudoMetricSpace {α β} [TopologicalSpace α] [m : PseudoMetricSpace β] {f : α → β} + (hf : Inducing f) : PseudoMetricSpace α := + .replaceTopology (.induced f m) hf.induced +#align inducing.comap_pseudo_metric_space Inducing.comapPseudoMetricSpace + +/-- Pull back a pseudometric space structure by a uniform inducing map. This is a version of +`PseudoMetricSpace.induced` useful in case if the domain already has a `UniformSpace` +structure. -/ +def UniformInducing.comapPseudoMetricSpace {α β} [UniformSpace α] [m : PseudoMetricSpace β] + (f : α → β) (h : UniformInducing f) : PseudoMetricSpace α := + .replaceUniformity (.induced f m) h.comap_uniformity.symm +#align uniform_inducing.comap_pseudo_metric_space UniformInducing.comapPseudoMetricSpace + +instance Subtype.pseudoMetricSpace {p : α → Prop} : PseudoMetricSpace (Subtype p) := + PseudoMetricSpace.induced Subtype.val ‹_› +#align subtype.pseudo_metric_space Subtype.pseudoMetricSpace + +theorem Subtype.dist_eq {p : α → Prop} (x y : Subtype p) : dist x y = dist (x : α) y := + rfl +#align subtype.dist_eq Subtype.dist_eq + +theorem Subtype.nndist_eq {p : α → Prop} (x y : Subtype p) : nndist x y = nndist (x : α) y := + rfl +#align subtype.nndist_eq Subtype.nndist_eq + +namespace MulOpposite + +@[to_additive] +instance instPseudoMetricSpaceMulOpposite : PseudoMetricSpace αᵐᵒᵖ := + PseudoMetricSpace.induced MulOpposite.unop ‹_› + +@[to_additive (attr := simp)] +theorem dist_unop (x y : αᵐᵒᵖ) : dist (unop x) (unop y) = dist x y := rfl +#align mul_opposite.dist_unop MulOpposite.dist_unop +#align add_opposite.dist_unop AddOpposite.dist_unop + +@[to_additive (attr := simp)] +theorem dist_op (x y : α) : dist (op x) (op y) = dist x y := rfl +#align mul_opposite.dist_op MulOpposite.dist_op +#align add_opposite.dist_op AddOpposite.dist_op + +@[to_additive (attr := simp)] +theorem nndist_unop (x y : αᵐᵒᵖ) : nndist (unop x) (unop y) = nndist x y := rfl +#align mul_opposite.nndist_unop MulOpposite.nndist_unop +#align add_opposite.nndist_unop AddOpposite.nndist_unop + +@[to_additive (attr := simp)] +theorem nndist_op (x y : α) : nndist (op x) (op y) = nndist x y := rfl +#align mul_opposite.nndist_op MulOpposite.nndist_op +#align add_opposite.nndist_op AddOpposite.nndist_op + +end MulOpposite + +section NNReal + +instance : PseudoMetricSpace ℝ≥0 := Subtype.pseudoMetricSpace + +theorem NNReal.dist_eq (a b : ℝ≥0) : dist a b = |(a : ℝ) - b| := rfl +#align nnreal.dist_eq NNReal.dist_eq + +theorem NNReal.nndist_eq (a b : ℝ≥0) : nndist a b = max (a - b) (b - a) := + eq_of_forall_ge_iff fun _ => by + simp only [← NNReal.coe_le_coe, coe_nndist, dist_eq, max_le_iff, abs_sub_le_iff, + tsub_le_iff_right, NNReal.coe_add] +#align nnreal.nndist_eq NNReal.nndist_eq + +@[simp] +theorem NNReal.nndist_zero_eq_val (z : ℝ≥0) : nndist 0 z = z := by + simp only [NNReal.nndist_eq, max_eq_right, tsub_zero, zero_tsub, zero_le'] +#align nnreal.nndist_zero_eq_val NNReal.nndist_zero_eq_val + +@[simp] +theorem NNReal.nndist_zero_eq_val' (z : ℝ≥0) : nndist z 0 = z := by + rw [nndist_comm] + exact NNReal.nndist_zero_eq_val z +#align nnreal.nndist_zero_eq_val' NNReal.nndist_zero_eq_val' + +theorem NNReal.le_add_nndist (a b : ℝ≥0) : a ≤ b + nndist a b := by + suffices (a : ℝ) ≤ (b : ℝ) + dist a b by + rwa [← NNReal.coe_le_coe, NNReal.coe_add, coe_nndist] + rw [← sub_le_iff_le_add'] + exact le_of_abs_le (dist_eq a b).ge +#align nnreal.le_add_nndist NNReal.le_add_nndist + +end NNReal + +section ULift + +variable [PseudoMetricSpace β] + +instance : PseudoMetricSpace (ULift β) := + PseudoMetricSpace.induced ULift.down ‹_› + +theorem ULift.dist_eq (x y : ULift β) : dist x y = dist x.down y.down := rfl +#align ulift.dist_eq ULift.dist_eq + +theorem ULift.nndist_eq (x y : ULift β) : nndist x y = nndist x.down y.down := rfl +#align ulift.nndist_eq ULift.nndist_eq + +@[simp] +theorem ULift.dist_up_up (x y : β) : dist (ULift.up x) (ULift.up y) = dist x y := rfl +#align ulift.dist_up_up ULift.dist_up_up + +@[simp] +theorem ULift.nndist_up_up (x y : β) : nndist (ULift.up x) (ULift.up y) = nndist x y := rfl +#align ulift.nndist_up_up ULift.nndist_up_up + +end ULift + +section Prod + +variable [PseudoMetricSpace β] + +-- porting note: added `let`, otherwise `simp` failed +instance Prod.pseudoMetricSpaceMax : PseudoMetricSpace (α × β) := + let i := PseudoEMetricSpace.toPseudoMetricSpaceOfDist + (fun x y : α × β => dist x.1 y.1 ⊔ dist x.2 y.2) + (fun x y => (max_lt (edist_lt_top _ _) (edist_lt_top _ _)).ne) fun x y => by + simp only [sup_eq_max, dist_edist, ← ENNReal.toReal_max (edist_ne_top _ _) (edist_ne_top _ _), + Prod.edist_eq] + i.replaceBornology fun s => by + simp only [← isBounded_image_fst_and_snd, isBounded_iff_eventually, ball_image_iff, ← + eventually_and, ← forall_and, ← max_le_iff] + rfl +#align prod.pseudo_metric_space_max Prod.pseudoMetricSpaceMax + +theorem Prod.dist_eq {x y : α × β} : dist x y = max (dist x.1 y.1) (dist x.2 y.2) := rfl +#align prod.dist_eq Prod.dist_eq + +@[simp] +theorem dist_prod_same_left {x : α} {y₁ y₂ : β} : dist (x, y₁) (x, y₂) = dist y₁ y₂ := by + simp [Prod.dist_eq, dist_nonneg] +#align dist_prod_same_left dist_prod_same_left + +@[simp] +theorem dist_prod_same_right {x₁ x₂ : α} {y : β} : dist (x₁, y) (x₂, y) = dist x₁ x₂ := by + simp [Prod.dist_eq, dist_nonneg] +#align dist_prod_same_right dist_prod_same_right + +theorem ball_prod_same (x : α) (y : β) (r : ℝ) : ball x r ×ˢ ball y r = ball (x, y) r := + ext fun z => by simp [Prod.dist_eq] +#align ball_prod_same ball_prod_same + +theorem closedBall_prod_same (x : α) (y : β) (r : ℝ) : + closedBall x r ×ˢ closedBall y r = closedBall (x, y) r := + ext fun z => by simp [Prod.dist_eq] +#align closed_ball_prod_same closedBall_prod_same + +theorem sphere_prod (x : α × β) (r : ℝ) : + sphere x r = sphere x.1 r ×ˢ closedBall x.2 r ∪ closedBall x.1 r ×ˢ sphere x.2 r := by + obtain hr | rfl | hr := lt_trichotomy r 0 + · simp [hr] + · cases x + simp_rw [← closedBall_eq_sphere_of_nonpos le_rfl, union_self, closedBall_prod_same] + · ext ⟨x', y'⟩ + simp_rw [Set.mem_union, Set.mem_prod, Metric.mem_closedBall, Metric.mem_sphere, Prod.dist_eq, + max_eq_iff] + refine' or_congr (and_congr_right _) (and_comm.trans (and_congr_left _)) + all_goals rintro rfl; rfl +#align sphere_prod sphere_prod + +end Prod + +-- porting note: 3 new lemmas +theorem dist_dist_dist_le_left (x y z : α) : dist (dist x z) (dist y z) ≤ dist x y := + abs_dist_sub_le .. + +theorem dist_dist_dist_le_right (x y z : α) : dist (dist x y) (dist x z) ≤ dist y z := by + simpa only [dist_comm x] using dist_dist_dist_le_left y z x + +theorem dist_dist_dist_le (x y x' y' : α) : dist (dist x y) (dist x' y') ≤ dist x x' + dist y y' := + (dist_triangle _ _ _).trans <| + add_le_add (dist_dist_dist_le_left _ _ _) (dist_dist_dist_le_right _ _ _) + +theorem uniformContinuous_dist : UniformContinuous fun p : α × α => dist p.1 p.2 := + Metric.uniformContinuous_iff.2 fun ε ε0 => + ⟨ε / 2, half_pos ε0, fun {a b} h => + calc dist (dist a.1 a.2) (dist b.1 b.2) ≤ dist a.1 b.1 + dist a.2 b.2 := + dist_dist_dist_le _ _ _ _ + _ ≤ dist a b + dist a b := add_le_add (le_max_left _ _) (le_max_right _ _) + _ < ε / 2 + ε / 2 := add_lt_add h h + _ = ε := add_halves ε⟩ +#align uniform_continuous_dist uniformContinuous_dist + +protected theorem UniformContinuous.dist [UniformSpace β] {f g : β → α} (hf : UniformContinuous f) + (hg : UniformContinuous g) : UniformContinuous fun b => dist (f b) (g b) := + uniformContinuous_dist.comp (hf.prod_mk hg) +#align uniform_continuous.dist UniformContinuous.dist + +@[continuity] +theorem continuous_dist : Continuous fun p : α × α => dist p.1 p.2 := + uniformContinuous_dist.continuous +#align continuous_dist continuous_dist + +@[continuity] +protected theorem Continuous.dist [TopologicalSpace β] {f g : β → α} (hf : Continuous f) + (hg : Continuous g) : Continuous fun b => dist (f b) (g b) := + continuous_dist.comp (hf.prod_mk hg : _) +#align continuous.dist Continuous.dist + +protected theorem Filter.Tendsto.dist {f g : β → α} {x : Filter β} {a b : α} + (hf : Tendsto f x (𝓝 a)) (hg : Tendsto g x (𝓝 b)) : + Tendsto (fun x => dist (f x) (g x)) x (𝓝 (dist a b)) := + (continuous_dist.tendsto (a, b)).comp (hf.prod_mk_nhds hg) +#align filter.tendsto.dist Filter.Tendsto.dist + +theorem nhds_comap_dist (a : α) : ((𝓝 (0 : ℝ)).comap (dist · a)) = 𝓝 a := by + simp only [@nhds_eq_comap_uniformity α, Metric.uniformity_eq_comap_nhds_zero, comap_comap, + (· ∘ ·), dist_comm] +#align nhds_comap_dist nhds_comap_dist + +theorem tendsto_iff_dist_tendsto_zero {f : β → α} {x : Filter β} {a : α} : + Tendsto f x (𝓝 a) ↔ Tendsto (fun b => dist (f b) a) x (𝓝 0) := by + rw [← nhds_comap_dist a, tendsto_comap_iff]; rfl +#align tendsto_iff_dist_tendsto_zero tendsto_iff_dist_tendsto_zero + +theorem continuous_iff_continuous_dist [TopologicalSpace β] {f : β → α} : + Continuous f ↔ Continuous fun x : β × β => dist (f x.1) (f x.2) := + ⟨fun h => h.fst'.dist h.snd', fun h => + continuous_iff_continuousAt.2 fun _ => tendsto_iff_dist_tendsto_zero.2 <| + (h.comp (continuous_id.prod_mk continuous_const)).tendsto' _ _ <| dist_self _⟩ +#align continuous_iff_continuous_dist continuous_iff_continuous_dist + +theorem uniformContinuous_nndist : UniformContinuous fun p : α × α => nndist p.1 p.2 := + uniformContinuous_dist.subtype_mk _ +#align uniform_continuous_nndist uniformContinuous_nndist + +protected theorem UniformContinuous.nndist [UniformSpace β] {f g : β → α} (hf : UniformContinuous f) + (hg : UniformContinuous g) : UniformContinuous fun b => nndist (f b) (g b) := + uniformContinuous_nndist.comp (hf.prod_mk hg) +#align uniform_continuous.nndist UniformContinuous.nndist + +theorem continuous_nndist : Continuous fun p : α × α => nndist p.1 p.2 := + uniformContinuous_nndist.continuous +#align continuous_nndist continuous_nndist + +protected theorem Continuous.nndist [TopologicalSpace β] {f g : β → α} (hf : Continuous f) + (hg : Continuous g) : Continuous fun b => nndist (f b) (g b) := + continuous_nndist.comp (hf.prod_mk hg : _) +#align continuous.nndist Continuous.nndist + +protected theorem Filter.Tendsto.nndist {f g : β → α} {x : Filter β} {a b : α} + (hf : Tendsto f x (𝓝 a)) (hg : Tendsto g x (𝓝 b)) : + Tendsto (fun x => nndist (f x) (g x)) x (𝓝 (nndist a b)) := + (continuous_nndist.tendsto (a, b)).comp (hf.prod_mk_nhds hg) +#align filter.tendsto.nndist Filter.Tendsto.nndist + +namespace Metric + +variable {x y z : α} {ε ε₁ ε₂ : ℝ} {s : Set α} + +theorem isClosed_ball : IsClosed (closedBall x ε) := + isClosed_le (continuous_id.dist continuous_const) continuous_const +#align metric.is_closed_ball Metric.isClosed_ball + +theorem isClosed_sphere : IsClosed (sphere x ε) := + isClosed_eq (continuous_id.dist continuous_const) continuous_const +#align metric.is_closed_sphere Metric.isClosed_sphere + +@[simp] +theorem closure_closedBall : closure (closedBall x ε) = closedBall x ε := + isClosed_ball.closure_eq +#align metric.closure_closed_ball Metric.closure_closedBall + +@[simp] +theorem closure_sphere : closure (sphere x ε) = sphere x ε := + isClosed_sphere.closure_eq +#align metric.closure_sphere Metric.closure_sphere + +theorem closure_ball_subset_closedBall : closure (ball x ε) ⊆ closedBall x ε := + closure_minimal ball_subset_closedBall isClosed_ball +#align metric.closure_ball_subset_closed_ball Metric.closure_ball_subset_closedBall + +theorem frontier_ball_subset_sphere : frontier (ball x ε) ⊆ sphere x ε := + frontier_lt_subset_eq (continuous_id.dist continuous_const) continuous_const +#align metric.frontier_ball_subset_sphere Metric.frontier_ball_subset_sphere + +theorem frontier_closedBall_subset_sphere : frontier (closedBall x ε) ⊆ sphere x ε := + frontier_le_subset_eq (continuous_id.dist continuous_const) continuous_const +#align metric.frontier_closed_ball_subset_sphere Metric.frontier_closedBall_subset_sphere + +theorem ball_subset_interior_closedBall : ball x ε ⊆ interior (closedBall x ε) := + interior_maximal ball_subset_closedBall isOpen_ball +#align metric.ball_subset_interior_closed_ball Metric.ball_subset_interior_closedBall + +/-- ε-characterization of the closure in pseudometric spaces-/ +theorem mem_closure_iff {s : Set α} {a : α} : a ∈ closure s ↔ ∀ ε > 0, ∃ b ∈ s, dist a b < ε := + (mem_closure_iff_nhds_basis nhds_basis_ball).trans <| by simp only [mem_ball, dist_comm] +#align metric.mem_closure_iff Metric.mem_closure_iff + +theorem mem_closure_range_iff {e : β → α} {a : α} : + a ∈ closure (range e) ↔ ∀ ε > 0, ∃ k : β, dist a (e k) < ε := by + simp only [mem_closure_iff, exists_range_iff] +#align metric.mem_closure_range_iff Metric.mem_closure_range_iff + +theorem mem_closure_range_iff_nat {e : β → α} {a : α} : + a ∈ closure (range e) ↔ ∀ n : ℕ, ∃ k : β, dist a (e k) < 1 / ((n : ℝ) + 1) := + (mem_closure_iff_nhds_basis nhds_basis_ball_inv_nat_succ).trans <| by + simp only [mem_ball, dist_comm, exists_range_iff, forall_const] +#align metric.mem_closure_range_iff_nat Metric.mem_closure_range_iff_nat + +theorem mem_of_closed' {s : Set α} (hs : IsClosed s) {a : α} : + a ∈ s ↔ ∀ ε > 0, ∃ b ∈ s, dist a b < ε := by + simpa only [hs.closure_eq] using @mem_closure_iff _ _ s a +#align metric.mem_of_closed' Metric.mem_of_closed' + +theorem closedBall_zero' (x : α) : closedBall x 0 = closure {x} := + Subset.antisymm + (fun _y hy => + mem_closure_iff.2 fun _ε ε0 => ⟨x, mem_singleton x, (mem_closedBall.1 hy).trans_lt ε0⟩) + (closure_minimal (singleton_subset_iff.2 (dist_self x).le) isClosed_ball) +#align metric.closed_ball_zero' Metric.closedBall_zero' + +lemma eventually_isCompact_closedBall [LocallyCompactSpace α] (x : α) : + ∀ᶠ r in 𝓝 (0 : ℝ), IsCompact (closedBall x r) := by + rcases local_compact_nhds (x := x) (n := univ) univ_mem with ⟨s, hs, -, s_compact⟩ + filter_upwards [eventually_closedBall_subset hs] with r hr + exact IsCompact.of_isClosed_subset s_compact isClosed_ball hr + +lemma exists_isCompact_closedBall [LocallyCompactSpace α] (x : α) : + ∃ r, 0 < r ∧ IsCompact (closedBall x r) := by + have : ∀ᶠ r in 𝓝[>] 0, IsCompact (closedBall x r) := + eventually_nhdsWithin_of_eventually_nhds (eventually_isCompact_closedBall x) + simpa only [and_comm] using (this.and self_mem_nhdsWithin).exists + +theorem dense_iff {s : Set α} : Dense s ↔ ∀ x, ∀ r > 0, (ball x r ∩ s).Nonempty := + forall_congr' fun x => by + simp only [mem_closure_iff, Set.Nonempty, exists_prop, mem_inter_iff, mem_ball', and_comm] +#align metric.dense_iff Metric.dense_iff + +theorem denseRange_iff {f : β → α} : DenseRange f ↔ ∀ x, ∀ r > 0, ∃ y, dist x (f y) < r := + forall_congr' fun x => by simp only [mem_closure_iff, exists_range_iff] +#align metric.dense_range_iff Metric.denseRange_iff + +-- porting note: `TopologicalSpace.IsSeparable.separableSpace` moved to `EMetricSpace` + +/-- The preimage of a separable set by an inducing map is separable. -/ +protected theorem _root_.Inducing.isSeparable_preimage {f : β → α} [TopologicalSpace β] + (hf : Inducing f) {s : Set α} (hs : IsSeparable s) : IsSeparable (f ⁻¹' s) := by + have : SeparableSpace s := hs.separableSpace + have : SecondCountableTopology s := UniformSpace.secondCountable_of_separable _ + have : Inducing ((mapsTo_preimage f s).restrict _ _ _) := + (hf.comp inducing_subtype_val).codRestrict _ + have := this.secondCountableTopology + exact isSeparable_of_separableSpace_subtype _ +#align inducing.is_separable_preimage Inducing.isSeparable_preimage + +protected theorem _root_.Embedding.isSeparable_preimage {f : β → α} [TopologicalSpace β] + (hf : Embedding f) {s : Set α} (hs : IsSeparable s) : IsSeparable (f ⁻¹' s) := + hf.toInducing.isSeparable_preimage hs +#align embedding.is_separable_preimage Embedding.isSeparable_preimage + +/-- If a map is continuous on a separable set `s`, then the image of `s` is also separable. -/ +theorem _root_.ContinuousOn.isSeparable_image [TopologicalSpace β] {f : α → β} {s : Set α} + (hf : ContinuousOn f s) (hs : IsSeparable s) : IsSeparable (f '' s) := by + rw [image_eq_range, ← image_univ] + exact (isSeparable_univ_iff.2 hs.separableSpace).image hf.restrict +#align continuous_on.is_separable_image ContinuousOn.isSeparable_image + +end Metric + +section Pi + +open Finset + +variable {π : β → Type*} [Fintype β] [∀ b, PseudoMetricSpace (π b)] + +/-- A finite product of pseudometric spaces is a pseudometric space, with the sup distance. -/ +instance pseudoMetricSpacePi : PseudoMetricSpace (∀ b, π b) := by + /- we construct the instance from the pseudoemetric space instance to avoid checking again that + the uniformity is the same as the product uniformity, but we register nevertheless a nice + formula for the distance -/ + let i := PseudoEMetricSpace.toPseudoMetricSpaceOfDist + (fun f g : ∀ b, π b => ((sup univ fun b => nndist (f b) (g b) : ℝ≥0) : ℝ)) + (fun f g => ((Finset.sup_lt_iff bot_lt_top).2 fun b _ => edist_lt_top _ _).ne) + (fun f g => by + simp only [edist_pi_def, edist_nndist, ← ENNReal.coe_finset_sup, ENNReal.coe_toReal]) + refine i.replaceBornology fun s => ?_ + simp only [← isBounded_def, isBounded_iff_eventually, ← forall_isBounded_image_eval_iff, + ball_image_iff, ← Filter.eventually_all, Function.eval_apply, @dist_nndist (π _)] + refine' eventually_congr ((eventually_ge_atTop 0).mono fun C hC => _) + lift C to ℝ≥0 using hC + refine' ⟨fun H x hx y hy => NNReal.coe_le_coe.2 <| Finset.sup_le fun b _ => H b x hx y hy, + fun H b x hx y hy => NNReal.coe_le_coe.2 _⟩ + simpa only using Finset.sup_le_iff.1 (NNReal.coe_le_coe.1 <| H hx hy) b (Finset.mem_univ b) +#align pseudo_metric_space_pi pseudoMetricSpacePi + +theorem nndist_pi_def (f g : ∀ b, π b) : nndist f g = sup univ fun b => nndist (f b) (g b) := + NNReal.eq rfl +#align nndist_pi_def nndist_pi_def + +theorem dist_pi_def (f g : ∀ b, π b) : dist f g = (sup univ fun b => nndist (f b) (g b) : ℝ≥0) := + rfl +#align dist_pi_def dist_pi_def + +theorem nndist_pi_le_iff {f g : ∀ b, π b} {r : ℝ≥0} : + nndist f g ≤ r ↔ ∀ b, nndist (f b) (g b) ≤ r := by simp [nndist_pi_def] +#align nndist_pi_le_iff nndist_pi_le_iff + +theorem nndist_pi_lt_iff {f g : ∀ b, π b} {r : ℝ≥0} (hr : 0 < r) : + nndist f g < r ↔ ∀ b, nndist (f b) (g b) < r := by + simp [nndist_pi_def, Finset.sup_lt_iff (show ⊥ < r from hr)] +#align nndist_pi_lt_iff nndist_pi_lt_iff + +theorem nndist_pi_eq_iff {f g : ∀ b, π b} {r : ℝ≥0} (hr : 0 < r) : + nndist f g = r ↔ (∃ i, nndist (f i) (g i) = r) ∧ ∀ b, nndist (f b) (g b) ≤ r := by + rw [eq_iff_le_not_lt, nndist_pi_lt_iff hr, nndist_pi_le_iff, not_forall, and_comm] + simp_rw [not_lt, and_congr_left_iff, le_antisymm_iff] + intro h + refine' exists_congr fun b => _ + apply (and_iff_right <| h _).symm +#align nndist_pi_eq_iff nndist_pi_eq_iff + +theorem dist_pi_lt_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 < r) : + dist f g < r ↔ ∀ b, dist (f b) (g b) < r := by + lift r to ℝ≥0 using hr.le + exact nndist_pi_lt_iff hr +#align dist_pi_lt_iff dist_pi_lt_iff + +theorem dist_pi_le_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 ≤ r) : + dist f g ≤ r ↔ ∀ b, dist (f b) (g b) ≤ r := by + lift r to ℝ≥0 using hr + exact nndist_pi_le_iff +#align dist_pi_le_iff dist_pi_le_iff + +theorem dist_pi_eq_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 < r) : + dist f g = r ↔ (∃ i, dist (f i) (g i) = r) ∧ ∀ b, dist (f b) (g b) ≤ r := by + lift r to ℝ≥0 using hr.le + simp_rw [← coe_nndist, NNReal.coe_eq, nndist_pi_eq_iff hr, NNReal.coe_le_coe] +#align dist_pi_eq_iff dist_pi_eq_iff + +theorem dist_pi_le_iff' [Nonempty β] {f g : ∀ b, π b} {r : ℝ} : + dist f g ≤ r ↔ ∀ b, dist (f b) (g b) ≤ r := by + by_cases hr : 0 ≤ r + · exact dist_pi_le_iff hr + · exact iff_of_false (fun h => hr <| dist_nonneg.trans h) fun h => + hr <| dist_nonneg.trans <| h <| Classical.arbitrary _ +#align dist_pi_le_iff' dist_pi_le_iff' + +theorem dist_pi_const_le (a b : α) : (dist (fun _ : β => a) fun _ => b) ≤ dist a b := + (dist_pi_le_iff dist_nonneg).2 fun _ => le_rfl +#align dist_pi_const_le dist_pi_const_le + +theorem nndist_pi_const_le (a b : α) : (nndist (fun _ : β => a) fun _ => b) ≤ nndist a b := + nndist_pi_le_iff.2 fun _ => le_rfl +#align nndist_pi_const_le nndist_pi_const_le + +@[simp] +theorem dist_pi_const [Nonempty β] (a b : α) : (dist (fun _ : β => a) fun _ => b) = dist a b := by + simpa only [dist_edist] using congr_arg ENNReal.toReal (edist_pi_const a b) +#align dist_pi_const dist_pi_const + +@[simp] +theorem nndist_pi_const [Nonempty β] (a b : α) : + (nndist (fun _ : β => a) fun _ => b) = nndist a b := + NNReal.eq <| dist_pi_const a b +#align nndist_pi_const nndist_pi_const + +theorem nndist_le_pi_nndist (f g : ∀ b, π b) (b : β) : nndist (f b) (g b) ≤ nndist f g := by + rw [← ENNReal.coe_le_coe, ← edist_nndist, ← edist_nndist] + exact edist_le_pi_edist f g b +#align nndist_le_pi_nndist nndist_le_pi_nndist + +theorem dist_le_pi_dist (f g : ∀ b, π b) (b : β) : dist (f b) (g b) ≤ dist f g := by + simp only [dist_nndist, NNReal.coe_le_coe, nndist_le_pi_nndist f g b] +#align dist_le_pi_dist dist_le_pi_dist + +/-- An open ball in a product space is a product of open balls. See also `ball_pi'` +for a version assuming `Nonempty β` instead of `0 < r`. -/ +theorem ball_pi (x : ∀ b, π b) {r : ℝ} (hr : 0 < r) : + ball x r = Set.pi univ fun b => ball (x b) r := by + ext p + simp [dist_pi_lt_iff hr] +#align ball_pi ball_pi + +/-- An open ball in a product space is a product of open balls. See also `ball_pi` +for a version assuming `0 < r` instead of `Nonempty β`. -/ +theorem ball_pi' [Nonempty β] (x : ∀ b, π b) (r : ℝ) : + ball x r = Set.pi univ fun b => ball (x b) r := + (lt_or_le 0 r).elim (ball_pi x) fun hr => by simp [ball_eq_empty.2 hr] +#align ball_pi' ball_pi' + +/-- A closed ball in a product space is a product of closed balls. See also `closedBall_pi'` +for a version assuming `Nonempty β` instead of `0 ≤ r`. -/ +theorem closedBall_pi (x : ∀ b, π b) {r : ℝ} (hr : 0 ≤ r) : + closedBall x r = Set.pi univ fun b => closedBall (x b) r := by + ext p + simp [dist_pi_le_iff hr] +#align closed_ball_pi closedBall_pi + +/-- A closed ball in a product space is a product of closed balls. See also `closedBall_pi` +for a version assuming `0 ≤ r` instead of `Nonempty β`. -/ +theorem closedBall_pi' [Nonempty β] (x : ∀ b, π b) (r : ℝ) : + closedBall x r = Set.pi univ fun b => closedBall (x b) r := + (le_or_lt 0 r).elim (closedBall_pi x) fun hr => by simp [closedBall_eq_empty.2 hr] +#align closed_ball_pi' closedBall_pi' + +/-- A sphere in a product space is a union of spheres on each component restricted to the closed +ball. -/ +theorem sphere_pi (x : ∀ b, π b) {r : ℝ} (h : 0 < r ∨ Nonempty β) : + sphere x r = (⋃ i : β, Function.eval i ⁻¹' sphere (x i) r) ∩ closedBall x r := by + obtain hr | rfl | hr := lt_trichotomy r 0 + · simp [hr] + · rw [closedBall_eq_sphere_of_nonpos le_rfl, eq_comm, Set.inter_eq_right] + letI := h.resolve_left (lt_irrefl _) + inhabit β + refine' subset_iUnion_of_subset default _ + intro x hx + replace hx := hx.le + rw [dist_pi_le_iff le_rfl] at hx + exact le_antisymm (hx default) dist_nonneg + · ext + simp [dist_pi_eq_iff hr, dist_pi_le_iff hr.le] +#align sphere_pi sphere_pi + +@[simp] +theorem Fin.nndist_insertNth_insertNth {n : ℕ} {α : Fin (n + 1) → Type*} + [∀ i, PseudoMetricSpace (α i)] (i : Fin (n + 1)) (x y : α i) (f g : ∀ j, α (i.succAbove j)) : + nndist (i.insertNth x f) (i.insertNth y g) = max (nndist x y) (nndist f g) := + eq_of_forall_ge_iff fun c => by simp [nndist_pi_le_iff, i.forall_iff_succAbove] +#align fin.nndist_insert_nth_insert_nth Fin.nndist_insertNth_insertNth + +@[simp] +theorem Fin.dist_insertNth_insertNth {n : ℕ} {α : Fin (n + 1) → Type*} + [∀ i, PseudoMetricSpace (α i)] (i : Fin (n + 1)) (x y : α i) (f g : ∀ j, α (i.succAbove j)) : + dist (i.insertNth x f) (i.insertNth y g) = max (dist x y) (dist f g) := by + simp only [dist_nndist, Fin.nndist_insertNth_insertNth, NNReal.coe_max] +#align fin.dist_insert_nth_insert_nth Fin.dist_insertNth_insertNth + +theorem Real.dist_le_of_mem_pi_Icc {x y x' y' : β → ℝ} (hx : x ∈ Icc x' y') (hy : y ∈ Icc x' y') : + dist x y ≤ dist x' y' := by + refine' (dist_pi_le_iff dist_nonneg).2 fun b => + (Real.dist_le_of_mem_uIcc _ _).trans (dist_le_pi_dist x' y' b) <;> refine' Icc_subset_uIcc _ + exacts [⟨hx.1 _, hx.2 _⟩, ⟨hy.1 _, hy.2 _⟩] +#align real.dist_le_of_mem_pi_Icc Real.dist_le_of_mem_pi_Icc + +end Pi + +section Compact + +/-- Any compact set in a pseudometric space can be covered by finitely many balls of a given +positive radius -/ +theorem finite_cover_balls_of_compact {α : Type u} [PseudoMetricSpace α] {s : Set α} + (hs : IsCompact s) {e : ℝ} (he : 0 < e) : + ∃ t, t ⊆ s ∧ Set.Finite t ∧ s ⊆ ⋃ x ∈ t, ball x e := + let ⟨t, hts, ht⟩ := hs.elim_nhds_subcover _ (fun x _ => ball_mem_nhds x he) + ⟨t, hts, t.finite_toSet, ht⟩ +#align finite_cover_balls_of_compact finite_cover_balls_of_compact + +alias IsCompact.finite_cover_balls := finite_cover_balls_of_compact +#align is_compact.finite_cover_balls IsCompact.finite_cover_balls + +end Compact + +section ProperSpace + +open Metric + +/-- A pseudometric space is proper if all closed balls are compact. -/ +class ProperSpace (α : Type u) [PseudoMetricSpace α] : Prop where + isCompact_closedBall : ∀ x : α, ∀ r, IsCompact (closedBall x r) +#align proper_space ProperSpace + +export ProperSpace (isCompact_closedBall) + +/-- In a proper pseudometric space, all spheres are compact. -/ +theorem isCompact_sphere {α : Type*} [PseudoMetricSpace α] [ProperSpace α] (x : α) (r : ℝ) : + IsCompact (sphere x r) := + (isCompact_closedBall x r).of_isClosed_subset isClosed_sphere sphere_subset_closedBall +#align is_compact_sphere isCompact_sphere + +/-- In a proper pseudometric space, any sphere is a `CompactSpace` when considered as a subtype. -/ +instance Metric.sphere.compactSpace {α : Type*} [PseudoMetricSpace α] [ProperSpace α] + (x : α) (r : ℝ) : CompactSpace (sphere x r) := + isCompact_iff_compactSpace.mp (isCompact_sphere _ _) + +-- see Note [lower instance priority] +/-- A proper pseudo metric space is sigma compact, and therefore second countable. -/ +instance (priority := 100) secondCountable_of_proper [ProperSpace α] : + SecondCountableTopology α := by + -- We already have `sigmaCompactSpace_of_locallyCompact_secondCountable`, so we don't + -- add an instance for `SigmaCompactSpace`. + suffices SigmaCompactSpace α by exact EMetric.secondCountable_of_sigmaCompact α + rcases em (Nonempty α) with (⟨⟨x⟩⟩ | hn) + · exact ⟨⟨fun n => closedBall x n, fun n => isCompact_closedBall _ _, iUnion_closedBall_nat _⟩⟩ + · exact ⟨⟨fun _ => ∅, fun _ => isCompact_empty, iUnion_eq_univ_iff.2 fun x => (hn ⟨x⟩).elim⟩⟩ +#align second_countable_of_proper secondCountable_of_proper + +/-- If all closed balls of large enough radius are compact, then the space is proper. Especially +useful when the lower bound for the radius is 0. -/ +theorem properSpace_of_compact_closedBall_of_le (R : ℝ) + (h : ∀ x : α, ∀ r, R ≤ r → IsCompact (closedBall x r)) : ProperSpace α := + ⟨fun x r => IsCompact.of_isClosed_subset (h x (max r R) (le_max_right _ _)) isClosed_ball + (closedBall_subset_closedBall <| le_max_left _ _)⟩ +#align proper_space_of_compact_closed_ball_of_le properSpace_of_compact_closedBall_of_le + +-- A compact pseudometric space is proper +-- see Note [lower instance priority] +instance (priority := 100) proper_of_compact [CompactSpace α] : ProperSpace α := + ⟨fun _ _ => isClosed_ball.isCompact⟩ +#align proper_of_compact proper_of_compact + +-- see Note [lower instance priority] +/-- A proper space is locally compact -/ +instance (priority := 100) locally_compact_of_proper [ProperSpace α] : LocallyCompactSpace α := + locallyCompactSpace_of_hasBasis (fun _ => nhds_basis_closedBall) fun _ _ _ => + isCompact_closedBall _ _ +#align locally_compact_of_proper locally_compact_of_proper + +-- see Note [lower instance priority] +/-- A proper space is complete -/ +instance (priority := 100) complete_of_proper [ProperSpace α] : CompleteSpace α := + ⟨fun {f} hf => by + /- We want to show that the Cauchy filter `f` is converging. It suffices to find a closed + ball (therefore compact by properness) where it is nontrivial. -/ + obtain ⟨t, t_fset, ht⟩ : ∃ t ∈ f, ∀ x ∈ t, ∀ y ∈ t, dist x y < 1 := + (Metric.cauchy_iff.1 hf).2 1 zero_lt_one + rcases hf.1.nonempty_of_mem t_fset with ⟨x, xt⟩ + have : closedBall x 1 ∈ f := mem_of_superset t_fset fun y yt => (ht y yt x xt).le + rcases (isCompact_iff_totallyBounded_isComplete.1 (isCompact_closedBall x 1)).2 f hf + (le_principal_iff.2 this) with + ⟨y, -, hy⟩ + exact ⟨y, hy⟩⟩ +#align complete_of_proper complete_of_proper + +/-- A binary product of proper spaces is proper. -/ +instance prod_properSpace {α : Type*} {β : Type*} [PseudoMetricSpace α] [PseudoMetricSpace β] + [ProperSpace α] [ProperSpace β] : ProperSpace (α × β) where + isCompact_closedBall := by + rintro ⟨x, y⟩ r + rw [← closedBall_prod_same x y] + exact (isCompact_closedBall x r).prod (isCompact_closedBall y r) +#align prod_proper_space prod_properSpace + +/-- A finite product of proper spaces is proper. -/ +instance pi_properSpace {π : β → Type*} [Fintype β] [∀ b, PseudoMetricSpace (π b)] + [h : ∀ b, ProperSpace (π b)] : ProperSpace (∀ b, π b) := by + refine' properSpace_of_compact_closedBall_of_le 0 fun x r hr => _ + rw [closedBall_pi _ hr] + exact isCompact_univ_pi fun _ => isCompact_closedBall _ _ +#align pi_proper_space pi_properSpace + +variable [ProperSpace α] {x : α} {r : ℝ} {s : Set α} + +/-- If a nonempty ball in a proper space includes a closed set `s`, then there exists a nonempty +ball with the same center and a strictly smaller radius that includes `s`. -/ +theorem exists_pos_lt_subset_ball (hr : 0 < r) (hs : IsClosed s) (h : s ⊆ ball x r) : + ∃ r' ∈ Ioo 0 r, s ⊆ ball x r' := by + rcases eq_empty_or_nonempty s with (rfl | hne) + · exact ⟨r / 2, ⟨half_pos hr, half_lt_self hr⟩, empty_subset _⟩ + have : IsCompact s := + (isCompact_closedBall x r).of_isClosed_subset hs (h.trans ball_subset_closedBall) + obtain ⟨y, hys, hy⟩ : ∃ y ∈ s, s ⊆ closedBall x (dist y x) := + this.exists_forall_ge hne (continuous_id.dist continuous_const).continuousOn + have hyr : dist y x < r := h hys + rcases exists_between hyr with ⟨r', hyr', hrr'⟩ + exact ⟨r', ⟨dist_nonneg.trans_lt hyr', hrr'⟩, hy.trans <| closedBall_subset_ball hyr'⟩ +#align exists_pos_lt_subset_ball exists_pos_lt_subset_ball + +/-- If a ball in a proper space includes a closed set `s`, then there exists a ball with the same +center and a strictly smaller radius that includes `s`. -/ +theorem exists_lt_subset_ball (hs : IsClosed s) (h : s ⊆ ball x r) : ∃ r' < r, s ⊆ ball x r' := by + cases' le_or_lt r 0 with hr hr + · rw [ball_eq_empty.2 hr, subset_empty_iff] at h + subst s + exact (exists_lt r).imp fun r' hr' => ⟨hr', empty_subset _⟩ + · exact (exists_pos_lt_subset_ball hr hs h).imp fun r' hr' => ⟨hr'.1.2, hr'.2⟩ +#align exists_lt_subset_ball exists_lt_subset_ball + +end ProperSpace + +theorem IsCompact.isSeparable {s : Set α} (hs : IsCompact s) : IsSeparable s := + haveI : CompactSpace s := isCompact_iff_compactSpace.mp hs + isSeparable_of_separableSpace_subtype s +#align is_compact.is_separable IsCompact.isSeparable + +namespace Metric + +section SecondCountable + +open TopologicalSpace + +/-- A pseudometric space is second countable if, for every `ε > 0`, there is a countable set which +is `ε`-dense. -/ +theorem secondCountable_of_almost_dense_set + (H : ∀ ε > (0 : ℝ), ∃ s : Set α, s.Countable ∧ ∀ x, ∃ y ∈ s, dist x y ≤ ε) : + SecondCountableTopology α := by + refine' EMetric.secondCountable_of_almost_dense_set fun ε ε0 => _ + rcases ENNReal.lt_iff_exists_nnreal_btwn.1 ε0 with ⟨ε', ε'0, ε'ε⟩ + choose s hsc y hys hyx using H ε' (by exact_mod_cast ε'0) + refine' ⟨s, hsc, iUnion₂_eq_univ_iff.2 fun x => ⟨y x, hys _, le_trans _ ε'ε.le⟩⟩ + exact_mod_cast hyx x +#align metric.second_countable_of_almost_dense_set Metric.secondCountable_of_almost_dense_set + +end SecondCountable + +end Metric + +theorem lebesgue_number_lemma_of_metric {s : Set α} {ι : Sort*} {c : ι → Set α} (hs : IsCompact s) + (hc₁ : ∀ i, IsOpen (c i)) (hc₂ : s ⊆ ⋃ i, c i) : ∃ δ > 0, ∀ x ∈ s, ∃ i, ball x δ ⊆ c i := + let ⟨_n, en, hn⟩ := lebesgue_number_lemma hs hc₁ hc₂ + let ⟨δ, δ0, hδ⟩ := mem_uniformity_dist.1 en + ⟨δ, δ0, fun x hx => let ⟨i, hi⟩ := hn x hx; ⟨i, fun _y hy => hi (hδ (mem_ball'.mp hy))⟩⟩ +#align lebesgue_number_lemma_of_metric lebesgue_number_lemma_of_metric + +theorem lebesgue_number_lemma_of_metric_sUnion {s : Set α} {c : Set (Set α)} (hs : IsCompact s) + (hc₁ : ∀ t ∈ c, IsOpen t) (hc₂ : s ⊆ ⋃₀ c) : ∃ δ > 0, ∀ x ∈ s, ∃ t ∈ c, ball x δ ⊆ t := by + rw [sUnion_eq_iUnion] at hc₂; simpa using lebesgue_number_lemma_of_metric hs (by simpa) hc₂ +#align lebesgue_number_lemma_of_metric_sUnion lebesgue_number_lemma_of_metric_sUnion + +namespace Metric +theorem exists_isLocalMin_mem_ball [ProperSpace α] [TopologicalSpace β] + [ConditionallyCompleteLinearOrder β] [OrderTopology β] {f : α → β} {a z : α} {r : ℝ} + (hf : ContinuousOn f (closedBall a r)) (hz : z ∈ closedBall a r) + (hf1 : ∀ z' ∈ sphere a r, f z < f z') : ∃ z ∈ ball a r, IsLocalMin f z := by + simp_rw [← closedBall_diff_ball] at hf1 + exact (isCompact_closedBall a r).exists_isLocalMin_mem_open ball_subset_closedBall hf hz hf1 + isOpen_ball +#align metric.exists_local_min_mem_ball Metric.exists_isLocalMin_mem_ball + +end Metric diff --git a/Mathlib/Topology/MetricSpace/Metrizable.lean b/Mathlib/Topology/Metrizable/Basic.lean similarity index 51% rename from Mathlib/Topology/MetricSpace/Metrizable.lean rename to Mathlib/Topology/Metrizable/Basic.lean index 591f5728f3811..05e67be4a4ec9 100644 --- a/Mathlib/Topology/MetricSpace/Metrizable.lean +++ b/Mathlib/Topology/Metrizable/Basic.lean @@ -3,10 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Analysis.SpecificLimits.Basic -import Mathlib.Topology.UrysohnsLemma -import Mathlib.Topology.ContinuousFunction.Bounded -import Mathlib.Topology.UniformSpace.Cauchy +import Mathlib.Topology.MetricSpace.Basic #align_import topology.metric_space.metrizable from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" @@ -15,17 +12,11 @@ import Mathlib.Topology.UniformSpace.Cauchy In this file we define metrizable topological spaces, i.e., topological spaces for which there exists a metric space structure that generates the same topology. - -We also show that a T₃ topological space with second countable topology `X` is metrizable. - -First we prove that `X` can be embedded into `l^∞`, then use this embedding to pull back the metric -space structure. -/ open Set Filter Metric - -open BoundedContinuousFunction Filter Topology +open scoped Filter Topology namespace TopologicalSpace @@ -145,101 +136,4 @@ theorem IsSeparable.secondCountableTopology [PseudoMetrizableSpace X] {s : Set X have := hs.separableSpace exact UniformSpace.secondCountable_of_separable s -variable (X) -variable [T3Space X] [SecondCountableTopology X] - -/-- A T₃ topological space with second countable topology can be embedded into `l^∞ = ℕ →ᵇ ℝ`. -/ -theorem exists_embedding_l_infty : ∃ f : X → ℕ →ᵇ ℝ, Embedding f := by - -- Choose a countable basis, and consider the set `s` of pairs of set `(U, V)` such that `U ∈ B`, - -- `V ∈ B`, and `closure U ⊆ V`. - rcases exists_countable_basis X with ⟨B, hBc, -, hB⟩ - let s : Set (Set X × Set X) := { UV ∈ B ×ˢ B | closure UV.1 ⊆ UV.2 } - -- `s` is a countable set. - haveI : Encodable s := ((hBc.prod hBc).mono (inter_subset_left _ _)).toEncodable - -- We don't have the space of bounded (possibly discontinuous) functions, so we equip `s` - -- with the discrete topology and deal with `s →ᵇ ℝ` instead. - letI : TopologicalSpace s := ⊥ - haveI : DiscreteTopology s := ⟨rfl⟩ - rsuffices ⟨f, hf⟩ : ∃ f : X → s →ᵇ ℝ, Embedding f - · exact ⟨fun x => (f x).extend (Encodable.encode' s) 0, - (BoundedContinuousFunction.isometry_extend (Encodable.encode' s) (0 : ℕ →ᵇ ℝ)).embedding.comp - hf⟩ - have hd : ∀ UV : s, Disjoint (closure UV.1.1) UV.1.2ᶜ := - fun UV => disjoint_compl_right.mono_right (compl_subset_compl.2 UV.2.2) - -- Choose a sequence of `εₙ > 0`, `n : s`, that is bounded above by `1` and tends to zero - -- along the `cofinite` filter. - obtain ⟨ε, ε01, hε⟩ : ∃ ε : s → ℝ, (∀ UV, ε UV ∈ Ioc (0 : ℝ) 1) ∧ Tendsto ε cofinite (𝓝 0) := by - rcases posSumOfEncodable zero_lt_one s with ⟨ε, ε0, c, hεc, hc1⟩ - refine' ⟨ε, fun UV => ⟨ε0 UV, _⟩, hεc.summable.tendsto_cofinite_zero⟩ - exact (le_hasSum hεc UV fun _ _ => (ε0 _).le).trans hc1 - /- For each `UV = (U, V) ∈ s` we use Urysohn's lemma to choose a function `f UV` that is equal to - zero on `U` and is equal to `ε UV` on the complement to `V`. -/ - have : ∀ UV : s, ∃ f : C(X, ℝ), - EqOn f 0 UV.1.1 ∧ EqOn f (fun _ => ε UV) UV.1.2ᶜ ∧ ∀ x, f x ∈ Icc 0 (ε UV) := by - intro UV - rcases exists_continuous_zero_one_of_closed isClosed_closure - (hB.isOpen UV.2.1.2).isClosed_compl (hd UV) with - ⟨f, hf₀, hf₁, hf01⟩ - exact ⟨ε UV • f, fun x hx => by simp [hf₀ (subset_closure hx)], fun x hx => by simp [hf₁ hx], - fun x => ⟨mul_nonneg (ε01 _).1.le (hf01 _).1, mul_le_of_le_one_right (ε01 _).1.le (hf01 _).2⟩⟩ - choose f hf0 hfε hf0ε using this - have hf01 : ∀ UV x, f UV x ∈ Icc (0 : ℝ) 1 := - fun UV x => Icc_subset_Icc_right (ε01 _).2 (hf0ε _ _) - -- The embedding is given by `F x UV = f UV x`. - set F : X → s →ᵇ ℝ := fun x => - ⟨⟨fun UV => f UV x, continuous_of_discreteTopology⟩, 1, - fun UV₁ UV₂ => Real.dist_le_of_mem_Icc_01 (hf01 _ _) (hf01 _ _)⟩ - have hF : ∀ x UV, F x UV = f UV x := fun _ _ => rfl - refine' ⟨F, Embedding.mk' _ (fun x y hxy => _) fun x => le_antisymm _ _⟩ - · /- First we prove that `F` is injective. Indeed, if `F x = F y` and `x ≠ y`, then we can find - `(U, V) ∈ s` such that `x ∈ U` and `y ∉ V`, hence `F x UV = 0 ≠ ε UV = F y UV`. -/ - by_contra Hne - rcases hB.mem_nhds_iff.1 (isOpen_ne.mem_nhds Hne) with ⟨V, hVB, hxV, hVy⟩ - rcases hB.exists_closure_subset (hB.mem_nhds hVB hxV) with ⟨U, hUB, hxU, hUV⟩ - set UV : ↥s := ⟨(U, V), ⟨hUB, hVB⟩, hUV⟩ - apply (ε01 UV).1.ne - calc - (0 : ℝ) = F x UV := (hf0 UV hxU).symm - _ = F y UV := by rw [hxy] - _ = ε UV := hfε UV fun h : y ∈ V => hVy h rfl - · /- Now we prove that each neighborhood `V` of `x : X` include a preimage of a neighborhood of - `F x` under `F`. Without loss of generality, `V` belongs to `B`. Choose `U ∈ B` such that - `x ∈ V` and `closure V ⊆ U`. Then the preimage of the `(ε (U, V))`-neighborhood of `F x` - is included by `V`. -/ - refine' ((nhds_basis_ball.comap _).le_basis_iff hB.nhds_hasBasis).2 _ - rintro V ⟨hVB, hxV⟩ - rcases hB.exists_closure_subset (hB.mem_nhds hVB hxV) with ⟨U, hUB, hxU, hUV⟩ - set UV : ↥s := ⟨(U, V), ⟨hUB, hVB⟩, hUV⟩ - refine' ⟨ε UV, (ε01 UV).1, fun y (hy : dist (F y) (F x) < ε UV) => _⟩ - replace hy : dist (F y UV) (F x UV) < ε UV - exact (BoundedContinuousFunction.dist_coe_le_dist _).trans_lt hy - contrapose! hy - rw [hF, hF, hfε UV hy, hf0 UV hxU, Pi.zero_apply, dist_zero_right] - exact le_abs_self _ - · /- Finally, we prove that `F` is continuous. Given `δ > 0`, consider the set `T` of `(U, V) ∈ s` - such that `ε (U, V) ≥ δ`. Since `ε` tends to zero, `T` is finite. Since each `f` is continuous, - we can choose a neighborhood such that `dist (F y (U, V)) (F x (U, V)) ≤ δ` for any - `(U, V) ∈ T`. For `(U, V) ∉ T`, the same inequality is true because both `F y (U, V)` and - `F x (U, V)` belong to the interval `[0, ε (U, V)]`. -/ - refine' (nhds_basis_closedBall.comap _).ge_iff.2 fun δ δ0 => _ - have h_fin : { UV : s | δ ≤ ε UV }.Finite := by simpa only [← not_lt] using hε (gt_mem_nhds δ0) - have : ∀ᶠ y in 𝓝 x, ∀ UV, δ ≤ ε UV → dist (F y UV) (F x UV) ≤ δ := by - refine' (eventually_all_finite h_fin).2 fun UV _ => _ - exact (f UV).continuous.tendsto x (closedBall_mem_nhds _ δ0) - refine' this.mono fun y hy => (BoundedContinuousFunction.dist_le δ0.le).2 fun UV => _ - cases' le_total δ (ε UV) with hle hle - exacts [hy _ hle, (Real.dist_le_of_mem_Icc (hf0ε _ _) (hf0ε _ _)).trans (by rwa [sub_zero])] -#align topological_space.exists_embedding_l_infty TopologicalSpace.exists_embedding_l_infty - -/-- *Urysohn's metrization theorem* (Tychonoff's version): a T₃ topological space with second -countable topology `X` is metrizable, i.e., there exists a metric space structure that generates the -same topology. -/ -theorem metrizableSpace_of_t3_second_countable : MetrizableSpace X := - let ⟨_, hf⟩ := exists_embedding_l_infty X - hf.metrizableSpace -#align topological_space.metrizable_space_of_t3_second_countable TopologicalSpace.metrizableSpace_of_t3_second_countable - -instance : MetrizableSpace ENNReal := - metrizableSpace_of_t3_second_countable ENNReal - end TopologicalSpace diff --git a/Mathlib/Topology/MetricSpace/MetrizableUniformity.lean b/Mathlib/Topology/Metrizable/Uniformity.lean similarity index 99% rename from Mathlib/Topology/MetricSpace/MetrizableUniformity.lean rename to Mathlib/Topology/Metrizable/Uniformity.lean index 23b1d4879db35..787d9b2e19fc5 100644 --- a/Mathlib/Topology/MetricSpace/MetrizableUniformity.lean +++ b/Mathlib/Topology/Metrizable/Uniformity.lean @@ -3,7 +3,8 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Topology.MetricSpace.Metrizable +import Mathlib.Topology.Metrizable.Basic +import Mathlib.Data.Nat.Lattice #align_import topology.metric_space.metrizable_uniformity from "leanprover-community/mathlib"@"195fcd60ff2bfe392543bceb0ec2adcdb472db4c" diff --git a/Mathlib/Topology/Metrizable/Urysohn.lean b/Mathlib/Topology/Metrizable/Urysohn.lean new file mode 100644 index 0000000000000..49c6c06e2e9ec --- /dev/null +++ b/Mathlib/Topology/Metrizable/Urysohn.lean @@ -0,0 +1,131 @@ +/- +Copyright (c) 2021 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Analysis.SpecificLimits.Basic +import Mathlib.Topology.UrysohnsLemma +import Mathlib.Topology.ContinuousFunction.Bounded +import Mathlib.Topology.Metrizable.Basic + +#align_import topology.metric_space.metrizable from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" +/-! +# Urysohn's Metrization Theorem + +In this file we prove Urysohn's Metrization Theorem: +a T₃ topological space with second countable topology `X` is metrizable. + +First we prove that `X` can be embedded into `l^∞`, then use this embedding to pull back the metric +space structure. + +## Implementation notes + +We use `ℕ →ᵇ ℝ`, not `lpSpace` for `l^∞` to avoid heavy imports. +-/ + +open Set Filter Metric +open scoped Topology BoundedContinuousFunction + +namespace TopologicalSpace + +section RegularSpace + +variable (X : Type*) [TopologicalSpace X] [RegularSpace X] [SecondCountableTopology X] + +/-- For a regular topological space with second countable topology, +there exists an inducing map to `l^∞ = ℕ →ᵇ ℝ`. -/ +theorem exists_inducing_l_infty : ∃ f : X → ℕ →ᵇ ℝ, Inducing f := by + -- Choose a countable basis, and consider the set `s` of pairs of set `(U, V)` such that `U ∈ B`, + -- `V ∈ B`, and `closure U ⊆ V`. + rcases exists_countable_basis X with ⟨B, hBc, -, hB⟩ + let s : Set (Set X × Set X) := { UV ∈ B ×ˢ B | closure UV.1 ⊆ UV.2 } + -- `s` is a countable set. + haveI : Encodable s := ((hBc.prod hBc).mono (inter_subset_left _ _)).toEncodable + -- We don't have the space of bounded (possibly discontinuous) functions, so we equip `s` + -- with the discrete topology and deal with `s →ᵇ ℝ` instead. + letI : TopologicalSpace s := ⊥ + haveI : DiscreteTopology s := ⟨rfl⟩ + rsuffices ⟨f, hf⟩ : ∃ f : X → s →ᵇ ℝ, Inducing f + · exact ⟨fun x => (f x).extend (Encodable.encode' s) 0, + (BoundedContinuousFunction.isometry_extend (Encodable.encode' s) + (0 : ℕ →ᵇ ℝ)).embedding.toInducing.comp hf⟩ + have hd : ∀ UV : s, Disjoint (closure UV.1.1) UV.1.2ᶜ := + fun UV => disjoint_compl_right.mono_right (compl_subset_compl.2 UV.2.2) + -- Choose a sequence of `εₙ > 0`, `n : s`, that is bounded above by `1` and tends to zero + -- along the `cofinite` filter. + obtain ⟨ε, ε01, hε⟩ : ∃ ε : s → ℝ, (∀ UV, ε UV ∈ Ioc (0 : ℝ) 1) ∧ Tendsto ε cofinite (𝓝 0) := by + rcases posSumOfEncodable zero_lt_one s with ⟨ε, ε0, c, hεc, hc1⟩ + refine' ⟨ε, fun UV => ⟨ε0 UV, _⟩, hεc.summable.tendsto_cofinite_zero⟩ + exact (le_hasSum hεc UV fun _ _ => (ε0 _).le).trans hc1 + /- For each `UV = (U, V) ∈ s` we use Urysohn's lemma to choose a function `f UV` that is equal to + zero on `U` and is equal to `ε UV` on the complement to `V`. -/ + have : ∀ UV : s, ∃ f : C(X, ℝ), + EqOn f 0 UV.1.1 ∧ EqOn f (fun _ => ε UV) UV.1.2ᶜ ∧ ∀ x, f x ∈ Icc 0 (ε UV) := by + intro UV + rcases exists_continuous_zero_one_of_closed isClosed_closure + (hB.isOpen UV.2.1.2).isClosed_compl (hd UV) with + ⟨f, hf₀, hf₁, hf01⟩ + exact ⟨ε UV • f, fun x hx => by simp [hf₀ (subset_closure hx)], fun x hx => by simp [hf₁ hx], + fun x => ⟨mul_nonneg (ε01 _).1.le (hf01 _).1, mul_le_of_le_one_right (ε01 _).1.le (hf01 _).2⟩⟩ + choose f hf0 hfε hf0ε using this + have hf01 : ∀ UV x, f UV x ∈ Icc (0 : ℝ) 1 := + fun UV x => Icc_subset_Icc_right (ε01 _).2 (hf0ε _ _) + -- The embedding is given by `F x UV = f UV x`. + set F : X → s →ᵇ ℝ := fun x => + ⟨⟨fun UV => f UV x, continuous_of_discreteTopology⟩, 1, + fun UV₁ UV₂ => Real.dist_le_of_mem_Icc_01 (hf01 _ _) (hf01 _ _)⟩ + have hF : ∀ x UV, F x UV = f UV x := fun _ _ => rfl + refine' ⟨F, inducing_iff_nhds.2 fun x => le_antisymm _ _⟩ + · /- First we prove that `F` is continuous. Given `δ > 0`, consider the set `T` of `(U, V) ∈ s` + such that `ε (U, V) ≥ δ`. Since `ε` tends to zero, `T` is finite. Since each `f` is continuous, + we can choose a neighborhood such that `dist (F y (U, V)) (F x (U, V)) ≤ δ` for any + `(U, V) ∈ T`. For `(U, V) ∉ T`, the same inequality is true because both `F y (U, V)` and + `F x (U, V)` belong to the interval `[0, ε (U, V)]`. -/ + refine' (nhds_basis_closedBall.comap _).ge_iff.2 fun δ δ0 => _ + have h_fin : { UV : s | δ ≤ ε UV }.Finite := by simpa only [← not_lt] using hε (gt_mem_nhds δ0) + have : ∀ᶠ y in 𝓝 x, ∀ UV, δ ≤ ε UV → dist (F y UV) (F x UV) ≤ δ := by + refine' (eventually_all_finite h_fin).2 fun UV _ => _ + exact (f UV).continuous.tendsto x (closedBall_mem_nhds _ δ0) + refine' this.mono fun y hy => (BoundedContinuousFunction.dist_le δ0.le).2 fun UV => _ + cases' le_total δ (ε UV) with hle hle + exacts [hy _ hle, (Real.dist_le_of_mem_Icc (hf0ε _ _) (hf0ε _ _)).trans (by rwa [sub_zero])] + · /- Finally, we prove that each neighborhood `V` of `x : X` + includes a preimage of a neighborhood of `F x` under `F`. + Without loss of generality, `V` belongs to `B`. + Choose `U ∈ B` such that `x ∈ V` and `closure V ⊆ U`. + Then the preimage of the `(ε (U, V))`-neighborhood of `F x` is included by `V`. -/ + refine' ((nhds_basis_ball.comap _).le_basis_iff hB.nhds_hasBasis).2 _ + rintro V ⟨hVB, hxV⟩ + rcases hB.exists_closure_subset (hB.mem_nhds hVB hxV) with ⟨U, hUB, hxU, hUV⟩ + set UV : ↥s := ⟨(U, V), ⟨hUB, hVB⟩, hUV⟩ + refine' ⟨ε UV, (ε01 UV).1, fun y (hy : dist (F y) (F x) < ε UV) => _⟩ + replace hy : dist (F y UV) (F x UV) < ε UV + exact (BoundedContinuousFunction.dist_coe_le_dist _).trans_lt hy + contrapose! hy + rw [hF, hF, hfε UV hy, hf0 UV hxU, Pi.zero_apply, dist_zero_right] + exact le_abs_self _ +#align topological_space.exists_embedding_l_infty TopologicalSpace.exists_inducing_l_infty + +/-- *Urysohn's metrization theorem* (Tychonoff's version): +a regular topological space with second countable topology `X` is metrizable, +i.e., there exists a pseudometric space structure that generates the same topology. -/ +lemma PseudoMetrizableSpace.of_regularSpace_secondCountableTopology : PseudoMetrizableSpace X := + let ⟨_, hf⟩ := exists_inducing_l_infty X + hf.pseudoMetrizableSpace + +end RegularSpace + +variable (X : Type*) [TopologicalSpace X] [T3Space X] [SecondCountableTopology X] + +/-- A T₃ topological space with second countable topology can be embedded into `l^∞ = ℕ →ᵇ ℝ`. -/ +theorem exists_embedding_l_infty : ∃ f : X → ℕ →ᵇ ℝ, Embedding f := + let ⟨f, hf⟩ := exists_inducing_l_infty X; ⟨f, hf.embedding⟩ + +/-- *Urysohn's metrization theorem* (Tychonoff's version): a T₃ topological space with second +countable topology `X` is metrizable, i.e., there exists a metric space structure that generates the +same topology. -/ +theorem metrizableSpace_of_t3_second_countable : MetrizableSpace X := + let ⟨_, hf⟩ := exists_embedding_l_infty X + hf.metrizableSpace +#align topological_space.metrizable_space_of_t3_second_countable TopologicalSpace.metrizableSpace_of_t3_second_countable + diff --git a/Mathlib/Topology/NhdsSet.lean b/Mathlib/Topology/NhdsSet.lean index b453054e4a5fc..82ae64bbfa696 100644 --- a/Mathlib/Topology/NhdsSet.lean +++ b/Mathlib/Topology/NhdsSet.lean @@ -71,6 +71,17 @@ theorem mem_nhdsSet_iff_exists : s ∈ 𝓝ˢ t ↔ ∃ U : Set α, IsOpen U ∧ rw [← subset_interior_iff_mem_nhdsSet, subset_interior_iff] #align mem_nhds_set_iff_exists mem_nhdsSet_iff_exists +/-- A proposition is true on a set neighborhood of `s` iff it is true on a larger open set -/ +theorem eventually_nhdsSet_iff_exists {p : α → Prop} : + (∀ᶠ x in 𝓝ˢ s, p x) ↔ ∃ t, IsOpen t ∧ s ⊆ t ∧ ∀ x, x ∈ t → p x := + mem_nhdsSet_iff_exists + +/-- A proposition is true on a set neighborhood of `s` +iff it is eventually true near each point in the set. -/ +theorem eventually_nhdsSet_iff_forall {p : α → Prop} : + (∀ᶠ x in 𝓝ˢ s, p x) ↔ ∀ x, x ∈ s → ∀ᶠ y in 𝓝 x, p y := + mem_nhdsSet_iff_forall + theorem hasBasis_nhdsSet (s : Set α) : (𝓝ˢ s).HasBasis (fun U => IsOpen U ∧ s ⊆ U) fun U => U := ⟨fun t => by simp [mem_nhdsSet_iff_exists, and_assoc]⟩ #align has_basis_nhds_set hasBasis_nhdsSet @@ -79,6 +90,9 @@ theorem IsOpen.mem_nhdsSet (hU : IsOpen s) : s ∈ 𝓝ˢ t ↔ t ⊆ s := by rw [← subset_interior_iff_mem_nhdsSet, hU.interior_eq] #align is_open.mem_nhds_set IsOpen.mem_nhdsSet +/-- An open set belongs to its own set neighborhoods filter. -/ +theorem IsOpen.mem_nhdsSet_self (ho : IsOpen s) : s ∈ 𝓝ˢ s := ho.mem_nhdsSet.mpr Subset.rfl + theorem principal_le_nhdsSet : 𝓟 s ≤ 𝓝ˢ s := fun _s hs => (subset_interior_iff_mem_nhdsSet.mpr hs).trans interior_subset #align principal_le_nhds_set principal_le_nhdsSet @@ -157,3 +171,9 @@ theorem Continuous.tendsto_nhdsSet {f : α → β} {t : Set β} (hf : Continuous ((hasBasis_nhdsSet s).tendsto_iff (hasBasis_nhdsSet t)).mpr fun U hU => ⟨f ⁻¹' U, ⟨hU.1.preimage hf, hst.mono Subset.rfl hU.2⟩, fun _ => id⟩ #align continuous.tendsto_nhds_set Continuous.tendsto_nhdsSet + +lemma Continuous.tendsto_nhdsSet_nhds {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] + {s : Set X} {y : Y} {f : X → Y} (h : Continuous f) (h' : EqOn f (fun _ ↦ y) s) : + Tendsto f (𝓝ˢ s) (𝓝 y) := by + rw [←nhdsSet_singleton] + exact h.tendsto_nhdsSet h' diff --git a/Mathlib/Topology/NoetherianSpace.lean b/Mathlib/Topology/NoetherianSpace.lean index bd0e95026b08e..276d36fbc66d3 100644 --- a/Mathlib/Topology/NoetherianSpace.lean +++ b/Mathlib/Topology/NoetherianSpace.lean @@ -90,7 +90,7 @@ theorem noetherianSpace_TFAE : ∀ s : Set α, IsCompact s, ∀ s : Opens α, IsCompact (s : Set α)] := by tfae_have 1 ↔ 2 - · refine' (noetherianSpace_iff α).trans (Opens.compl_bijective.2.wellFounded_iff _) + · refine' (noetherianSpace_iff α).trans (Opens.compl_bijective.2.wellFounded_iff _) exact (@OrderIso.compl (Set α)).lt_iff_lt.symm tfae_have 1 ↔ 4 · exact noetherianSpace_iff_opens α diff --git a/Mathlib/Topology/Order/Basic.lean b/Mathlib/Topology/Order/Basic.lean index 981359bbd932a..c7440bdebff1c 100644 --- a/Mathlib/Topology/Order/Basic.lean +++ b/Mathlib/Topology/Order/Basic.lean @@ -276,25 +276,6 @@ theorem IsClosed.hypograph [TopologicalSpace β] {f : β → α} {s : Set β} (h (hs.preimage continuous_fst).isClosed_le continuousOn_snd (hf.comp continuousOn_fst Subset.rfl) #align is_closed.hypograph IsClosed.hypograph --- Porting note: todo: move these lemmas to `Topology.Algebra.Order.LeftRight` -theorem nhdsWithin_Ici_neBot {a b : α} (H₂ : a ≤ b) : NeBot (𝓝[Ici a] b) := - nhdsWithin_neBot_of_mem H₂ -#align nhds_within_Ici_ne_bot nhdsWithin_Ici_neBot - -@[instance] -theorem nhdsWithin_Ici_self_neBot (a : α) : NeBot (𝓝[≥] a) := - nhdsWithin_Ici_neBot (le_refl a) -#align nhds_within_Ici_self_ne_bot nhdsWithin_Ici_self_neBot - -theorem nhdsWithin_Iic_neBot {a b : α} (H : a ≤ b) : NeBot (𝓝[Iic b] a) := - nhdsWithin_neBot_of_mem H -#align nhds_within_Iic_ne_bot nhdsWithin_Iic_neBot - -@[instance] -theorem nhdsWithin_Iic_self_neBot (a : α) : NeBot (𝓝[≤] a) := - nhdsWithin_Iic_neBot (le_refl a) -#align nhds_within_Iic_self_ne_bot nhdsWithin_Iic_self_neBot - end Preorder section PartialOrder @@ -2477,12 +2458,6 @@ instance nhdsWithin_Ioi_self_neBot [NoMaxOrder α] (a : α) : NeBot (𝓝[>] a) nhdsWithin_Ioi_neBot (le_refl a) #align nhds_within_Ioi_self_ne_bot nhdsWithin_Ioi_self_neBot -theorem Filter.Eventually.exists_gt [NoMaxOrder α] {a : α} {p : α → Prop} (h : ∀ᶠ x in 𝓝 a, p x) : - ∃ b > a, p b := by - simpa only [exists_prop, gt_iff_lt, and_comm] using - ((h.filter_mono (@nhdsWithin_le_nhds _ _ a (Ioi a))).and self_mem_nhdsWithin).exists -#align filter.eventually.exists_gt Filter.Eventually.exists_gt - theorem nhdsWithin_Iio_neBot' {b c : α} (H₁ : (Iio c).Nonempty) (H₂ : b ≤ c) : NeBot (𝓝[Iio c] b) := mem_closure_iff_nhdsWithin_neBot.1 <| by rwa [closure_Iio' H₁] @@ -2500,11 +2475,6 @@ instance nhdsWithin_Iio_self_neBot [NoMinOrder α] (a : α) : NeBot (𝓝[<] a) nhdsWithin_Iio_neBot (le_refl a) #align nhds_within_Iio_self_ne_bot nhdsWithin_Iio_self_neBot -theorem Filter.Eventually.exists_lt [NoMinOrder α] {a : α} {p : α → Prop} (h : ∀ᶠ x in 𝓝 a, p x) : - ∃ b < a, p b := - Filter.Eventually.exists_gt (α := αᵒᵈ) h -#align filter.eventually.exists_lt Filter.Eventually.exists_lt - theorem right_nhdsWithin_Ico_neBot {a b : α} (H : a < b) : NeBot (𝓝[Ico a b] b) := (isLUB_Ico H).nhdsWithin_neBot (nonempty_Ico.2 H) #align right_nhds_within_Ico_ne_bot right_nhdsWithin_Ico_neBot diff --git a/Mathlib/Topology/SeparatedMap.lean b/Mathlib/Topology/SeparatedMap.lean new file mode 100644 index 0000000000000..0b3befa648bab --- /dev/null +++ b/Mathlib/Topology/SeparatedMap.lean @@ -0,0 +1,218 @@ +/- +Copyright (c) 2023 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.Topology.Connected.Basic +import Mathlib.Topology.Separation +/-! +# Separated maps and locally injective maps out of a topological space. + +This module introduces a pair of dual notions `IsSeparatedMap` and `IsLocallyInjective`. + +A function from a topological space `X` to a type `Y` is a separated map if any two distinct +points in `X` with the same image in `Y` can be separated by open neighborhoods. +A constant function is a separated map if and only if `X` is a `T2Space`. + +A function from a topological space `X` is locally injective if every point of `X` +has a neighborhood on which `f` is injective. +A constant function is locally injective if and only if `X` is discrete. + +Given `f : X → Y` we can form the pullback $X \times_Y X$; the diagonal map +$\Delta: X \to X \times_Y X$ is always an embedding. It is a closed embedding +iff `f` is a separated map, iff the equal locus of any two continuous maps +coequalized by `f` is closed. It is an open embedding iff `f` is locally injective, +iff any such equal locus is open. Therefore, if `f` is a locally injective separated map, +the equal locus of two continuous maps coequalized by `f` is clopen, so if the two maps +agree on a point, then they agree on the whole connected component. + +The analogue of separated maps and locally injective maps in algebraic geometry are +separated morphisms and unramified morphisms, respectively. + +## Reference + +https://stacks.math.columbia.edu/tag/0CY0 +-/ + +open scoped Topology + +variable {X Y A} [TopologicalSpace X] [TopologicalSpace A] + +theorem embedding_toPullbackDiag (f : X → Y) : Embedding (toPullbackDiag f) := + Embedding.mk' _ (injective_toPullbackDiag f) fun x ↦ by + rw [toPullbackDiag, nhds_induced, Filter.comap_comap, nhds_prod_eq, Filter.comap_prod] + erw [Filter.comap_id, inf_idem] + +lemma Continuous.mapPullback {X₁ X₂ Y₁ Y₂ Z₁ Z₂} + [TopologicalSpace X₁] [TopologicalSpace X₂] [TopologicalSpace Z₁] [TopologicalSpace Z₂] + {f₁ : X₁ → Y₁} {g₁ : Z₁ → Y₁} {f₂ : X₂ → Y₂} {g₂ : Z₂ → Y₂} + {mapX : X₁ → X₂} (contX : Continuous mapX) {mapY : Y₁ → Y₂} + {mapZ : Z₁ → Z₂} (contZ : Continuous mapZ) + {commX : f₂ ∘ mapX = mapY ∘ f₁} {commZ : g₂ ∘ mapZ = mapY ∘ g₁} : + Continuous (Function.mapPullback mapX mapY mapZ commX commZ) := by + refine continuous_induced_rng.mpr (continuous_prod_mk.mpr ⟨?_, ?_⟩) <;> + apply_rules [continuous_fst, continuous_snd, continuous_subtype_val, Continuous.comp] + +/-- A function from a topological space `X` to a type `Y` is a separated map if any two distinct + points in `X` with the same image in `Y` can be separated by open neighborhoods. -/ +def IsSeparatedMap (f : X → Y) : Prop := ∀ x₁ x₂, f x₁ = f x₂ → + x₁ ≠ x₂ → ∃ s₁ s₂, IsOpen s₁ ∧ IsOpen s₂ ∧ x₁ ∈ s₁ ∧ x₂ ∈ s₂ ∧ Disjoint s₁ s₂ + +lemma t2space_iff_isSeparatedMap (y : Y) : T2Space X ↔ IsSeparatedMap fun _ : X ↦ y := + ⟨fun ⟨t2⟩ x₁ x₂ _ hne ↦ t2 x₁ x₂ hne, fun sep ↦ ⟨fun x₁ x₂ hne ↦ sep x₁ x₂ rfl hne⟩⟩ + +lemma T2Space.isSeparatedMap [T2Space X] (f : X → Y) : IsSeparatedMap f := fun _ _ _ ↦ t2_separation + +lemma Function.Injective.isSeparatedMap {f : X → Y} (inj : f.Injective) : IsSeparatedMap f := + fun _ _ he hne ↦ (hne (inj he)).elim + +lemma isSeparatedMap_iff_disjoint_nhds {f : X → Y} : IsSeparatedMap f ↔ + ∀ x₁ x₂, f x₁ = f x₂ → x₁ ≠ x₂ → Disjoint (𝓝 x₁) (𝓝 x₂) := + forall₃_congr fun x x' _ ↦ by simp only [(nhds_basis_opens x).disjoint_iff (nhds_basis_opens x'), + exists_prop, ← exists_and_left, and_assoc, and_comm, and_left_comm] + +lemma isSeparatedMap_iff_nhds {f : X → Y} : IsSeparatedMap f ↔ + ∀ x₁ x₂, f x₁ = f x₂ → x₁ ≠ x₂ → ∃ s₁ ∈ 𝓝 x₁, ∃ s₂ ∈ 𝓝 x₂, Disjoint s₁ s₂ := by + simp_rw [isSeparatedMap_iff_disjoint_nhds, Filter.disjoint_iff] + +open Set Filter in +theorem isSeparatedMap_iff_isClosed_diagonal {f : X → Y} : + IsSeparatedMap f ↔ IsClosed f.pullbackDiagonal := by + simp_rw [isSeparatedMap_iff_nhds, ← isOpen_compl_iff, isOpen_iff_mem_nhds, + Subtype.forall, Prod.forall, nhds_induced, nhds_prod_eq] + refine forall₄_congr fun x₁ x₂ _ _ ↦ ⟨fun h ↦ ?_, fun ⟨t, ht, t_sub⟩ ↦ ?_⟩ + · simp_rw [← Filter.disjoint_iff, ← compl_diagonal_mem_prod] at h + exact ⟨_, h, subset_rfl⟩ + · obtain ⟨s₁, h₁, s₂, h₂, s_sub⟩ := mem_prod_iff.mp ht + exact ⟨s₁, h₁, s₂, h₂, disjoint_left.2 fun x h₁ h₂ ↦ @t_sub ⟨(x, x), rfl⟩ (s_sub ⟨h₁, h₂⟩) rfl⟩ + +theorem isSeparatedMap_iff_closedEmbedding {f : X → Y} : + IsSeparatedMap f ↔ ClosedEmbedding (toPullbackDiag f) := by + rw [isSeparatedMap_iff_isClosed_diagonal, ← range_toPullbackDiag] + exact ⟨fun h ↦ ⟨embedding_toPullbackDiag f, h⟩, fun h ↦ h.2⟩ + +theorem isSeparatedMap_iff_isClosedMap {f : X → Y} : + IsSeparatedMap f ↔ IsClosedMap (toPullbackDiag f) := + isSeparatedMap_iff_closedEmbedding.trans + ⟨ClosedEmbedding.isClosedMap, closedEmbedding_of_continuous_injective_closed + (embedding_toPullbackDiag f).continuous (injective_toPullbackDiag f)⟩ + +open Function.Pullback in +theorem IsSeparatedMap.pullback {f : X → Y} (sep : IsSeparatedMap f) (g : A → Y) : + IsSeparatedMap (@snd X Y A f g) := by + rw [isSeparatedMap_iff_isClosed_diagonal] at sep ⊢ + rw [← preimage_map_fst_pullbackDiagonal] + refine sep.preimage (Continuous.mapPullback ?_ ?_) <;> + apply_rules [continuous_fst, continuous_subtype_val, Continuous.comp] + +theorem IsSeparatedMap.comp_left {f : X → Y} (sep : IsSeparatedMap f) {g : Y → A} + (inj : g.Injective) : IsSeparatedMap (g ∘ f) := fun x₁ x₂ he ↦ sep x₁ x₂ (inj he) + +theorem IsSeparatedMap.comp_right {f : X → Y} (sep : IsSeparatedMap f) {g : A → X} + (cont : Continuous g) (inj : g.Injective) : IsSeparatedMap (f ∘ g) := by + rw [isSeparatedMap_iff_isClosed_diagonal] at sep ⊢ + rw [← inj.preimage_pullbackDiagonal] + exact sep.preimage (cont.mapPullback cont) + +/-- A function from a topological space `X` is locally injective if every point of `X` + has a neighborhood on which `f` is injective. -/ +def IsLocallyInjective (f : X → Y) : Prop := ∀ x : X, ∃ U, IsOpen U ∧ x ∈ U ∧ U.InjOn f + +lemma Function.Injective.IsLocallyInjective {f : X → Y} (inj : f.Injective) : + IsLocallyInjective f := fun _ ↦ ⟨_, isOpen_univ, trivial, fun _ _ _ _ ↦ @inj _ _⟩ + +lemma isLocallyInjective_iff_nhds {f : X → Y} : + IsLocallyInjective f ↔ ∀ x : X, ∃ U ∈ 𝓝 x, U.InjOn f := by + constructor <;> intro h x + · obtain ⟨U, ho, hm, hi⟩ := h x; exact ⟨U, ho.mem_nhds hm, hi⟩ + · obtain ⟨U, hn, hi⟩ := h x + exact ⟨interior U, isOpen_interior, mem_interior_iff_mem_nhds.mpr hn, hi.mono interior_subset⟩ + +theorem isLocallyInjective_iff_isOpen_diagonal {f : X → Y} : + IsLocallyInjective f ↔ IsOpen f.pullbackDiagonal := by + simp_rw [isLocallyInjective_iff_nhds, isOpen_iff_mem_nhds, + Subtype.forall, Prod.forall, nhds_induced, nhds_prod_eq, Filter.mem_comap] + refine ⟨?_, fun h x ↦ ?_⟩ + · rintro h x x' hx (rfl : x = x') + obtain ⟨U, hn, hi⟩ := h x + exact ⟨_, Filter.prod_mem_prod hn hn, fun {p} hp ↦ hi hp.1 hp.2 p.2⟩ + · obtain ⟨t, ht, t_sub⟩ := h x x rfl rfl + obtain ⟨t₁, h₁, t₂, h₂, prod_sub⟩ := Filter.mem_prod_iff.mp ht + exact ⟨t₁ ∩ t₂, Filter.inter_mem h₁ h₂, + fun x₁ h₁ x₂ h₂ he ↦ @t_sub ⟨(x₁, x₂), he⟩ (prod_sub ⟨h₁.1, h₂.2⟩)⟩ + +theorem IsLocallyInjective_iff_openEmbedding {f : X → Y} : + IsLocallyInjective f ↔ OpenEmbedding (toPullbackDiag f) := by + rw [isLocallyInjective_iff_isOpen_diagonal, ← range_toPullbackDiag] + exact ⟨fun h ↦ ⟨embedding_toPullbackDiag f, h⟩, fun h ↦ h.2⟩ + +theorem isLocallyInjective_iff_isOpenMap {f : X → Y} : + IsLocallyInjective f ↔ IsOpenMap (toPullbackDiag f) := + IsLocallyInjective_iff_openEmbedding.trans + ⟨OpenEmbedding.isOpenMap, openEmbedding_of_continuous_injective_open + (embedding_toPullbackDiag f).continuous (injective_toPullbackDiag f)⟩ + +theorem discreteTopology_iff_locallyInjective (y : Y) : + DiscreteTopology X ↔ IsLocallyInjective fun _ : X ↦ y := by + rw [discreteTopology_iff_singleton_mem_nhds, isLocallyInjective_iff_nhds] + refine forall_congr' fun x ↦ ⟨fun h ↦ ⟨{x}, h, Set.injOn_singleton _ _⟩, fun ⟨U, hU, inj⟩ ↦ ?_⟩ + convert hU; ext x'; refine ⟨?_, fun h ↦ inj h (mem_of_mem_nhds hU) rfl⟩ + rintro rfl; exact mem_of_mem_nhds hU + +theorem IsLocallyInjective.comp_left {f : X → Y} (hf : IsLocallyInjective f) {g : Y → A} + (hg : g.Injective) : IsLocallyInjective (g ∘ f) := + fun x ↦ let ⟨U, hU, hx, inj⟩ := hf x; ⟨U, hU, hx, hg.comp_injOn inj⟩ + +theorem IsLocallyInjective.comp_right {f : X → Y} (hf : IsLocallyInjective f) {g : A → X} + (cont : Continuous g) (hg : g.Injective) : IsLocallyInjective (f ∘ g) := by + rw [isLocallyInjective_iff_isOpen_diagonal] at hf ⊢ + rw [← hg.preimage_pullbackDiagonal] + apply hf.preimage (cont.mapPullback cont) + +section eqLocus + +variable {f : X → Y} (sep : IsSeparatedMap f) (inj : IsLocallyInjective f) + {g₁ g₂ : A → X} (h₁ : Continuous g₁) (h₂ : Continuous g₂) + +theorem IsSeparatedMap.isClosed_eqLocus (he : f ∘ g₁ = f ∘ g₂) : IsClosed {a | g₁ a = g₂ a} := + let g : A → f.Pullback f := fun a ↦ ⟨⟨g₁ a, g₂ a⟩, congr_fun he a⟩ + (isSeparatedMap_iff_isClosed_diagonal.mp sep).preimage (by continuity : Continuous g) + +theorem IsLocallyInjective.isOpen_eqLocus (he : f ∘ g₁ = f ∘ g₂) : IsOpen {a | g₁ a = g₂ a} := + let g : A → f.Pullback f := fun a ↦ ⟨⟨g₁ a, g₂ a⟩, congr_fun he a⟩ + (isLocallyInjective_iff_isOpen_diagonal.mp inj).preimage (by continuity : Continuous g) + +end eqLocus + +variable {E A : Type*} [TopologicalSpace E] [TopologicalSpace A] {p : E → X} + +namespace IsSeparatedMap + +variable (sep : IsSeparatedMap p) (inj : IsLocallyInjective p) {s : Set A} (hs : IsPreconnected s) + {g g₁ g₂ : A → E} + +/-- If `p` is a locally injective separated map, and `A` is a connected space, + then two lifts `g₁, g₂ : A → E` of a map `f : A → X` are equal if they agree at one point. -/ +theorem eq_of_comp_eq [PreconnectedSpace A] (h₁ : Continuous g₁) (h₂ : Continuous g₂) + (he : p ∘ g₁ = p ∘ g₂) (a : A) (ha : g₁ a = g₂ a) : g₁ = g₂ := funext fun a' ↦ by + apply (IsClopen.eq_univ ⟨inj.isOpen_eqLocus h₁ h₂ he, sep.isClosed_eqLocus h₁ h₂ he⟩ ⟨a, ha⟩).symm + ▸ Set.mem_univ a' + +theorem eqOn_of_comp_eqOn (h₁ : ContinuousOn g₁ s) (h₂ : ContinuousOn g₂ s) + (he : s.EqOn (p ∘ g₁) (p ∘ g₂)) {a : A} (has : a ∈ s) (ha : g₁ a = g₂ a) : s.EqOn g₁ g₂ := by + rw [← Set.restrict_eq_restrict_iff] at he ⊢ + rw [continuousOn_iff_continuous_restrict] at h₁ h₂ + rw [isPreconnected_iff_preconnectedSpace] at hs + exact sep.eq_of_comp_eq inj h₁ h₂ he ⟨a, has⟩ ha + +theorem const_of_comp [PreconnectedSpace A] (cont : Continuous g) + (he : ∀ a a', p (g a) = p (g a')) (a a') : g a = g a' := + congr_fun (sep.eq_of_comp_eq inj cont continuous_const (funext fun a ↦ he a a') a' rfl) a + +theorem constOn_of_comp (cont : ContinuousOn g s) + (he : ∀ a ∈ s, ∀ a' ∈ s, p (g a) = p (g a')) + {a a'} (ha : a ∈ s) (ha' : a' ∈ s) : g a = g a' := + sep.eqOn_of_comp_eqOn inj hs cont continuous_const.continuousOn + (fun a ha ↦ he a ha a' ha') ha' rfl ha + +end IsSeparatedMap diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 0cb4cf38cb1e4..d5116154d14a0 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -29,12 +29,16 @@ This file defines the predicate `SeparatedNhds`, and common separation axioms there is two disjoint open sets, one containing `x`, and the other `y`. * `T25Space`: A T₂.₅/Urysohn space is a space where, for every two points `x ≠ y`, there is two open sets, one containing `x`, and the other `y`, whose closures are disjoint. -* `T3Space`: A T₃ space, is one where given any closed `C` and `x ∉ C`, - there is disjoint open sets containing `x` and `C` respectively. In `mathlib`, T₃ implies T₂.₅. +* `RegularSpace`: A regular space is one where, given any closed `C` and `x ∉ C`, + there are disjoint open sets containing `x` and `C` respectively. Such a space is not necessarily + Hausdorff. +* `T3Space`: A T₃ space is a T0 regular space. In `mathlib`, T₃ implies T₂.₅. * `NormalSpace`: A normal space, is one where given two disjoint closed sets, we can find two open sets that separate them. * `T4Space`: A T₄ space is a normal T₁ space. T₄ implies T₃. -* `T5Space`: A T₅ space, also known as a *completely normal Hausdorff space* +* `T5Space`: A T₅ space, also known as a *completely normal Hausdorff space*, is a space in which, + given two sets `s` and `t` such that the closure of `s` is disjoint from `t`, and conversely, + then `s` and `t` have disjoint neighborhoods. ## Main results @@ -53,8 +57,8 @@ This file defines the predicate `SeparatedNhds`, and common separation axioms ### T₂ spaces * `t2_iff_nhds`: A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter. -* `t2_iff_isClosed_diagonal`: A space is T₂ iff the `diagonal` of `α` (that is, the set of all - points of the form `(a, a) : α × α`) is closed under the product topology. +* `t2_iff_isClosed_diagonal`: A space is T₂ iff the `diagonal` of `X` (that is, the set of all + points of the form `(a, a) : X × X`) is closed under the product topology. * `finset_disjoint_finset_opens_of_t2`: Any two disjoint finsets are `SeparatedNhds`. * Most topological constructions preserve Hausdorffness; these results are part of the typeclass inference system (e.g. `Embedding.t2Space`) @@ -63,7 +67,7 @@ This file defines the predicate `SeparatedNhds`, and common separation axioms * `WeaklyLocallyCompactSpace.locallyCompactSpace`: If a topological space is both weakly locally compact (i.e., each point has a compact neighbourhood) and is T₂, then it is locally compact. -* `totallySeparatedSpace_of_t1_of_basis_clopen`: If `α` has a clopen basis, then +* `totallySeparatedSpace_of_t1_of_basis_clopen`: If `X` has a clopen basis, then it is a `TotallySeparatedSpace`. * `loc_compact_t2_tot_disc_iff_tot_sep`: A locally compact T₂ space is totally disconnected iff it is totally separated. @@ -75,7 +79,7 @@ If the space is also compact: is the intersection of all its clopen neighbourhoods. * `compact_t2_tot_disc_iff_tot_sep`: Being a `TotallyDisconnectedSpace` is equivalent to being a `TotallySeparatedSpace`. -* `ConnectedComponents.t2`: `ConnectedComponents α` is T₂ for `α` T₂ and compact. +* `ConnectedComponents.t2`: `ConnectedComponents X` is T₂ for `X` T₂ and compact. ### T₃ spaces @@ -93,7 +97,7 @@ open scoped Classical universe u v -variable {α : Type u} {β : Type v} [TopologicalSpace α] +variable {X : Type*} {Y : Type*} [TopologicalSpace X] section Separation @@ -101,11 +105,11 @@ section Separation `SeparatedNhds` is a predicate on pairs of sub`Set`s of a topological space. It holds if the two sub`Set`s are contained in disjoint open sets. -/ -def SeparatedNhds : Set α → Set α → Prop := fun s t : Set α => - ∃ U V : Set α, IsOpen U ∧ IsOpen V ∧ s ⊆ U ∧ t ⊆ V ∧ Disjoint U V +def SeparatedNhds : Set X → Set X → Prop := fun s t : Set X => + ∃ U V : Set X, IsOpen U ∧ IsOpen V ∧ s ⊆ U ∧ t ⊆ V ∧ Disjoint U V #align separated_nhds SeparatedNhds -theorem separatedNhds_iff_disjoint {s t : Set α} : SeparatedNhds s t ↔ Disjoint (𝓝ˢ s) (𝓝ˢ t) := by +theorem separatedNhds_iff_disjoint {s t : Set X} : SeparatedNhds s t ↔ Disjoint (𝓝ˢ s) (𝓝ˢ t) := by simp only [(hasBasis_nhdsSet s).disjoint_iff (hasBasis_nhdsSet t), SeparatedNhds, exists_prop, ← exists_and_left, and_assoc, and_comm, and_left_comm] #align separated_nhds_iff_disjoint separatedNhds_iff_disjoint @@ -114,18 +118,18 @@ alias ⟨SeparatedNhds.disjoint_nhdsSet, _⟩ := separatedNhds_iff_disjoint namespace SeparatedNhds -variable {s s₁ s₂ t t₁ t₂ u : Set α} +variable {s s₁ s₂ t t₁ t₂ u : Set X} @[symm] theorem symm : SeparatedNhds s t → SeparatedNhds t s := fun ⟨U, V, oU, oV, aU, bV, UV⟩ => ⟨V, U, oV, oU, bV, aU, Disjoint.symm UV⟩ #align separated_nhds.symm SeparatedNhds.symm -theorem comm (s t : Set α) : SeparatedNhds s t ↔ SeparatedNhds t s := +theorem comm (s t : Set X) : SeparatedNhds s t ↔ SeparatedNhds t s := ⟨symm, symm⟩ #align separated_nhds.comm SeparatedNhds.comm -theorem preimage [TopologicalSpace β] {f : α → β} {s t : Set β} (h : SeparatedNhds s t) +theorem preimage [TopologicalSpace Y] {f : X → Y} {s t : Set Y} (h : SeparatedNhds s t) (hf : Continuous f) : SeparatedNhds (f ⁻¹' s) (f ⁻¹' t) := let ⟨U, V, oU, oV, sU, tV, UV⟩ := h ⟨f ⁻¹' U, f ⁻¹' V, oU.preimage hf, oV.preimage hf, preimage_mono sU, preimage_mono tV, @@ -145,11 +149,11 @@ theorem disjoint_closure_right (h : SeparatedNhds s t) : Disjoint s (closure t) h.symm.disjoint_closure_left.symm #align separated_nhds.disjoint_closure_right SeparatedNhds.disjoint_closure_right -theorem empty_right (s : Set α) : SeparatedNhds s ∅ := +theorem empty_right (s : Set X) : SeparatedNhds s ∅ := ⟨_, _, isOpen_univ, isOpen_empty, fun a _ => mem_univ a, Subset.rfl, disjoint_empty _⟩ #align separated_nhds.empty_right SeparatedNhds.empty_right -theorem empty_left (s : Set α) : SeparatedNhds ∅ s := +theorem empty_left (s : Set X) : SeparatedNhds ∅ s := (empty_right _).symm #align separated_nhds.empty_left SeparatedNhds.empty_left @@ -171,99 +175,99 @@ end SeparatedNhds /-- A T₀ space, also known as a Kolmogorov space, is a topological space such that for every pair `x ≠ y`, there is an open set containing one but not the other. We formulate the definition in terms of the `Inseparable` relation. -/ -class T0Space (α : Type u) [TopologicalSpace α] : Prop where +class T0Space (X : Type u) [TopologicalSpace X] : Prop where /-- Two inseparable points in a T₀ space are equal. -/ - t0 : ∀ ⦃x y : α⦄, Inseparable x y → x = y + t0 : ∀ ⦃x y : X⦄, Inseparable x y → x = y #align t0_space T0Space -theorem t0Space_iff_inseparable (α : Type u) [TopologicalSpace α] : - T0Space α ↔ ∀ x y : α, Inseparable x y → x = y := +theorem t0Space_iff_inseparable (X : Type u) [TopologicalSpace X] : + T0Space X ↔ ∀ x y : X, Inseparable x y → x = y := ⟨fun ⟨h⟩ => h, fun h => ⟨h⟩⟩ #align t0_space_iff_inseparable t0Space_iff_inseparable -theorem t0Space_iff_not_inseparable (α : Type u) [TopologicalSpace α] : - T0Space α ↔ ∀ x y : α, x ≠ y → ¬Inseparable x y := by +theorem t0Space_iff_not_inseparable (X : Type u) [TopologicalSpace X] : + T0Space X ↔ ∀ x y : X, x ≠ y → ¬Inseparable x y := by simp only [t0Space_iff_inseparable, Ne.def, not_imp_not] #align t0_space_iff_not_inseparable t0Space_iff_not_inseparable -theorem Inseparable.eq [T0Space α] {x y : α} (h : Inseparable x y) : x = y := +theorem Inseparable.eq [T0Space X] {x y : X} (h : Inseparable x y) : x = y := T0Space.t0 h #align inseparable.eq Inseparable.eq /-- A topology `Inducing` map from a T₀ space is injective. -/ -protected theorem Inducing.injective [TopologicalSpace β] [T0Space α] {f : α → β} +protected theorem Inducing.injective [TopologicalSpace Y] [T0Space X] {f : X → Y} (hf : Inducing f) : Injective f := fun _ _ h => (hf.inseparable_iff.1 <| .of_eq h).eq #align inducing.injective Inducing.injective /-- A topology `Inducing` map from a T₀ space is a topological embedding. -/ -protected theorem Inducing.embedding [TopologicalSpace β] [T0Space α] {f : α → β} +protected theorem Inducing.embedding [TopologicalSpace Y] [T0Space X] {f : X → Y} (hf : Inducing f) : Embedding f := ⟨hf, hf.injective⟩ #align inducing.embedding Inducing.embedding -lemma embedding_iff_inducing [TopologicalSpace β] [T0Space α] {f : α → β} : +lemma embedding_iff_inducing [TopologicalSpace Y] [T0Space X] {f : X → Y} : Embedding f ↔ Inducing f := ⟨Embedding.toInducing, Inducing.embedding⟩ #align embedding_iff_inducing embedding_iff_inducing -theorem t0Space_iff_nhds_injective (α : Type u) [TopologicalSpace α] : - T0Space α ↔ Injective (𝓝 : α → Filter α) := - t0Space_iff_inseparable α +theorem t0Space_iff_nhds_injective (X : Type u) [TopologicalSpace X] : + T0Space X ↔ Injective (𝓝 : X → Filter X) := + t0Space_iff_inseparable X #align t0_space_iff_nhds_injective t0Space_iff_nhds_injective -theorem nhds_injective [T0Space α] : Injective (𝓝 : α → Filter α) := - (t0Space_iff_nhds_injective α).1 ‹_› +theorem nhds_injective [T0Space X] : Injective (𝓝 : X → Filter X) := + (t0Space_iff_nhds_injective X).1 ‹_› #align nhds_injective nhds_injective -theorem inseparable_iff_eq [T0Space α] {x y : α} : Inseparable x y ↔ x = y := +theorem inseparable_iff_eq [T0Space X] {x y : X} : Inseparable x y ↔ x = y := nhds_injective.eq_iff #align inseparable_iff_eq inseparable_iff_eq @[simp] -theorem nhds_eq_nhds_iff [T0Space α] {a b : α} : 𝓝 a = 𝓝 b ↔ a = b := +theorem nhds_eq_nhds_iff [T0Space X] {a b : X} : 𝓝 a = 𝓝 b ↔ a = b := nhds_injective.eq_iff #align nhds_eq_nhds_iff nhds_eq_nhds_iff @[simp] -theorem inseparable_eq_eq [T0Space α] : Inseparable = @Eq α := +theorem inseparable_eq_eq [T0Space X] : Inseparable = @Eq X := funext₂ fun _ _ => propext inseparable_iff_eq #align inseparable_eq_eq inseparable_eq_eq -theorem TopologicalSpace.IsTopologicalBasis.inseparable_iff {b : Set (Set α)} - (hb : IsTopologicalBasis b) {x y : α} : Inseparable x y ↔ ∀ s ∈ b, (x ∈ s ↔ y ∈ s) := +theorem TopologicalSpace.IsTopologicalBasis.inseparable_iff {b : Set (Set X)} + (hb : IsTopologicalBasis b) {x y : X} : Inseparable x y ↔ ∀ s ∈ b, (x ∈ s ↔ y ∈ s) := ⟨fun h s hs ↦ inseparable_iff_forall_open.1 h _ (hb.isOpen hs), fun h ↦ hb.nhds_hasBasis.eq_of_same_basis $ by convert hb.nhds_hasBasis using 2 exact and_congr_right (h _)⟩ -theorem TopologicalSpace.IsTopologicalBasis.eq_iff [T0Space α] {b : Set (Set α)} - (hb : IsTopologicalBasis b) {x y : α} : x = y ↔ ∀ s ∈ b, (x ∈ s ↔ y ∈ s) := +theorem TopologicalSpace.IsTopologicalBasis.eq_iff [T0Space X] {b : Set (Set X)} + (hb : IsTopologicalBasis b) {x y : X} : x = y ↔ ∀ s ∈ b, (x ∈ s ↔ y ∈ s) := inseparable_iff_eq.symm.trans hb.inseparable_iff -theorem t0Space_iff_exists_isOpen_xor'_mem (α : Type u) [TopologicalSpace α] : - T0Space α ↔ ∀ x y, x ≠ y → ∃ U : Set α, IsOpen U ∧ Xor' (x ∈ U) (y ∈ U) := by +theorem t0Space_iff_exists_isOpen_xor'_mem (X : Type u) [TopologicalSpace X] : + T0Space X ↔ ∀ x y, x ≠ y → ∃ U : Set X, IsOpen U ∧ Xor' (x ∈ U) (y ∈ U) := by simp only [t0Space_iff_not_inseparable, xor_iff_not_iff, not_forall, exists_prop, inseparable_iff_forall_open] #align t0_space_iff_exists_is_open_xor_mem t0Space_iff_exists_isOpen_xor'_mem -theorem exists_isOpen_xor'_mem [T0Space α] {x y : α} (h : x ≠ y) : - ∃ U : Set α, IsOpen U ∧ Xor' (x ∈ U) (y ∈ U) := - (t0Space_iff_exists_isOpen_xor'_mem α).1 ‹_› x y h +theorem exists_isOpen_xor'_mem [T0Space X] {x y : X} (h : x ≠ y) : + ∃ U : Set X, IsOpen U ∧ Xor' (x ∈ U) (y ∈ U) := + (t0Space_iff_exists_isOpen_xor'_mem X).1 ‹_› x y h #align exists_is_open_xor_mem exists_isOpen_xor'_mem /-- Specialization forms a partial order on a t0 topological space. -/ -def specializationOrder (α : Type*) [TopologicalSpace α] [T0Space α] : PartialOrder α := - { specializationPreorder α, PartialOrder.lift (OrderDual.toDual ∘ 𝓝) nhds_injective with } +def specializationOrder (X) [TopologicalSpace X] [T0Space X] : PartialOrder X := + { specializationPreorder X, PartialOrder.lift (OrderDual.toDual ∘ 𝓝) nhds_injective with } #align specialization_order specializationOrder -instance : T0Space (SeparationQuotient α) := +instance : T0Space (SeparationQuotient X) := ⟨fun x y => Quotient.inductionOn₂' x y fun _ _ h => SeparationQuotient.mk_eq_mk.2 <| SeparationQuotient.inducing_mk.inseparable_iff.1 h⟩ -theorem minimal_nonempty_closed_subsingleton [T0Space α] {s : Set α} (hs : IsClosed s) +theorem minimal_nonempty_closed_subsingleton [T0Space X] {s : Set X} (hs : IsClosed s) (hmin : ∀ t, t ⊆ s → t.Nonempty → IsClosed t → t = s) : s.Subsingleton := by - clear β -- Porting note: added + clear Y -- Porting note: added refine' fun x hx y hy => of_not_not fun hxy => _ rcases exists_isOpen_xor'_mem hxy with ⟨U, hUo, hU⟩ wlog h : x ∈ U ∧ y ∉ U @@ -273,7 +277,7 @@ theorem minimal_nonempty_closed_subsingleton [T0Space α] {s : Set α} (hs : IsC exact (this.symm.subset hx).2 hxU #align minimal_nonempty_closed_subsingleton minimal_nonempty_closed_subsingleton -theorem minimal_nonempty_closed_eq_singleton [T0Space α] {s : Set α} (hs : IsClosed s) +theorem minimal_nonempty_closed_eq_singleton [T0Space X] {s : Set X} (hs : IsClosed s) (hne : s.Nonempty) (hmin : ∀ t, t ⊆ s → t.Nonempty → IsClosed t → t = s) : ∃ x, s = {x} := exists_eq_singleton_iff_nonempty_subsingleton.2 ⟨hne, minimal_nonempty_closed_subsingleton hs hmin⟩ @@ -281,17 +285,16 @@ theorem minimal_nonempty_closed_eq_singleton [T0Space α] {s : Set α} (hs : IsC /-- Given a closed set `S` in a compact T₀ space, there is some `x ∈ S` such that `{x}` is closed. -/ -theorem IsClosed.exists_closed_singleton {α : Type*} [TopologicalSpace α] [T0Space α] - [CompactSpace α] {S : Set α} (hS : IsClosed S) (hne : S.Nonempty) : - ∃ x : α, x ∈ S ∧ IsClosed ({x} : Set α) := by +theorem IsClosed.exists_closed_singleton [T0Space X] [CompactSpace X] {S : Set X} + (hS : IsClosed S) (hne : S.Nonempty) : ∃ x : X, x ∈ S ∧ IsClosed ({x} : Set X) := by obtain ⟨V, Vsub, Vne, Vcls, hV⟩ := hS.exists_minimal_nonempty_closed_subset hne rcases minimal_nonempty_closed_eq_singleton Vcls Vne hV with ⟨x, rfl⟩ exact ⟨x, Vsub (mem_singleton x), Vcls⟩ #align is_closed.exists_closed_singleton IsClosed.exists_closed_singleton -theorem minimal_nonempty_open_subsingleton [T0Space α] {s : Set α} (hs : IsOpen s) +theorem minimal_nonempty_open_subsingleton [T0Space X] {s : Set X} (hs : IsOpen s) (hmin : ∀ t, t ⊆ s → t.Nonempty → IsOpen t → t = s) : s.Subsingleton := by - clear β -- Porting note: added + clear Y -- Porting note: added refine' fun x hx y hy => of_not_not fun hxy => _ rcases exists_isOpen_xor'_mem hxy with ⟨U, hUo, hU⟩ wlog h : x ∈ U ∧ y ∉ U @@ -301,17 +304,17 @@ theorem minimal_nonempty_open_subsingleton [T0Space α] {s : Set α} (hs : IsOpe exact hyU (this.symm.subset hy).2 #align minimal_nonempty_open_subsingleton minimal_nonempty_open_subsingleton -theorem minimal_nonempty_open_eq_singleton [T0Space α] {s : Set α} (hs : IsOpen s) +theorem minimal_nonempty_open_eq_singleton [T0Space X] {s : Set X} (hs : IsOpen s) (hne : s.Nonempty) (hmin : ∀ t, t ⊆ s → t.Nonempty → IsOpen t → t = s) : ∃ x, s = {x} := exists_eq_singleton_iff_nonempty_subsingleton.2 ⟨hne, minimal_nonempty_open_subsingleton hs hmin⟩ #align minimal_nonempty_open_eq_singleton minimal_nonempty_open_eq_singleton /-- Given an open finite set `S` in a T₀ space, there is some `x ∈ S` such that `{x}` is open. -/ -theorem exists_open_singleton_of_open_finite [T0Space α] {s : Set α} (hfin : s.Finite) - (hne : s.Nonempty) (ho : IsOpen s) : ∃ x ∈ s, IsOpen ({x} : Set α) := by - lift s to Finset α using hfin +theorem exists_open_singleton_of_open_finite [T0Space X] {s : Set X} (hfin : s.Finite) + (hne : s.Nonempty) (ho : IsOpen s) : ∃ x ∈ s, IsOpen ({x} : Set X) := by + lift s to Finset X using hfin induction' s using Finset.strongInductionOn with s ihs - rcases em (∃ t, t ⊂ s ∧ t.Nonempty ∧ IsOpen (t : Set α)) with (⟨t, hts, htne, hto⟩ | ht) + rcases em (∃ t, t ⊂ s ∧ t.Nonempty ∧ IsOpen (t : Set X)) with (⟨t, hts, htne, hto⟩ | ht) · rcases ihs t hts htne hto with ⟨x, hxt, hxo⟩ exact ⟨x, hts.1 hxt, hxo⟩ · -- Porting note: was `rcases minimal_nonempty_open_eq_singleton ho hne _ with ⟨x, hx⟩` @@ -320,46 +323,46 @@ theorem exists_open_singleton_of_open_finite [T0Space α] {s : Set α} (hfin : s · exact ⟨x, hx.symm ▸ rfl, hx ▸ ho⟩ refine minimal_nonempty_open_eq_singleton ho hne ?_ refine' fun t hts htne hto => of_not_not fun hts' => ht _ - lift t to Finset α using s.finite_toSet.subset hts + lift t to Finset X using s.finite_toSet.subset hts exact ⟨t, ssubset_iff_subset_ne.2 ⟨hts, mt Finset.coe_inj.2 hts'⟩, htne, hto⟩ #align exists_open_singleton_of_open_finite exists_open_singleton_of_open_finite -theorem exists_open_singleton_of_finite [T0Space α] [Finite α] [Nonempty α] : - ∃ x : α, IsOpen ({x} : Set α) := +theorem exists_open_singleton_of_finite [T0Space X] [Finite X] [Nonempty X] : + ∃ x : X, IsOpen ({x} : Set X) := let ⟨x, _, h⟩ := exists_open_singleton_of_open_finite (Set.toFinite _) univ_nonempty isOpen_univ ⟨x, h⟩ #align exists_open_singleton_of_fintype exists_open_singleton_of_finite -theorem t0Space_of_injective_of_continuous [TopologicalSpace β] {f : α → β} - (hf : Function.Injective f) (hf' : Continuous f) [T0Space β] : T0Space α := +theorem t0Space_of_injective_of_continuous [TopologicalSpace Y] {f : X → Y} + (hf : Function.Injective f) (hf' : Continuous f) [T0Space Y] : T0Space X := ⟨fun _ _ h => hf <| (h.map hf').eq⟩ #align t0_space_of_injective_of_continuous t0Space_of_injective_of_continuous -protected theorem Embedding.t0Space [TopologicalSpace β] [T0Space β] {f : α → β} - (hf : Embedding f) : T0Space α := +protected theorem Embedding.t0Space [TopologicalSpace Y] [T0Space Y] {f : X → Y} + (hf : Embedding f) : T0Space X := t0Space_of_injective_of_continuous hf.inj hf.continuous #align embedding.t0_space Embedding.t0Space -instance Subtype.t0Space [T0Space α] {p : α → Prop} : T0Space (Subtype p) := +instance Subtype.t0Space [T0Space X] {p : X → Prop} : T0Space (Subtype p) := embedding_subtype_val.t0Space #align subtype.t0_space Subtype.t0Space -theorem t0Space_iff_or_not_mem_closure (α : Type u) [TopologicalSpace α] : - T0Space α ↔ ∀ a b : α, a ≠ b → a ∉ closure ({b} : Set α) ∨ b ∉ closure ({a} : Set α) := by +theorem t0Space_iff_or_not_mem_closure (X : Type u) [TopologicalSpace X] : + T0Space X ↔ ∀ a b : X, a ≠ b → a ∉ closure ({b} : Set X) ∨ b ∉ closure ({a} : Set X) := by simp only [t0Space_iff_not_inseparable, inseparable_iff_mem_closure, not_and_or] #align t0_space_iff_or_not_mem_closure t0Space_iff_or_not_mem_closure -instance Prod.instT0Space [TopologicalSpace β] [T0Space α] [T0Space β] : T0Space (α × β) := +instance Prod.instT0Space [TopologicalSpace Y] [T0Space X] [T0Space Y] : T0Space (X × Y) := ⟨fun _ _ h => Prod.ext (h.map continuous_fst).eq (h.map continuous_snd).eq⟩ -instance Pi.instT0Space {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - [∀ i, T0Space (π i)] : - T0Space (∀ i, π i) := +instance Pi.instT0Space {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] + [∀ i, T0Space (X i)] : + T0Space (∀ i, X i) := ⟨fun _ _ h => funext fun i => (h.map (continuous_apply i)).eq⟩ #align pi.t0_space Pi.instT0Space -theorem T0Space.of_cover (h : ∀ x y, Inseparable x y → ∃ s : Set α, x ∈ s ∧ y ∈ s ∧ T0Space s) : - T0Space α := by +theorem T0Space.of_cover (h : ∀ x y, Inseparable x y → ∃ s : Set X, x ∈ s ∧ y ∈ s ∧ T0Space s) : + T0Space X := by refine' ⟨fun x y hxy => _⟩ rcases h x y hxy with ⟨s, hxs, hys, hs⟩ lift x to s using hxs; lift y to s using hys @@ -367,7 +370,7 @@ theorem T0Space.of_cover (h : ∀ x y, Inseparable x y → ∃ s : Set α, x ∈ exact congr_arg Subtype.val hxy.eq #align t0_space.of_cover T0Space.of_cover -theorem T0Space.of_open_cover (h : ∀ x, ∃ s : Set α, x ∈ s ∧ IsOpen s ∧ T0Space s) : T0Space α := +theorem T0Space.of_open_cover (h : ∀ x, ∃ s : Set X, x ∈ s ∧ IsOpen s ∧ T0Space s) : T0Space X := T0Space.of_cover fun x _ hxy => let ⟨s, hxs, hso, hs⟩ := h x ⟨s, hxs, (hxy.mem_open_iff hso).1 hxs, hs⟩ @@ -376,47 +379,47 @@ theorem T0Space.of_open_cover (h : ∀ x, ∃ s : Set α, x ∈ s ∧ IsOpen s /-- A T₁ space, also known as a Fréchet space, is a topological space where every singleton set is closed. Equivalently, for every pair `x ≠ y`, there is an open set containing `x` and not `y`. -/ -class T1Space (α : Type u) [TopologicalSpace α] : Prop where +class T1Space (X : Type u) [TopologicalSpace X] : Prop where /-- A singleton in a T₁ space is a closed set. -/ - t1 : ∀ x, IsClosed ({x} : Set α) + t1 : ∀ x, IsClosed ({x} : Set X) #align t1_space T1Space -theorem isClosed_singleton [T1Space α] {x : α} : IsClosed ({x} : Set α) := +theorem isClosed_singleton [T1Space X] {x : X} : IsClosed ({x} : Set X) := T1Space.t1 x #align is_closed_singleton isClosed_singleton -theorem isOpen_compl_singleton [T1Space α] {x : α} : IsOpen ({x}ᶜ : Set α) := +theorem isOpen_compl_singleton [T1Space X] {x : X} : IsOpen ({x}ᶜ : Set X) := isClosed_singleton.isOpen_compl #align is_open_compl_singleton isOpen_compl_singleton -theorem isOpen_ne [T1Space α] {x : α} : IsOpen { y | y ≠ x } := +theorem isOpen_ne [T1Space X] {x : X} : IsOpen { y | y ≠ x } := isOpen_compl_singleton #align is_open_ne isOpen_ne @[to_additive] -theorem Continuous.isOpen_mulSupport [T1Space α] [One α] [TopologicalSpace β] {f : β → α} +theorem Continuous.isOpen_mulSupport [T1Space X] [One X] [TopologicalSpace Y] {f : Y → X} (hf : Continuous f) : IsOpen (mulSupport f) := isOpen_ne.preimage hf #align continuous.is_open_mul_support Continuous.isOpen_mulSupport #align continuous.is_open_support Continuous.isOpen_support -theorem Ne.nhdsWithin_compl_singleton [T1Space α] {x y : α} (h : x ≠ y) : 𝓝[{y}ᶜ] x = 𝓝 x := +theorem Ne.nhdsWithin_compl_singleton [T1Space X] {x y : X} (h : x ≠ y) : 𝓝[{y}ᶜ] x = 𝓝 x := isOpen_ne.nhdsWithin_eq h #align ne.nhds_within_compl_singleton Ne.nhdsWithin_compl_singleton -theorem Ne.nhdsWithin_diff_singleton [T1Space α] {x y : α} (h : x ≠ y) (s : Set α) : +theorem Ne.nhdsWithin_diff_singleton [T1Space X] {x y : X} (h : x ≠ y) (s : Set X) : 𝓝[s \ {y}] x = 𝓝[s] x := by rw [diff_eq, inter_comm, nhdsWithin_inter_of_mem] exact mem_nhdsWithin_of_mem_nhds (isOpen_ne.mem_nhds h) #align ne.nhds_within_diff_singleton Ne.nhdsWithin_diff_singleton -lemma nhdsWithin_compl_singleton_le [T1Space α] (x y : α) : 𝓝[{x}ᶜ] x ≤ 𝓝[{y}ᶜ] x := by +lemma nhdsWithin_compl_singleton_le [T1Space X] (x y : X) : 𝓝[{x}ᶜ] x ≤ 𝓝[{y}ᶜ] x := by rcases eq_or_ne x y with rfl|hy · exact Eq.le rfl · rw [Ne.nhdsWithin_compl_singleton hy] exact nhdsWithin_le_nhds -theorem isOpen_setOf_eventually_nhdsWithin [T1Space α] {p : α → Prop} : +theorem isOpen_setOf_eventually_nhdsWithin [T1Space X] {p : X → Prop} : IsOpen { x | ∀ᶠ y in 𝓝[≠] x, p y } := by refine' isOpen_iff_mem_nhds.mpr fun a ha => _ filter_upwards [eventually_nhds_nhdsWithin.mpr ha] with b hb @@ -426,37 +429,37 @@ theorem isOpen_setOf_eventually_nhdsWithin [T1Space α] {p : α → Prop} : exact hb.filter_mono nhdsWithin_le_nhds #align is_open_set_of_eventually_nhds_within isOpen_setOf_eventually_nhdsWithin -protected theorem Set.Finite.isClosed [T1Space α] {s : Set α} (hs : Set.Finite s) : IsClosed s := by +protected theorem Set.Finite.isClosed [T1Space X] {s : Set X} (hs : Set.Finite s) : IsClosed s := by rw [← biUnion_of_singleton s] exact hs.isClosed_biUnion fun i _ => isClosed_singleton #align set.finite.is_closed Set.Finite.isClosed -theorem TopologicalSpace.IsTopologicalBasis.exists_mem_of_ne [T1Space α] {b : Set (Set α)} - (hb : IsTopologicalBasis b) {x y : α} (h : x ≠ y) : ∃ a ∈ b, x ∈ a ∧ y ∉ a := by +theorem TopologicalSpace.IsTopologicalBasis.exists_mem_of_ne [T1Space X] {b : Set (Set X)} + (hb : IsTopologicalBasis b) {x y : X} (h : x ≠ y) : ∃ a ∈ b, x ∈ a ∧ y ∉ a := by rcases hb.isOpen_iff.1 isOpen_ne x h with ⟨a, ab, xa, ha⟩ exact ⟨a, ab, xa, fun h => ha h rfl⟩ #align topological_space.is_topological_basis.exists_mem_of_ne TopologicalSpace.IsTopologicalBasis.exists_mem_of_ne -theorem Filter.coclosedCompact_le_cofinite [T1Space α] : - Filter.coclosedCompact α ≤ Filter.cofinite := fun s hs => +theorem Filter.coclosedCompact_le_cofinite [T1Space X] : + Filter.coclosedCompact X ≤ Filter.cofinite := fun s hs => compl_compl s ▸ hs.isCompact.compl_mem_coclosedCompact_of_isClosed hs.isClosed #align filter.coclosed_compact_le_cofinite Filter.coclosedCompact_le_cofinite -variable (α) +variable (X) /-- In a `T1Space`, relatively compact sets form a bornology. Its cobounded filter is `Filter.coclosedCompact`. See also `Bornology.inCompact` the bornology of sets contained in a compact set. -/ -def Bornology.relativelyCompact [T1Space α] : Bornology α where - cobounded' := Filter.coclosedCompact α +def Bornology.relativelyCompact [T1Space X] : Bornology X where + cobounded' := Filter.coclosedCompact X le_cofinite' := Filter.coclosedCompact_le_cofinite #align bornology.relatively_compact Bornology.relativelyCompact -variable {α} +variable {X} -theorem Bornology.relativelyCompact.isBounded_iff [T1Space α] {s : Set α} : - @Bornology.IsBounded _ (Bornology.relativelyCompact α) s ↔ IsCompact (closure s) := by - change sᶜ ∈ Filter.coclosedCompact α ↔ _ +theorem Bornology.relativelyCompact.isBounded_iff [T1Space X] {s : Set X} : + @Bornology.IsBounded _ (Bornology.relativelyCompact X) s ↔ IsCompact (closure s) := by + change sᶜ ∈ Filter.coclosedCompact X ↔ _ rw [Filter.mem_coclosedCompact] constructor · rintro ⟨t, ht₁, ht₂, hst⟩ @@ -466,21 +469,21 @@ theorem Bornology.relativelyCompact.isBounded_iff [T1Space α] {s : Set α} : exact ⟨closure s, isClosed_closure, h, compl_subset_compl.mpr subset_closure⟩ #align bornology.relatively_compact.is_bounded_iff Bornology.relativelyCompact.isBounded_iff -protected theorem Finset.isClosed [T1Space α] (s : Finset α) : IsClosed (s : Set α) := +protected theorem Finset.isClosed [T1Space X] (s : Finset X) : IsClosed (s : Set X) := s.finite_toSet.isClosed #align finset.is_closed Finset.isClosed -theorem t1Space_TFAE (α : Type u) [ TopologicalSpace α ] : - List.TFAE [T1Space α, - ∀ x, IsClosed ({ x } : Set α), - ∀ x, IsOpen ({ x }ᶜ : Set α), - Continuous (@CofiniteTopology.of α), - ∀ ⦃x y : α⦄, x ≠ y → {y}ᶜ ∈ 𝓝 x, - ∀ ⦃x y : α⦄, x ≠ y → ∃ s ∈ 𝓝 x, y ∉ s, - ∀ ⦃x y : α⦄, x ≠ y → ∃ U : Set α, IsOpen U ∧ x ∈ U ∧ y ∉ U, - ∀ ⦃x y : α⦄, x ≠ y → Disjoint (𝓝 x) (pure y), - ∀ ⦃x y : α⦄, x ≠ y → Disjoint (pure x) (𝓝 y), - ∀ ⦃x y : α⦄, x ⤳ y → x = y] := by +theorem t1Space_TFAE (X : Type u) [TopologicalSpace X] : + List.TFAE [T1Space X, + ∀ x, IsClosed ({ x } : Set X), + ∀ x, IsOpen ({ x }ᶜ : Set X), + Continuous (@CofiniteTopology.of X), + ∀ ⦃x y : X⦄, x ≠ y → {y}ᶜ ∈ 𝓝 x, + ∀ ⦃x y : X⦄, x ≠ y → ∃ s ∈ 𝓝 x, y ∉ s, + ∀ ⦃x y : X⦄, x ≠ y → ∃ U : Set X, IsOpen U ∧ x ∈ U ∧ y ∉ U, + ∀ ⦃x y : X⦄, x ≠ y → Disjoint (𝓝 x) (pure y), + ∀ ⦃x y : X⦄, x ≠ y → Disjoint (pure x) (𝓝 y), + ∀ ⦃x y : X⦄, x ⤳ y → x = y] := by tfae_have 1 ↔ 2 · exact ⟨fun h => h.1, fun h => ⟨h⟩⟩ tfae_have 2 ↔ 3 @@ -509,85 +512,85 @@ theorem t1Space_TFAE (α : Type u) [ TopologicalSpace α ] : tfae_finish #align t1_space_tfae t1Space_TFAE -theorem t1Space_iff_continuous_cofinite_of {α : Type*} [TopologicalSpace α] : - T1Space α ↔ Continuous (@CofiniteTopology.of α) := - (t1Space_TFAE α).out 0 3 +theorem t1Space_iff_continuous_cofinite_of : T1Space X ↔ Continuous (@CofiniteTopology.of X) := + (t1Space_TFAE X).out 0 3 #align t1_space_iff_continuous_cofinite_of t1Space_iff_continuous_cofinite_of -theorem CofiniteTopology.continuous_of [T1Space α] : Continuous (@CofiniteTopology.of α) := +theorem CofiniteTopology.continuous_of [T1Space X] : Continuous (@CofiniteTopology.of X) := t1Space_iff_continuous_cofinite_of.mp ‹_› #align cofinite_topology.continuous_of CofiniteTopology.continuous_of theorem t1Space_iff_exists_open : - T1Space α ↔ ∀ x y, x ≠ y → ∃ U : Set α, IsOpen U ∧ x ∈ U ∧ y ∉ U := - (t1Space_TFAE α).out 0 6 + T1Space X ↔ ∀ x y, x ≠ y → ∃ U : Set X, IsOpen U ∧ x ∈ U ∧ y ∉ U := + (t1Space_TFAE X).out 0 6 #align t1_space_iff_exists_open t1Space_iff_exists_open -theorem t1Space_iff_disjoint_pure_nhds : T1Space α ↔ ∀ ⦃x y : α⦄, x ≠ y → Disjoint (pure x) (𝓝 y) := - (t1Space_TFAE α).out 0 8 +theorem t1Space_iff_disjoint_pure_nhds : T1Space X ↔ ∀ ⦃x y : X⦄, x ≠ y → Disjoint (pure x) (𝓝 y) := + (t1Space_TFAE X).out 0 8 #align t1_space_iff_disjoint_pure_nhds t1Space_iff_disjoint_pure_nhds -theorem t1Space_iff_disjoint_nhds_pure : T1Space α ↔ ∀ ⦃x y : α⦄, x ≠ y → Disjoint (𝓝 x) (pure y) := - (t1Space_TFAE α).out 0 7 +theorem t1Space_iff_disjoint_nhds_pure : T1Space X ↔ ∀ ⦃x y : X⦄, x ≠ y → Disjoint (𝓝 x) (pure y) := + (t1Space_TFAE X).out 0 7 #align t1_space_iff_disjoint_nhds_pure t1Space_iff_disjoint_nhds_pure -theorem t1Space_iff_specializes_imp_eq : T1Space α ↔ ∀ ⦃x y : α⦄, x ⤳ y → x = y := - (t1Space_TFAE α).out 0 9 +theorem t1Space_iff_specializes_imp_eq : T1Space X ↔ ∀ ⦃x y : X⦄, x ⤳ y → x = y := + (t1Space_TFAE X).out 0 9 #align t1_space_iff_specializes_imp_eq t1Space_iff_specializes_imp_eq -theorem disjoint_pure_nhds [T1Space α] {x y : α} (h : x ≠ y) : Disjoint (pure x) (𝓝 y) := +theorem disjoint_pure_nhds [T1Space X] {x y : X} (h : x ≠ y) : Disjoint (pure x) (𝓝 y) := t1Space_iff_disjoint_pure_nhds.mp ‹_› h #align disjoint_pure_nhds disjoint_pure_nhds -theorem disjoint_nhds_pure [T1Space α] {x y : α} (h : x ≠ y) : Disjoint (𝓝 x) (pure y) := +theorem disjoint_nhds_pure [T1Space X] {x y : X} (h : x ≠ y) : Disjoint (𝓝 x) (pure y) := t1Space_iff_disjoint_nhds_pure.mp ‹_› h #align disjoint_nhds_pure disjoint_nhds_pure -theorem Specializes.eq [T1Space α] {x y : α} (h : x ⤳ y) : x = y := +theorem Specializes.eq [T1Space X] {x y : X} (h : x ⤳ y) : x = y := t1Space_iff_specializes_imp_eq.1 ‹_› h #align specializes.eq Specializes.eq -theorem specializes_iff_eq [T1Space α] {x y : α} : x ⤳ y ↔ x = y := +theorem specializes_iff_eq [T1Space X] {x y : X} : x ⤳ y ↔ x = y := ⟨Specializes.eq, fun h => h ▸ specializes_rfl⟩ #align specializes_iff_eq specializes_iff_eq -@[simp] theorem specializes_eq_eq [T1Space α] : (· ⤳ ·) = @Eq α := +@[simp] theorem specializes_eq_eq [T1Space X] : (· ⤳ ·) = @Eq X := funext₂ fun _ _ => propext specializes_iff_eq #align specializes_eq_eq specializes_eq_eq @[simp] -theorem pure_le_nhds_iff [T1Space α] {a b : α} : pure a ≤ 𝓝 b ↔ a = b := +theorem pure_le_nhds_iff [T1Space X] {a b : X} : pure a ≤ 𝓝 b ↔ a = b := specializes_iff_pure.symm.trans specializes_iff_eq #align pure_le_nhds_iff pure_le_nhds_iff @[simp] -theorem nhds_le_nhds_iff [T1Space α] {a b : α} : 𝓝 a ≤ 𝓝 b ↔ a = b := +theorem nhds_le_nhds_iff [T1Space X] {a b : X} : 𝓝 a ≤ 𝓝 b ↔ a = b := specializes_iff_eq #align nhds_le_nhds_iff nhds_le_nhds_iff -instance {α : Type*} : T1Space (CofiniteTopology α) := +instance : T1Space (CofiniteTopology X) := t1Space_iff_continuous_cofinite_of.mpr continuous_id -theorem t1Space_antitone {α : Type*} : Antitone (@T1Space α) := fun a _ h _ => +theorem t1Space_antitone : Antitone (@T1Space X) := fun a _ h _ => @T1Space.mk _ a fun x => (T1Space.t1 x).mono h #align t1_space_antitone t1Space_antitone -theorem continuousWithinAt_update_of_ne [T1Space α] [DecidableEq α] [TopologicalSpace β] {f : α → β} - {s : Set α} {x y : α} {z : β} (hne : y ≠ x) : - ContinuousWithinAt (Function.update f x z) s y ↔ ContinuousWithinAt f s y := +theorem continuousWithinAt_update_of_ne [T1Space X] [DecidableEq X] [TopologicalSpace Y] {f : X → Y} + {s : Set X} {x x' : X} {y : Y} (hne : x' ≠ x) : + ContinuousWithinAt (Function.update f x y) s x' ↔ ContinuousWithinAt f s x' := EventuallyEq.congr_continuousWithinAt (mem_nhdsWithin_of_mem_nhds <| mem_of_superset (isOpen_ne.mem_nhds hne) fun _y' hy' => Function.update_noteq hy' _ _) (Function.update_noteq hne _ _) #align continuous_within_at_update_of_ne continuousWithinAt_update_of_ne -theorem continuousAt_update_of_ne [T1Space α] [DecidableEq α] [TopologicalSpace β] {f : α → β} - {x y : α} {z : β} (hne : y ≠ x) : ContinuousAt (Function.update f x z) y ↔ ContinuousAt f y := +theorem continuousAt_update_of_ne [T1Space X] [DecidableEq X] [TopologicalSpace Y] + {f : X → Y} {x x' : X} {y : Y} (hne : x' ≠ x) : + ContinuousAt (Function.update f x y) x' ↔ ContinuousAt f x' := by simp only [← continuousWithinAt_univ, continuousWithinAt_update_of_ne hne] #align continuous_at_update_of_ne continuousAt_update_of_ne -theorem continuousOn_update_iff [T1Space α] [DecidableEq α] [TopologicalSpace β] {f : α → β} - {s : Set α} {x : α} {y : β} : +theorem continuousOn_update_iff [T1Space X] [DecidableEq X] [TopologicalSpace Y] {f : X → Y} + {s : Set X} {x : X} {y : Y} : ContinuousOn (Function.update f x y) s ↔ ContinuousOn f (s \ {x}) ∧ (x ∈ s → Tendsto f (𝓝[s \ {x}] x) (𝓝 y)) := by rw [ContinuousOn, ← and_forall_ne x, and_comm] @@ -601,64 +604,64 @@ theorem continuousOn_update_iff [T1Space α] [DecidableEq α] [TopologicalSpace · exact continuousWithinAt_update_same #align continuous_on_update_iff continuousOn_update_iff -theorem t1Space_of_injective_of_continuous [TopologicalSpace β] {f : α → β} - (hf : Function.Injective f) (hf' : Continuous f) [T1Space β] : T1Space α := +theorem t1Space_of_injective_of_continuous [TopologicalSpace Y] {f : X → Y} + (hf : Function.Injective f) (hf' : Continuous f) [T1Space Y] : T1Space X := t1Space_iff_specializes_imp_eq.2 fun _ _ h => hf (h.map hf').eq #align t1_space_of_injective_of_continuous t1Space_of_injective_of_continuous -protected theorem Embedding.t1Space [TopologicalSpace β] [T1Space β] {f : α → β} - (hf : Embedding f) : T1Space α := +protected theorem Embedding.t1Space [TopologicalSpace Y] [T1Space Y] {f : X → Y} + (hf : Embedding f) : T1Space X := t1Space_of_injective_of_continuous hf.inj hf.continuous #align embedding.t1_space Embedding.t1Space -instance Subtype.t1Space {α : Type u} [TopologicalSpace α] [T1Space α] {p : α → Prop} : +instance Subtype.t1Space {X : Type u} [TopologicalSpace X] [T1Space X] {p : X → Prop} : T1Space (Subtype p) := embedding_subtype_val.t1Space #align subtype.t1_space Subtype.t1Space -instance [TopologicalSpace β] [T1Space α] [T1Space β] : T1Space (α × β) := +instance [TopologicalSpace Y] [T1Space X] [T1Space Y] : T1Space (X × Y) := ⟨fun ⟨a, b⟩ => @singleton_prod_singleton _ _ a b ▸ isClosed_singleton.prod isClosed_singleton⟩ -instance {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] [∀ i, T1Space (π i)] : - T1Space (∀ i, π i) := +instance {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, T1Space (X i)] : + T1Space (∀ i, X i) := ⟨fun f => univ_pi_singleton f ▸ isClosed_set_pi fun _ _ => isClosed_singleton⟩ -- see Note [lower instance priority] -instance (priority := 100) T1Space.t0Space [T1Space α] : T0Space α := +instance (priority := 100) T1Space.t0Space [T1Space X] : T0Space X := ⟨fun _ _ h => h.specializes.eq⟩ #align t1_space.t0_space T1Space.t0Space @[simp] -theorem compl_singleton_mem_nhds_iff [T1Space α] {x y : α} : {x}ᶜ ∈ 𝓝 y ↔ y ≠ x := +theorem compl_singleton_mem_nhds_iff [T1Space X] {x y : X} : {x}ᶜ ∈ 𝓝 y ↔ y ≠ x := isOpen_compl_singleton.mem_nhds_iff #align compl_singleton_mem_nhds_iff compl_singleton_mem_nhds_iff -theorem compl_singleton_mem_nhds [T1Space α] {x y : α} (h : y ≠ x) : {x}ᶜ ∈ 𝓝 y := +theorem compl_singleton_mem_nhds [T1Space X] {x y : X} (h : y ≠ x) : {x}ᶜ ∈ 𝓝 y := compl_singleton_mem_nhds_iff.mpr h #align compl_singleton_mem_nhds compl_singleton_mem_nhds @[simp] -theorem closure_singleton [T1Space α] {a : α} : closure ({a} : Set α) = {a} := +theorem closure_singleton [T1Space X] {x : X} : closure ({x} : Set X) = {x} := isClosed_singleton.closure_eq #align closure_singleton closure_singleton -- porting note: todo: the proof was `hs.induction_on (by simp) fun x => by simp` -theorem Set.Subsingleton.closure [T1Space α] {s : Set α} (hs : s.Subsingleton) : +theorem Set.Subsingleton.closure [T1Space X] {s : Set X} (hs : s.Subsingleton) : (closure s).Subsingleton := by rcases hs.eq_empty_or_singleton with (rfl | ⟨x, rfl⟩) <;> simp #align set.subsingleton.closure Set.Subsingleton.closure @[simp] -theorem subsingleton_closure [T1Space α] {s : Set α} : (closure s).Subsingleton ↔ s.Subsingleton := +theorem subsingleton_closure [T1Space X] {s : Set X} : (closure s).Subsingleton ↔ s.Subsingleton := ⟨fun h => h.anti subset_closure, fun h => h.closure⟩ #align subsingleton_closure subsingleton_closure -theorem isClosedMap_const {α β} [TopologicalSpace α] [TopologicalSpace β] [T1Space β] {y : β} : - IsClosedMap (Function.const α y) := +theorem isClosedMap_const {X Y} [TopologicalSpace X] [TopologicalSpace Y] [T1Space Y] {y : Y} : + IsClosedMap (Function.const X y) := IsClosedMap.of_nonempty fun s _ h2s => by simp_rw [const, h2s.image_const, isClosed_singleton] #align is_closed_map_const isClosedMap_const -theorem nhdsWithin_insert_of_ne [T1Space α] {x y : α} {s : Set α} (hxy : x ≠ y) : +theorem nhdsWithin_insert_of_ne [T1Space X] {x y : X} {s : Set X} (hxy : x ≠ y) : 𝓝[insert y s] x = 𝓝[s] x := by refine' le_antisymm (Filter.le_def.2 fun t ht => _) (nhdsWithin_mono x <| subset_insert y s) obtain ⟨o, ho, hxo, host⟩ := mem_nhdsWithin.mp ht @@ -669,7 +672,7 @@ theorem nhdsWithin_insert_of_ne [T1Space α] {x y : α} {s : Set α} (hxy : x /-- If `t` is a subset of `s`, except for one point, then `insert x s` is a neighborhood of `x` within `t`. -/ -theorem insert_mem_nhdsWithin_of_subset_insert [T1Space α] {x y : α} {s t : Set α} +theorem insert_mem_nhdsWithin_of_subset_insert [T1Space X] {x y : X} {s t : Set X} (hu : t ⊆ insert y s) : insert x s ∈ 𝓝[t] x := by rcases eq_or_ne x y with (rfl | h) · exact mem_of_superset self_mem_nhdsWithin hu @@ -679,21 +682,21 @@ theorem insert_mem_nhdsWithin_of_subset_insert [T1Space α] {x y : α} {s t : Se #align insert_mem_nhds_within_of_subset_insert insert_mem_nhdsWithin_of_subset_insert @[simp] -theorem ker_nhds [T1Space α] (x : α) : (𝓝 x).ker = {x} := by +theorem ker_nhds [T1Space X] (x : X) : (𝓝 x).ker = {x} := by simp [ker_nhds_eq_specializes] -theorem biInter_basis_nhds [T1Space α] {ι : Sort*} {p : ι → Prop} {s : ι → Set α} {x : α} +theorem biInter_basis_nhds [T1Space X] {ι : Sort*} {p : ι → Prop} {s : ι → Set X} {x : X} (h : (𝓝 x).HasBasis p s) : ⋂ (i) (_ : p i), s i = {x} := by rw [← h.ker, ker_nhds] #align bInter_basis_nhds biInter_basis_nhds @[simp] -theorem compl_singleton_mem_nhdsSet_iff [T1Space α] {x : α} {s : Set α} : {x}ᶜ ∈ 𝓝ˢ s ↔ x ∉ s := by +theorem compl_singleton_mem_nhdsSet_iff [T1Space X] {x : X} {s : Set X} : {x}ᶜ ∈ 𝓝ˢ s ↔ x ∉ s := by rw [isOpen_compl_singleton.mem_nhdsSet, subset_compl_singleton_iff] #align compl_singleton_mem_nhds_set_iff compl_singleton_mem_nhdsSet_iff @[simp] -theorem nhdsSet_le_iff [T1Space α] {s t : Set α} : 𝓝ˢ s ≤ 𝓝ˢ t ↔ s ⊆ t := by +theorem nhdsSet_le_iff [T1Space X] {s t : Set X} : 𝓝ˢ s ≤ 𝓝ˢ t ↔ s ⊆ t := by refine' ⟨_, fun h => monotone_nhdsSet h⟩ simp_rw [Filter.le_def]; intro h x hx specialize h {x}ᶜ @@ -703,34 +706,34 @@ theorem nhdsSet_le_iff [T1Space α] {s t : Set α} : 𝓝ˢ s ≤ 𝓝ˢ t ↔ s #align nhds_set_le_iff nhdsSet_le_iff @[simp] -theorem nhdsSet_inj_iff [T1Space α] {s t : Set α} : 𝓝ˢ s = 𝓝ˢ t ↔ s = t := by +theorem nhdsSet_inj_iff [T1Space X] {s t : Set X} : 𝓝ˢ s = 𝓝ˢ t ↔ s = t := by simp_rw [le_antisymm_iff] exact and_congr nhdsSet_le_iff nhdsSet_le_iff #align nhds_set_inj_iff nhdsSet_inj_iff -theorem injective_nhdsSet [T1Space α] : Function.Injective (𝓝ˢ : Set α → Filter α) := fun _ _ hst => +theorem injective_nhdsSet [T1Space X] : Function.Injective (𝓝ˢ : Set X → Filter X) := fun _ _ hst => nhdsSet_inj_iff.mp hst #align injective_nhds_set injective_nhdsSet -theorem strictMono_nhdsSet [T1Space α] : StrictMono (𝓝ˢ : Set α → Filter α) := +theorem strictMono_nhdsSet [T1Space X] : StrictMono (𝓝ˢ : Set X → Filter X) := monotone_nhdsSet.strictMono_of_injective injective_nhdsSet #align strict_mono_nhds_set strictMono_nhdsSet @[simp] -theorem nhds_le_nhdsSet_iff [T1Space α] {s : Set α} {x : α} : 𝓝 x ≤ 𝓝ˢ s ↔ x ∈ s := by +theorem nhds_le_nhdsSet_iff [T1Space X] {s : Set X} {x : X} : 𝓝 x ≤ 𝓝ˢ s ↔ x ∈ s := by rw [← nhdsSet_singleton, nhdsSet_le_iff, singleton_subset_iff] #align nhds_le_nhds_set_iff nhds_le_nhdsSet_iff /-- Removing a non-isolated point from a dense set, one still obtains a dense set. -/ -theorem Dense.diff_singleton [T1Space α] {s : Set α} (hs : Dense s) (x : α) [NeBot (𝓝[≠] x)] : +theorem Dense.diff_singleton [T1Space X] {s : Set X} (hs : Dense s) (x : X) [NeBot (𝓝[≠] x)] : Dense (s \ {x}) := hs.inter_of_open_right (dense_compl_singleton x) isOpen_compl_singleton #align dense.diff_singleton Dense.diff_singleton /-- Removing a finset from a dense set in a space without isolated points, one still obtains a dense set. -/ -theorem Dense.diff_finset [T1Space α] [∀ x : α, NeBot (𝓝[≠] x)] {s : Set α} (hs : Dense s) - (t : Finset α) : Dense (s \ t) := by +theorem Dense.diff_finset [T1Space X] [∀ x : X, NeBot (𝓝[≠] x)] {s : Set X} (hs : Dense s) + (t : Finset X) : Dense (s \ t) := by induction t using Finset.induction_on with | empty => simpa using hs | insert _ ih => @@ -740,54 +743,54 @@ theorem Dense.diff_finset [T1Space α] [∀ x : α, NeBot (𝓝[≠] x)] {s : Se /-- Removing a finite set from a dense set in a space without isolated points, one still obtains a dense set. -/ -theorem Dense.diff_finite [T1Space α] [∀ x : α, NeBot (𝓝[≠] x)] {s : Set α} (hs : Dense s) - {t : Set α} (ht : t.Finite) : Dense (s \ t) := by +theorem Dense.diff_finite [T1Space X] [∀ x : X, NeBot (𝓝[≠] x)] {s : Set X} (hs : Dense s) + {t : Set X} (ht : t.Finite) : Dense (s \ t) := by convert hs.diff_finset ht.toFinset exact (Finite.coe_toFinset _).symm #align dense.diff_finite Dense.diff_finite -/-- If a function to a `T1Space` tends to some limit `b` at some point `a`, then necessarily -`b = f a`. -/ -theorem eq_of_tendsto_nhds [TopologicalSpace β] [T1Space β] {f : α → β} {a : α} {b : β} - (h : Tendsto f (𝓝 a) (𝓝 b)) : f a = b := - by_contra fun hfa : f a ≠ b => - have fact₁ : {f a}ᶜ ∈ 𝓝 b := compl_singleton_mem_nhds hfa.symm - have fact₂ : Tendsto f (pure a) (𝓝 b) := h.comp (tendsto_id'.2 <| pure_le_nhds a) - fact₂ fact₁ (Eq.refl <| f a) +/-- If a function to a `T1Space` tends to some limit `y` at some point `x`, then necessarily +`y = f x`. -/ +theorem eq_of_tendsto_nhds [TopologicalSpace Y] [T1Space Y] {f : X → Y} {x : X} {y : Y} + (h : Tendsto f (𝓝 x) (𝓝 y)) : f x = y := + by_contra fun hfa : f x ≠ y => + have fact₁ : {f x}ᶜ ∈ 𝓝 y := compl_singleton_mem_nhds hfa.symm + have fact₂ : Tendsto f (pure x) (𝓝 y) := h.comp (tendsto_id'.2 <| pure_le_nhds x) + fact₂ fact₁ (Eq.refl <| f x) #align eq_of_tendsto_nhds eq_of_tendsto_nhds -theorem Filter.Tendsto.eventually_ne [TopologicalSpace β] [T1Space β] {α : Type*} {g : α → β} - {l : Filter α} {b₁ b₂ : β} (hg : Tendsto g l (𝓝 b₁)) (hb : b₁ ≠ b₂) : ∀ᶠ z in l, g z ≠ b₂ := +theorem Filter.Tendsto.eventually_ne [TopologicalSpace Y] [T1Space Y] {g : X → Y} + {l : Filter X} {b₁ b₂ : Y} (hg : Tendsto g l (𝓝 b₁)) (hb : b₁ ≠ b₂) : ∀ᶠ z in l, g z ≠ b₂ := hg.eventually (isOpen_compl_singleton.eventually_mem hb) #align filter.tendsto.eventually_ne Filter.Tendsto.eventually_ne -theorem ContinuousAt.eventually_ne [TopologicalSpace β] [T1Space β] {g : α → β} {a : α} {b : β} - (hg1 : ContinuousAt g a) (hg2 : g a ≠ b) : ∀ᶠ z in 𝓝 a, g z ≠ b := +theorem ContinuousAt.eventually_ne [TopologicalSpace Y] [T1Space Y] {g : X → Y} {x : X} {y : Y} + (hg1 : ContinuousAt g x) (hg2 : g x ≠ y) : ∀ᶠ z in 𝓝 x, g z ≠ y := hg1.tendsto.eventually_ne hg2 #align continuous_at.eventually_ne ContinuousAt.eventually_ne -theorem eventually_ne_nhds [T1Space α] {a b : α} (h : a ≠ b) : ∀ᶠ x in 𝓝 a, x ≠ b := +theorem eventually_ne_nhds [T1Space X] {a b : X} (h : a ≠ b) : ∀ᶠ x in 𝓝 a, x ≠ b := IsOpen.eventually_mem isOpen_ne h -theorem eventually_ne_nhdsWithin [T1Space α] {a b : α} {s : Set α} (h : a ≠ b) : +theorem eventually_ne_nhdsWithin [T1Space X] {a b : X} {s : Set X} (h : a ≠ b) : ∀ᶠ x in 𝓝[s] a, x ≠ b := Filter.Eventually.filter_mono nhdsWithin_le_nhds <| eventually_ne_nhds h -/-- To prove a function to a `T1Space` is continuous at some point `a`, it suffices to prove that -`f` admits *some* limit at `a`. -/ -theorem continuousAt_of_tendsto_nhds [TopologicalSpace β] [T1Space β] {f : α → β} {a : α} {b : β} - (h : Tendsto f (𝓝 a) (𝓝 b)) : ContinuousAt f a := by +/-- To prove a function to a `T1Space` is continuous at some point `x`, it suffices to prove that +`f` admits *some* limit at `x`. -/ +theorem continuousAt_of_tendsto_nhds [TopologicalSpace Y] [T1Space Y] {f : X → Y} {x : X} {y : Y} + (h : Tendsto f (𝓝 x) (𝓝 y)) : ContinuousAt f x := by rwa [ContinuousAt, eq_of_tendsto_nhds h] #align continuous_at_of_tendsto_nhds continuousAt_of_tendsto_nhds @[simp] -theorem tendsto_const_nhds_iff [T1Space α] {l : Filter β} [NeBot l] {c d : α} : +theorem tendsto_const_nhds_iff [T1Space X] {l : Filter Y} [NeBot l] {c d : X} : Tendsto (fun _ => c) l (𝓝 d) ↔ c = d := by simp_rw [Tendsto, Filter.map_const, pure_le_nhds_iff] #align tendsto_const_nhds_iff tendsto_const_nhds_iff /-- A point with a finite neighborhood has to be isolated. -/ -theorem isOpen_singleton_of_finite_mem_nhds {α : Type*} [TopologicalSpace α] [T1Space α] (x : α) - {s : Set α} (hs : s ∈ 𝓝 x) (hsf : s.Finite) : IsOpen ({x} : Set α) := by +theorem isOpen_singleton_of_finite_mem_nhds [T1Space X] (x : X) + {s : Set X} (hs : s ∈ 𝓝 x) (hsf : s.Finite) : IsOpen ({x} : Set X) := by have A : {x} ⊆ s := by simp only [singleton_subset_iff, mem_of_mem_nhds hs] have B : IsClosed (s \ {x}) := (hsf.subset (diff_subset _ _)).isClosed have C : (s \ {x})ᶜ ∈ 𝓝 x := B.isOpen_compl.mem_nhds fun h => h.2 rfl @@ -797,14 +800,14 @@ theorem isOpen_singleton_of_finite_mem_nhds {α : Type*} [TopologicalSpace α] [ /-- If the punctured neighborhoods of a point form a nontrivial filter, then any neighborhood is infinite. -/ -theorem infinite_of_mem_nhds {α} [TopologicalSpace α] [T1Space α] (x : α) [hx : NeBot (𝓝[≠] x)] - {s : Set α} (hs : s ∈ 𝓝 x) : Set.Infinite s := by +theorem infinite_of_mem_nhds {X} [TopologicalSpace X] [T1Space X] (x : X) [hx : NeBot (𝓝[≠] x)] + {s : Set X} (hs : s ∈ 𝓝 x) : Set.Infinite s := by refine' fun hsf => hx.1 _ rw [← isOpen_singleton_iff_punctured_nhds] exact isOpen_singleton_of_finite_mem_nhds x hs hsf #align infinite_of_mem_nhds infinite_of_mem_nhds -theorem discrete_of_t1_of_finite {X : Type*} [TopologicalSpace X] [T1Space X] [Finite X] : +theorem discrete_of_t1_of_finite [T1Space X] [Finite X] : DiscreteTopology X := by apply singletons_open_iff_discrete.mp intro x @@ -812,28 +815,28 @@ theorem discrete_of_t1_of_finite {X : Type*} [TopologicalSpace X] [T1Space X] [F exact (Set.toFinite _).isClosed #align discrete_of_t1_of_finite discrete_of_t1_of_finite -theorem PreconnectedSpace.trivial_of_discrete [PreconnectedSpace α] [DiscreteTopology α] : - Subsingleton α := by +theorem PreconnectedSpace.trivial_of_discrete [PreconnectedSpace X] [DiscreteTopology X] : + Subsingleton X := by rw [← not_nontrivial_iff_subsingleton] rintro ⟨x, y, hxy⟩ rw [Ne.def, ← mem_singleton_iff, (isClopen_discrete _).eq_univ <| singleton_nonempty y] at hxy exact hxy (mem_univ x) #align preconnected_space.trivial_of_discrete PreconnectedSpace.trivial_of_discrete -theorem IsPreconnected.infinite_of_nontrivial [T1Space α] {s : Set α} (h : IsPreconnected s) +theorem IsPreconnected.infinite_of_nontrivial [T1Space X] {s : Set X} (h : IsPreconnected s) (hs : s.Nontrivial) : s.Infinite := by refine' mt (fun hf => (subsingleton_coe s).mp _) (not_subsingleton_iff.mpr hs) haveI := @discrete_of_t1_of_finite s _ _ hf.to_subtype exact @PreconnectedSpace.trivial_of_discrete _ _ (Subtype.preconnectedSpace h) _ #align is_preconnected.infinite_of_nontrivial IsPreconnected.infinite_of_nontrivial -theorem ConnectedSpace.infinite [ConnectedSpace α] [Nontrivial α] [T1Space α] : Infinite α := +theorem ConnectedSpace.infinite [ConnectedSpace X] [Nontrivial X] [T1Space X] : Infinite X := infinite_univ_iff.mp <| isPreconnected_univ.infinite_of_nontrivial nontrivial_univ #align connected_space.infinite ConnectedSpace.infinite /-- A non-trivial connected T1 space has no isolated points. -/ instance ConnectedSpace.neBot_nhdsWithin_compl_of_nontrivial_of_t1space - [ConnectedSpace α] [Nontrivial α] [T1Space α] (x : α) : + [ConnectedSpace X] [Nontrivial X] [T1Space X] (x : X) : NeBot (𝓝[≠] x) := by by_contra contra rw [not_neBot, ← isOpen_singleton_iff_punctured_nhds] at contra @@ -841,22 +844,22 @@ instance ConnectedSpace.neBot_nhdsWithin_compl_of_nontrivial_of_t1space contra (compl_union_self _) (Set.nonempty_compl_of_nontrivial _) (singleton_nonempty _) simp [compl_inter_self {x}] at contra -theorem singleton_mem_nhdsWithin_of_mem_discrete {s : Set α} [DiscreteTopology s] {x : α} +theorem singleton_mem_nhdsWithin_of_mem_discrete {s : Set X} [DiscreteTopology s] {x : X} (hx : x ∈ s) : {x} ∈ 𝓝[s] x := by have : ({⟨x, hx⟩} : Set s) ∈ 𝓝 (⟨x, hx⟩ : s) := by simp [nhds_discrete] simpa only [nhdsWithin_eq_map_subtype_coe hx, image_singleton] using - @image_mem_map _ _ _ ((↑) : s → α) _ this + @image_mem_map _ _ _ ((↑) : s → X) _ this #align singleton_mem_nhds_within_of_mem_discrete singleton_mem_nhdsWithin_of_mem_discrete /-- The neighbourhoods filter of `x` within `s`, under the discrete topology, is equal to the pure `x` filter (which is the principal filter at the singleton `{x}`.) -/ -theorem nhdsWithin_of_mem_discrete {s : Set α} [DiscreteTopology s] {x : α} (hx : x ∈ s) : +theorem nhdsWithin_of_mem_discrete {s : Set X} [DiscreteTopology s] {x : X} (hx : x ∈ s) : 𝓝[s] x = pure x := le_antisymm (le_pure_iff.2 <| singleton_mem_nhdsWithin_of_mem_discrete hx) (pure_le_nhdsWithin hx) #align nhds_within_of_mem_discrete nhdsWithin_of_mem_discrete theorem Filter.HasBasis.exists_inter_eq_singleton_of_mem_discrete {ι : Type*} {p : ι → Prop} - {t : ι → Set α} {s : Set α} [DiscreteTopology s] {x : α} (hb : (𝓝 x).HasBasis p t) + {t : ι → Set X} {s : Set X} [DiscreteTopology s] {x : X} (hb : (𝓝 x).HasBasis p t) (hx : x ∈ s) : ∃ i, p i ∧ t i ∩ s = {x} := by rcases (nhdsWithin_hasBasis hb s).mem_iff.1 (singleton_mem_nhdsWithin_of_mem_discrete hx) with ⟨i, hi, hix⟩ @@ -865,7 +868,7 @@ theorem Filter.HasBasis.exists_inter_eq_singleton_of_mem_discrete {ι : Type*} { /-- A point `x` in a discrete subset `s` of a topological space admits a neighbourhood that only meets `s` at `x`. -/ -theorem nhds_inter_eq_singleton_of_mem_discrete {s : Set α} [DiscreteTopology s] {x : α} +theorem nhds_inter_eq_singleton_of_mem_discrete {s : Set X} [DiscreteTopology s] {x : X} (hx : x ∈ s) : ∃ U ∈ 𝓝 x, U ∩ s = {x} := by simpa using (𝓝 x).basis_sets.exists_inter_eq_singleton_of_mem_discrete hx #align nhds_inter_eq_singleton_of_mem_discrete nhds_inter_eq_singleton_of_mem_discrete @@ -875,7 +878,7 @@ such that 1. `U` is a punctured neighborhood of `x` (ie. `U ∪ {x}` is a neighbourhood of `x`), 2. `U` is disjoint from `s`. -/ -theorem disjoint_nhdsWithin_of_mem_discrete {s : Set α} [DiscreteTopology s] {x : α} (hx : x ∈ s) : +theorem disjoint_nhdsWithin_of_mem_discrete {s : Set X} [DiscreteTopology s] {x : X} (hx : x ∈ s) : ∃ U ∈ 𝓝[≠] x, Disjoint U s := let ⟨V, h, h'⟩ := nhds_inter_eq_singleton_of_mem_discrete hx ⟨{x}ᶜ ∩ V, inter_mem_nhdsWithin _ h, @@ -886,8 +889,7 @@ theorem disjoint_nhdsWithin_of_mem_discrete {s : Set α} [DiscreteTopology s] {x `t ⊆ s`, then the topological space structure on `t` induced by `X` is the same as the one obtained by the induced topological space structure on `s`. Use `embedding_inclusion` instead. -/ @[deprecated embedding_inclusion] -theorem TopologicalSpace.subset_trans {X : Type*} [TopologicalSpace X] {s t : Set X} - (ts : t ⊆ s) : +theorem TopologicalSpace.subset_trans {s t : Set X} (ts : t ⊆ s) : (instTopologicalSpaceSubtype : TopologicalSpace t) = (instTopologicalSpaceSubtype : TopologicalSpace s).induced (Set.inclusion ts) := (embedding_inclusion ts).induced @@ -897,44 +899,44 @@ theorem TopologicalSpace.subset_trans {X : Type*} [TopologicalSpace X] {s t : Se `x ≠ y` there exists disjoint open sets around `x` and `y`. This is the most widely used of the separation axioms. -/ @[mk_iff t2Space_iff] -class T2Space (α : Type u) [TopologicalSpace α] : Prop where +class T2Space (X : Type u) [TopologicalSpace X] : Prop where /-- Every two points in a Hausdorff space admit disjoint open neighbourhoods. -/ - t2 : ∀ x y, x ≠ y → ∃ u v : Set α, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v + t2 : ∀ x y, x ≠ y → ∃ u v : Set X, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v #align t2_space T2Space /-- Two different points can be separated by open sets. -/ -theorem t2_separation [T2Space α] {x y : α} (h : x ≠ y) : - ∃ u v : Set α, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v := +theorem t2_separation [T2Space X] {x y : X} (h : x ≠ y) : + ∃ u v : Set X, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v := T2Space.t2 x y h #align t2_separation t2_separation -- todo: use this as a definition? -theorem t2Space_iff_disjoint_nhds : T2Space α ↔ ∀ x y : α, x ≠ y → Disjoint (𝓝 x) (𝓝 y) := by - refine (t2Space_iff α).trans (forall₃_congr fun x y _ => ?_) +theorem t2Space_iff_disjoint_nhds : T2Space X ↔ ∀ x y : X, x ≠ y → Disjoint (𝓝 x) (𝓝 y) := by + refine (t2Space_iff X).trans (forall₃_congr fun x y _ => ?_) simp only [(nhds_basis_opens x).disjoint_iff (nhds_basis_opens y), exists_prop, ← exists_and_left, and_assoc, and_comm, and_left_comm] #align t2_space_iff_disjoint_nhds t2Space_iff_disjoint_nhds @[simp] -theorem disjoint_nhds_nhds [T2Space α] {x y : α} : Disjoint (𝓝 x) (𝓝 y) ↔ x ≠ y := +theorem disjoint_nhds_nhds [T2Space X] {x y : X} : Disjoint (𝓝 x) (𝓝 y) ↔ x ≠ y := ⟨fun hd he => by simp [he, nhds_neBot.ne] at hd, t2Space_iff_disjoint_nhds.mp ‹_› x y⟩ #align disjoint_nhds_nhds disjoint_nhds_nhds -theorem pairwise_disjoint_nhds [T2Space α] : Pairwise (Disjoint on (𝓝 : α → Filter α)) := fun _ _ => +theorem pairwise_disjoint_nhds [T2Space X] : Pairwise (Disjoint on (𝓝 : X → Filter X)) := fun _ _ => disjoint_nhds_nhds.2 #align pairwise_disjoint_nhds pairwise_disjoint_nhds -protected theorem Set.pairwiseDisjoint_nhds [T2Space α] (s : Set α) : s.PairwiseDisjoint 𝓝 := +protected theorem Set.pairwiseDisjoint_nhds [T2Space X] (s : Set X) : s.PairwiseDisjoint 𝓝 := pairwise_disjoint_nhds.set_pairwise s #align set.pairwise_disjoint_nhds Set.pairwiseDisjoint_nhds /-- Points of a finite set can be separated by open sets from each other. -/ -theorem Set.Finite.t2_separation [T2Space α] {s : Set α} (hs : s.Finite) : - ∃ U : α → Set α, (∀ x, x ∈ U x ∧ IsOpen (U x)) ∧ s.PairwiseDisjoint U := +theorem Set.Finite.t2_separation [T2Space X] {s : Set X} (hs : s.Finite) : + ∃ U : X → Set X, (∀ x, x ∈ U x ∧ IsOpen (U x)) ∧ s.PairwiseDisjoint U := s.pairwiseDisjoint_nhds.exists_mem_filter_basis hs nhds_basis_opens #align set.finite.t2_separation Set.Finite.t2_separation -theorem isOpen_setOf_disjoint_nhds_nhds : IsOpen { p : α × α | Disjoint (𝓝 p.1) (𝓝 p.2) } := by +theorem isOpen_setOf_disjoint_nhds_nhds : IsOpen { p : X × X | Disjoint (𝓝 p.1) (𝓝 p.2) } := by simp only [isOpen_iff_mem_nhds, Prod.forall, mem_setOf_eq] intro x y h obtain ⟨U, hU, V, hV, hd⟩ := ((nhds_basis_opens x).disjoint_iff (nhds_basis_opens y)).mp h @@ -945,103 +947,139 @@ theorem isOpen_setOf_disjoint_nhds_nhds : IsOpen { p : α × α | Disjoint (𝓝 #align is_open_set_of_disjoint_nhds_nhds isOpen_setOf_disjoint_nhds_nhds -- see Note [lower instance priority] -instance (priority := 100) T2Space.t1Space [T2Space α] : T1Space α := +instance (priority := 100) T2Space.t1Space [T2Space X] : T1Space X := t1Space_iff_disjoint_pure_nhds.mpr fun _ _ hne => (disjoint_nhds_nhds.2 hne).mono_left <| pure_le_nhds _ #align t2_space.t1_space T2Space.t1Space /-- A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter. -/ -theorem t2_iff_nhds : T2Space α ↔ ∀ {x y : α}, NeBot (𝓝 x ⊓ 𝓝 y) → x = y := by +theorem t2_iff_nhds : T2Space X ↔ ∀ {x y : X}, NeBot (𝓝 x ⊓ 𝓝 y) → x = y := by simp only [t2Space_iff_disjoint_nhds, disjoint_iff, neBot_iff, Ne.def, not_imp_comm] #align t2_iff_nhds t2_iff_nhds -theorem eq_of_nhds_neBot [T2Space α] {x y : α} (h : NeBot (𝓝 x ⊓ 𝓝 y)) : x = y := +theorem eq_of_nhds_neBot [T2Space X] {x y : X} (h : NeBot (𝓝 x ⊓ 𝓝 y)) : x = y := t2_iff_nhds.mp ‹_› h #align eq_of_nhds_ne_bot eq_of_nhds_neBot -theorem t2Space_iff_nhds : T2Space α ↔ ∀ {x y : α}, x ≠ y → ∃ U ∈ 𝓝 x, ∃ V ∈ 𝓝 y, Disjoint U V := by +theorem t2Space_iff_nhds : T2Space X ↔ ∀ {x y : X}, x ≠ y → ∃ U ∈ 𝓝 x, ∃ V ∈ 𝓝 y, Disjoint U V := by simp only [t2Space_iff_disjoint_nhds, Filter.disjoint_iff] #align t2_space_iff_nhds t2Space_iff_nhds -theorem t2_separation_nhds [T2Space α] {x y : α} (h : x ≠ y) : +theorem t2_separation_nhds [T2Space X] {x y : X} (h : x ≠ y) : ∃ u v, u ∈ 𝓝 x ∧ v ∈ 𝓝 y ∧ Disjoint u v := let ⟨u, v, open_u, open_v, x_in, y_in, huv⟩ := t2_separation h ⟨u, v, open_u.mem_nhds x_in, open_v.mem_nhds y_in, huv⟩ #align t2_separation_nhds t2_separation_nhds -theorem t2_separation_compact_nhds [LocallyCompactSpace α] [T2Space α] {x y : α} (h : x ≠ y) : +theorem t2_separation_compact_nhds [LocallyCompactSpace X] [T2Space X] {x y : X} (h : x ≠ y) : ∃ u v, u ∈ 𝓝 x ∧ v ∈ 𝓝 y ∧ IsCompact u ∧ IsCompact v ∧ Disjoint u v := by simpa only [exists_prop, ← exists_and_left, and_comm, and_assoc, and_left_comm] using ((compact_basis_nhds x).disjoint_iff (compact_basis_nhds y)).1 (disjoint_nhds_nhds.2 h) #align t2_separation_compact_nhds t2_separation_compact_nhds theorem t2_iff_ultrafilter : - T2Space α ↔ ∀ {x y : α} (f : Ultrafilter α), ↑f ≤ 𝓝 x → ↑f ≤ 𝓝 y → x = y := + T2Space X ↔ ∀ {x y : X} (f : Ultrafilter X), ↑f ≤ 𝓝 x → ↑f ≤ 𝓝 y → x = y := t2_iff_nhds.trans <| by simp only [← exists_ultrafilter_iff, and_imp, le_inf_iff, exists_imp] #align t2_iff_ultrafilter t2_iff_ultrafilter -theorem t2_iff_isClosed_diagonal : T2Space α ↔ IsClosed (diagonal α) := by +theorem t2_iff_isClosed_diagonal : T2Space X ↔ IsClosed (diagonal X) := by simp only [t2Space_iff_disjoint_nhds, ← isOpen_compl_iff, isOpen_iff_mem_nhds, Prod.forall, nhds_prod_eq, compl_diagonal_mem_prod, mem_compl_iff, mem_diagonal_iff] #align t2_iff_is_closed_diagonal t2_iff_isClosed_diagonal -theorem isClosed_diagonal [T2Space α] : IsClosed (diagonal α) := +theorem isClosed_diagonal [T2Space X] : IsClosed (diagonal X) := t2_iff_isClosed_diagonal.mp ‹_› #align is_closed_diagonal isClosed_diagonal -- porting note: 2 lemmas moved below -theorem tendsto_nhds_unique [T2Space α] {f : β → α} {l : Filter β} {a b : α} [NeBot l] +theorem tendsto_nhds_unique [T2Space X] {f : Y → X} {l : Filter Y} {a b : X} [NeBot l] (ha : Tendsto f l (𝓝 a)) (hb : Tendsto f l (𝓝 b)) : a = b := eq_of_nhds_neBot <| neBot_of_le <| le_inf ha hb #align tendsto_nhds_unique tendsto_nhds_unique -theorem tendsto_nhds_unique' [T2Space α] {f : β → α} {l : Filter β} {a b : α} (_ : NeBot l) +theorem tendsto_nhds_unique' [T2Space X] {f : Y → X} {l : Filter Y} {a b : X} (_ : NeBot l) (ha : Tendsto f l (𝓝 a)) (hb : Tendsto f l (𝓝 b)) : a = b := eq_of_nhds_neBot <| neBot_of_le <| le_inf ha hb #align tendsto_nhds_unique' tendsto_nhds_unique' -theorem tendsto_nhds_unique_of_eventuallyEq [T2Space α] {f g : β → α} {l : Filter β} {a b : α} +theorem tendsto_nhds_unique_of_eventuallyEq [T2Space X] {f g : Y → X} {l : Filter Y} {a b : X} [NeBot l] (ha : Tendsto f l (𝓝 a)) (hb : Tendsto g l (𝓝 b)) (hfg : f =ᶠ[l] g) : a = b := tendsto_nhds_unique (ha.congr' hfg) hb #align tendsto_nhds_unique_of_eventually_eq tendsto_nhds_unique_of_eventuallyEq -theorem tendsto_nhds_unique_of_frequently_eq [T2Space α] {f g : β → α} {l : Filter β} {a b : α} +theorem tendsto_nhds_unique_of_frequently_eq [T2Space X] {f g : Y → X} {l : Filter Y} {a b : X} (ha : Tendsto f l (𝓝 a)) (hb : Tendsto g l (𝓝 b)) (hfg : ∃ᶠ x in l, f x = g x) : a = b := - have : ∃ᶠ z : α × α in 𝓝 (a, b), z.1 = z.2 := (ha.prod_mk_nhds hb).frequently hfg + have : ∃ᶠ z : X × X in 𝓝 (a, b), z.1 = z.2 := (ha.prod_mk_nhds hb).frequently hfg not_not.1 fun hne => this (isClosed_diagonal.isOpen_compl.mem_nhds hne) #align tendsto_nhds_unique_of_frequently_eq tendsto_nhds_unique_of_frequently_eq +/-- If a function `f` is + +- injective on a compact set `s`; +- continuous at every point of this set; +- injective on a neighborhood of each point of this set, + +then it is injective on a neighborhood of this set. -/ +theorem Set.InjOn.exists_mem_nhdsSet {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] + [T2Space Y] {f : X → Y} {s : Set X} (inj : InjOn f s) (sc : IsCompact s) + (fc : ∀ x ∈ s, ContinuousAt f x) (loc : ∀ x ∈ s, ∃ u ∈ 𝓝 x, InjOn f u) : + ∃ t ∈ 𝓝ˢ s, InjOn f t := by + have : ∀ x ∈ s ×ˢ s, ∀ᶠ y in 𝓝 x, f y.1 = f y.2 → y.1 = y.2 := fun (x, y) ⟨hx, hy⟩ ↦ by + rcases eq_or_ne x y with rfl | hne + · rcases loc x hx with ⟨u, hu, hf⟩ + exact Filter.mem_of_superset (prod_mem_nhds hu hu) <| forall_prod_set.2 hf + · suffices ∀ᶠ z in 𝓝 (x, y), f z.1 ≠ f z.2 from this.mono fun _ hne h ↦ absurd h hne + refine (fc x hx).prod_map' (fc y hy) <| isClosed_diagonal.isOpen_compl.mem_nhds ?_ + exact inj.ne hx hy hne + rw [← eventually_nhdsSet_iff_forall, sc.nhdsSet_prod_eq sc] at this + exact eventually_prod_self_iff.1 this + +/-- If a function `f` is + +- injective on a compact set `s`; +- continuous at every point of this set; +- injective on a neighborhood of each point of this set, + +then it is injective on an open neighborhood of this set. -/ +theorem Set.InjOn.exists_isOpen_superset {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] + [T2Space Y] {f : X → Y} {s : Set X} (inj : InjOn f s) (sc : IsCompact s) + (fc : ∀ x ∈ s, ContinuousAt f x) (loc : ∀ x ∈ s, ∃ u ∈ 𝓝 x, InjOn f u) : + ∃ t, IsOpen t ∧ s ⊆ t ∧ InjOn f t := + let ⟨_t, hst, ht⟩ := inj.exists_mem_nhdsSet sc fc loc + let ⟨u, huo, hsu, hut⟩ := mem_nhdsSet_iff_exists.1 hst + ⟨u, huo, hsu, ht.mono hut⟩ + /-- A T₂.₅ space, also known as a Urysohn space, is a topological space where for every pair `x ≠ y`, there are two open sets, with the intersection of closures empty, one containing `x` and the other `y` . -/ -class T25Space (α : Type u) [TopologicalSpace α] : Prop where +class T25Space (X : Type u) [TopologicalSpace X] : Prop where /-- Given two distinct points in a T₂.₅ space, their filters of closed neighborhoods are disjoint. -/ - t2_5 : ∀ ⦃x y : α⦄, x ≠ y → Disjoint ((𝓝 x).lift' closure) ((𝓝 y).lift' closure) + t2_5 : ∀ ⦃x y : X⦄, x ≠ y → Disjoint ((𝓝 x).lift' closure) ((𝓝 y).lift' closure) #align t2_5_space T25Space @[simp] -theorem disjoint_lift'_closure_nhds [T25Space α] {x y : α} : +theorem disjoint_lift'_closure_nhds [T25Space X] {x y : X} : Disjoint ((𝓝 x).lift' closure) ((𝓝 y).lift' closure) ↔ x ≠ y := ⟨fun h hxy => by simp [hxy, nhds_neBot.ne] at h, fun h => T25Space.t2_5 h⟩ #align disjoint_lift'_closure_nhds disjoint_lift'_closure_nhds -- see Note [lower instance priority] -instance (priority := 100) T25Space.t2Space [T25Space α] : T2Space α := +instance (priority := 100) T25Space.t2Space [T25Space X] : T2Space X := t2Space_iff_disjoint_nhds.2 fun _ _ hne => (disjoint_lift'_closure_nhds.2 hne).mono (le_lift'_closure _) (le_lift'_closure _) #align t2_5_space.t2_space T25Space.t2Space -theorem exists_nhds_disjoint_closure [T25Space α] {x y : α} (h : x ≠ y) : +theorem exists_nhds_disjoint_closure [T25Space X] {x y : X} (h : x ≠ y) : ∃ s ∈ 𝓝 x, ∃ t ∈ 𝓝 y, Disjoint (closure s) (closure t) := ((𝓝 x).basis_sets.lift'_closure.disjoint_iff (𝓝 y).basis_sets.lift'_closure).1 <| disjoint_lift'_closure_nhds.2 h #align exists_nhds_disjoint_closure exists_nhds_disjoint_closure -theorem exists_open_nhds_disjoint_closure [T25Space α] {x y : α} (h : x ≠ y) : - ∃ u : Set α, - x ∈ u ∧ IsOpen u ∧ ∃ v : Set α, y ∈ v ∧ IsOpen v ∧ Disjoint (closure u) (closure v) := by +theorem exists_open_nhds_disjoint_closure [T25Space X] {x y : X} (h : x ≠ y) : + ∃ u : Set X, + x ∈ u ∧ IsOpen u ∧ ∃ v : Set X, y ∈ v ∧ IsOpen v ∧ Disjoint (closure u) (closure v) := by simpa only [exists_prop, and_assoc] using ((nhds_basis_opens x).lift'_closure.disjoint_iff (nhds_basis_opens y).lift'_closure).1 (disjoint_lift'_closure_nhds.2 h) @@ -1049,34 +1087,34 @@ theorem exists_open_nhds_disjoint_closure [T25Space α] {x y : α} (h : x ≠ y) section limUnder -variable [T2Space α] {f : Filter α} +variable [T2Space X] {f : Filter X} /-! ### Properties of `lim` and `limUnder` -In this section we use explicit `Nonempty α` instances for `lim` and `limUnder`. This way the lemmas -are useful without a `Nonempty α` instance. +In this section we use explicit `Nonempty X` instances for `lim` and `limUnder`. This way the lemmas +are useful without a `Nonempty X` instance. -/ -theorem lim_eq {a : α} [NeBot f] (h : f ≤ 𝓝 a) : @lim _ _ ⟨a⟩ f = a := - tendsto_nhds_unique (le_nhds_lim ⟨a, h⟩) h +theorem lim_eq {x : X} [NeBot f] (h : f ≤ 𝓝 x) : @lim _ _ ⟨x⟩ f = x := + tendsto_nhds_unique (le_nhds_lim ⟨x, h⟩) h set_option linter.uppercaseLean3 false in #align Lim_eq lim_eq -theorem lim_eq_iff [NeBot f] (h : ∃ a : α, f ≤ nhds a) {a} : @lim _ _ ⟨a⟩ f = a ↔ f ≤ 𝓝 a := +theorem lim_eq_iff [NeBot f] (h : ∃ x : X, f ≤ 𝓝 x) {x} : @lim _ _ ⟨x⟩ f = x ↔ f ≤ 𝓝 x := ⟨fun c => c ▸ le_nhds_lim h, lim_eq⟩ set_option linter.uppercaseLean3 false in #align Lim_eq_iff lim_eq_iff -theorem Ultrafilter.lim_eq_iff_le_nhds [CompactSpace α] {x : α} {F : Ultrafilter α} : +theorem Ultrafilter.lim_eq_iff_le_nhds [CompactSpace X] {x : X} {F : Ultrafilter X} : F.lim = x ↔ ↑F ≤ 𝓝 x := ⟨fun h => h ▸ F.le_nhds_lim, lim_eq⟩ set_option linter.uppercaseLean3 false in #align ultrafilter.Lim_eq_iff_le_nhds Ultrafilter.lim_eq_iff_le_nhds -theorem isOpen_iff_ultrafilter' [CompactSpace α] (U : Set α) : - IsOpen U ↔ ∀ F : Ultrafilter α, F.lim ∈ U → U ∈ F.1 := by +theorem isOpen_iff_ultrafilter' [CompactSpace X] (U : Set X) : + IsOpen U ↔ ∀ F : Ultrafilter X, F.lim ∈ U → U ∈ F.1 := by rw [isOpen_iff_ultrafilter] refine' ⟨fun h F hF => h F.lim hF F F.le_nhds_lim, _⟩ intro cond x hx f h @@ -1084,42 +1122,42 @@ theorem isOpen_iff_ultrafilter' [CompactSpace α] (U : Set α) : exact cond _ hx #align is_open_iff_ultrafilter' isOpen_iff_ultrafilter' -theorem Filter.Tendsto.limUnder_eq {a : α} {f : Filter β} [NeBot f] {g : β → α} - (h : Tendsto g f (𝓝 a)) : @limUnder _ _ _ ⟨a⟩ f g = a := +theorem Filter.Tendsto.limUnder_eq {x : X} {f : Filter Y} [NeBot f] {g : Y → X} + (h : Tendsto g f (𝓝 x)) : @limUnder _ _ _ ⟨x⟩ f g = x := lim_eq h #align filter.tendsto.lim_eq Filter.Tendsto.limUnder_eq -theorem Filter.limUnder_eq_iff {f : Filter β} [NeBot f] {g : β → α} (h : ∃ a, Tendsto g f (𝓝 a)) - {a} : @limUnder _ _ _ ⟨a⟩ f g = a ↔ Tendsto g f (𝓝 a) := +theorem Filter.limUnder_eq_iff {f : Filter Y} [NeBot f] {g : Y → X} (h : ∃ x, Tendsto g f (𝓝 x)) + {x} : @limUnder _ _ _ ⟨x⟩ f g = x ↔ Tendsto g f (𝓝 x) := ⟨fun c => c ▸ tendsto_nhds_limUnder h, Filter.Tendsto.limUnder_eq⟩ #align filter.lim_eq_iff Filter.limUnder_eq_iff -theorem Continuous.limUnder_eq [TopologicalSpace β] {f : β → α} (h : Continuous f) (a : β) : - @limUnder _ _ _ ⟨f a⟩ (𝓝 a) f = f a := - (h.tendsto a).limUnder_eq +theorem Continuous.limUnder_eq [TopologicalSpace Y] {f : Y → X} (h : Continuous f) (y : Y) : + @limUnder _ _ _ ⟨f y⟩ (𝓝 y) f = f y := + (h.tendsto y).limUnder_eq #align continuous.lim_eq Continuous.limUnder_eq @[simp] -theorem lim_nhds (a : α) : @lim _ _ ⟨a⟩ (𝓝 a) = a := +theorem lim_nhds (x : X) : @lim _ _ ⟨x⟩ (𝓝 x) = x := lim_eq le_rfl set_option linter.uppercaseLean3 false in #align Lim_nhds lim_nhds @[simp] -theorem limUnder_nhds_id (a : α) : @limUnder _ _ _ ⟨a⟩ (𝓝 a) id = a := - lim_nhds a +theorem limUnder_nhds_id (x : X) : @limUnder _ _ _ ⟨x⟩ (𝓝 x) id = x := + lim_nhds x #align lim_nhds_id limUnder_nhds_id @[simp] -theorem lim_nhdsWithin {a : α} {s : Set α} (h : a ∈ closure s) : @lim _ _ ⟨a⟩ (𝓝[s] a) = a := - haveI : NeBot (𝓝[s] a) := mem_closure_iff_clusterPt.1 h +theorem lim_nhdsWithin {x : X} {s : Set X} (h : x ∈ closure s) : @lim _ _ ⟨x⟩ (𝓝[s] x) = x := + haveI : NeBot (𝓝[s] x) := mem_closure_iff_clusterPt.1 h lim_eq inf_le_left set_option linter.uppercaseLean3 false in #align Lim_nhds_within lim_nhdsWithin @[simp] -theorem limUnder_nhdsWithin_id {a : α} {s : Set α} (h : a ∈ closure s) : - @limUnder _ _ _ ⟨a⟩ (𝓝[s] a) id = a := +theorem limUnder_nhdsWithin_id {x : X} {s : Set X} (h : x ∈ closure s) : + @limUnder _ _ _ ⟨x⟩ (𝓝[s] x) id = x := lim_nhdsWithin h #align lim_nhds_within_id limUnder_nhdsWithin_id @@ -1131,62 +1169,61 @@ end limUnder We use two lemmas to prove that various standard constructions generate Hausdorff spaces from Hausdorff spaces: -* `separated_by_continuous` says that two points `x y : α` can be separated by open neighborhoods - provided that there exists a continuous map `f : α → β` with a Hausdorff codomain such that +* `separated_by_continuous` says that two points `x y : X` can be separated by open neighborhoods + provided that there exists a continuous map `f : X → Y` with a Hausdorff codomain such that `f x ≠ f y`. We use this lemma to prove that topological spaces defined using `induced` are Hausdorff spaces. -* `separated_by_openEmbedding` says that for an open embedding `f : α → β` of a Hausdorff space - `α`, the images of two distinct points `x y : α`, `x ≠ y` can be separated by open neighborhoods. +* `separated_by_openEmbedding` says that for an open embedding `f : X → Y` of a Hausdorff space + `X`, the images of two distinct points `x y : X`, `x ≠ y` can be separated by open neighborhoods. We use this lemma to prove that topological spaces defined using `coinduced` are Hausdorff spaces. -/ -- see Note [lower instance priority] -instance (priority := 100) DiscreteTopology.toT2Space {α : Type*} [TopologicalSpace α] - [DiscreteTopology α] : T2Space α := +instance (priority := 100) DiscreteTopology.toT2Space + [DiscreteTopology X] : T2Space X := ⟨fun x y h => ⟨{x}, {y}, isOpen_discrete _, isOpen_discrete _, rfl, rfl, disjoint_singleton.2 h⟩⟩ #align discrete_topology.to_t2_space DiscreteTopology.toT2Space -theorem separated_by_continuous {α : Type*} {β : Type*} [TopologicalSpace α] [TopologicalSpace β] - [T2Space β] {f : α → β} (hf : Continuous f) {x y : α} (h : f x ≠ f y) : - ∃ u v : Set α, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v := +theorem separated_by_continuous [TopologicalSpace Y] [T2Space Y] + {f : X → Y} (hf : Continuous f) {x y : X} (h : f x ≠ f y) : + ∃ u v : Set X, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v := let ⟨u, v, uo, vo, xu, yv, uv⟩ := t2_separation h ⟨f ⁻¹' u, f ⁻¹' v, uo.preimage hf, vo.preimage hf, xu, yv, uv.preimage _⟩ #align separated_by_continuous separated_by_continuous -theorem separated_by_openEmbedding {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] - [T2Space α] {f : α → β} (hf : OpenEmbedding f) {x y : α} (h : x ≠ y) : - ∃ u v : Set β, IsOpen u ∧ IsOpen v ∧ f x ∈ u ∧ f y ∈ v ∧ Disjoint u v := +theorem separated_by_openEmbedding [TopologicalSpace Y] [T2Space X] + {f : X → Y} (hf : OpenEmbedding f) {x y : X} (h : x ≠ y) : + ∃ u v : Set Y, IsOpen u ∧ IsOpen v ∧ f x ∈ u ∧ f y ∈ v ∧ Disjoint u v := let ⟨u, v, uo, vo, xu, yv, uv⟩ := t2_separation h ⟨f '' u, f '' v, hf.isOpenMap _ uo, hf.isOpenMap _ vo, mem_image_of_mem _ xu, mem_image_of_mem _ yv, disjoint_image_of_injective hf.inj uv⟩ #align separated_by_open_embedding separated_by_openEmbedding -instance {α : Type*} {p : α → Prop} [TopologicalSpace α] [T2Space α] : T2Space (Subtype p) := +instance {p : X → Prop} [T2Space X] : T2Space (Subtype p) := ⟨fun _ _ h => separated_by_continuous continuous_subtype_val (mt Subtype.eq h)⟩ -instance Prod.t2Space {α : Type*} {β : Type*} [TopologicalSpace α] [T2Space α] - [TopologicalSpace β] [T2Space β] : T2Space (α × β) := +instance Prod.t2Space [T2Space X] [TopologicalSpace Y] [T2Space Y] : T2Space (X × Y) := ⟨fun _ _ h => Or.elim (not_and_or.mp (mt Prod.ext_iff.mpr h)) (fun h₁ => separated_by_continuous continuous_fst h₁) fun h₂ => separated_by_continuous continuous_snd h₂⟩ /-- If the codomain of an injective continuous function is a Hausdorff space, then so is its domain. -/ -theorem T2Space.of_injective_continuous [TopologicalSpace β] [T2Space β] {f : α → β} - (hinj : Injective f) (hc : Continuous f) : T2Space α := +theorem T2Space.of_injective_continuous [TopologicalSpace Y] [T2Space Y] {f : X → Y} + (hinj : Injective f) (hc : Continuous f) : T2Space X := ⟨fun _ _ h => separated_by_continuous hc (hinj.ne h)⟩ /-- If the codomain of a topological embedding is a Hausdorff space, then so is its domain. See also `T2Space.of_continuous_injective`. -/ -theorem Embedding.t2Space [TopologicalSpace β] [T2Space β] {f : α → β} (hf : Embedding f) : - T2Space α := +theorem Embedding.t2Space [TopologicalSpace Y] [T2Space Y] {f : X → Y} (hf : Embedding f) : + T2Space X := .of_injective_continuous hf.inj hf.continuous #align embedding.t2_space Embedding.t2Space -instance {α β : Type*} [TopologicalSpace α] [T2Space α] [TopologicalSpace β] [T2Space β] : - T2Space (α ⊕ β) := by +instance [T2Space X] [TopologicalSpace Y] [T2Space Y] : + T2Space (X ⊕ Y) := by constructor rintro (x | x) (y | y) h · exact separated_by_openEmbedding openEmbedding_inl <| ne_of_apply_ne _ h @@ -1194,15 +1231,15 @@ instance {α β : Type*} [TopologicalSpace α] [T2Space α] [TopologicalSpace β · exact separated_by_continuous continuous_isLeft <| by simp · exact separated_by_openEmbedding openEmbedding_inr <| ne_of_apply_ne _ h -instance Pi.t2Space {α : Type*} {β : α → Type v} [∀ a, TopologicalSpace (β a)] - [∀ a, T2Space (β a)] : T2Space (∀ a, β a) := +instance Pi.t2Space {Y : X → Type v} [∀ a, TopologicalSpace (Y a)] + [∀ a, T2Space (Y a)] : T2Space (∀ a, Y a) := ⟨fun _ _ h => let ⟨i, hi⟩ := not_forall.mp (mt funext h) separated_by_continuous (continuous_apply i) hi⟩ #align Pi.t2_space Pi.t2Space -instance Sigma.t2Space {ι} {α : ι → Type*} [∀ i, TopologicalSpace (α i)] [∀ a, T2Space (α a)] : - T2Space (Σi, α i) := by +instance Sigma.t2Space {ι} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ a, T2Space (X a)] : + T2Space (Σi, X i) := by constructor rintro ⟨i, x⟩ ⟨j, y⟩ neq rcases eq_or_ne i j with (rfl | h) @@ -1212,32 +1249,32 @@ instance Sigma.t2Space {ι} {α : ι → Type*} [∀ i, TopologicalSpace (α i)] exact separated_by_continuous (continuous_def.2 fun u _ => isOpen_sigma_fst_preimage u) h #align sigma.t2_space Sigma.t2Space -variable {γ : Type*} [TopologicalSpace β] [TopologicalSpace γ] +variable {Z : Type*} [TopologicalSpace Y] [TopologicalSpace Z] -theorem isClosed_eq [T2Space α] {f g : β → α} (hf : Continuous f) (hg : Continuous g) : - IsClosed { x : β | f x = g x } := +theorem isClosed_eq [T2Space X] {f g : Y → X} (hf : Continuous f) (hg : Continuous g) : + IsClosed { y : Y | f y = g y } := continuous_iff_isClosed.mp (hf.prod_mk hg) _ isClosed_diagonal #align is_closed_eq isClosed_eq -theorem isOpen_ne_fun [T2Space α] {f g : β → α} (hf : Continuous f) (hg : Continuous g) : - IsOpen { x : β | f x ≠ g x } := +theorem isOpen_ne_fun [T2Space X] {f g : Y → X} (hf : Continuous f) (hg : Continuous g) : + IsOpen { y : Y | f y ≠ g y } := isOpen_compl_iff.mpr <| isClosed_eq hf hg #align is_open_ne_fun isOpen_ne_fun /-- If two continuous maps are equal on `s`, then they are equal on the closure of `s`. See also `Set.EqOn.of_subset_closure` for a more general version. -/ -protected theorem Set.EqOn.closure [T2Space α] {s : Set β} {f g : β → α} (h : EqOn f g s) +protected theorem Set.EqOn.closure [T2Space X] {s : Set Y} {f g : Y → X} (h : EqOn f g s) (hf : Continuous f) (hg : Continuous g) : EqOn f g (closure s) := closure_minimal h (isClosed_eq hf hg) #align set.eq_on.closure Set.EqOn.closure /-- If two continuous functions are equal on a dense set, then they are equal. -/ -theorem Continuous.ext_on [T2Space α] {s : Set β} (hs : Dense s) {f g : β → α} (hf : Continuous f) +theorem Continuous.ext_on [T2Space X] {s : Set Y} (hs : Dense s) {f g : Y → X} (hf : Continuous f) (hg : Continuous g) (h : EqOn f g s) : f = g := funext fun x => h.closure hf hg (hs x) #align continuous.ext_on Continuous.ext_on -theorem eqOn_closure₂' [T2Space α] {s : Set β} {t : Set γ} {f g : β → γ → α} +theorem eqOn_closure₂' [T2Space Z] {s : Set X} {t : Set Y} {f g : X → Y → Z} (h : ∀ x ∈ s, ∀ y ∈ t, f x y = g x y) (hf₁ : ∀ x, Continuous (f x)) (hf₂ : ∀ y, Continuous fun x => f x y) (hg₁ : ∀ x, Continuous (g x)) (hg₂ : ∀ y, Continuous fun x => g x y) : ∀ x ∈ closure s, ∀ y ∈ closure t, f x y = g x y := @@ -1246,7 +1283,7 @@ theorem eqOn_closure₂' [T2Space α] {s : Set β} {t : Set γ} {f g : β → γ isClosed_biInter fun y _ => isClosed_eq (hf₂ _) (hg₂ _) #align eq_on_closure₂' eqOn_closure₂' -theorem eqOn_closure₂ [T2Space α] {s : Set β} {t : Set γ} {f g : β → γ → α} +theorem eqOn_closure₂ [T2Space Z] {s : Set X} {t : Set Y} {f g : X → Y → Z} (h : ∀ x ∈ s, ∀ y ∈ t, f x y = g x y) (hf : Continuous (uncurry f)) (hg : Continuous (uncurry g)) : ∀ x ∈ closure s, ∀ y ∈ closure t, f x y = g x y := eqOn_closure₂' h (fun x => continuous_uncurry_left x hf) (fun x => continuous_uncurry_right x hf) @@ -1255,7 +1292,7 @@ theorem eqOn_closure₂ [T2Space α] {s : Set β} {t : Set γ} {f g : β → γ /-- If `f x = g x` for all `x ∈ s` and `f`, `g` are continuous on `t`, `s ⊆ t ⊆ closure s`, then `f x = g x` for all `x ∈ t`. See also `Set.EqOn.closure`. -/ -theorem Set.EqOn.of_subset_closure [T2Space α] {s t : Set β} {f g : β → α} (h : EqOn f g s) +theorem Set.EqOn.of_subset_closure [T2Space Y] {s t : Set X} {f g : X → Y} (h : EqOn f g s) (hf : ContinuousOn f t) (hg : ContinuousOn g t) (hst : s ⊆ t) (hts : t ⊆ closure s) : EqOn f g t := by intro x hx @@ -1265,19 +1302,19 @@ theorem Set.EqOn.of_subset_closure [T2Space α] {s t : Set β} {f g : β → α} ((hg x hx).mono_left <| nhdsWithin_mono _ hst) (h.eventuallyEq_of_mem self_mem_nhdsWithin) #align set.eq_on.of_subset_closure Set.EqOn.of_subset_closure -theorem Function.LeftInverse.closed_range [T2Space α] {f : α → β} {g : β → α} +theorem Function.LeftInverse.closed_range [T2Space X] {f : X → Y} {g : Y → X} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : IsClosed (range g) := have : EqOn (g ∘ f) id (closure <| range g) := h.rightInvOn_range.eqOn.closure (hg.comp hf) continuous_id isClosed_of_closure_subset fun x hx => ⟨f x, this hx⟩ #align function.left_inverse.closed_range Function.LeftInverse.closed_range -theorem Function.LeftInverse.closedEmbedding [T2Space α] {f : α → β} {g : β → α} +theorem Function.LeftInverse.closedEmbedding [T2Space X] {f : X → Y} {g : Y → X} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : ClosedEmbedding g := ⟨h.embedding hf hg, h.closed_range hf hg⟩ #align function.left_inverse.closed_embedding Function.LeftInverse.closedEmbedding -theorem isCompact_isCompact_separated [T2Space α] {s t : Set α} (hs : IsCompact s) +theorem isCompact_isCompact_separated [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) (hst : Disjoint s t) : SeparatedNhds s t := by simp only [SeparatedNhds, prod_subset_compl_diagonal_iff_disjoint.symm] at hst ⊢ exact generalized_tube_lemma hs ht isClosed_diagonal.isOpen_compl hst @@ -1285,21 +1322,21 @@ theorem isCompact_isCompact_separated [T2Space α] {s t : Set α} (hs : IsCompac section SeparatedFinset -theorem finset_disjoint_finset_opens_of_t2 [T2Space α] (s t : Finset α) (h : Disjoint s t) : - SeparatedNhds (s : Set α) t := +theorem finset_disjoint_finset_opens_of_t2 [T2Space X] (s t : Finset X) (h : Disjoint s t) : + SeparatedNhds (s : Set X) t := isCompact_isCompact_separated s.finite_toSet.isCompact t.finite_toSet.isCompact <| by exact_mod_cast h #align finset_disjoint_finset_opens_of_t2 finset_disjoint_finset_opens_of_t2 -theorem point_disjoint_finset_opens_of_t2 [T2Space α] {x : α} {s : Finset α} (h : x ∉ s) : - SeparatedNhds ({x} : Set α) s := by +theorem point_disjoint_finset_opens_of_t2 [T2Space X] {x : X} {s : Finset X} (h : x ∉ s) : + SeparatedNhds ({x} : Set X) s := by exact_mod_cast finset_disjoint_finset_opens_of_t2 {x} s (Finset.disjoint_singleton_left.mpr h) #align point_disjoint_finset_opens_of_t2 point_disjoint_finset_opens_of_t2 end SeparatedFinset /-- In a `T2Space`, every compact set is closed. -/ -theorem IsCompact.isClosed [T2Space α] {s : Set α} (hs : IsCompact s) : IsClosed s := +theorem IsCompact.isClosed [T2Space X] {s : Set X} (hs : IsCompact s) : IsClosed s := isOpen_compl_iff.1 <| isOpen_iff_forall_mem_open.mpr fun x hx => let ⟨u, v, _, vo, su, xv, uv⟩ := isCompact_isCompact_separated hs isCompact_singleton (disjoint_singleton_right.2 hx) @@ -1307,62 +1344,62 @@ theorem IsCompact.isClosed [T2Space α] {s : Set α} (hs : IsCompact s) : IsClos #align is_compact.is_closed IsCompact.isClosed @[simp] -theorem Filter.coclosedCompact_eq_cocompact [T2Space α] : coclosedCompact α = cocompact α := by +theorem Filter.coclosedCompact_eq_cocompact [T2Space X] : coclosedCompact X = cocompact X := by simp only [coclosedCompact, cocompact, iInf_and', - and_iff_right_of_imp (@IsCompact.isClosed α _ _ _)] + and_iff_right_of_imp (@IsCompact.isClosed X _ _ _)] #align filter.coclosed_compact_eq_cocompact Filter.coclosedCompact_eq_cocompact @[simp] -theorem Bornology.relativelyCompact_eq_inCompact [T2Space α] : - Bornology.relativelyCompact α = Bornology.inCompact α := +theorem Bornology.relativelyCompact_eq_inCompact [T2Space X] : + Bornology.relativelyCompact X = Bornology.inCompact X := Bornology.ext _ _ Filter.coclosedCompact_eq_cocompact #align bornology.relatively_compact_eq_in_compact Bornology.relativelyCompact_eq_inCompact -theorem IsCompact.preimage_continuous [CompactSpace α] [T2Space β] {f : α → β} {s : Set β} +theorem IsCompact.preimage_continuous [CompactSpace X] [T2Space Y] {f : X → Y} {s : Set Y} (hs : IsCompact s) (hf : Continuous f) : IsCompact (f ⁻¹' s) := (hs.isClosed.preimage hf).isCompact -/-- If `V : ι → Set α` is a decreasing family of compact sets then any neighborhood of +/-- If `V : ι → Set X` is a decreasing family of compact sets then any neighborhood of `⋂ i, V i` contains some `V i`. This is a version of `exists_subset_nhds_of_isCompact'` where we -don't need to assume each `V i` closed because it follows from compactness since `α` is +don't need to assume each `V i` closed because it follows from compactness since `X` is assumed to be Hausdorff. -/ -theorem exists_subset_nhds_of_isCompact [T2Space α] {ι : Type*} [Nonempty ι] {V : ι → Set α} - (hV : Directed (· ⊇ ·) V) (hV_cpct : ∀ i, IsCompact (V i)) {U : Set α} +theorem exists_subset_nhds_of_isCompact [T2Space X] {ι : Type*} [Nonempty ι] {V : ι → Set X} + (hV : Directed (· ⊇ ·) V) (hV_cpct : ∀ i, IsCompact (V i)) {U : Set X} (hU : ∀ x ∈ ⋂ i, V i, U ∈ 𝓝 x) : ∃ i, V i ⊆ U := exists_subset_nhds_of_isCompact' hV hV_cpct (fun i => (hV_cpct i).isClosed) hU #align exists_subset_nhds_of_is_compact exists_subset_nhds_of_isCompact -theorem CompactExhaustion.isClosed [T2Space α] (K : CompactExhaustion α) (n : ℕ) : IsClosed (K n) := +theorem CompactExhaustion.isClosed [T2Space X] (K : CompactExhaustion X) (n : ℕ) : IsClosed (K n) := (K.isCompact n).isClosed #align compact_exhaustion.is_closed CompactExhaustion.isClosed -theorem IsCompact.inter [T2Space α] {s t : Set α} (hs : IsCompact s) (ht : IsCompact t) : +theorem IsCompact.inter [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) : IsCompact (s ∩ t) := hs.inter_right <| ht.isClosed #align is_compact.inter IsCompact.inter -theorem isCompact_closure_of_subset_compact [T2Space α] {s t : Set α} (ht : IsCompact t) +theorem isCompact_closure_of_subset_compact [T2Space X] {s t : Set X} (ht : IsCompact t) (h : s ⊆ t) : IsCompact (closure s) := ht.of_isClosed_subset isClosed_closure (closure_minimal h ht.isClosed) #align is_compact_closure_of_subset_compact isCompact_closure_of_subset_compact @[simp] -theorem exists_compact_superset_iff [T2Space α] {s : Set α} : +theorem exists_compact_superset_iff [T2Space X] {s : Set X} : (∃ K, IsCompact K ∧ s ⊆ K) ↔ IsCompact (closure s) := ⟨fun ⟨_K, hK, hsK⟩ => isCompact_closure_of_subset_compact hK hsK, fun h => ⟨closure s, h, subset_closure⟩⟩ #align exists_compact_superset_iff exists_compact_superset_iff -theorem image_closure_of_isCompact [T2Space β] {s : Set α} (hs : IsCompact (closure s)) {f : α → β} +theorem image_closure_of_isCompact [T2Space Y] {s : Set X} (hs : IsCompact (closure s)) {f : X → Y} (hf : ContinuousOn f (closure s)) : f '' closure s = closure (f '' s) := Subset.antisymm hf.image_closure <| closure_minimal (image_subset f subset_closure) (hs.image_of_continuousOn hf).isClosed #align image_closure_of_is_compact image_closure_of_isCompact /-- If a compact set is covered by two open sets, then we can cover it by two compact subsets. -/ -theorem IsCompact.binary_compact_cover [T2Space α] {K U V : Set α} (hK : IsCompact K) +theorem IsCompact.binary_compact_cover [T2Space X] {K U V : Set X} (hK : IsCompact K) (hU : IsOpen U) (hV : IsOpen V) (h2K : K ⊆ U ∪ V) : - ∃ K₁ K₂ : Set α, IsCompact K₁ ∧ IsCompact K₂ ∧ K₁ ⊆ U ∧ K₂ ⊆ V ∧ K = K₁ ∪ K₂ := by + ∃ K₁ K₂ : Set X, IsCompact K₁ ∧ IsCompact K₂ ∧ K₁ ⊆ U ∧ K₂ ⊆ V ∧ K = K₁ ∪ K₂ := by obtain ⟨O₁, O₂, h1O₁, h1O₂, h2O₁, h2O₂, hO⟩ := isCompact_isCompact_separated (hK.diff hU) (hK.diff hV) (by rwa [disjoint_iff_inter_eq_empty, diff_inter_diff, diff_eq_empty]) @@ -1372,18 +1409,18 @@ theorem IsCompact.binary_compact_cover [T2Space α] {K U V : Set α} (hK : IsCom #align is_compact.binary_compact_cover IsCompact.binary_compact_cover /-- A continuous map from a compact space to a Hausdorff space is a closed map. -/ -protected theorem Continuous.isClosedMap [CompactSpace α] [T2Space β] {f : α → β} +protected theorem Continuous.isClosedMap [CompactSpace X] [T2Space Y] {f : X → Y} (h : Continuous f) : IsClosedMap f := fun _s hs => (hs.isCompact.image h).isClosed #align continuous.is_closed_map Continuous.isClosedMap /-- A continuous injective map from a compact space to a Hausdorff space is a closed embedding. -/ -theorem Continuous.closedEmbedding [CompactSpace α] [T2Space β] {f : α → β} (h : Continuous f) +theorem Continuous.closedEmbedding [CompactSpace X] [T2Space Y] {f : X → Y} (h : Continuous f) (hf : Function.Injective f) : ClosedEmbedding f := closedEmbedding_of_continuous_injective_closed h hf h.isClosedMap #align continuous.closed_embedding Continuous.closedEmbedding /-- A continuous surjective map from a compact space to a Hausdorff space is a quotient map. -/ -theorem QuotientMap.of_surjective_continuous [CompactSpace α] [T2Space β] {f : α → β} +theorem QuotientMap.of_surjective_continuous [CompactSpace X] [T2Space Y] {f : X → Y} (hsurj : Surjective f) (hcont : Continuous f) : QuotientMap f := hcont.isClosedMap.to_quotientMap hcont hsurj #align quotient_map.of_surjective_continuous QuotientMap.of_surjective_continuous @@ -1393,9 +1430,9 @@ section open Finset Function /-- For every finite open cover `Uᵢ` of a compact set, there exists a compact cover `Kᵢ ⊆ Uᵢ`. -/ -theorem IsCompact.finite_compact_cover [T2Space α] {s : Set α} (hs : IsCompact s) {ι} (t : Finset ι) - (U : ι → Set α) (hU : ∀ i ∈ t, IsOpen (U i)) (hsC : s ⊆ ⋃ i ∈ t, U i) : - ∃ K : ι → Set α, (∀ i, IsCompact (K i)) ∧ (∀ i, K i ⊆ U i) ∧ s = ⋃ i ∈ t, K i := by +theorem IsCompact.finite_compact_cover [T2Space X] {s : Set X} (hs : IsCompact s) {ι} (t : Finset ι) + (U : ι → Set X) (hU : ∀ i ∈ t, IsOpen (U i)) (hsC : s ⊆ ⋃ i ∈ t, U i) : + ∃ K : ι → Set X, (∀ i, IsCompact (K i)) ∧ (∀ i, K i ⊆ U i) ∧ s = ⋃ i ∈ t, K i := by induction' t using Finset.induction with x t hx ih generalizing U s · refine' ⟨fun _ => ∅, fun _ => isCompact_empty, fun i => empty_subset _, _⟩ simpa only [subset_empty_iff, Finset.not_mem_empty, iUnion_false, iUnion_empty] using hsC @@ -1421,8 +1458,8 @@ end -- see Note [lower instance priority] /-- A weakly locally compact Hausdorff space is locally compact. -/ -instance WeaklyLocallyCompactSpace.locallyCompactSpace [WeaklyLocallyCompactSpace α] [T2Space α] : - LocallyCompactSpace α := +instance WeaklyLocallyCompactSpace.locallyCompactSpace [WeaklyLocallyCompactSpace X] [T2Space X] : + LocallyCompactSpace X := ⟨fun x _n hn => let ⟨_u, un, uo, xu⟩ := mem_nhds_iff.mp hn let ⟨k, kc, kx⟩ := exists_compact_mem_nhds x @@ -1439,15 +1476,15 @@ instance WeaklyLocallyCompactSpace.locallyCompactSpace [WeaklyLocallyCompactSpac #align locally_compact_of_compact_nhds WeaklyLocallyCompactSpace.locallyCompactSpace @[deprecated WeaklyLocallyCompactSpace.locallyCompactSpace] -theorem locally_compact_of_compact [T2Space α] [CompactSpace α] : - LocallyCompactSpace α := +theorem locally_compact_of_compact [T2Space X] [CompactSpace X] : + LocallyCompactSpace X := inferInstance #align locally_compact_of_compact locally_compact_of_compact /-- In a weakly locally compact T₂ space, every compact set has an open neighborhood with compact closure. -/ -theorem exists_open_superset_and_isCompact_closure [WeaklyLocallyCompactSpace α] [T2Space α] - {K : Set α} (hK : IsCompact K) : ∃ V, IsOpen V ∧ K ⊆ V ∧ IsCompact (closure V) := by +theorem exists_open_superset_and_isCompact_closure [WeaklyLocallyCompactSpace X] [T2Space X] + {K : Set X} (hK : IsCompact K) : ∃ V, IsOpen V ∧ K ⊆ V ∧ IsCompact (closure V) := by rcases exists_compact_superset hK with ⟨K', hK', hKK'⟩ exact ⟨interior K', isOpen_interior, hKK', isCompact_closure_of_subset_compact hK' interior_subset⟩ @@ -1455,8 +1492,8 @@ theorem exists_open_superset_and_isCompact_closure [WeaklyLocallyCompactSpace α /-- In a weakly locally compact T₂ space, every point has an open neighborhood with compact closure. -/ -theorem exists_open_with_compact_closure [WeaklyLocallyCompactSpace α] [T2Space α] (x : α) : - ∃ U : Set α, IsOpen U ∧ x ∈ U ∧ IsCompact (closure U) := by +theorem exists_open_with_compact_closure [WeaklyLocallyCompactSpace X] [T2Space X] (x : X) : + ∃ U : Set X, IsOpen U ∧ x ∈ U ∧ IsCompact (closure U) := by simpa only [singleton_subset_iff] using exists_open_superset_and_isCompact_closure isCompact_singleton #align exists_open_with_compact_closure exists_open_with_compact_closure @@ -1464,7 +1501,7 @@ theorem exists_open_with_compact_closure [WeaklyLocallyCompactSpace α] [T2Space /-- In a locally compact T₂ space, given a compact set `K` inside an open set `U`, we can find an open set `V` between these sets with compact closure: `K ⊆ V` and the closure of `V` is inside `U`. -/ -theorem exists_open_between_and_isCompact_closure [LocallyCompactSpace α] [T2Space α] {K U : Set α} +theorem exists_open_between_and_isCompact_closure [LocallyCompactSpace X] [T2Space X] {K U : Set X} (hK : IsCompact K) (hU : IsOpen U) (hKU : K ⊆ U) : ∃ V, IsOpen V ∧ K ⊆ V ∧ closure V ⊆ U ∧ IsCompact (closure V) := by rcases exists_compact_between hK hU hKU with ⟨V, hV, hKV, hVU⟩ @@ -1473,7 +1510,7 @@ theorem exists_open_between_and_isCompact_closure [LocallyCompactSpace α] [T2Sp isCompact_closure_of_subset_compact hV interior_subset⟩ #align exists_open_between_and_is_compact_closure exists_open_between_and_isCompact_closure -theorem isPreirreducible_iff_subsingleton [T2Space α] {S : Set α} : +theorem isPreirreducible_iff_subsingleton [T2Space X] {S : Set X} : IsPreirreducible S ↔ S.Subsingleton := by refine' ⟨fun h x hx y hy => _, Set.Subsingleton.isPreirreducible⟩ by_contra e @@ -1482,20 +1519,20 @@ theorem isPreirreducible_iff_subsingleton [T2Space α] {S : Set α} : #align is_preirreducible_iff_subsingleton isPreirreducible_iff_subsingleton -- todo: use `alias` + `attribute [protected]` once we get `attribute [protected]` -protected lemma IsPreirreducible.subsingleton [T2Space α] {S : Set α} (h : IsPreirreducible S) : +protected lemma IsPreirreducible.subsingleton [T2Space X] {S : Set X} (h : IsPreirreducible S) : S.Subsingleton := isPreirreducible_iff_subsingleton.1 h #align is_preirreducible.subsingleton IsPreirreducible.subsingleton -theorem isIrreducible_iff_singleton [T2Space α] {S : Set α} : IsIrreducible S ↔ ∃ x, S = {x} := by +theorem isIrreducible_iff_singleton [T2Space X] {S : Set X} : IsIrreducible S ↔ ∃ x, S = {x} := by rw [IsIrreducible, isPreirreducible_iff_subsingleton, exists_eq_singleton_iff_nonempty_subsingleton] #align is_irreducible_iff_singleton isIrreducible_iff_singleton /-- There does not exist a nontrivial preirreducible T₂ space. -/ -theorem not_preirreducible_nontrivial_t2 (α) [TopologicalSpace α] [PreirreducibleSpace α] - [Nontrivial α] [T2Space α] : False := - (PreirreducibleSpace.isPreirreducible_univ (α := α)).subsingleton.not_nontrivial nontrivial_univ +theorem not_preirreducible_nontrivial_t2 (X) [TopologicalSpace X] [PreirreducibleSpace X] + [Nontrivial X] [T2Space X] : False := + (PreirreducibleSpace.isPreirreducible_univ (α := X)).subsingleton.not_nontrivial nontrivial_univ #align not_preirreducible_nontrivial_t2 not_preirreducible_nontrivial_t2 end Separation @@ -1512,13 +1549,13 @@ class RegularSpace (X : Type u) [TopologicalSpace X] : Prop where regular : ∀ {s : Set X} {a}, IsClosed s → a ∉ s → Disjoint (𝓝ˢ s) (𝓝 a) #align regular_space RegularSpace -theorem regularSpace_TFAE (X : Type u) [ TopologicalSpace X ] : +theorem regularSpace_TFAE (X : Type u) [TopologicalSpace X] : List.TFAE [RegularSpace X, - ∀ (s : Set X) a, a ∉ closure s → Disjoint (𝓝ˢ s) (𝓝 a), - ∀ (a : X) (s : Set X), Disjoint (𝓝ˢ s) (𝓝 a) ↔ a ∉ closure s, - ∀ (a : X) (s : Set X), s ∈ 𝓝 a → ∃ t ∈ 𝓝 a, IsClosed t ∧ t ⊆ s, - ∀ a : X, (𝓝 a).lift' closure ≤ 𝓝 a, - ∀ a : X , (𝓝 a).lift' closure = 𝓝 a] := by + ∀ (s : Set X) x, x ∉ closure s → Disjoint (𝓝ˢ s) (𝓝 x), + ∀ (x : X) (s : Set X), Disjoint (𝓝ˢ s) (𝓝 x) ↔ x ∉ closure s, + ∀ (x : X) (s : Set X), s ∈ 𝓝 x → ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s, + ∀ x : X, (𝓝 x).lift' closure ≤ 𝓝 x, + ∀ x : X , (𝓝 x).lift' closure = 𝓝 x] := by tfae_have 1 ↔ 5 · rw [regularSpace_iff, (@compl_surjective (Set X) _).forall, forall_swap] simp only [isClosed_compl_iff, mem_compl_iff, Classical.not_not, @and_comm (_ ∈ _), @@ -1546,103 +1583,109 @@ theorem regularSpace_TFAE (X : Type u) [ TopologicalSpace X ] : tfae_finish #align regular_space_tfae regularSpace_TFAE -theorem RegularSpace.ofLift'_closure (h : ∀ a : α, (𝓝 a).lift' closure = 𝓝 a) : RegularSpace α := - Iff.mpr ((regularSpace_TFAE α).out 0 5) h +theorem RegularSpace.ofLift'_closure (h : ∀ x : X, (𝓝 x).lift' closure = 𝓝 x) : RegularSpace X := + Iff.mpr ((regularSpace_TFAE X).out 0 5) h #align regular_space.of_lift'_closure RegularSpace.ofLift'_closure -theorem RegularSpace.ofBasis {ι : α → Sort*} {p : ∀ a, ι a → Prop} {s : ∀ a, ι a → Set α} +theorem RegularSpace.ofBasis {ι : X → Sort*} {p : ∀ a, ι a → Prop} {s : ∀ a, ι a → Set X} (h₁ : ∀ a, (𝓝 a).HasBasis (p a) (s a)) (h₂ : ∀ a i, p a i → IsClosed (s a i)) : - RegularSpace α := + RegularSpace X := RegularSpace.ofLift'_closure fun a => (h₁ a).lift'_closure_eq_self (h₂ a) #align regular_space.of_basis RegularSpace.ofBasis theorem RegularSpace.ofExistsMemNhdsIsClosedSubset - (h : ∀ (a : α), ∀ s ∈ 𝓝 a, ∃ t ∈ 𝓝 a, IsClosed t ∧ t ⊆ s) : RegularSpace α := - Iff.mpr ((regularSpace_TFAE α).out 0 3) h + (h : ∀ (x : X), ∀ s ∈ 𝓝 x, ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s) : RegularSpace X := + Iff.mpr ((regularSpace_TFAE X).out 0 3) h #align regular_space.of_exists_mem_nhds_is_closed_subset RegularSpace.ofExistsMemNhdsIsClosedSubset -variable [RegularSpace α] {a : α} {s : Set α} +/-- A locally compact T2 space is regular. -/ +instance [LocallyCompactSpace X] [T2Space X] : RegularSpace X := by + apply RegularSpace.ofExistsMemNhdsIsClosedSubset (fun x s hx ↦ ?_) + rcases local_compact_nhds hx with ⟨k, kx, ks, hk⟩ + exact ⟨k, kx, hk.isClosed, ks⟩ + +variable [RegularSpace X] {x : X} {s : Set X} -theorem disjoint_nhdsSet_nhds : Disjoint (𝓝ˢ s) (𝓝 a) ↔ a ∉ closure s := by - have h := (regularSpace_TFAE α).out 0 2 +theorem disjoint_nhdsSet_nhds : Disjoint (𝓝ˢ s) (𝓝 x) ↔ x ∉ closure s := by + have h := (regularSpace_TFAE X).out 0 2 exact h.mp ‹_› _ _ #align disjoint_nhds_set_nhds disjoint_nhdsSet_nhds -theorem disjoint_nhds_nhdsSet : Disjoint (𝓝 a) (𝓝ˢ s) ↔ a ∉ closure s := +theorem disjoint_nhds_nhdsSet : Disjoint (𝓝 x) (𝓝ˢ s) ↔ x ∉ closure s := disjoint_comm.trans disjoint_nhdsSet_nhds #align disjoint_nhds_nhds_set disjoint_nhds_nhdsSet -theorem exists_mem_nhds_isClosed_subset {a : α} {s : Set α} (h : s ∈ 𝓝 a) : - ∃ t ∈ 𝓝 a, IsClosed t ∧ t ⊆ s := by - have h' := (regularSpace_TFAE α).out 0 3 +theorem exists_mem_nhds_isClosed_subset {x : X} {s : Set X} (h : s ∈ 𝓝 x) : + ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s := by + have h' := (regularSpace_TFAE X).out 0 3 exact h'.mp ‹_› _ _ h #align exists_mem_nhds_is_closed_subset exists_mem_nhds_isClosed_subset -theorem closed_nhds_basis (a : α) : (𝓝 a).HasBasis (fun s : Set α => s ∈ 𝓝 a ∧ IsClosed s) id := +theorem closed_nhds_basis (x : X) : (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsClosed s) id := hasBasis_self.2 fun _ => exists_mem_nhds_isClosed_subset #align closed_nhds_basis closed_nhds_basis -theorem lift'_nhds_closure (a : α) : (𝓝 a).lift' closure = 𝓝 a := - (closed_nhds_basis a).lift'_closure_eq_self fun _ => And.right +theorem lift'_nhds_closure (x : X) : (𝓝 x).lift' closure = 𝓝 x := + (closed_nhds_basis x).lift'_closure_eq_self fun _ => And.right #align lift'_nhds_closure lift'_nhds_closure -theorem Filter.HasBasis.nhds_closure {ι : Sort*} {a : α} {p : ι → Prop} {s : ι → Set α} - (h : (𝓝 a).HasBasis p s) : (𝓝 a).HasBasis p fun i => closure (s i) := - lift'_nhds_closure a ▸ h.lift'_closure +theorem Filter.HasBasis.nhds_closure {ι : Sort*} {x : X} {p : ι → Prop} {s : ι → Set X} + (h : (𝓝 x).HasBasis p s) : (𝓝 x).HasBasis p fun i => closure (s i) := + lift'_nhds_closure x ▸ h.lift'_closure #align filter.has_basis.nhds_closure Filter.HasBasis.nhds_closure -theorem hasBasis_nhds_closure (a : α) : (𝓝 a).HasBasis (fun s => s ∈ 𝓝 a) closure := - (𝓝 a).basis_sets.nhds_closure +theorem hasBasis_nhds_closure (x : X) : (𝓝 x).HasBasis (fun s => s ∈ 𝓝 x) closure := + (𝓝 x).basis_sets.nhds_closure #align has_basis_nhds_closure hasBasis_nhds_closure -theorem hasBasis_opens_closure (a : α) : (𝓝 a).HasBasis (fun s => a ∈ s ∧ IsOpen s) closure := - (nhds_basis_opens a).nhds_closure +theorem hasBasis_opens_closure (x : X) : (𝓝 x).HasBasis (fun s => x ∈ s ∧ IsOpen s) closure := + (nhds_basis_opens x).nhds_closure #align has_basis_opens_closure hasBasis_opens_closure -theorem TopologicalSpace.IsTopologicalBasis.nhds_basis_closure {B : Set (Set α)} - (hB : IsTopologicalBasis B) (a : α) : - (𝓝 a).HasBasis (fun s : Set α => a ∈ s ∧ s ∈ B) closure := by +theorem TopologicalSpace.IsTopologicalBasis.nhds_basis_closure {B : Set (Set X)} + (hB : IsTopologicalBasis B) (x : X) : + (𝓝 x).HasBasis (fun s : Set X => x ∈ s ∧ s ∈ B) closure := by simpa only [and_comm] using hB.nhds_hasBasis.nhds_closure #align topological_space.is_topological_basis.nhds_basis_closure TopologicalSpace.IsTopologicalBasis.nhds_basis_closure -theorem TopologicalSpace.IsTopologicalBasis.exists_closure_subset {B : Set (Set α)} - (hB : IsTopologicalBasis B) {a : α} {s : Set α} (h : s ∈ 𝓝 a) : - ∃ t ∈ B, a ∈ t ∧ closure t ⊆ s := by +theorem TopologicalSpace.IsTopologicalBasis.exists_closure_subset {B : Set (Set X)} + (hB : IsTopologicalBasis B) {x : X} {s : Set X} (h : s ∈ 𝓝 x) : + ∃ t ∈ B, x ∈ t ∧ closure t ⊆ s := by simpa only [exists_prop, and_assoc] using hB.nhds_hasBasis.nhds_closure.mem_iff.mp h #align topological_space.is_topological_basis.exists_closure_subset TopologicalSpace.IsTopologicalBasis.exists_closure_subset -theorem disjoint_nhds_nhds_iff_not_specializes {a b : α} : Disjoint (𝓝 a) (𝓝 b) ↔ ¬a ⤳ b := by +theorem disjoint_nhds_nhds_iff_not_specializes {a b : X} : Disjoint (𝓝 a) (𝓝 b) ↔ ¬a ⤳ b := by rw [← nhdsSet_singleton, disjoint_nhdsSet_nhds, specializes_iff_mem_closure] #align disjoint_nhds_nhds_iff_not_specializes disjoint_nhds_nhds_iff_not_specializes -theorem specializes_comm {a b : α} : a ⤳ b ↔ b ⤳ a := by - simp only [← (disjoint_nhds_nhds_iff_not_specializes (α := α)).not_left, disjoint_comm] +theorem specializes_comm {a b : X} : a ⤳ b ↔ b ⤳ a := by + simp only [← (disjoint_nhds_nhds_iff_not_specializes (X := X)).not_left, disjoint_comm] #align specializes_comm specializes_comm alias ⟨Specializes.symm, _⟩ := specializes_comm #align specializes.symm Specializes.symm -theorem specializes_iff_inseparable {a b : α} : a ⤳ b ↔ Inseparable a b := +theorem specializes_iff_inseparable {a b : X} : a ⤳ b ↔ Inseparable a b := ⟨fun h => h.antisymm h.symm, le_of_eq⟩ #align specializes_iff_inseparable specializes_iff_inseparable -theorem isClosed_setOf_specializes : IsClosed { p : α × α | p.1 ⤳ p.2 } := by +theorem isClosed_setOf_specializes : IsClosed { p : X × X | p.1 ⤳ p.2 } := by simp only [← isOpen_compl_iff, compl_setOf, ← disjoint_nhds_nhds_iff_not_specializes, isOpen_setOf_disjoint_nhds_nhds] #align is_closed_set_of_specializes isClosed_setOf_specializes -theorem isClosed_setOf_inseparable : IsClosed { p : α × α | Inseparable p.1 p.2 } := by +theorem isClosed_setOf_inseparable : IsClosed { p : X × X | Inseparable p.1 p.2 } := by simp only [← specializes_iff_inseparable, isClosed_setOf_specializes] #align is_closed_set_of_inseparable isClosed_setOf_inseparable -protected theorem Inducing.regularSpace [TopologicalSpace β] {f : β → α} (hf : Inducing f) : - RegularSpace β := +protected theorem Inducing.regularSpace [TopologicalSpace Y] {f : Y → X} (hf : Inducing f) : + RegularSpace Y := RegularSpace.ofBasis (fun b => by rw [hf.nhds_eq_comap b]; exact (closed_nhds_basis _).comap _) fun b s hs => by exact hs.2.preimage hf.continuous #align inducing.regular_space Inducing.regularSpace -theorem regularSpace_induced (f : β → α) : @RegularSpace β (induced f ‹_›) := +theorem regularSpace_induced (f : Y → X) : @RegularSpace Y (induced f ‹_›) := letI := induced f ‹_› Inducing.regularSpace ⟨rfl⟩ #align regular_space_induced regularSpace_induced @@ -1672,55 +1715,115 @@ theorem RegularSpace.inf {X} {t₁ t₂ : TopologicalSpace X} (h₁ : @RegularSp exact regularSpace_iInf (Bool.forall_bool.2 ⟨h₂, h₁⟩) #align regular_space.inf RegularSpace.inf -instance {p : α → Prop} : RegularSpace (Subtype p) := +instance {p : X → Prop} : RegularSpace (Subtype p) := embedding_subtype_val.toInducing.regularSpace -instance [TopologicalSpace β] [RegularSpace β] : RegularSpace (α × β) := - (regularSpace_induced (@Prod.fst α β)).inf (regularSpace_induced (@Prod.snd α β)) +instance [TopologicalSpace Y] [RegularSpace Y] : RegularSpace (X × Y) := + (regularSpace_induced (@Prod.fst X Y)).inf (regularSpace_induced (@Prod.snd X Y)) -instance {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] [∀ i, RegularSpace (π i)] : - RegularSpace (∀ i, π i) := +instance {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, RegularSpace (X i)] : + RegularSpace (∀ i, X i) := regularSpace_iInf fun _ => regularSpace_induced _ end RegularSpace +section ClosableCompactSubsetOpenSpace + +/-- A class of topological spaces in which, given a compact set included inside an open set, then +the closure of the compact set is also included in the open set. +Satisfied notably for T2 spaces and regular spaces, and useful +when discussing classes of regular measures. -/ +class ClosableCompactSubsetOpenSpace (X : Type*) [TopologicalSpace X] : Prop := + closure_subset_of_isOpen : ∀ (K U : Set X), IsCompact K → IsOpen U → K ⊆ U → closure K ⊆ U + +theorem IsCompact.closure_subset_of_isOpen [ClosableCompactSubsetOpenSpace X] + {s : Set X} (hs : IsCompact s) {u : Set X} (hu : IsOpen u) (h : s ⊆ u) : + closure s ⊆ u := + ClosableCompactSubsetOpenSpace.closure_subset_of_isOpen s u hs hu h + +instance [T2Space X] : ClosableCompactSubsetOpenSpace X := + ⟨fun K _U K_comp _U_open KU ↦ by rwa [K_comp.isClosed.closure_eq]⟩ + +/-- In a (possibly non-Hausdorff) regular space, if a compact set `s` is contained in an +open set `u`, then its closure is also contained in `u`. -/ +instance [RegularSpace X] : ClosableCompactSubsetOpenSpace X := by + refine ⟨fun s u hs hu h ↦ ?_⟩ + obtain ⟨F, sF, F_closed, Fu⟩ : ∃ F, s ⊆ F ∧ IsClosed F ∧ F ⊆ u := by + apply hs.induction_on (p := fun t ↦ ∃ F, t ⊆ F ∧ IsClosed F ∧ F ⊆ u) + · exact ⟨∅, by simp⟩ + · intro t' t ht't ⟨F, tF, F_closed, Fu⟩ + exact ⟨F, ht't.trans tF, F_closed, Fu⟩ + · intro t t' ⟨F, tF, F_closed, Fu⟩ ⟨F', t'F', F'_closed, F'u⟩ + exact ⟨F ∪ F', union_subset_union tF t'F', F_closed.union F'_closed, union_subset Fu F'u⟩ + · intro x hx + rcases exists_mem_nhds_isClosed_subset (hu.mem_nhds (h hx)) with ⟨F, F_mem, F_closed, Fu⟩ + exact ⟨F, nhdsWithin_le_nhds F_mem, F, Subset.rfl, F_closed, Fu⟩ + exact (closure_minimal sF F_closed).trans Fu + +/-- In a (possibly non-Hausdorff) locally compact space with the `ClosableCompactSubsetOpenSpace` + property (for instance regular spaces), for every containment `K ⊆ U` of a compact set `K` in an + open set `U`, there is a compact closed neighborhood `L` such that `K ⊆ L ⊆ U`: equivalently, + there is a compact closed set `L` such that `K ⊆ interior L` and `L ⊆ U`. -/ +theorem exists_compact_closed_between [LocallyCompactSpace X] [ClosableCompactSubsetOpenSpace X] + {K U : Set X} (hK : IsCompact K) (hU : IsOpen U) (h_KU : K ⊆ U) : + ∃ L, IsCompact L ∧ IsClosed L ∧ K ⊆ interior L ∧ L ⊆ U := by + rcases exists_compact_between hK hU h_KU with ⟨L, L_comp, KL, LU⟩ + rcases exists_compact_between hK isOpen_interior KL with ⟨M, M_comp, KM, ML⟩ + refine ⟨closure M, ?_, isClosed_closure, ?_, ?_⟩ + · have : closure M ∩ L = closure M := by + apply inter_eq_self_of_subset_left + exact (M_comp.closure_subset_of_isOpen isOpen_interior ML).trans interior_subset + rw [← this] + apply L_comp.inter_left isClosed_closure + · exact KM.trans (interior_mono subset_closure) + · apply M_comp.closure_subset_of_isOpen hU + exact ML.trans (interior_subset.trans LU) + +protected theorem IsCompact.closure [WeaklyLocallyCompactSpace X] [ClosableCompactSubsetOpenSpace X] + {K : Set X} (hK : IsCompact K) : IsCompact (closure K) := by + rcases exists_compact_superset hK with ⟨L, L_comp, hL⟩ + exact L_comp.of_isClosed_subset isClosed_closure + ((hK.closure_subset_of_isOpen isOpen_interior hL).trans interior_subset) + +end ClosableCompactSubsetOpenSpace + section T3 /-- A T₃ space is a T₀ space which is a regular space. Any T₃ space is a T₁ space, a T₂ space, and a T₂.₅ space. -/ -class T3Space (α : Type u) [TopologicalSpace α] extends T0Space α, RegularSpace α : Prop +class T3Space (X : Type u) [TopologicalSpace X] extends T0Space X, RegularSpace X : Prop #align t3_space T3Space -instance (priority := 90) [T0Space α] [RegularSpace α] : T3Space α := ⟨⟩ +instance (priority := 90) [T0Space X] [RegularSpace X] : T3Space X := ⟨⟩ -- see Note [lower instance priority] -instance (priority := 100) T3Space.t25Space [T3Space α] : T25Space α := by +instance (priority := 100) T3Space.t25Space [T3Space X] : T25Space X := by refine' ⟨fun x y hne => _⟩ rw [lift'_nhds_closure, lift'_nhds_closure] have : x ∉ closure {y} ∨ y ∉ closure {x} := - (t0Space_iff_or_not_mem_closure α).mp inferInstance x y hne + (t0Space_iff_or_not_mem_closure X).mp inferInstance x y hne simp only [← disjoint_nhds_nhdsSet, nhdsSet_singleton] at this exact this.elim id fun h => h.symm #align t3_space.t2_5_space T3Space.t25Space -protected theorem Embedding.t3Space [TopologicalSpace β] [T3Space β] {f : α → β} - (hf : Embedding f) : T3Space α := +protected theorem Embedding.t3Space [TopologicalSpace Y] [T3Space Y] {f : X → Y} + (hf : Embedding f) : T3Space X := { toT0Space := hf.t0Space toRegularSpace := hf.toInducing.regularSpace } #align embedding.t3_space Embedding.t3Space -instance Subtype.t3Space [T3Space α] {p : α → Prop} : T3Space (Subtype p) := +instance Subtype.t3Space [T3Space X] {p : X → Prop} : T3Space (Subtype p) := embedding_subtype_val.t3Space #align subtype.t3_space Subtype.t3Space -instance [TopologicalSpace β] [T3Space α] [T3Space β] : T3Space (α × β) := ⟨⟩ +instance [TopologicalSpace Y] [T3Space X] [T3Space Y] : T3Space (X × Y) := ⟨⟩ -instance {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] [∀ i, T3Space (π i)] : - T3Space (∀ i, π i) := ⟨⟩ +instance {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, T3Space (X i)] : + T3Space (∀ i, X i) := ⟨⟩ /-- Given two points `x ≠ y`, we can find neighbourhoods `x ∈ V₁ ⊆ U₁` and `y ∈ V₂ ⊆ U₂`, with the `Vₖ` closed and the `Uₖ` open, such that the `Uₖ` are disjoint. -/ -theorem disjoint_nested_nhds [T3Space α] {x y : α} (h : x ≠ y) : +theorem disjoint_nested_nhds [T3Space X] {x y : X} (h : x ≠ y) : ∃ U₁ ∈ 𝓝 x, ∃ V₁ ∈ 𝓝 x, ∃ U₂ ∈ 𝓝 y, ∃ V₂ ∈ 𝓝 y, IsClosed V₁ ∧ IsClosed V₂ ∧ IsOpen U₁ ∧ IsOpen U₂ ∧ V₁ ⊆ U₁ ∧ V₂ ⊆ U₂ ∧ Disjoint U₁ U₂ := by rcases t2_separation h with ⟨U₁, U₂, U₁_op, U₂_op, x_in, y_in, H⟩ @@ -1733,7 +1836,7 @@ theorem disjoint_nested_nhds [T3Space α] {x y : α} (h : x ≠ y) : open SeparationQuotient /-- The `SeparationQuotient` of a regular space is a T₃ space. -/ -instance [RegularSpace α] : T3Space (SeparationQuotient α) where +instance [RegularSpace X] : T3Space (SeparationQuotient X) where regular {s a} hs ha := by rcases surjective_mk a with ⟨a, rfl⟩ rw [← disjoint_comap_iff surjective_mk, comap_mk_nhds_mk, comap_mk_nhdsSet] @@ -1749,16 +1852,16 @@ class NormalSpace (X : Type u) [TopologicalSpace X] : Prop where /-- Two disjoint sets in a normal space admit disjoint neighbourhoods. -/ normal : ∀ s t : Set X, IsClosed s → IsClosed t → Disjoint s t → SeparatedNhds s t -theorem normal_separation [NormalSpace α] {s t : Set α} (H1 : IsClosed s) (H2 : IsClosed t) +theorem normal_separation [NormalSpace X] {s t : Set X} (H1 : IsClosed s) (H2 : IsClosed t) (H3 : Disjoint s t) : SeparatedNhds s t := NormalSpace.normal s t H1 H2 H3 #align normal_separation normal_separation -theorem disjoint_nhdsSet_nhdsSet [NormalSpace α] {s t : Set α} (hs : IsClosed s) (ht : IsClosed t) +theorem disjoint_nhdsSet_nhdsSet [NormalSpace X] {s t : Set X} (hs : IsClosed s) (ht : IsClosed t) (hd : Disjoint s t) : Disjoint (𝓝ˢ s) (𝓝ˢ t) := (normal_separation hs ht hd).disjoint_nhdsSet -theorem normal_exists_closure_subset [NormalSpace α] {s t : Set α} (hs : IsClosed s) (ht : IsOpen t) +theorem normal_exists_closure_subset [NormalSpace X] {s t : Set X} (hs : IsClosed s) (ht : IsOpen t) (hst : s ⊆ t) : ∃ u, IsOpen u ∧ s ⊆ u ∧ closure u ⊆ t := by have : Disjoint s tᶜ := Set.disjoint_left.mpr fun x hxs hxt => hxt (hst hxs) rcases normal_separation hs (isClosed_compl_iff.2 ht) this with @@ -1769,8 +1872,8 @@ theorem normal_exists_closure_subset [NormalSpace α] {s t : Set α} (hs : IsClo #align normal_exists_closure_subset normal_exists_closure_subset /-- If the codomain of a closed embedding is a normal space, then so is the domain. -/ -protected theorem ClosedEmbedding.normalSpace [TopologicalSpace β] [NormalSpace β] {f : α → β} - (hf : ClosedEmbedding f) : NormalSpace α where +protected theorem ClosedEmbedding.normalSpace [TopologicalSpace Y] [NormalSpace Y] {f : X → Y} + (hf : ClosedEmbedding f) : NormalSpace X where normal s t hs ht hst := by have H : SeparatedNhds (f '' s) (f '' t) := NormalSpace.normal (f '' s) (f '' t) (hf.isClosedMap s hs) (hf.isClosedMap t ht) @@ -1779,19 +1882,19 @@ protected theorem ClosedEmbedding.normalSpace [TopologicalSpace β] [NormalSpace /-- A regular topological space with second countable topology is a normal space. -/ instance (priority := 100) NormalSpace.of_regularSpace_secondCountableTopology - [RegularSpace α] [SecondCountableTopology α] : NormalSpace α := by - have key : ∀ {s t : Set α}, IsClosed t → Disjoint s t → - ∃ U : Set (countableBasis α), (s ⊆ ⋃ u ∈ U, ↑u) ∧ (∀ u ∈ U, Disjoint (closure ↑u) t) ∧ - ∀ n : ℕ, IsClosed (⋃ (u ∈ U) (_ : Encodable.encode u ≤ n), closure (u : Set α)) := by + [RegularSpace X] [SecondCountableTopology X] : NormalSpace X := by + have key : ∀ {s t : Set X}, IsClosed t → Disjoint s t → + ∃ U : Set (countableBasis X), (s ⊆ ⋃ u ∈ U, ↑u) ∧ (∀ u ∈ U, Disjoint (closure ↑u) t) ∧ + ∀ n : ℕ, IsClosed (⋃ (u ∈ U) (_ : Encodable.encode u ≤ n), closure (u : Set X)) := by intro s t hc hd rw [disjoint_left] at hd - have : ∀ x ∈ s, ∃ U ∈ countableBasis α, x ∈ U ∧ Disjoint (closure U) t := by + have : ∀ x ∈ s, ∃ U ∈ countableBasis X, x ∈ U ∧ Disjoint (closure U) t := by intro x hx - rcases (isBasis_countableBasis α).exists_closure_subset (hc.compl_mem_nhds (hd hx)) + rcases (isBasis_countableBasis X).exists_closure_subset (hc.compl_mem_nhds (hd hx)) with ⟨u, hu, hxu, hut⟩ exact ⟨u, hu, hxu, disjoint_left.2 hut⟩ choose! U hu hxu hd using this - set V : s → countableBasis α := MapsTo.restrict _ _ _ hu + set V : s → countableBasis X := MapsTo.restrict _ _ _ hu refine' ⟨range V, _, forall_range_iff.2 <| Subtype.forall.2 hd, fun n => _⟩ · rw [biUnion_range] exact fun x hx => mem_iUnion.2 ⟨⟨x, hx⟩, hxu x hx⟩ @@ -1827,25 +1930,25 @@ end NormalSpace section Normality /-- A T₄ space is a normal T₁ space. -/ -class T4Space (α : Type u) [TopologicalSpace α] extends T1Space α, NormalSpace α : Prop +class T4Space (X : Type u) [TopologicalSpace X] extends T1Space X, NormalSpace X : Prop #align normal_space NormalSpace -instance (priority := 100) [T1Space α] [NormalSpace α] : T4Space α := ⟨⟩ +instance (priority := 100) [T1Space X] [NormalSpace X] : T4Space X := ⟨⟩ -- see Note [lower instance priority] -instance (priority := 100) T4Space.t3Space [T4Space α] : T3Space α where +instance (priority := 100) T4Space.t3Space [T4Space X] : T3Space X where regular hs hxs := by simpa only [nhdsSet_singleton] using (normal_separation hs isClosed_singleton (disjoint_singleton_right.mpr hxs)).disjoint_nhdsSet #align normal_space.t3_space T4Space.t3Space -instance (priority := 100) T4Space.of_compactSpace_t2Space [CompactSpace α] [T2Space α] : - T4Space α where +instance (priority := 100) T4Space.of_compactSpace_t2Space [CompactSpace X] [T2Space X] : + T4Space X where normal _s _t hs ht := isCompact_isCompact_separated hs.isCompact ht.isCompact #align normal_of_compact_t2 T4Space.of_compactSpace_t2Space /-- If the codomain of a closed embedding is a T₄ space, then so is the domain. -/ -protected theorem ClosedEmbedding.t4Space [TopologicalSpace β] [T4Space β] {f : α → β} - (hf : ClosedEmbedding f) : T4Space α where +protected theorem ClosedEmbedding.t4Space [TopologicalSpace Y] [T4Space Y] {f : X → Y} + (hf : ClosedEmbedding f) : T4Space X where toT1Space := hf.toEmbedding.t1Space toNormalSpace := hf.normalSpace #align closed_embedding.normal_space ClosedEmbedding.t4Space @@ -1853,7 +1956,7 @@ protected theorem ClosedEmbedding.t4Space [TopologicalSpace β] [T4Space β] {f namespace SeparationQuotient /-- The `SeparationQuotient` of a normal space is a normal space. -/ -instance [NormalSpace α] : NormalSpace (SeparationQuotient α) where +instance [NormalSpace X] : NormalSpace (SeparationQuotient X) where normal s t hs ht hd := separatedNhds_iff_disjoint.2 <| by rw [← disjoint_comap_iff surjective_mk, comap_mk_nhdsSet, comap_mk_nhdsSet] exact disjoint_nhdsSet_nhdsSet (hs.preimage continuous_mk) (ht.preimage continuous_mk) @@ -1861,27 +1964,27 @@ instance [NormalSpace α] : NormalSpace (SeparationQuotient α) where end SeparationQuotient -variable (α) +variable (X) end Normality section CompletelyNormal -/-- A topological space `α` is a *completely normal Hausdorff space* if each subspace `s : Set α` is -a normal Hausdorff space. Equivalently, `α` is a `T₁` space and for any two sets `s`, `t` such that +/-- A topological space `X` is a *completely normal Hausdorff space* if each subspace `s : Set X` is +a normal Hausdorff space. Equivalently, `X` is a `T₁` space and for any two sets `s`, `t` such that `closure s` is disjoint with `t` and `s` is disjoint with `closure t`, there exist disjoint neighbourhoods of `s` and `t`. -/ -class T5Space (α : Type u) [TopologicalSpace α] extends T1Space α : Prop where +class T5Space (X : Type u) [TopologicalSpace X] extends T1Space X : Prop where /-- If `closure s` is disjoint with `t` and `s` is disjoint with `closure t`, then `s` and `t` admit disjoint neighbourhoods. -/ completely_normal : - ∀ ⦃s t : Set α⦄, Disjoint (closure s) t → Disjoint s (closure t) → Disjoint (𝓝ˢ s) (𝓝ˢ t) + ∀ ⦃s t : Set X⦄, Disjoint (closure s) t → Disjoint s (closure t) → Disjoint (𝓝ˢ s) (𝓝ˢ t) #align t5_space T5Space export T5Space (completely_normal) -theorem Embedding.t5Space [TopologicalSpace β] [T5Space β] {e : α → β} (he : Embedding e) : - T5Space α := by +theorem Embedding.t5Space [TopologicalSpace Y] [T5Space Y] {e : X → Y} (he : Embedding e) : + T5Space X := by haveI := he.t1Space refine' ⟨fun s t hd₁ hd₂ => _⟩ simp only [he.toInducing.nhdsSet_eq_comap] @@ -1893,12 +1996,12 @@ theorem Embedding.t5Space [TopologicalSpace β] [T5Space β] {e : α → β} (he #align embedding.t5_space Embedding.t5Space /-- A subspace of a `T₅` space is a `T₅` space. -/ -instance [T5Space α] {p : α → Prop} : T5Space { x // p x } := +instance [T5Space X] {p : X → Prop} : T5Space { x // p x } := embedding_subtype_val.t5Space -- see Note [lower instance priority] /-- A `T₅` space is a `T₄` space. -/ -instance (priority := 100) T5Space.toT4Space [T5Space α] : T4Space α where +instance (priority := 100) T5Space.toT4Space [T5Space X] : T4Space X where normal s t hs ht hd := separatedNhds_iff_disjoint.2 <| completely_normal (by rwa [hs.closure_eq]) (by rwa [ht.closure_eq]) #align t5_space.to_normal_space T5Space.toT4Space @@ -1909,10 +2012,10 @@ open SeparationQuotient We don't have typeclasses for completely normal spaces (without T₁ assumption) and R₀ spaces, so we use `T5Space` for assumption and for conclusion. -One can prove this using a homeomorphism between `α` and `SeparationQuotient α`. +One can prove this using a homeomorphism between `X` and `SeparationQuotient X`. We give an alternative proof of the `completely_normal` axiom -that works without assuming that `α` is a T₁ space. -/ -instance [T5Space α] : T5Space (SeparationQuotient α) where +that works without assuming that `X` is a T₁ space. -/ +instance [T5Space X] : T5Space (SeparationQuotient X) where completely_normal s t hd₁ hd₂ := by rw [← disjoint_comap_iff surjective_mk, comap_mk_nhdsSet, comap_mk_nhdsSet] apply T5Space.completely_normal <;> rw [← preimage_mk_closure] @@ -1920,40 +2023,40 @@ instance [T5Space α] : T5Space (SeparationQuotient α) where end CompletelyNormal -/-- In a compact t2 space, the connected component of a point equals the intersection of all +/-- In a compact T₂ space, the connected component of a point equals the intersection of all its clopen neighbourhoods. -/ -theorem connectedComponent_eq_iInter_clopen [T2Space α] [CompactSpace α] (x : α) : - connectedComponent x = ⋂ Z : { Z : Set α // IsClopen Z ∧ x ∈ Z }, Z := by +theorem connectedComponent_eq_iInter_clopen [T2Space X] [CompactSpace X] (x : X) : + connectedComponent x = ⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, s := by apply Subset.antisymm connectedComponent_subset_iInter_clopen -- Reduce to showing that the clopen intersection is connected. - refine' IsPreconnected.subset_connectedComponent _ (mem_iInter.2 fun Z => Z.2.2) + refine' IsPreconnected.subset_connectedComponent _ (mem_iInter.2 fun s => s.2.2) -- We do this by showing that any disjoint cover by two closed sets implies -- that one of these closed sets must contain our whole thing. - -- To reduce to the case where the cover is disjoint on all of `α` we need that `s` is closed - have hs : @IsClosed α _ (⋂ Z : { Z : Set α // IsClopen Z ∧ x ∈ Z }, Z) := - isClosed_iInter fun Z => Z.2.1.2 + -- To reduce to the case where the cover is disjoint on all of `X` we need that `s` is closed + have hs : @IsClosed X _ (⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, s) := + isClosed_iInter fun s => s.2.1.2 rw [isPreconnected_iff_subset_of_fully_disjoint_closed hs] intro a b ha hb hab ab_disj -- Since our space is normal, we get two larger disjoint open sets containing the disjoint -- closed sets. If we can show that our intersection is a subset of any of these we can then -- "descend" this to show that it is a subset of either a or b. rcases normal_separation ha hb ab_disj with ⟨u, v, hu, hv, hau, hbv, huv⟩ - obtain ⟨Z, H⟩ : ∃ Z : Set α, IsClopen Z ∧ x ∈ Z ∧ Z ⊆ u ∪ v - /- Now we find a clopen set `Z` around `x`, contained in `u ∪ v`. We utilize the fact that + obtain ⟨s, H⟩ : ∃ s : Set X, IsClopen s ∧ x ∈ s ∧ s ⊆ u ∪ v + /- Now we find a clopen set `s` around `x`, contained in `u ∪ v`. We utilize the fact that `X \ u ∪ v` will be compact, so there must be some finite intersection of clopen neighbourhoods of `X` disjoint to it, but a finite intersection of clopen sets is clopen so we let this be our - `Z`. -/ + `s`. -/ · have H1 := (hu.union hv).isClosed_compl.isCompact.inter_iInter_nonempty - (fun Z : { Z : Set α // IsClopen Z ∧ x ∈ Z } => Z) fun Z => Z.2.1.2 + (fun s : { s : Set X // IsClopen s ∧ x ∈ s } => s) fun s => s.2.1.2 rw [← not_disjoint_iff_nonempty_inter, imp_not_comm, not_forall] at H1 cases' H1 (disjoint_compl_left_iff_subset.2 <| hab.trans <| union_subset_union hau hbv) - with Zi H2 - refine' ⟨⋂ U ∈ Zi, Subtype.val U, _, _, _⟩ - · exact isClopen_biInter_finset fun Z _ => Z.2.1 - · exact mem_iInter₂.2 fun Z _ => Z.2.2 + with si H2 + refine' ⟨⋂ U ∈ si, Subtype.val U, _, _, _⟩ + · exact isClopen_biInter_finset fun s _ => s.2.1 + · exact mem_iInter₂.2 fun s _ => s.2.2 · rwa [← disjoint_compl_left_iff_subset, disjoint_iff_inter_eq_empty, ← not_nonempty_iff_eq_empty] - -- So, we get a disjoint decomposition `Z = Z ∩ u ∪ Z ∩ v` of clopen sets. The intersection of all + -- So, we get a disjoint decomposition `s = s ∩ u ∪ s ∩ v` of clopen sets. The intersection of all -- clopen neighbourhoods will then lie in whichever of u or v x lies in and hence will be a subset -- of either a or b. · have H1 := isClopen_inter_of_disjoint_cover_clopen H.1 H.2.2 hu hv huv @@ -1961,26 +2064,26 @@ theorem connectedComponent_eq_iInter_clopen [T2Space α] [CompactSpace α] (x : have H2 := isClopen_inter_of_disjoint_cover_clopen H.1 H.2.2 hv hu huv.symm by_cases hxu : x ∈ u <;> [left; right] -- The x ∈ u case. - · suffices ⋂ Z : { Z : Set α // IsClopen Z ∧ x ∈ Z }, ↑Z ⊆ u + · suffices ⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, ↑s ⊆ u from Disjoint.left_le_of_le_sup_right hab (huv.mono this hbv) - · apply Subset.trans _ (inter_subset_right Z u) - exact iInter_subset (fun Z : { Z : Set α // IsClopen Z ∧ x ∈ Z } => Z.1) - ⟨Z ∩ u, H1, mem_inter H.2.1 hxu⟩ + · apply Subset.trans _ (inter_subset_right s u) + exact iInter_subset (fun s : { s : Set X // IsClopen s ∧ x ∈ s } => s.1) + ⟨s ∩ u, H1, mem_inter H.2.1 hxu⟩ -- If x ∉ u, we get x ∈ v since x ∈ u ∪ v. The rest is then like the x ∈ u case. · have h1 : x ∈ v := (hab.trans (union_subset_union hau hbv) (mem_iInter.2 fun i => i.2.2)).resolve_left hxu - suffices ⋂ Z : { Z : Set α // IsClopen Z ∧ x ∈ Z }, ↑Z ⊆ v + suffices ⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, ↑s ⊆ v from (huv.symm.mono this hau).left_le_of_le_sup_left hab - · refine Subset.trans ?_ (inter_subset_right Z v) - exact iInter_subset (fun Z : { Z : Set α // IsClopen Z ∧ x ∈ Z } => Z.1) - ⟨Z ∩ v, H2, mem_inter H.2.1 h1⟩ + · refine Subset.trans ?_ (inter_subset_right s v) + exact iInter_subset (fun s : { s : Set X // IsClopen s ∧ x ∈ s } => s.1) + ⟨s ∩ v, H2, mem_inter H.2.1 h1⟩ #align connected_component_eq_Inter_clopen connectedComponent_eq_iInter_clopen section Profinite /-- A T1 space with a clopen basis is totally separated. -/ -theorem totallySeparatedSpace_of_t1_of_basis_clopen [T1Space α] - (h : IsTopologicalBasis { s : Set α | IsClopen s }) : TotallySeparatedSpace α := by +theorem totallySeparatedSpace_of_t1_of_basis_clopen [T1Space X] + (h : IsTopologicalBasis { s : Set X | IsClopen s }) : TotallySeparatedSpace X := by constructor rintro x - y - hxy rcases h.mem_nhds_iff.mp (isOpen_ne.mem_nhds hxy) with ⟨U, hU, hxU, hyU⟩ @@ -1988,44 +2091,44 @@ theorem totallySeparatedSpace_of_t1_of_basis_clopen [T1Space α] disjoint_compl_right⟩ #align totally_separated_space_of_t1_of_basis_clopen totallySeparatedSpace_of_t1_of_basis_clopen -variable [T2Space α] [CompactSpace α] +variable [T2Space X] [CompactSpace X] /-- A compact Hausdorff space is totally disconnected if and only if it is totally separated, this is also true for locally compact spaces. -/ -theorem compact_t2_tot_disc_iff_tot_sep : TotallyDisconnectedSpace α ↔ TotallySeparatedSpace α := by +theorem compact_t2_tot_disc_iff_tot_sep : TotallyDisconnectedSpace X ↔ TotallySeparatedSpace X := by refine ⟨fun h => ⟨fun x _ y _ => ?_⟩, @TotallySeparatedSpace.totallyDisconnectedSpace _ _⟩ contrapose! intro hyp suffices x ∈ connectedComponent y by simpa [totallyDisconnectedSpace_iff_connectedComponent_singleton.1 h y, mem_singleton_iff] rw [connectedComponent_eq_iInter_clopen, mem_iInter] - rintro ⟨w : Set α, hw : IsClopen w, hy : y ∈ w⟩ + rintro ⟨w : Set X, hw : IsClopen w, hy : y ∈ w⟩ by_contra hx exact hyp wᶜ w hw.2.isOpen_compl hw.1 hx hy (@isCompl_compl _ w _).symm.codisjoint.top_le disjoint_compl_left #align compact_t2_tot_disc_iff_tot_sep compact_t2_tot_disc_iff_tot_sep -variable [TotallyDisconnectedSpace α] +variable [TotallyDisconnectedSpace X] /-- A totally disconnected compact Hausdorff space is totally separated. -/ -instance : TotallySeparatedSpace α := compact_t2_tot_disc_iff_tot_sep.mp inferInstance +instance : TotallySeparatedSpace X := compact_t2_tot_disc_iff_tot_sep.mp inferInstance -theorem nhds_basis_clopen (x : α) : (𝓝 x).HasBasis (fun s : Set α => x ∈ s ∧ IsClopen s) id := +theorem nhds_basis_clopen (x : X) : (𝓝 x).HasBasis (fun s : Set X => x ∈ s ∧ IsClopen s) id := ⟨fun U => by constructor · have hx : connectedComponent x = {x} := totallyDisconnectedSpace_iff_connectedComponent_singleton.mp ‹_› x rw [connectedComponent_eq_iInter_clopen] at hx intro hU - let N := { Z // IsClopen Z ∧ x ∈ Z } - suffices : ∃ Z : N, Z.val ⊆ U + let N := { s // IsClopen s ∧ x ∈ s } + suffices : ∃ s : N, s.val ⊆ U · rcases this with ⟨⟨s, hs, hs'⟩, hs''⟩; exact ⟨s, ⟨hs', hs⟩, hs''⟩ haveI : Nonempty N := ⟨⟨univ, isClopen_univ, mem_univ x⟩⟩ - have hNcl : ∀ Z : N, IsClosed Z.val := fun Z => Z.property.1.2 - have hdir : Directed Superset fun Z : N => Z.val := by + have hNcl : ∀ s : N, IsClosed s.val := fun s => s.property.1.2 + have hdir : Directed Superset fun s : N => s.val := by rintro ⟨s, hs, hxs⟩ ⟨t, ht, hxt⟩ exact ⟨⟨s ∩ t, hs.inter ht, ⟨hxs, hxt⟩⟩, inter_subset_left s t, inter_subset_right s t⟩ - have h_nhd : ∀ y ∈ ⋂ Z : N, Z.val, U ∈ 𝓝 y := fun y y_in => by + have h_nhd : ∀ y ∈ ⋂ s : N, s.val, U ∈ 𝓝 y := fun y y_in => by erw [hx, mem_singleton_iff] at y_in rwa [y_in] exact exists_subset_nhds_of_compactSpace hdir hNcl h_nhd @@ -2034,7 +2137,7 @@ theorem nhds_basis_clopen (x : α) : (𝓝 x).HasBasis (fun s : Set α => x ∈ exact ⟨V, hUV, V_op, hxV⟩⟩ #align nhds_basis_clopen nhds_basis_clopen -theorem isTopologicalBasis_clopen : IsTopologicalBasis { s : Set α | IsClopen s } := by +theorem isTopologicalBasis_clopen : IsTopologicalBasis { s : Set X | IsClopen s } := by apply isTopologicalBasis_of_open_of_nhds fun U (hU : IsClopen U) => hU.1 intro x U hxU U_op have : U ∈ 𝓝 x := IsOpen.mem_nhds U_op hxU @@ -2045,8 +2148,8 @@ theorem isTopologicalBasis_clopen : IsTopologicalBasis { s : Set α | IsClopen s /-- Every member of an open set in a compact Hausdorff totally disconnected space is contained in a clopen set contained in the open set. -/ -theorem compact_exists_clopen_in_open {x : α} {U : Set α} (is_open : IsOpen U) (memU : x ∈ U) : - ∃ V : Set α, IsClopen V ∧ x ∈ V ∧ V ⊆ U := +theorem compact_exists_clopen_in_open {x : X} {U : Set X} (is_open : IsOpen U) (memU : x ∈ U) : + ∃ V : Set X, IsClopen V ∧ x ∈ V ∧ V ⊆ U := isTopologicalBasis_clopen.mem_nhds_iff.1 (is_open.mem_nhds memU) #align compact_exists_clopen_in_open compact_exists_clopen_in_open @@ -2100,8 +2203,8 @@ theorem loc_compact_t2_tot_disc_iff_tot_sep : end LocallyCompact -/-- `ConnectedComponents α` is Hausdorff when `α` is Hausdorff and compact -/ -instance ConnectedComponents.t2 [T2Space α] [CompactSpace α] : T2Space (ConnectedComponents α) := by +/-- `ConnectedComponents X` is Hausdorff when `X` is Hausdorff and compact -/ +instance ConnectedComponents.t2 [T2Space X] [CompactSpace X] : T2Space (ConnectedComponents X) := by -- Proof follows that of: https://stacks.math.columbia.edu/tag/0900 -- Fix 2 distinct connected components, with points a and b refine ⟨ConnectedComponents.surjective_coe.forall₂.2 fun a b ne => ?_⟩ @@ -2110,16 +2213,16 @@ instance ConnectedComponents.t2 [T2Space α] [CompactSpace α] : T2Space (Connec -- write ↑b as the intersection of all clopen subsets containing it rw [connectedComponent_eq_iInter_clopen b, disjoint_iff_inter_eq_empty] at h -- Now we show that this can be reduced to some clopen containing `↑b` being disjoint to `↑a` - obtain ⟨U, V, hU, ha, hb, rfl⟩ : ∃ (U : Set α) (V : Set (ConnectedComponents α)), + obtain ⟨U, V, hU, ha, hb, rfl⟩ : ∃ (U : Set X) (V : Set (ConnectedComponents X)), IsClopen U ∧ connectedComponent a ∩ U = ∅ ∧ connectedComponent b ⊆ U ∧ (↑) ⁻¹' V = U := by have h := - (isClosed_connectedComponent (α := α)).isCompact.elim_finite_subfamily_closed - _ (fun Z : { Z : Set α // IsClopen Z ∧ b ∈ Z } => Z.2.1.2) h + (isClosed_connectedComponent (α := X)).isCompact.elim_finite_subfamily_closed + _ (fun s : { s : Set X // IsClopen s ∧ b ∈ s } => s.2.1.2) h cases' h with fin_a ha -- This clopen and its complement will separate the connected components of `a` and `b` - set U : Set α := ⋂ (i : { Z // IsClopen Z ∧ b ∈ Z }) (_ : i ∈ fin_a), i + set U : Set X := ⋂ (i : { s // IsClopen s ∧ b ∈ s }) (_ : i ∈ fin_a), i have hU : IsClopen U := isClopen_biInter_finset fun i _ => i.2.1 - exact ⟨U, (↑) '' U, hU, ha, subset_iInter₂ fun Z _ => Z.2.1.connectedComponent_subset Z.2.2, + exact ⟨U, (↑) '' U, hU, ha, subset_iInter₂ fun s _ => s.2.1.connectedComponent_subset s.2.2, (connectedComponents_preimage_image U).symm ▸ hU.biUnion_connectedComponent_eq⟩ rw [ConnectedComponents.quotientMap_coe.isClopen_preimage] at hU refine' ⟨Vᶜ, V, hU.compl.isOpen, hU.isOpen, _, hb mem_connectedComponent, disjoint_compl_left⟩ diff --git a/Mathlib/Topology/Sequences.lean b/Mathlib/Topology/Sequences.lean index 4c9800fd09780..c91386ce66cea 100644 --- a/Mathlib/Topology/Sequences.lean +++ b/Mathlib/Topology/Sequences.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Jan-David Salchow. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jan-David Salchow, Patrick Massot, Yury Kudryashov -/ -import Mathlib.Topology.MetricSpace.Basic +import Mathlib.Topology.MetricSpace.Bounded #align_import topology.sequences from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" @@ -163,7 +163,8 @@ theorem FrechetUrysohnSpace.of_seq_tendsto_imp_tendsto · exact subset_seqClosure hx · obtain ⟨u, hux, hus⟩ : ∃ u : ℕ → X, Tendsto u atTop (𝓝 x) ∧ ∃ᶠ x in atTop, u x ∈ s · simpa only [ContinuousAt, hx, tendsto_nhds_true, (· ∘ ·), ← not_frequently, exists_prop, - ← mem_closure_iff_frequently, hcx, imp_false, not_forall, not_not] using h (· ∉ s) x + ← mem_closure_iff_frequently, hcx, imp_false, not_forall, not_not, not_false_eq_true, + not_true_eq_false] using h (· ∉ s) x rcases extraction_of_frequently_atTop hus with ⟨φ, φ_mono, hφ⟩ exact ⟨u ∘ φ, hφ, hux.comp φ_mono.tendsto_atTop⟩ #align frechet_urysohn_space.of_seq_tendsto_imp_tendsto FrechetUrysohnSpace.of_seq_tendsto_imp_tendsto diff --git a/Mathlib/Topology/Sheaves/Presheaf.lean b/Mathlib/Topology/Sheaves/Presheaf.lean index 37d1f9c6b500f..ede11350b74cb 100644 --- a/Mathlib/Topology/Sheaves/Presheaf.lean +++ b/Mathlib/Topology/Sheaves/Presheaf.lean @@ -335,7 +335,7 @@ def pullbackObjObjOfImageOpen {X Y : TopCat.{v}} (f : X ⟶ Y) (ℱ : Y.Presheaf · refine' (homOfLE _).op apply (Set.image_subset f s.pt.hom.unop.le).trans exact Set.image_preimage.l_u_le (SetLike.coe s.pt.left.unop) - · simp + · simp [autoParam] -- porting note : add `fac`, `uniq` manually fac := fun _ _ => by ext; simp uniq := fun _ _ _ => by ext; simp } diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 66fc011284afb..cc4aec89b10f0 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -156,6 +156,14 @@ theorem HasCompactMulSupport.intro [T2Space α] {K : Set α} (hK : IsCompact K) #align has_compact_mul_support.intro HasCompactMulSupport.intro #align has_compact_support.intro HasCompactSupport.intro +@[to_additive] +theorem HasCompactMulSupport.intro' {K : Set α} (hK : IsCompact K) (h'K : IsClosed K) + (hfK : ∀ x, x ∉ K → f x = 1) : HasCompactMulSupport f := by + have : mulTSupport f ⊆ K := by + rw [← h'K.closure_eq] + apply closure_mono (mulSupport_subset_iff'.2 hfK) + exact IsCompact.of_isClosed_subset hK ( isClosed_mulTSupport f) this + @[to_additive] theorem HasCompactMulSupport.of_mulSupport_subset_isCompact [T2Space α] {K : Set α} (hK : IsCompact K) (h : mulSupport f ⊆ K) : HasCompactMulSupport f := @@ -179,11 +187,17 @@ theorem hasCompactMulSupport_iff_eventuallyEq : #align has_compact_mul_support_iff_eventually_eq hasCompactMulSupport_iff_eventuallyEq #align has_compact_support_iff_eventually_eq hasCompactSupport_iff_eventuallyEq +@[to_additive] +theorem isCompact_range_of_mulSupport_subset_isCompact [TopologicalSpace β] + (hf : Continuous f) {k : Set α} (hk : IsCompact k) (h'f : mulSupport f ⊆ k) : + IsCompact (range f) := by + cases' range_eq_image_or_of_mulSupport_subset h'f with h2 h2 <;> rw [h2] + exacts [hk.image hf, (hk.image hf).insert 1] + @[to_additive] theorem HasCompactMulSupport.isCompact_range [TopologicalSpace β] (h : HasCompactMulSupport f) (hf : Continuous f) : IsCompact (range f) := by - cases' range_eq_image_mulTSupport_or f with h2 h2 <;> rw [h2] - exacts [h.image hf, (h.image hf).insert 1] + apply isCompact_range_of_mulSupport_subset_isCompact hf h (subset_mulTSupport f) #align has_compact_mul_support.is_compact_range HasCompactMulSupport.isCompact_range #align has_compact_support.is_compact_range HasCompactSupport.isCompact_range diff --git a/Mathlib/Topology/UniformSpace/Compact.lean b/Mathlib/Topology/UniformSpace/Compact.lean index 78a829727c238..6d2608d55e4cd 100644 --- a/Mathlib/Topology/UniformSpace/Compact.lean +++ b/Mathlib/Topology/UniformSpace/Compact.lean @@ -261,6 +261,33 @@ theorem Continuous.tendstoUniformly [WeaklyLocallyCompactSpace α] [CompactSpace this.tendstoUniformly hxK #align continuous.tendsto_uniformly Continuous.tendstoUniformly +/-- In a product space `α × β`, assume that a function `f` is continuous on `s × k` where `k` is +compact. Then, along the fiber above any `q ∈ s`, `f` is transversely uniformly continuous, i.e., +if `p ∈ s` is close enough to `q`, then `f p x` is uniformly close to `f q x` for all `x ∈ k`. -/ +lemma IsCompact.mem_uniformity_of_prod + {α β E : Type*} [TopologicalSpace α] [TopologicalSpace β] [UniformSpace E] + {f : α → β → E} {s : Set α} {k : Set β} {q : α} {u : Set (E × E)} + (hk : IsCompact k) (hf : ContinuousOn f.uncurry (s ×ˢ k)) (hq : q ∈ s) (hu : u ∈ 𝓤 E) : + ∃ v ∈ 𝓝[s] q, ∀ p ∈ v, ∀ x ∈ k, (f p x, f q x) ∈ u := by + apply hk.induction_on (p := fun t ↦ ∃ v ∈ 𝓝[s] q, ∀ p ∈ v, ∀ x ∈ t, (f p x, f q x) ∈ u) + · exact ⟨univ, univ_mem, by simp⟩ + · intro t' t ht't ⟨v, v_mem, hv⟩ + exact ⟨v, v_mem, fun p hp x hx ↦ hv p hp x (ht't hx)⟩ + · intro t t' ⟨v, v_mem, hv⟩ ⟨v', v'_mem, hv'⟩ + refine ⟨v ∩ v', inter_mem v_mem v'_mem, fun p hp x hx ↦ ?_⟩ + rcases hx with h'x|h'x + · exact hv p hp.1 x h'x + · exact hv' p hp.2 x h'x + · rcases comp_symm_of_uniformity hu with ⟨u', u'_mem, u'_symm, hu'⟩ + intro x hx + obtain ⟨v, hv, w, hw, hvw⟩ : + ∃ v ∈ 𝓝[s] q, ∃ w ∈ 𝓝[k] x, v ×ˢ w ⊆ f.uncurry ⁻¹' {z | (f q x, z) ∈ u'} := + mem_nhdsWithin_prod_iff.1 (hf (q, x) ⟨hq, hx⟩ (mem_nhds_left (f q x) u'_mem)) + refine ⟨w, hw, v, hv, fun p hp y hy ↦ ?_⟩ + have A : (f q x, f p y) ∈ u' := hvw (⟨hp, hy⟩ : (p, y) ∈ v ×ˢ w) + have B : (f q x, f q y) ∈ u' := hvw (⟨mem_of_mem_nhdsWithin hq hv, hy⟩ : (q, y) ∈ v ×ˢ w) + exact hu' (prod_mk_mem_compRel (u'_symm A) B) + section UniformConvergence /-- An equicontinuous family of functions defined on a compact uniform space is automatically diff --git a/Mathlib/Topology/UnitInterval.lean b/Mathlib/Topology/UnitInterval.lean index ed47470081007..7187c5021eeda 100644 --- a/Mathlib/Topology/UnitInterval.lean +++ b/Mathlib/Topology/UnitInterval.lean @@ -163,6 +163,8 @@ theorem le_one' {t : I} : t ≤ 1 := t.2.2 #align unit_interval.le_one' unitInterval.le_one' +instance : NeZero (1 : I) := ⟨fun h ↦ one_ne_zero <| congrArg Subtype.val h⟩ + theorem mul_pos_mem_iff {a t : ℝ} (ha : 0 < a) : a * t ∈ I ↔ t ∈ Set.Icc (0 : ℝ) (1 / a) := by constructor <;> rintro ⟨h₁, h₂⟩ <;> constructor · exact nonneg_of_mul_nonneg_right h₁ ha diff --git a/Mathlib/Util/PiNotation.lean b/Mathlib/Util/Delaborators.lean similarity index 70% rename from Mathlib/Util/PiNotation.lean rename to Mathlib/Util/Delaborators.lean index 8856801e45d79..e14d724484811 100644 --- a/Mathlib/Util/PiNotation.lean +++ b/Mathlib/Util/Delaborators.lean @@ -23,7 +23,7 @@ also be written as `(x : α) → β x`. -/ -- A direct copy of forall notation but with `Π`/`Pi` instead of `∀`/`Forall`. @[term_parser] def piNotation := leading_parser:leadPrec - unicodeSymbol "Π" "Pi" >> + unicodeSymbol "Π" "PiType" >> many1 (ppSpace >> (binderIdent <|> bracketedBinder)) >> optType >> ", " >> termParser @@ -60,6 +60,16 @@ def delabPi : Delab := whenPPOption Lean.getPPNotation do if x == y then `(∀ $x:ident ≤ $z, $body) else pure stx | `(Π ($i:ident : $_), $j:ident ∈ $s → $body) => if i == j then `(Π $i:ident ∈ $s, $body) else pure stx + | `(∀ ($i:ident : $_), $j:ident ∉ $s → $body) => + if i == j then `(∀ $i:ident ∉ $s, $body) else pure stx + | `(∀ ($i:ident : $_), $j:ident ⊆ $s → $body) => + if i == j then `(∀ $i:ident ⊆ $s, $body) else pure stx + | `(∀ ($i:ident : $_), $j:ident ⊂ $s → $body) => + if i == j then `(∀ $i:ident ⊂ $s, $body) else pure stx + | `(∀ ($i:ident : $_), $j:ident ⊇ $s → $body) => + if i == j then `(∀ $i:ident ⊇ $s, $body) else pure stx + | `(∀ ($i:ident : $_), $j:ident ⊃ $s → $body) => + if i == j then `(∀ $i:ident ⊃ $s, $body) else pure stx | _ => pure stx /-- Override the Lean 4 pi notation delaborator with one that uses `Π` and prints @@ -122,9 +132,35 @@ def exists_delab : Delab := whenPPOption Lean.getPPNotation do | `(∃ $x:ident, $y:ident ≤ $z ∧ $body) | `(∃ ($x:ident : $_), $y:ident ≤ $z ∧ $body) => if x == y then `(∃ $x:ident ≤ $z, $body) else pure stx + | `(∃ $x:ident, $y:ident ∉ $z ∧ $body) + | `(∃ ($x:ident : $_), $y:ident ∉ $z ∧ $body) => do + if x == y then `(∃ $x:ident ∉ $z, $body) else pure stx + | `(∃ $x:ident, $y:ident ⊆ $z ∧ $body) + | `(∃ ($x:ident : $_), $y:ident ⊆ $z ∧ $body) => + if x == y then `(∃ $x:ident ⊆ $z, $body) else pure stx + | `(∃ $x:ident, $y:ident ⊂ $z ∧ $body) + | `(∃ ($x:ident : $_), $y:ident ⊂ $z ∧ $body) => + if x == y then `(∃ $x:ident ⊂ $z, $body) else pure stx + | `(∃ $x:ident, $y:ident ⊇ $z ∧ $body) + | `(∃ ($x:ident : $_), $y:ident ⊇ $z ∧ $body) => + if x == y then `(∃ $x:ident ⊇ $z, $body) else pure stx + | `(∃ $x:ident, $y:ident ⊃ $z ∧ $body) + | `(∃ ($x:ident : $_), $y:ident ⊃ $z ∧ $body) => + if x == y then `(∃ $x:ident ⊃ $z, $body) else pure stx | _ => pure stx - -- Merging match stx with - | `(∃ $group:bracketedExplicitBinders, ∃ $groups*, $body) => `(∃ $group $groups*, $body) + | `(∃ $group:bracketedExplicitBinders, ∃ $[$groups:bracketedExplicitBinders]*, $body) => + `(∃ $group $groups*, $body) + | `(∃ $b:binderIdent, ∃ $[$bs:binderIdent]*, $body) => `(∃ $b:binderIdent $[$bs]*, $body) | _ => pure stx end existential + +open Lean Lean.PrettyPrinter.Delaborator + +/-- Delaborator for `∉`. -/ +@[delab app.Not] def delab_not_in := whenPPOption Lean.getPPNotation do + let #[f] := (← SubExpr.getExpr).getAppArgs | failure + guard <| f.isAppOfArity ``Membership.mem 5 + let stx₁ ← SubExpr.withAppArg <| SubExpr.withNaryArg 3 delab + let stx₂ ← SubExpr.withAppArg <| SubExpr.withNaryArg 4 delab + return ← `($stx₁ ∉ $stx₂) diff --git a/Mathlib/Util/DischargerAsTactic.lean b/Mathlib/Util/DischargerAsTactic.lean index 6c7b95a8562a9..1b5e7e447e8de 100644 --- a/Mathlib/Util/DischargerAsTactic.lean +++ b/Mathlib/Util/DischargerAsTactic.lean @@ -3,6 +3,8 @@ Copyright (c) 2023 Alex J. Best. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best -/ +import Lean.Meta.Tactic.Simp.Main +import Lean.Elab.Tactic.Basic import Std.Tactic.Exact /-! diff --git a/docs/references.bib b/docs/references.bib index 286547b72e0c6..416951e60acc0 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -425,6 +425,17 @@ @Book{ bourbaki2023 url = {https://doi.org/10.1007/978-3-031-19505-1} } +@Article{ bowlerGeschke2015, + author = {Bowler, Nathan and Geschke, Stefan}, + year = {2015}, + month = {10}, + pages = {1}, + title = {Self-dual uniform matroids on infinite sets}, + volume = {144}, + journal = {Proceedings of the American Mathematical Society}, + doi = {10.1090/proc/12667} +} + @Book{ boydVandenberghe2004, author = {Stephen P. Boyd and Lieven Vandenberghe}, title = {Convex Optimization}, @@ -448,6 +459,19 @@ @Book{ brodmannsharp13 mrclass = {13D45 (13-01)} } +@Article{ bruhnDiestelKriesselPendavinghWollan2013, + author = {Henning Bruhn and Reinhard Diestel and Matthias Kriesell + and Rudi Pendavingh and Paul Wollan}, + title = {Axioms for infinite matroids}, + journal = {Advances in Mathematics}, + volume = {239}, + pages = {18-46}, + year = {2013}, + issn = {0001-8708}, + doi = {https://doi.org/10.1016/j.aim.2013.01.011}, + url = {https://www.sciencedirect.com/science/article/pii/S0001870813000261} +} + @Book{ cabreragarciarodriguezpalacios2014, author = {Miguel {Cabrera Garc\'{\i}a} and \'Angel {Rodr\'{\i}guez Palacios}}, @@ -2119,6 +2143,17 @@ @Article{ orosi2018faulhaber zbl = {1411.41023} } +@Book{ oxley2011, + author = {Oxley, James}, + title = {Matroid Theory}, + publisher = {Oxford University Press}, + year = {2011}, + month = {02}, + isbn = {9780198566946}, + doi = {10.1093/acprof:oso/9780198566946.001.0001}, + url = {https://doi.org/10.1093/acprof:oso/9780198566946.001.0001} +} + @InCollection{ petridis2014, author = {Petridis, G.}, title = {The {Pl{\"u}nnecke}-{Ruzsa} inequality: an overview}, diff --git a/lake-manifest.json b/lake-manifest.json index d453e32198622..df78cb9375b9e 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -2,9 +2,17 @@ "packagesDir": "lake-packages", "packages": [{"git": + {"url": "https://github.com/leanprover/std4", + "subDir?": null, + "rev": "4167e5e4f60a47906d6fc42000b64cd93205f710", + "opts": {}, + "name": "std", + "inputRev?": "nightly-testing", + "inherited": false}}, + {"git": {"url": "https://github.com/leanprover-community/quote4", "subDir?": null, - "rev": "a387c0eb611857e2460cf97a8e861c944286e6b2", + "rev": "511eb96eca98a7474683b8ae55193a7e7c51bc39", "opts": {}, "name": "Qq", "inputRev?": "master", @@ -12,10 +20,10 @@ {"git": {"url": "https://github.com/leanprover-community/aesop", "subDir?": null, - "rev": "9dc4a1097a690216eaa7cf2d2290efd447e60d7a", + "rev": "a76881668efab6e48a4f3333ad7ba894e8095b04", "opts": {}, "name": "aesop", - "inputRev?": "master", + "inputRev?": "nightly-testing", "inherited": false}}, {"git": {"url": "https://github.com/leanprover/lean4-cli", @@ -28,17 +36,9 @@ {"git": {"url": "https://github.com/leanprover-community/ProofWidgets4", "subDir?": null, - "rev": "5382e38eca1e2537d75d4c4705a9e744424b0037", + "rev": "f1a5c7808b001305ba07d8626f45ee054282f589", "opts": {}, "name": "proofwidgets", - "inputRev?": "v0.0.19", - "inherited": false}}, - {"git": - {"url": "https://github.com/leanprover/std4", - "subDir?": null, - "rev": "b8693516737c0e47c528cf801d53534244ea81a5", - "opts": {}, - "name": "std", - "inputRev?": "nightly-testing", + "inputRev?": "v0.0.21", "inherited": false}}], "name": "mathlib"} diff --git a/lakefile.lean b/lakefile.lean index 09ae705d31893..87b65e2ed5d2b 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -48,9 +48,9 @@ require «doc-gen4» from git "https://github.com/leanprover/doc-gen4" @ "main" require std from git "https://github.com/leanprover/std4" @ "nightly-testing" require Qq from git "https://github.com/leanprover-community/quote4" @ "master" -require aesop from git "https://github.com/leanprover-community/aesop" @ "master" +require aesop from git "https://github.com/leanprover-community/aesop" @ "nightly-testing" require Cli from git "https://github.com/leanprover/lean4-cli" @ "nightly" -require proofwidgets from git "https://github.com/leanprover-community/ProofWidgets4" @ "v0.0.19" +require proofwidgets from git "https://github.com/leanprover-community/ProofWidgets4" @ "v0.0.21" lean_lib Cache where moreLeanArgs := moreLeanArgs diff --git a/scripts/bench/temci-config.run.yml b/scripts/bench/temci-config.run.yml index 6a26fc0b61c0a..3ed508e561d04 100644 --- a/scripts/bench/temci-config.run.yml +++ b/scripts/bench/temci-config.run.yml @@ -18,6 +18,15 @@ properties: ['wall-clock', 'instructions'] cmd: | make lint +- attributes: + description: open Mathlib + run_config: + runner: perf_stat + perf_stat: + properties: ['wall-clock', 'instructions'] + rusage_properties: ['maxrss'] + cmd: | + LEAN_PATH=$(lake print-paths --no-build Mathlib | jq -r '.oleanPath | join(":")') lean Mathlib.lean - attributes: description: size run_config: diff --git a/scripts/fix-by-linebreaks.sh b/scripts/fix-by-linebreaks.sh index 3dbce4e0a3742..f0fdd8aa12f08 100755 --- a/scripts/fix-by-linebreaks.sh +++ b/scripts/fix-by-linebreaks.sh @@ -1,5 +1,5 @@ #!/usr/bin/env bash -# Modify all lean files in mathlib to put the "by" in lines that only contain " by" at the end of the previous line, +# Modify all lean files in mathlib to put the "by" in lines that only contain " by" at the end of the previous line, # when the previous line with " by" appended is not longer than 100 characters. grep -lr "^ by\$" Mathlib | xargs -n 1 awk -i inplace '{do {{if (match($0, "^ by$") && length(p) < 98 && (!(match(p, "^[ \t]*--.*$")))) {p=p " by";} else {if (NR!=1) {print p}; p=$0}}} while (getline == 1) if (getline==0) print p}' diff --git a/scripts/lean-pr-testing-comments.sh b/scripts/lean-pr-testing-comments.sh index f1d346dc61d8a..0311c871b3544 100755 --- a/scripts/lean-pr-testing-comments.sh +++ b/scripts/lean-pr-testing-comments.sh @@ -66,11 +66,11 @@ if [[ "$branch_name" =~ ^lean-pr-testing-([0-9]+)$ ]]; then existing_comment=$(curl -L -s -H "Authorization: token $TOKEN" \ -H "Accept: application/vnd.github.v3+json" \ "https://api.github.com/repos/leanprover/lean4/issues/$pr_number/comments" \ - | jq '.[] | select(.body | startswith("- ✅ Mathlib") or startswith("- ❌ Mathlib") or startswith("- 💥 Mathlib") or startswith("- 🟡 Mathlib"))') + | jq '.[] | select(.body | startswith("- ❗ Mathlib") or startswith("- ✅ Mathlib") or startswith("- ❌ Mathlib") or startswith("- 💥 Mathlib") or startswith("- 🟡 Mathlib"))') existing_comment_id=$(echo "$existing_comment" | jq -r .id) existing_comment_body=$(echo "$existing_comment" | jq -r .body) - branch="[lean-pr-testing-$pr_number](https://github.com/leanprover-community/mathlib4/compare/master...lean-pr-testing-$pr_number)" + branch="[lean-pr-testing-$pr_number](https://github.com/leanprover-community/mathlib4/compare/nightly-testing...lean-pr-testing-$pr_number)" # Depending on the success/failure, set the appropriate message if [ "$LINT_OUTCOME" == "cancelled" ] || [ "$TEST_OUTCOME" == "cancelled" ] || [ "$BUILD_OUTCOME" == "cancelled" ]; then message="- 🟡 Mathlib branch $branch build against this PR was cancelled. ($current_time) [View Log]($WORKFLOW_URL)" diff --git a/scripts/make_port_status.py b/scripts/make_port_status.py index b353d07c6c871..1fbb5f15d1fc8 100755 --- a/scripts/make_port_status.py +++ b/scripts/make_port_status.py @@ -1,5 +1,6 @@ #!/usr/bin/env python3 +import pytz import datetime import github import os @@ -120,7 +121,7 @@ def get_mathlib4_module_commit_info(contents): sync_prs = defaultdict(set) mathlib4repo = github.Github(github_token).get_repo("leanprover-community/mathlib4") for pr in mathlib4repo.get_pulls(state='open'): - if pr.created_at < datetime.datetime(2022, 12, 1, 0, 0, 0): + if pr.created_at < datetime.datetime(2022, 12, 1, 0, 0, 0, tzinfo=pytz.UTC): continue if 'no-source-header' in (l.name for l in pr.labels): continue @@ -145,7 +146,7 @@ def get_mathlib4_module_commit_info(contents): f = subprocess.run( ['git', 'cat-file', 'blob', f'port-status-pull/{num}:{l}'], capture_output=True) - import_, repo, commit = get_mathlib4_module_commit_info(f.stdout.decode()) + import_, repo, commit = get_mathlib4_module_commit_info(f.stdout.decode(encoding='utf8', errors='replace')) prs_of_import.setdefault(import_, []).append({'pr': num, 'repo': repo, 'commit': commit, 'fname': l}) COMMENTS_URL = "https://raw.githubusercontent.com/wiki/leanprover-community/mathlib4/port-comments.md" diff --git a/test/FieldSimp.lean b/test/FieldSimp.lean index f6df29b27ec1f..0a7ef8bbcdb8d 100644 --- a/test/FieldSimp.lean +++ b/test/FieldSimp.lean @@ -16,20 +16,17 @@ variable {R : Type _} [CommRing R] (a b c d e f g : R) (u₁ u₂ : Rˣ) /-- Check that `divp_add_divp_same` takes priority over `divp_add_divp`. -/ -example : a /ₚ u₁ + b /ₚ u₁ = (a + b) /ₚ u₁ := -by field_simp +example : a /ₚ u₁ + b /ₚ u₁ = (a + b) /ₚ u₁ := by field_simp /-- Check that `divp_sub_divp_same` takes priority over `divp_sub_divp`. -/ -example : a /ₚ u₁ - b /ₚ u₁ = (a - b) /ₚ u₁ := -by field_simp +example : a /ₚ u₁ - b /ₚ u₁ = (a - b) /ₚ u₁ := by field_simp /- Combining `eq_divp_iff_mul_eq` and `divp_eq_iff_mul_eq`. -/ -example : a /ₚ u₁ = b /ₚ u₂ ↔ a * u₂ = b * u₁ := -by field_simp +example : a /ₚ u₁ = b /ₚ u₂ ↔ a * u₂ = b * u₁ := by field_simp /-- Making sure inverses of units are rewritten properly. @@ -40,17 +37,14 @@ example : ↑u₁⁻¹ = 1 /ₚ u₁ := by field_simp Checking arithmetic expressions. -/ example : (f - (e + c * -(a /ₚ u₁) * b + d) - g) = - (f * u₁ - (e * u₁ + c * (-a) * b + d * u₁) - g * u₁) /ₚ u₁ := -by field_simp + (f * u₁ - (e * u₁ + c * (-a) * b + d * u₁) - g * u₁) /ₚ u₁ := by field_simp /-- Division of units. -/ -example : a /ₚ (u₁ / u₂) = a * u₂ /ₚ u₁ := -by field_simp +example : a /ₚ (u₁ / u₂) = a * u₂ /ₚ u₁ := by field_simp -example : a /ₚ u₁ /ₚ u₂ = a /ₚ (u₂ * u₁) := -by field_simp +example : a /ₚ u₁ /ₚ u₂ = a /ₚ (u₂ * u₁) := by field_simp /-- Test that the discharger can clear nontrivial denominators in ℚ. @@ -60,6 +54,12 @@ example (x : ℚ) (h₀ : x ≠ 0) : field_simp ring +/-- Use a custom discharger -/ +example (x : ℚ) (h₀ : x ≠ 0) : + (4 / x)⁻¹ * ((3 * x^3) / x)^2 * ((1 / (2 * x))⁻¹)^3 = 18 * x^8 := by + field_simp (discharger := simp; assumption) + ring + example {x y z w : ℚ} (h : x / y = z / w) (hy : y ≠ 0) (hw : w ≠ 0) : x * w = z * y := by field_simp at h assumption diff --git a/test/MLList.lean b/test/MLList.lean index 031f47dfb0895..490d201c7691d 100644 --- a/test/MLList.lean +++ b/test/MLList.lean @@ -5,6 +5,7 @@ Authors: Scott Morrison -/ import Std.Data.MLList.Basic import Mathlib.Control.Basic +import Mathlib.Tactic.RunCmd @[reducible] def S (α : Type) := StateT (List Nat) Option α def append (x : Nat) : S Unit := @@ -18,7 +19,7 @@ def F : Nat → S Nat open Lean -#eval show MetaM Unit from do +run_cmd Lean.Elab.Command.liftTermElabM do -- Note that `fix` fails if any invocation of `F` fails. -- This is different from previous behaviour, where it just terminated the lazy list. -- Hence we must use `.takeAsList 11` here rather than `.force`. @@ -36,7 +37,7 @@ def l1 : MLList S Nat := MLList.ofList [0,1,2] def l2 : MLList S Nat := MLList.ofList [3,4,5] def ll : MLList S Nat := (MLList.ofList [l1, l2]).join -#eval show MetaM Unit from do +run_cmd Lean.Elab.Command.liftTermElabM do let x := ll.force.run [] guard $ x = some ([0, 1, 2, 3, 4, 5], []) @@ -44,13 +45,13 @@ def half_or_fail (n : Nat) : MetaM Nat := do guard (n % 2 = 0) pure (n / 2) -#eval do +run_cmd Lean.Elab.Command.liftTermElabM do let x : MLList MetaM Nat := MLList.range let y := x.filterMapM fun n => try? <| half_or_fail n let z ← y.takeAsList 10 guard $ z.length = 10 -#eval do +run_cmd Lean.Elab.Command.liftTermElabM do let R : MLList MetaM Nat := MLList.range let S : MLList MetaM Nat := R.filterMapM fun n => try? do guard (n % 5 = 0) @@ -60,7 +61,7 @@ do guard (n % 2 = 0) guard $ n = [0,1,2,3,4] guard $ m = 0 -#eval do +run_cmd Lean.Elab.Command.liftTermElabM do let R : MLList MetaM Nat := MLList.range let n ← R.firstM fun n => try? do guard (n = 5) diff --git a/test/Real.lean b/test/Real.lean index f8dcefa64d7d5..8548b17d0f2b7 100644 --- a/test/Real.lean +++ b/test/Real.lean @@ -7,7 +7,7 @@ import Mathlib.Data.Real.Basic private axiom test_sorry : ∀ {α}, α -unsafe def testRepr (r : ℝ) (s : String) : Lean.Elab.Command.CommandElabM Unit := +unsafe def testRepr (r : ℝ) (s : String) : Lean.Elab.Command.CommandElabM Unit := unless toString (repr r) = s do throwError "got {repr r}" run_cmd unsafe testRepr 0 "Real.ofCauchy (sorry /- 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, ... -/)" diff --git a/test/Replace.lean b/test/Replace.lean index a5039e4914ff2..0cdcc3576dee1 100644 --- a/test/Replace.lean +++ b/test/Replace.lean @@ -4,9 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino -/ import Mathlib.Tactic.Replace +import Std.Tactic.GuardExpr set_option linter.unusedVariables false +private axiom test_sorry : ∀ {α}, α + /-- Test the `:=` syntax works -/ example {A B : Type} (h : A) (f : A → B) : B := by replace h := f h @@ -24,3 +27,15 @@ example : True := by replace : 2 + 2 = 4 simp_arith trivial + +-- Regression test. `replace h` used to close goal and leave metavariables. +-- Note that `replace h` does *not* delete `h` in this case because the type of the new `h` +-- is a metavariable whose context includes the old `h`. +example (h : True) : False := by + guard_hyp h : True + replace h + · exact true + guard_hyp h : Bool + rename_i h' + guard_hyp h' : True + exact test_sorry diff --git a/test/RewriteSearch/Basic.lean b/test/RewriteSearch/Basic.lean new file mode 100644 index 0000000000000..d06ecc1d270af --- /dev/null +++ b/test/RewriteSearch/Basic.lean @@ -0,0 +1,55 @@ +import Mathlib.Tactic.RewriteSearch + +set_option autoImplicit true + +-- You can enable tracing of the `rw_search` algorithm using +-- set_option trace.rw_search true + +-- You can get timing information (very useful if tweaking the search algorithm!) using +-- set_option profiler true + +/-- info: Try this: rw [@List.length_append, Nat.add_comm] -/ +#guard_msgs in +example (xs ys : List α) : (xs ++ ys).length = ys.length + xs.length := by + rw_search + +/-- info: Try this: rw [← @add_assoc, @add_right_comm, @add_assoc, @add_add_add_comm, ← @add_assoc, @add_right_comm] -/ +#guard_msgs in +example [AddCommMonoid α] {a b c d : α} : (a + b) + (c + d) = a + d + c + b := by + rw_search + +/-- info: Try this: rw [@List.length_append, @List.length_append, @add_rotate, @Nat.add_right_cancel_iff, Nat.two_mul] -/ +#guard_msgs in +example (xs ys : List α) : + (xs ++ ys ++ ys).length = 2 * ys.length + xs.length := by + rw_search + +/-- info: Try this: rw [@List.length_append, @List.length_append, Nat.two_mul, Nat.add_assoc, @add_rotate', Nat.add_assoc] -/ +#guard_msgs in +example (xs ys : List α) : + (xs ++ ys ++ ys).length = 2 * ys.length + xs.length := by + rw_search [-add_rotate] + +/-- info: Try this: rw [Int.add_right_comm, @add_right_cancel_iff, @add_sub, @add_sub_cancel'] -/ +#guard_msgs in +example {a b c : Int} : a + b = c + b + (a - c) := by + rw_search + +/-! A test of the current tokenization scheme. -/ +/-- info: ["(", "[", "5", ",", "3", "]", ",", "4", "+", "(", "2", "*", "1", ")", ")"] -/ +#guard_msgs in +open Mathlib.Tactic.RewriteSearch in +#eval ("([5, 3], 4 + (2 * 1))".splitOn.map splitDelimiters).join + +-- Function that always constructs `[0]`. Used in the following example. +def makeSingleton : Nat → List Nat + | 0 => [0] + | b + 1 => makeSingleton b + +/-- info: Try this: rw [← ih] -/ +#guard_msgs in +example (n : Nat) : makeSingleton n = [0] := by + induction' n with n' ih + · simp only [makeSingleton] + · -- At one point, this failed with: unknown free variable '_uniq.62770' + rw_search diff --git a/test/RewriteSearch/Polynomial.lean b/test/RewriteSearch/Polynomial.lean new file mode 100644 index 0000000000000..a3cabf23500dc --- /dev/null +++ b/test/RewriteSearch/Polynomial.lean @@ -0,0 +1,203 @@ +import Mathlib.Data.Polynomial.Eval +import Mathlib.Data.Polynomial.Inductions +import Mathlib.Tactic.RewriteSearch + +set_option autoImplicit true + +open Polynomial + +/-- +info: Try this: rw [@sub_eq_add_neg, ← @C_neg, @natDegree_add_C] +-/ +#guard_msgs in +example {R : Type*} [Ring R] {p : Polynomial R} {a : R} : + natDegree (p - C a) = natDegree p := by + rw_search [-Polynomial.natDegree_sub_C] + + +-- This one works, but is very slow: + +-- /-- +-- info: Try this: rw [@X_mul, @eq_sub_iff_add_eq, @divX_mul_X_add] +-- -/ +-- #guard_msgs in +-- set_option maxHeartbeats 5000000 in +-- theorem Polynomial.X_mul_divX [Field F] (p : Polynomial F) : +-- Polynomial.X * Polynomial.divX p = p - Polynomial.C (Polynomial.coeff p 0) := by +-- -- Manual proof: rw [eq_sub_iff_add_eq, Polynomial.X_mul_divX_add] +-- rw_search + + +-- All rewrite-only lemmas from `Mathlib.Data.Polynomial.Degree.Definitions`, +-- whose statements are equalities. +-- TODO: `rw_search` should handle `iff` as well. + +universe u v + +open + BigOperators + Finset + Finsupp + Polynomial + +-- Polynomial.degree_of_subsingleton.{u} +#guard_msgs(drop info) in +example {R : Type u} [Semiring R] {p : Polynomial R} [Subsingleton R] : + Polynomial.degree p = ⊥ := by + rw_search [-Polynomial.degree_of_subsingleton] + -- Mathlib proof: + -- rw [Subsingleton.elim p 0, degree_zero] + done + +-- Fails: +-- -- Polynomial.natDegree_of_subsingleton.{u} +-- example {R : Type u} [Semiring R] {p : Polynomial R} +-- [Subsingleton R] : Polynomial.natDegree p = 0 := by +-- rw_search [-Polynomial.natDegree_of_subsingleton] +-- -- Mathlib proof: +-- -- rw [Subsingleton.elim p 0, natDegree_zero] +-- done + +-- Polynomial.degree_C_mul_X_pow.{u} +#guard_msgs(drop info) in +example {R : Type u} {a : R} [Semiring R] (n : ℕ) (ha : a ≠ 0) : + Polynomial.degree (Polynomial.C a * Polynomial.X ^ n) = n := by + rw_search [-Polynomial.degree_C_mul_X_pow] + -- Mathlib proof: + -- rw [C_mul_X_pow_eq_monomial, degree_monomial n ha] + done + +-- Fails: +-- -- Polynomial.Monic.eq_X_add_C.{u} +-- example {R : Type u} [Semiring R] {p : Polynomial R} (hm : Polynomial.Monic p) +-- (hnd : Polynomial.natDegree p = 1) : p = Polynomial.X + Polynomial.C (Polynomial.coeff p 0) := by +-- rw_search [-Polynomial.Monic.eq_X_add_C] +-- -- Mathlib proof: +-- -- rw [← one_mul X, ← C_1, ← hm.coeff_natDegree, hnd, ← eq_X_add_C_of_natDegree_le_one hnd.le] +-- done + +-- Fails: +-- -- Polynomial.natDegree_int_cast.{u} +-- example {R : Type u} [Ring R] (n : ℤ) : Polynomial.natDegree (n : R[X]) = 0 := by +-- rw_search [-Polynomial.natDegree_int_cast] +-- -- Mathlib proof: +-- -- rw [← C_eq_int_cast, natDegree_C] +-- done + +-- Fails: +-- -- Polynomial.leadingCoeff_neg.{u} +-- example {R : Type u} [Ring R] (p : Polynomial R) : +-- Polynomial.leadingCoeff (-p) = -Polynomial.leadingCoeff p := by +-- rw_search [-Polynomial.leadingCoeff_neg] +-- -- Mathlib proof: +-- -- rw [leadingCoeff, leadingCoeff, natDegree_neg, coeff_neg] +-- done + +-- Polynomial.degree_add_eq_right_of_degree_lt.{u} +#guard_msgs(drop info) in +example {R : Type u} [Semiring R] {p q : Polynomial R} + (h : Polynomial.degree p < Polynomial.degree q) : Polynomial.degree (p + q) = Polynomial.degree q := by + rw_search [-Polynomial.degree_add_eq_right_of_degree_lt] + -- Mathlib proof: + -- rw [add_comm, degree_add_eq_left_of_degree_lt h] + done + +-- Polynomial.leadingCoeff_C_mul_X_pow.{u} +#guard_msgs(drop info) in +example {R : Type u} [Semiring R] (a : R) (n : ℕ) : + Polynomial.leadingCoeff (Polynomial.C a * Polynomial.X ^ n) = a := by + rw_search [-Polynomial.leadingCoeff_C_mul_X_pow] + -- Mathlib proof: + -- rw [C_mul_X_pow_eq_monomial, leadingCoeff_monomial] + done + +-- Polynomial.C_mul_X_pow_eq_self.{u} +#guard_msgs(drop info) in +example {R : Type u} [Semiring R] {p : Polynomial R} + (h : Finset.card (Polynomial.support p) ≤ 1) : + Polynomial.C (Polynomial.leadingCoeff p) * Polynomial.X ^ Polynomial.natDegree p = p := by + rw_search [-Polynomial.C_mul_X_pow_eq_self] + -- Mathlib proof: + -- rw [C_mul_X_pow_eq_monomial, monomial_natDegree_leadingCoeff_eq_self h] + done + +-- Fails: +-- -- Polynomial.degree_linear.{u} +-- example {R : Type u} {a b : R} [Semiring R] (ha : a ≠ 0) : +-- Polynomial.degree (Polynomial.C a * Polynomial.X + Polynomial.C b) = 1 := by +-- rw_search [-Polynomial.degree_linear] +-- -- Mathlib proof: +-- -- rw [degree_add_eq_left_of_degree_lt <| degree_C_lt_degree_C_mul_X ha, degree_C_mul_X ha] +-- done + +-- Polynomial.natDegree_linear.{u} +#guard_msgs(drop info) in +example {R : Type u} {a b : R} [Semiring R] (ha : a ≠ 0) : + Polynomial.natDegree (Polynomial.C a * Polynomial.X + Polynomial.C b) = 1 := by + rw_search [-Polynomial.natDegree_linear] + -- Mathlib proof: + -- rw [natDegree_add_C, natDegree_C_mul_X a ha] + done + +-- Fails: +-- -- Polynomial.degree_X_pow.{u} +-- example {R : Type u} [Semiring R] [Nontrivial R] (n : ℕ) : +-- Polynomial.degree (Polynomial.X ^ n : R[X]) = n := by +-- rw_search [-Polynomial.degree_X_pow] +-- -- Mathlib proof: +-- -- rw [X_pow_eq_monomial, degree_monomial _ (one_ne_zero' R)] +-- done + +-- Polynomial.degree_X_sub_C.{u} +#guard_msgs(drop info) in +example {R : Type u} [Ring R] [Nontrivial R] (a : R) : + Polynomial.degree (Polynomial.X - Polynomial.C a) = 1 := by + rw_search [-Polynomial.degree_X_sub_C] + -- Mathlib proof: + -- rw [sub_eq_add_neg, ← map_neg C a, degree_X_add_C] + done + +-- Polynomial.natDegree_X_sub_C.{u} +#guard_msgs(drop info) in +example {R : Type u} [Ring R] [Nontrivial R] (x : R) : + Polynomial.natDegree (Polynomial.X - Polynomial.C x) = 1 := by + rw_search [-Polynomial.natDegree_X_sub_C] + -- Mathlib proof: + -- rw [natDegree_sub_C, natDegree_X] + done + +-- Polynomial.nextCoeff_X_sub_C.{v} +#guard_msgs(drop info) in +example {S : Type v} [Ring S] (c : S) : + Polynomial.nextCoeff (Polynomial.X - Polynomial.C c) = -c := by + rw_search [-Polynomial.nextCoeff_X_sub_C] + -- Mathlib proof: + -- rw [sub_eq_add_neg, ← map_neg C c, nextCoeff_X_add_C] + done + +-- Polynomial.degree_X_pow_sub_C.{u} +#guard_msgs(drop info) in +example {R : Type u} [Ring R] [Nontrivial R] {n : ℕ} (hn : 0 < n) (a : R) : + Polynomial.degree (Polynomial.X ^ n - Polynomial.C a) = n := by + rw_search [-Polynomial.degree_X_pow_sub_C] + -- Mathlib proof: + -- rw [sub_eq_add_neg, ← map_neg C a, degree_X_pow_add_C hn] + done + +-- Polynomial.natDegree_X_pow_sub_C.{u} +#guard_msgs(drop info) in +example {R : Type u} [Ring R] [Nontrivial R] {n : ℕ} {r : R} : + Polynomial.natDegree (Polynomial.X ^ n - Polynomial.C r) = n := by + rw_search [-Polynomial.natDegree_X_pow_sub_C] + -- Mathlib proof: + -- rw [sub_eq_add_neg, ← map_neg C r, natDegree_X_pow_add_C] + done + +-- Fails: +-- -- Polynomial.leadingCoeff_X_sub_C.{v} +-- example {S : Type v} [Ring S] (r : S) : +-- Polynomial.leadingCoeff (Polynomial.X - Polynomial.C r) = 1 := by +-- rw_search [-Polynomial.leadingCoeff_X_sub_C] +-- -- Mathlib proof: +-- -- rw [sub_eq_add_neg, ← map_neg C r, leadingCoeff_X_add_C] +-- done diff --git a/test/Zify.lean b/test/Zify.lean index 783858ba7b8fa..66963e7b37df1 100644 --- a/test/Zify.lean +++ b/test/Zify.lean @@ -5,6 +5,7 @@ Authors: Moritz Doll, Robert Y. Lewis -/ import Mathlib.Tactic.Zify +import Mathlib.Data.Int.Basic import Std.Tactic.GuardExpr private axiom test_sorry : ∀ {α}, α diff --git a/test/apply_fun.lean b/test/apply_fun.lean index 09838b1023c90..5b741700653d7 100644 --- a/test/apply_fun.lean +++ b/test/apply_fun.lean @@ -258,3 +258,14 @@ example : 1 = 1 := by · exact test_sorry apply_fun f using (g f) rfl + + +def funFamily (_i : ℕ) : Bool → Bool := id + +-- `apply_fun` should not silence errors in `assumption` +/-- +error: maximum recursion depth has been reached (use `set_option maxRecDepth ` to increase limit) +-/ +#guard_msgs (error) in +example (_h₁ : Function.Injective (funFamily ((List.range 128).map (fun _ => 0)).sum)) : true = true := by + apply_fun funFamily 0 diff --git a/test/convert2.lean b/test/convert2.lean new file mode 100644 index 0000000000000..b70bada2d9358 --- /dev/null +++ b/test/convert2.lean @@ -0,0 +1,11 @@ +import Mathlib.Data.List.Defs +import Mathlib.Data.Nat.Basic + +-- Prior to #7945 this failed with `(kernel) declaration has metavariables '_example'`. +/-- +error: maximum recursion depth has been reached (use `set_option maxRecDepth ` to increase limit) +-/ +#guard_msgs (error) in +example (_h₁ : ((List.range 128).map (fun _ => 0)).sum = 0) : 0 ∣ 1 := by + apply Nat.dvd_of_mul_dvd_mul_left Nat.zero_lt_one + convert Nat.dvd_mul_left 0 1 diff --git a/test/delaborators.lean b/test/delaborators.lean index 740aea9832830..086fe3a9513c1 100644 --- a/test/delaborators.lean +++ b/test/delaborators.lean @@ -1,5 +1,5 @@ import Std.Tactic.GuardMsgs -import Mathlib.Util.PiNotation +import Mathlib.Util.Delaborators import Mathlib.Data.Set.Lattice section PiNotation @@ -45,6 +45,48 @@ variable (P : Nat → Prop) (α : Nat → Type) (s : Set ℕ) #guard_msgs in #check ∀ x, x ∈ s → P x +/-- info: ∀ x ∉ s, P x : Prop -/ +#guard_msgs in +#check ∀ x ∉ s,P x + +/-- info: ∀ x ∉ s, P x : Prop -/ +#guard_msgs in +#check ∀ x, x ∉ s → P x + +variable (Q : Set ℕ → Prop) + +/-- info: ∀ t ⊆ s, Q t : Prop -/ +#guard_msgs in +#check ∀ t ⊆ s, Q t + +/-- info: ∀ t ⊆ s, Q t : Prop -/ +#guard_msgs in +#check ∀ t, t ⊆ s → Q t + +/-- info: ∀ t ⊂ s, Q t : Prop -/ +#guard_msgs in +#check ∀ t ⊂ s, Q t + +/-- info: ∀ t ⊂ s, Q t : Prop -/ +#guard_msgs in +#check ∀ t, t ⊂ s → Q t + +/-- info: ∀ t ⊇ s, Q t : Prop -/ +#guard_msgs in +#check ∀ t ⊇ s, Q t + +/-- info: ∀ t ⊇ s, Q t : Prop -/ +#guard_msgs in +#check ∀ t, t ⊇ s → Q t + +/-- info: ∀ t ⊃ s, Q t : Prop -/ +#guard_msgs in +#check ∀ t ⊃ s, Q t + +/-- info: ∀ t ⊃ s, Q t : Prop -/ +#guard_msgs in +#check ∀ t, t ⊃ s → Q t + /-- info: (x : ℕ) → α x : Type -/ #guard_msgs in #check (x : Nat) → α x @@ -103,6 +145,78 @@ section existential #guard_msgs in #check ∃ (i : Nat), i < 3 ∧ i = i +variable (s : Set ℕ) (P : ℕ → Prop) (Q : Set ℕ → Prop) + +/-- info: ∃ x ∉ s, P x : Prop -/ +#guard_msgs in +#check ∃ x ∉ s, P x + +/-- info: ∃ x ∉ s, P x : Prop -/ +#guard_msgs in +#check ∃ x, x ∉ s ∧ P x + +variable (Q : Set ℕ → Prop) + +/-- info: ∃ t ⊆ s, Q t : Prop -/ +#guard_msgs in +#check ∃ t ⊆ s, Q t + +/-- info: ∃ t ⊆ s, Q t : Prop -/ +#guard_msgs in +#check ∃ t, t ⊆ s ∧ Q t + +/-- info: ∃ t ⊂ s, Q t : Prop -/ +#guard_msgs in +#check ∃ t ⊂ s, Q t + +/-- info: ∃ t ⊂ s, Q t : Prop -/ +#guard_msgs in +#check ∃ t, t ⊂ s ∧ Q t + +/-- info: ∃ t ⊇ s, Q t : Prop -/ +#guard_msgs in +#check ∃ t ⊇ s, Q t + +/-- info: ∃ t ⊇ s, Q t : Prop -/ +#guard_msgs in +#check ∃ t, t ⊇ s ∧ Q t + +/-- info: ∃ t ⊃ s, Q t : Prop -/ +#guard_msgs in +#check ∃ t ⊃ s, Q t + +/-- info: ∃ t ⊃ s, Q t : Prop -/ +#guard_msgs in +#check ∃ t, t ⊃ s ∧ Q t + +/-- info: ∃ n k, n = k : Prop -/ +#guard_msgs in +#check ∃ n k, n = k + +/-- info: ∃ n k, n = k : Prop -/ +#guard_msgs in +#check ∃ n, ∃ k, n = k + +section merging + +/-- info: ∃ (_ : True), True : Prop -/ +#guard_msgs in #check ∃ (_ : True), True + +/-- info: ∃ (_ : True), ∃ x, x = x : Prop -/ +#guard_msgs in #check ∃ (_ : True) (x : Nat), x = x + +/-- info: ∃ (_ : True), ∃ x y, x = y : Prop -/ +#guard_msgs in #check ∃ (_ : True) (x y : Nat), x = y + +/-- info: ∃ (_ : True) (_ : False), True : Prop -/ +#guard_msgs in #check ∃ (_ : True) (_ : False), True + +set_option pp.funBinderTypes true in +/-- info: ∃ (x : ℕ) (x : ℕ), True : Prop -/ +#guard_msgs in #check ∃ (_ : Nat) (_ : Nat), True + +end merging + end existential section prod diff --git a/test/finset_repr.lean b/test/finset_repr.lean index fc7f94ebbb11c..3a82f90fa69eb 100644 --- a/test/finset_repr.lean +++ b/test/finset_repr.lean @@ -1,11 +1,11 @@ import Mathlib.Data.Finset.Sort -#eval show Lean.MetaM Unit from guard <| - (repr (0 : Multiset (List ℕ)) |>.pretty 15) = "0" -#eval show Lean.MetaM Unit from guard <| - (repr ({[1, 2, 3], [4, 5, 6]} : Multiset (List ℕ)) |>.pretty 15) = "{[1, 2, 3],\n [4, 5, 6]}" - -#eval show Lean.MetaM Unit from guard <| - (repr (∅ : Finset (List ℕ)) |>.pretty 15) = "∅" -#eval show Lean.MetaM Unit from guard <| - (repr ({[1, 2, 3], [4, 5, 6]} : Finset (List ℕ)) |>.pretty 15) = "{[1, 2, 3],\n [4, 5, 6]}" +run_cmd Lean.Elab.Command.liftTermElabM do + unsafe guard <| + (repr (0 : Multiset (List ℕ)) |>.pretty 15) = "0" + unsafe guard <| + (repr ({[1, 2, 3], [4, 5, 6]} : Multiset (List ℕ)) |>.pretty 15) = "{[1, 2, 3],\n [4, 5, 6]}" + unsafe guard <| + (repr (∅ : Finset (List ℕ)) |>.pretty 15) = "∅" + unsafe guard <| + (repr ({[1, 2, 3], [4, 5, 6]} : Finset (List ℕ)) |>.pretty 15) = "{[1, 2, 3],\n [4, 5, 6]}" diff --git a/test/left_right.lean b/test/left_right.lean deleted file mode 100644 index fb395221b19ea..0000000000000 --- a/test/left_right.lean +++ /dev/null @@ -1,28 +0,0 @@ -import Mathlib.Tactic.LeftRight - -def zero : Nat := by - left - -#eval zero -- 0 - -def two : Nat := by - right - exact 1 - -#eval two -- 2 - -example : Sum Nat (List Nat) := by - left - exact zero - -example : Sum Nat (List Nat) := by - right - exact [0] - -example : (1 = 1) ∨ (2 = 3) := by - left - rfl - -example : (1 = 2) ∨ (3 = 3) := by - right - rfl diff --git a/test/lift.lean b/test/lift.lean index c8def191b7e18..a04dfb3cecd15 100644 --- a/test/lift.lean +++ b/test/lift.lean @@ -206,7 +206,7 @@ example (n : ℕ) : n = 0 ∨ ∃ p : ℕ+, n = p := by right exact ⟨n, rfl⟩ · left - exact Nat.eq_zero_of_nonpos _ hn + exact Nat.eq_zero_of_not_pos hn example (n : ℤ) (hn : 0 < n) : True := by lift n to ℕ+ diff --git a/test/linear_combination.lean b/test/linear_combination.lean index 55e5adab860ed..05d7b32149fd0 100644 --- a/test/linear_combination.lean +++ b/test/linear_combination.lean @@ -140,7 +140,7 @@ example (a b : ℝ) (ha : 2 * a = 4) (hab : 2 * b = a - b) : b = 2 / 3 := by example (x y : ℤ) (h1 : x = -3) (_h2 : y = 10) : 2 * x = -6 := by linear_combination (norm := skip) 2 * h1 - simp + simp (config := {decide := true}) /-! ### Cases without any arguments provided -/ diff --git a/test/matrix.lean b/test/matrix.lean index e13ff85cceab2..deac900999aea 100644 --- a/test/matrix.lean +++ b/test/matrix.lean @@ -139,9 +139,10 @@ example {α : Type _} [CommRing α] {a b c d : α} : Matrix.det !![a, b; c, d] = a * d - b * c := by simp? [Matrix.det_succ_row_zero, Fin.sum_univ_succ] says simp only [det_succ_row_zero, Nat.odd_iff_not_even, of_apply, cons_val', empty_val', - cons_val_fin_one, cons_val_zero, det_unique, Fin.default_eq_zero, submatrix_apply, Fin.succ_zero_eq_one, ne_eq, - cons_val_one, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, pow_zero, one_mul, Fin.zero_succAbove, head_cons, - Finset.univ_unique, Fin.val_succ, Fin.coe_fin_one, zero_add, pow_one, cons_val_succ, neg_mul, + cons_val_fin_one, cons_val_zero, det_unique, Fin.default_eq_zero, submatrix_apply, + Fin.succ_zero_eq_one, ne_eq, cons_val_one, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, + pow_zero, one_mul, not_true_eq_false, Fin.zero_succAbove, head_cons, Finset.univ_unique, + Fin.val_succ, Fin.coe_fin_one, zero_add, pow_one, cons_val_succ, neg_mul, Fin.succ_succAbove_zero, Finset.sum_const, Finset.card_singleton, smul_neg, one_smul] ring @@ -150,10 +151,10 @@ example {α : Type _} [CommRing α] {a b c d e f g h i : α} : a * e * i - a * f * h - b * d * i + b * f * g + c * d * h - c * e * g := by simp? [Matrix.det_succ_row_zero, Fin.sum_univ_succ] says simp only [det_succ_row_zero, Nat.odd_iff_not_even, of_apply, cons_val', empty_val', - cons_val_fin_one, cons_val_zero, submatrix_apply, Fin.succ_zero_eq_one, cons_val_one, - head_cons, submatrix_submatrix, det_unique, Fin.default_eq_zero, Function.comp_apply, - Fin.succ_one_eq_two, ne_eq, cons_val_two, tail_cons, head_fin_const, Fin.sum_univ_succ, - Fin.val_zero, pow_zero, one_mul, Fin.zero_succAbove, Finset.univ_unique, Fin.val_succ, + cons_val_fin_one, cons_val_zero, submatrix_apply, Fin.succ_zero_eq_one, cons_val_one, head_cons, + submatrix_submatrix, det_unique, Fin.default_eq_zero, Function.comp_apply, Fin.succ_one_eq_two, + ne_eq, cons_val_two, tail_cons, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, pow_zero, + one_mul, not_true_eq_false, Fin.zero_succAbove, Finset.univ_unique, Fin.val_succ, Fin.coe_fin_one, zero_add, pow_one, neg_mul, Fin.succ_succAbove_zero, Finset.sum_neg_distrib, Finset.sum_singleton, cons_val_succ, Fin.succ_succAbove_one, even_add_self, Even.neg_pow, one_pow, Finset.sum_const, Finset.card_singleton, one_smul] diff --git a/test/norm_num.lean b/test/norm_num.lean index 9d3bc05c99bf2..e9037f30f44b9 100644 --- a/test/norm_num.lean +++ b/test/norm_num.lean @@ -37,7 +37,7 @@ example : (6:ℝ) < 10 := by norm_num1 example : (7:ℝ)/2 > 3 := by norm_num1 example : (4:ℝ)⁻¹ < 1 := by norm_num1 example : ((1:ℝ) / 2)⁻¹ = 2 := by norm_num1 --- example : 2 ^ 17 - 1 = 131071 := by norm_num1 +example : 2 ^ 17 - 1 = 131071 := by norm_num1 -- example : (3 : ℝ) ^ (-2 : ℤ) = 1/9 := by norm_num1 -- example : (3 : ℝ) ^ (-2 : ℤ) = 1/9 := by norm_num1 -- example : (-3 : ℝ) ^ (0 : ℤ) = 1 := by norm_num1 diff --git a/test/norm_num_ext.lean b/test/norm_num_ext.lean index e1cccc8619ce4..d06b0a36a5e4d 100644 --- a/test/norm_num_ext.lean +++ b/test/norm_num_ext.lean @@ -335,18 +335,24 @@ variable {α : Type _} [CommRing α] open BigOperators -- Lists: +-- `by decide` closes the three goals below. +/- example : ([1, 2, 1, 3]).sum = 7 := by norm_num only --- example : (([1, 2, 1, 3] : List ℚ).map (fun i => i^2)).sum = 15 := by norm_num [-List.map] --TODO example : (List.range 10).sum = 45 := by norm_num only example : (List.finRange 10).sum = 45 := by norm_num only +-/ +-- example : (([1, 2, 1, 3] : List ℚ).map (fun i => i^2)).sum = 15 := by norm_num [-List.map] --TODO -- Multisets: +-- `by decide` closes the three goals below. +/- example : (1 ::ₘ 2 ::ₘ 1 ::ₘ 3 ::ₘ {}).sum = 7 := by norm_num only example : ((1 ::ₘ 2 ::ₘ 1 ::ₘ 3 ::ₘ {}).map (fun i => i^2)).sum = 15 := by norm_num only --- example : (({1, 2, 1, 3} : Multiset ℚ).map (fun i => i^2)).sum = 15 := by -- TODO --- norm_num [-Multiset.map_cons] example : (Multiset.range 10).sum = 45 := by norm_num only example : (↑[1, 2, 1, 3] : Multiset ℕ).sum = 7 := by norm_num only +-/ +-- example : (({1, 2, 1, 3} : Multiset ℚ).map (fun i => i^2)).sum = 15 := by -- TODO +-- norm_num [-Multiset.map_cons] -- Finsets: example : Finset.prod (Finset.cons 2 ∅ (Finset.not_mem_empty _)) (λ x => x) = 2 := by norm_num1 diff --git a/test/notation3.lean b/test/notation3.lean index 778dc57922a0c..292e92fea3b69 100644 --- a/test/notation3.lean +++ b/test/notation3.lean @@ -20,6 +20,11 @@ notation3 "∀ᶠ " (...) " in " f ", " r:(scoped p => Filter.eventually p f) => /-- info: ∀ᶠ (x : ℕ) in Filter.atTop, x < 3 : Prop -/ #guard_msgs in #check ∀ᶠ x in Filter.atTop, x < 3 +-- Testing lambda expressions: +notation3 "∀ᶠ' " f ", " p => Filter.eventually (fun x => (p : _ → _) x) f +/-- info: ∀ᶠ' Filter.atTop, fun x ↦ x < 3 : Prop -/ +#guard_msgs in #check ∀ᶠ' Filter.atTop, fun x => x < 3 + def foobar (p : α → Prop) (f : Prop) := ∀ x, p x = f notation3 "∀ᶠᶠ " (...) " in " f ": " @@ -46,13 +51,21 @@ notation3 "func! " (...) ", " r:(scoped p => func p) => r #guard_msgs in #check (func! (x : Nat → Nat), x) (· * 2) structure MyUnit where -notation3 "~{" (x"; "* => foldl (a b => Prod.mk a b) MyUnit) "}~" => x +notation3 "~{" (x"; "* => foldl (a b => (a, b)) MyUnit) "}~" => x /-- info: ~{1; true; ~{2}~}~ : ((Type × ℕ) × Bool) × Type × ℕ -/ #guard_msgs in #check ~{1; true; ~{2}~}~ /-- info: ~{}~ : Type -/ #guard_msgs in #check ~{}~ -notation3 "%[" (x", "* => foldr (a b => List.cons a b) List.nil) "]" => x +structure MyUnit' where +instance : OfNat MyUnit' (nat_lit 0) := ⟨{}⟩ +notation3 "MyUnit'0" => (0 : MyUnit') +/-- info: MyUnit'0 : MyUnit' -/ +#guard_msgs in #check (0 : MyUnit') +/-- info: 0 : ℕ -/ +#guard_msgs in #check 0 + +notation3 "%[" (x", "* => foldr (a b => a :: b) []) "]" => x /-- info: %[1, 2, 3] : List ℕ -/ #guard_msgs in #check %[1, 2, 3] @@ -73,9 +86,9 @@ notation3 "*'[" x "] " (...) ", " v:(scoped c => bar' x <| foo' x c) => v /-- info: bar' 1 : ℕ → ℕ -/ #guard_msgs in #check bar' 1 --- Currently does not pretty print due to pi type -notation3 (prettyPrint := false) "MyPi " (...) ", " r:(scoped p => (x : _) → p x) => r -/-- info: ∀ (x : ℕ), (fun x ↦ ∀ (x_1 : ℕ), (fun y ↦ x < y) x_1) x : Prop -/ +-- Need to give type ascription to `p` so that `p x` elaborates when making matcher +notation3 "MyPi " (...) ", " r:(scoped p => (x : _) → (p : _ → _) x) => r +/-- info: MyPi (x : ℕ) (y : ℕ), x < y : Prop -/ #guard_msgs in #check MyPi (x : Nat) (y : Nat), x < y -- The notation parses fine, but the delaborator never succeeds, which is expected @@ -84,11 +97,57 @@ notation3 "BAD " c "; " (x", "* => foldl (a b => b) c) " DAB" => myId x /-- info: myId 3 : ℕ -/ #guard_msgs in #check BAD 1; 2, 3 DAB +section +variable (x : Nat) +local notation3 "y" => x + 1 +/-- info: y : ℕ -/ +#guard_msgs in #check y +/-- info: y : ℕ -/ +#guard_msgs in #check x + 1 +end + +section +variable (α : Type u) [Add α] +local notation3 x " +α " y => (x + y : α) +variable (x y : α) +/-- info: x +α y : α -/ +#guard_msgs in #check x +α y +/-- info: x +α y : α -/ +#guard_msgs in #check x + y +/-- info: 1 + 2 : ℕ -/ +#guard_msgs in #check 1 + 2 +end + +def idStr : String → String := id + +/-- +error: application type mismatch + idStr Nat.zero +argument + Nat.zero +has type + ℕ : Type +but is expected to have type + String : Type +--- +warning: Was not able to generate a pretty printer for this notation. If you do not expect it to be +pretty printable, then you can use `notation3 (prettyPrint := false)`. If the notation expansion +refers to section variables, be sure to do `local notation3`. Otherwise, you might be able to adjust +the notation expansion to make it matchable; pretty printing relies on deriving an expression +matcher from the expansion. (Use `set_option trace.notation3 true` to get some debug information.) +-/ +#guard_msgs in +notation3 "error" => idStr Nat.zero + section /-- -warning: Was not able to generate a pretty printer for this notation. If you do not expect it to be pretty printable, then you can use `notation3 (prettyPrint := false)`. If the notation expansion refers to section variables, be sure to do `local notation3`. Otherwise, you might be able to adjust the notation expansion to make it matchable; pretty printing relies on deriving an expression matcher from the expansion. (Use `set_option trace.notation3 true` to get some debug information.) +warning: Was not able to generate a pretty printer for this notation. If you do not expect it to be +pretty printable, then you can use `notation3 (prettyPrint := false)`. If the notation expansion +refers to section variables, be sure to do `local notation3`. Otherwise, you might be able to adjust +the notation expansion to make it matchable; pretty printing relies on deriving an expression +matcher from the expansion. (Use `set_option trace.notation3 true` to get some debug information.) -/ -#guard_msgs in local notation3 "#" n => Fin.mk n (by decide) +#guard_msgs (warning, drop error) in local notation3 "#" n => Fin.mk n (by decide) end section @@ -103,3 +162,17 @@ error: failed to reduce to 'true' #guard_msgs in example : Fin 5 := #6 end + +section test_scoped + +scoped[MyNotation] notation3 "π" => (3 : Nat) + +/-- error: unknown identifier 'π' -/ +#guard_msgs in #check π + +open scoped MyNotation + +/-- info: π : ℕ -/ +#guard_msgs in #check π + +end test_scoped diff --git a/test/peel.lean b/test/peel.lean new file mode 100644 index 0000000000000..bcb2e64677a2b --- /dev/null +++ b/test/peel.lean @@ -0,0 +1,271 @@ +import Mathlib.Tactic.Peel +import Mathlib.Topology.Instances.Real + +open Filter Topology + +/-! ## General usage -/ + +example (p q : Nat → Prop) (h₁ : ∀ x, p x) (h₂ : ∀ x, p x → q x) : ∀ y, q y := by + peel 1 h₁ + rename_i y + guard_target =ₐ q y + exact h₂ _ this + +example (p q : Nat → Prop) (h₁ : ∀ {x}, p x) (h₂ : ∀ x, p x → q x) : ∀ y, q y := by + peel 1 h₁ + rename_i y + guard_target =ₐ q y + exact h₂ _ this + +example (p q : Nat → Nat → Prop) (h₁ : ∀ {x y}, p x y) (h₂ : ∀ x y, p x y → q x y) : + ∀ u v, q u v := by + peel 2 h₁ + rename_i u v + guard_target =ₐ q u v + exact h₂ _ _ this + +example (p q : Nat → Prop) (h₁ : ∀ x, p x) (h₂ : ∀ x, p x → q x) : ∀ y, q y := by + peel h₁ + rename_i y + guard_target =ₐ q y + exact h₂ _ this + +example (p q : Nat → Prop) (h₁ : ∀ x, p x) (h₂ : ∀ x, p x → q x) : ∀ y, q y := by + peel h₁ using h₂ _ this + +example (p q : Nat → Nat → Prop) (h₁ : ∀ x y, p x y) (h₂ : ∀ x y, p x y → q x y) : + ∀ u v, q u v := by + peel h₁ + rename_i u v + guard_target =ₐ q u v + exact h₂ _ _ this + +example (p q : Nat → Prop) (h₁ : ∀ x, p x) (h₂ : ∀ x, p x → q x) : ∀ y, q y := by + peel h₁ with foo + rename_i y + guard_target =ₐ q y + exact h₂ _ foo + +example (p q : Nat → Prop) (h₁ : ∀ x, p x) (h₂ : ∀ x, p x → q x) : ∀ y, q y := by + peel h₁ with foo w + guard_target =ₐ q w + exact h₂ w foo + +example (p q : Nat → Prop) (h₁ : ∀ x, p x) (h₂ : ∀ x, p x → q x) : ∀ y, q y := by + peel h₁ with foo _ + rename_i w + guard_target =ₐ q w + exact h₂ w foo + +example (p q : Nat → Prop) (h₁ : ∀ x, p x) (h₂ : ∀ x, p x → q x) : ∀ y, q y := by + peel h₁ with _ w + guard_target =ₐ q w + exact h₂ w this + +example (p q : Nat → Nat → Prop) (h₁ : ∀ x y, p x y) (h₂ : ∀ x y, p x y → q x y) : + ∀ u v, q u v := by + peel h₁ with h_peel s t + guard_target =ₐ q s t + exact h₂ s t h_peel + +example (p q : Nat → Prop) (h : ∀ y, p y) (h₁ : ∀ z, p z → q z) : ∀ x, q x := by + peel h + rename_i x + guard_target =ₐ q x + exact h₁ _ <| by assumption + +example (p q : Nat → Prop) (h : ∃ y, p y) (h₁ : ∀ z, p z → q z) : ∃ x, q x := by + peel h with h a + guard_target =ₐ q a + exact h₁ a h + +example (x y : ℝ) (h : ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, x + n = y + ε) : + ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, x - ε = y - n := by + peel h with h_peel ε hε N n hn + guard_target =ₐ x - ε = y - n + linarith + +example (x y : ℝ) (h : ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, x + n = y + ε) : + ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, x - ε = y - n := by + peel h + fail_if_success peel 2 this + peel 0 this + fail_if_success peel this + linarith + +example (p q : ℝ → ℝ → Prop) (h : ∀ ε > 0, ∃ δ > 0, p ε δ) + (hpq : ∀ x y, x > 0 → y > 0 → p x y → q x y) : + ∀ ε > 0, ∃ δ > 0, q ε δ := by + peel h with h ε hε δ hδ + guard_target =ₐ q ε δ + exact hpq ε δ hε hδ h + +example (p q : ℝ → ℝ → Prop) (h : ∀ ε > 0, ∃ δ > 0, p ε δ) + (hpq : ∀ x y, x > 0 → y > 0 → p x y → q x y) : + ∀ ε > 0, ∃ δ > 0, q ε δ := by + peel h with h ε hε δ hδ using hpq ε δ hε hδ h + +example (x y : ℝ) : (∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, x + n = y + ε) ↔ + ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, x - ε = y - n := by + peel 5 + constructor + all_goals + intro + linarith + +/-! ## Usage after other tactics and with multiple goals -/ + +example : (∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℚ) < ε) ↔ + ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℚ) ≤ ε := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · peel 5 h + exact this.le + · intro ε hε + peel 3 h (ε / 2) (half_pos hε) + exact this.trans_lt (half_lt_self hε) + +example : (∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℚ) < ε) ↔ + ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℚ) ≤ ε := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · peel 5 h using this.le + · intro ε hε + peel 3 h (ε / 2) (half_pos hε) using this.trans_lt (half_lt_self hε) + +/-! ## Use with `↔` goals -/ + +example (x y : ℚ) (h : ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, x + n = y + ε) : + ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, x - ε = y - n := by + intro ε hε + peel 3 (h ε hε) + rename_i _ n _ + guard_hyp this : x + ↑n = y + ε + guard_target =ₐ x - ε = y - n + linarith + +example (x y : ℝ) : (∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, x + n = y + ε) ↔ + ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, x - ε = y - n := by + peel with ε hε N n hn + constructor + all_goals + intro + linarith + +example : (∃ k > 0, ∃ n ≥ k, n = k) ↔ ∃ k > 0, ∃ n ≥ k, k = n := by + peel 4 + exact eq_comm + +/-! ## Eventually and frequently -/ + +example {f : ℝ → ℝ} (h : ∀ x : ℝ, ∀ᶠ y in 𝓝 x, |f y - f x| ≤ |y - x|) : + ∀ x : ℝ, ∀ᶠ y in 𝓝 x, |f y - f x| ^ 2 ≤ |y - x| ^ 2 := by + peel h with h_peel x y + gcongr + +example (α : Type*) (f g : Filter α) (p q : α → α → Prop) (h : ∀ᶠ x in f, ∃ᶠ y in g, p x y) + (h₁ : ∀ x y, p x y → q x y) : ∀ᶠ x in f, ∃ᶠ y in g, q x y := by + peel h with h_peel x y + exact h₁ x y h_peel + +example (α : Type*) (f : Filter α) (p q : α → Prop) (h : ∀ᶠ x in f, p x) (h₁ : ∀ x, p x → q x) : + ∀ᶠ x in f, q x := by + peel h with h_peel x + exact h₁ x h_peel + +/-! ## Type classes -/ + +example {R : Type*} [CommRing R] (h : ∀ x : R, ∃ y : R, x + y = 2) : + ∀ x : R, ∃ y : R, (x + y) ^ 2 = 4 := by + peel 2 h + rw [this] + ring + +example {G : Type*} [Group G] [TopologicalSpace G] [TopologicalGroup G] + (h : ∀ᶠ x in 𝓝 (1 : G), ∃ᶠ y in 𝓝 x, x * y⁻¹ = 1) : + ∀ᶠ x in 𝓝 (1 : G), ∃ᶠ y in 𝓝 x, x ^ 2 = y ^ 2 := by + peel h with h_peel a b + observe : a = b⁻¹⁻¹ + simp [this] + +/-! ## Unfolding definitions -/ + +example {α β γ : Type*} {f : α → β} {g : β → γ} (h : Function.Injective (g ∘ f)) : + Function.Injective f := by + peel 2 h with _ x y + intro hf + apply this + congrm(g $hf) + +example {α β γ : Type*} {f : α → β} {g : β → γ} (h : Function.Surjective (g ∘ f)) : + Function.Surjective g := by + peel 1 h with _ y + fail_if_success peel this + obtain ⟨x, rfl⟩ := this + exact ⟨f x, rfl⟩ + +def toInf (f : ℕ → ℕ) : Prop := ∀ m, ∃ n, ∀ n' ≥ n, m ≤ f n' + +example (f : ℕ → ℕ) (h : toInf f) : toInf (fun n => 2 * f n) := by + peel h with this m n n' h + dsimp + linarith + +/-! ## Error messages -/ + +/-- +error: Tactic 'peel' could not match quantifiers in + x = y +and + x = y +-/ +#guard_msgs in example (x y : ℝ) (h : x = y) : x = y := by + peel h + +/-- +error: Tactic 'peel' could not match quantifiers in + ∃ y, ∀ (x : ℕ), x ≠ y +and + ∀ (x : ℕ), ∃ y, x ≠ y +-/ +#guard_msgs in +example (h : ∃ y : ℕ, ∀ x, x ≠ y) : ∀ x : ℕ, ∃ y, x ≠ y := by + peel h + +/-- +error: Tactic 'peel' could not match quantifiers in + ∀ (n : ℕ), 0 ≤ n +and + ∀ (n : ℤ), 0 ≤ n +-/ +#guard_msgs in +example (h : ∀ n : ℕ, 0 ≤ n) : ∀ n : ℤ, 0 ≤ n := by + peel h + +/-- +error: Tactic 'peel' could not match quantifiers in + ∃ n, 0 ≤ n +and + ∃ n, 0 ≤ n +-/ +#guard_msgs in +example (h : ∃ n : ℕ, 0 ≤ n) : ∃ n : ℤ, 0 ≤ n := by + peel 1 h + +/-- +error: Tactic 'peel' could not match quantifiers in + ∃ᶠ (n : ℕ) in atTop, 0 ≤ n +and + ∃ᶠ (n : ℕ) in atBot, 0 ≤ n +-/ +#guard_msgs in +example (h : ∃ᶠ n : ℕ in atTop, 0 ≤ n) : ∃ᶠ n : ℕ in atBot, 0 ≤ n := by + peel 1 h + +/-- +error: Tactic 'peel' could not match quantifiers in + ∀ᶠ (n : ℕ) in atTop, 0 ≤ n +and + ∀ᶠ (n : ℕ) in atBot, 0 ≤ n +-/ +#guard_msgs in +example (h : ∀ᶠ n : ℕ in atTop, 0 ≤ n) : ∀ᶠ n : ℕ in atBot, 0 ≤ n := by + peel 1 h diff --git a/test/positivity.lean b/test/positivity.lean index 8309c39ed70b5..631a878ae8c1c 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -216,6 +216,7 @@ example {a : ℤ} (ha : 1 < a) : 0 < |(3:ℤ) + a| := by positivity example : 0 < NNReal.sqrt 5 := by positivity example : 0 ≤ Real.sqrt (-5) := by positivity +example (x : ℝ) : 0 ≤ Real.sqrt x := by positivity example : 0 < Real.sqrt 5 := by positivity example {a : ℝ} (_ha : 0 ≤ a) : 0 ≤ Real.sqrt a := by positivity @@ -246,7 +247,7 @@ example : 0 ≤ max (-3 : ℤ) 5 := by positivity example (n : ℕ) : 0 < n.succ := by positivity example (n : ℕ) : 0 < n ! := by positivity --- example (n k : ℕ) : 0 < n.asc_factorial k := by positivity +example (n k : ℕ) : 0 < n.ascFactorial k := by positivity -- example {α : Type _} (s : Finset α) (hs : s.Nonempty) : 0 < s.card := by positivity -- example {α : Type _} [Fintype α] [Nonempty α] : 0 < Fintype.card α := by positivity @@ -295,9 +296,9 @@ example {a : ℕ} (ha : 0 < a) : (0 : ℚ) < a := by positivity example {a : ℤ} (ha : a ≠ 0) : (a : ℚ) ≠ 0 := by positivity example {a : ℤ} (ha : 0 ≤ a) : (0 : ℚ) ≤ a := by positivity example {a : ℤ} (ha : 0 < a) : (0 : ℚ) < a := by positivity --- example {a : ℚ} (ha : a ≠ 0) : (a : ℝ) ≠ 0 := by positivity --- example {a : ℚ} (ha : 0 ≤ a) : (0 : ℝ) ≤ a := by positivity --- example {a : ℚ} (ha : 0 < a) : (0 : ℝ) < a := by positivity +example {a : ℚ} (ha : a ≠ 0) : (a : ℝ) ≠ 0 := by positivity +example {a : ℚ} (ha : 0 ≤ a) : (0 : ℝ) ≤ a := by positivity +example {a : ℚ} (ha : 0 < a) : (0 : ℝ) < a := by positivity example {r : ℝ≥0} : (0 : ℝ) ≤ r := by positivity example {r : ℝ≥0} (hr : 0 < r) : (0 : ℝ) < r := by positivity -- example {r : ℝ≥0} (hr : 0 < r) : (0 : ℝ≥0∞) < r := by positivity diff --git a/test/propose.lean b/test/propose.lean index 52ce9bb2eaa6f..4150e87c4305c 100644 --- a/test/propose.lean +++ b/test/propose.lean @@ -11,9 +11,9 @@ set_option autoImplicit true theorem foo (L M : List α) (w : L.Disjoint M) (m : a ∈ L) : a ∉ M := fun h => w m h /-- -info: Try this: have : List.Disjoint M L := List.disjoint_symm w ---- info: Try this: have : List.Disjoint K M := List.disjoint_of_subset_left m w +--- +info: Try this: have : List.Disjoint M L := List.disjoint_symm w -/ #guard_msgs in example (K L M : List α) (w : L.Disjoint M) (m : K ⊆ L) : True := by @@ -57,7 +57,7 @@ example (p : Nat × String) : True := by /-- info: Try this: have : List.Disjoint M L := List.disjoint_symm w --- -info: Try this: have : ¬a ∈ M := foo L M w m +info: Try this: have : a ∉ M := foo L M w m -/ #guard_msgs in example (K L M : List α) (w : L.Disjoint M) (m : a ∈ L) : True := by @@ -69,15 +69,15 @@ example (K L M : List α) (w : L.Disjoint M) (m : a ∈ L) : True := by /-- info: Try this: have : IsUnit p := isUnit_of_dvd_one h --- -info: Try this: have : p ≠ 1 := ne_one hp +info: Try this: have : ¬IsUnit p := not_unit hp --- -info: Try this: have : p ∣ p * p ↔ p ∣ p ∨ p ∣ p := dvd_mul hp +info: Try this: have : ¬p ∣ 1 := not_dvd_one hp --- info: Try this: have : p ≠ 0 := ne_zero hp --- -info: Try this: have : ¬p ∣ 1 := not_dvd_one hp +info: Try this: have : p ∣ p * p ↔ p ∣ p ∨ p ∣ p := dvd_mul hp --- -info: Try this: have : ¬IsUnit p := not_unit hp +info: Try this: have : p ≠ 1 := ne_one hp -/ #guard_msgs in -- From Mathlib.Algebra.Associated: diff --git a/test/push_neg.lean b/test/push_neg.lean index c376f6a0a628c..c6d24c94b233e 100644 --- a/test/push_neg.lean +++ b/test/push_neg.lean @@ -155,3 +155,40 @@ end use_distrib example (a : α) (o : Option α) (h : ¬∀ hs, o.get hs ≠ a) : ∃ hs, o.get hs = a := by push_neg at h exact h + +example (s : Set α) (h : ¬s.Nonempty) : s = ∅ := by + push_neg at h + exact h + +example (s : Set α) (h : ¬ s = ∅) : s.Nonempty := by + push_neg at h + exact h + +example (s : Set α) (h : s ≠ ∅) : s.Nonempty := by + push_neg at h + exact h + +example (s : Set α) (h : ∅ ≠ s) : s.Nonempty := by + push_neg at h + exact h + +namespace no_proj + +structure G (V : Type) where + Adj : V → V → Prop + +def g : G Nat where + Adj a b := (a ≠ b) ∧ ((a ∣ b) ∨ (b ∣ a)) + +example {p q : Nat} : ¬ g.Adj p q := by + rw [g] + guard_target =ₛ ¬ G.Adj { Adj := fun a b => (a ≠ b) ∧ ((a ∣ b) ∨ (b ∣ a)) } p q + push_neg + guard_target =ₛ ¬ G.Adj { Adj := fun a b => (a ≠ b) ∧ ((a ∣ b) ∨ (b ∣ a)) } p q + dsimp only + guard_target =ₛ ¬ ((p ≠ q) ∧ ((p ∣ q) ∨ (q ∣ p))) + push_neg + guard_target =ₛ p ≠ q → ¬p ∣ q ∧ ¬q ∣ p + exact test_sorry + +end no_proj diff --git a/test/rewrites.lean b/test/rewrites.lean index ba58073a864ce..8ebc70b0c00e0 100644 --- a/test/rewrites.lean +++ b/test/rewrites.lean @@ -109,6 +109,7 @@ axiom f_eq (n) : f n = z -- It be lovely if `rw?` could produce two *different* rewrites by `f_eq` here! #guard_msgs(drop info) in lemma test : f n = f m := by + fail_if_success rw? [-f_eq] -- Check that we can forbid lemmas. rw? rw [f_eq] diff --git a/test/ring.lean b/test/ring.lean index e0cfd276faa0c..96dfc230368dc 100644 --- a/test/ring.lean +++ b/test/ring.lean @@ -12,10 +12,12 @@ notation "ℝ" => Real example (x y : ℕ) : x + y = y + x := by ring example (x y : ℕ) : x + y + y = 2 * y + x := by ring example (x y : ℕ) : x + id y = y + id x := by ring! +example (x y : ℕ+) : x + y = y + x := by ring example {α} [CommRing α] (x y : α) : x + y + y - x = 2 * y := by ring example {α} [CommSemiring α] (x y : α) : (x + y)^2 = x^2 + 2 • x * y + y^2 := by ring example (x y : ℕ) : (x + y) ^ 3 = x ^ 3 + y ^ 3 + 3 * (x * y ^ 2 + x ^ 2 * y) := by ring +example (x y : ℕ+) : (x + y) ^ 3 = x ^ 3 + y ^ 3 + 3 * (x * y ^ 2 + x ^ 2 * y) := by ring example (x y : ℝ) : (x + y) ^ 3 = x ^ 3 + y ^ 3 + 3 * (x * y ^ 2 + x ^ 2 * y) := by ring example {α} [CommSemiring α] (x : α) : (x + 1) ^ 6 = (1 + x) ^ 6 := by ring example (n : ℕ) : (n / 2) + (n / 2) = 2 * (n / 2) := by ring diff --git a/test/search/BestFirst.lean b/test/search/BestFirst.lean index 4fcfd06730706..f402524e23c7a 100644 --- a/test/search/BestFirst.lean +++ b/test/search/BestFirst.lean @@ -7,43 +7,46 @@ import Mathlib.Data.Nat.Sqrt We check that `bestFirstSearch` can find its way around a wall. -/ -open Lean MLList +open Lean MLList Function -def wall : Int × Int → Bool := +def Point := Int × Int +deriving Repr + +def wall : Point → Bool := fun ⟨x, y⟩ => x ≤ 3 || y ≤ 3 || x ≥ 20 || y ≥ 20 || (x ≥ 6 && y ≥ 6) -def nbhd : Int × Int → MLList MetaM (Int × Int) := +def nbhd : Point → MLList MetaM Point := fun ⟨x, y⟩ => MLList.ofList ([(x+1,y), (x-1,y), (x,y+1), (x,y-1)].filter wall) -instance : Ord (Int × Int) where - compare p q := - let rp := p.1^2 + p.2^2 - let rq := q.1^2 + q.2^2 - if rp < rq then .lt - else if rq < rp then .gt - else if p.1 < q.1 then .lt - else if q.1 < p.1 then .gt - else if p.2 < q.2 then .lt - else if q.2 < p.2 then .gt - else .eq +def emb : Point → Nat ×ₗ (Int ×ₗ Int) + | (x, y) => (x.natAbs^2 + y.natAbs^2, x, y) + +lemma emb_injective : Injective emb := + fun ⟨x, y⟩ ⟨w, z⟩ h => by injection h + +instance : LinearOrder Point := LinearOrder.lift' _ emb_injective run_cmd Elab.Command.liftTermElabM do - if ! ((← bestFirstSearch nbhd (10, 10) (maxQueued := some 4) |>.takeUpToFirst (· = (0,0)) |>.force) = - [(10, 10), (11, 10), (9, 10), (8, 10), (7, 10), (6, 10), (6, 11), (6, 9), (7, 9), (6, 8), - (7, 8), (6, 7), (7, 7), (6, 6), (7, 6), (8, 6), (8, 7), (9, 6), (9, 7), (8, 8), (10, 6), (9, 8), - (8, 9), (10, 7), (11, 6), (9, 9), (11, 7), (10, 8), (12, 6), (10, 9), (11, 8), (13, 6), (12, 7), - (11, 9), (12, 8), (13, 7), (12, 9), (13, 8), (14, 7), (13, 9), (12, 10), (14, 8), (13, 10), - (12, 11), (15, 7), (14, 6), (15, 6), (15, 8), (14, 9), (16, 6), (15, 9), (14, 10), (16, 8), - (17, 6), (16, 7), (15, 10), (14, 11), (17, 7), (15, 11), (13, 11), (13, 12), (14, 12), (12, 12), - (11, 12), (10, 12), (9, 12), (8, 12), (7, 12), (6, 12), (6, 13), (7, 13), (7, 11), (8, 11), - (9, 11), (10, 11), (6, 14), (11, 11), (7, 14), (6, 15), (8, 14), (7, 15), (9, 14), (8, 15), - (8, 13), (9, 13), (10, 13), (6, 16), (11, 13), (10, 14), (12, 13), (11, 14), (7, 16), (6, 17), - (10, 15), (8, 16), (7, 17), (13, 13), (12, 14), (9, 16), (8, 17), (11, 15), (9, 15), (10, 16), - (9, 17), (14, 13), (13, 14), (11, 16), (10, 17), (14, 14), (13, 15), (15, 13), (12, 16), (11, 17), - (15, 14), (14, 15), (16, 13), (15, 12), (16, 12), (13, 16), (12, 17), (12, 15), (17, 12), - (16, 11), (17, 11), (16, 10), (17, 10), (16, 9), (17, 9), (18, 9), (17, 8), (18, 8), (19, 8), - (18, 7), (19, 7), (18, 6), (19, 6), (20, 6), (21, 6), (20, 7), (20, 5), (21, 5), (20, 4), (21, 4), - (20, 3), (21, 3), (19, 3), (18, 3), (17, 3), (16, 3), (15, 3), (14, 3), (13, 3), (12, 3), (11, 3), - (10, 3), (9, 3), (8, 3), (7, 3), (6, 3), (5, 3), (4, 3), (3, 3), (2, 3), (1, 3), (0, 3), (-1, 3), - (0, 4), (0, 2), (1, 2), (-1, 2), (0, 1), (1, 1), (-1, 1), (0, 0)]) then throwError "Test failed!" + let r := + (← bestFirstSearch nbhd (10, 10) (maxQueued := some 4) |>.takeUpToFirst (· = (0,0)) |>.force) + if r ≠ + [(10, 10), (11, 10), (9, 10), (8, 10), (7, 10), (6, 10), (6, 11), (6, 9), (7, 9), (6, 8), + (7, 8), (6, 7), (7, 7), (6, 6), (7, 6), (8, 6), (8, 7), (9, 6), (9, 7), (8, 8), (10, 6), + (9, 8), (8, 9), (10, 7), (11, 6), (9, 9), (11, 7), (10, 8), (12, 6), (10, 9), (11, 8), + (13, 6), (12, 7), (11, 9), (12, 8), (13, 7), (12, 9), (13, 8), (14, 7), (13, 9), (12, 10), + (14, 8), (13, 10), (12, 11), (15, 7), (14, 6), (15, 6), (15, 8), (14, 9), (16, 6), (15, 9), + (14, 10), (16, 8), (17, 6), (16, 7), (15, 10), (14, 11), (17, 7), (15, 11), (13, 11), + (13, 12), (14, 12), (12, 12), (11, 12), (10, 12), (9, 12), (8, 12), (7, 12), (6, 12), (6, 13), + (7, 13), (7, 11), (8, 11), (9, 11), (10, 11), (6, 14), (11, 11), (7, 14), (6, 15), (8, 14), + (7, 15), (9, 14), (8, 15), (8, 13), (9, 13), (10, 13), (6, 16), (11, 13), (10, 14), (12, 13), + (11, 14), (7, 16), (6, 17), (10, 15), (8, 16), (7, 17), (9, 16), (8, 17), (6, 18), (10, 16), + (9, 17), (9, 15), (8, 18), (11, 16), (10, 17), (12, 16), (11, 17), (11, 15), (12, 15), + (13, 15), (12, 14), (13, 14), (14, 14), (13, 13), (14, 13), (15, 13), (16, 13), (15, 14), + (15, 12), (16, 12), (17, 12), (16, 11), (17, 11), (16, 10), (17, 10), (16, 9), (17, 9), + (18, 9), (17, 8), (18, 8), (19, 8), (18, 7), (19, 7), (18, 6), (19, 6), (20, 6), (21, 6), + (20, 7), (20, 5), (21, 5), (20, 4), (21, 4), (20, 3), (21, 3), (19, 3), (18, 3), (17, 3), + (16, 3), (15, 3), (14, 3), (13, 3), (12, 3), (11, 3), (10, 3), (9, 3), (8, 3), (7, 3), + (6, 3), (5, 3), (4, 3), (3, 3), (2, 3), (1, 3), (0, 3), (-1, 3), (0, 4), (0, 2), (1, 2), + (-1, 2), (0, 1), (1, 1), (-1, 1), (0, 0)] + then throwError "Test failed!" diff --git a/test/symm.lean b/test/symm.lean index c0a0d07ac82ae..cb718ae754038 100644 --- a/test/symm.lean +++ b/test/symm.lean @@ -1,6 +1,6 @@ -import Mathlib.Tactic.Relation.Symm -import Mathlib.Algebra.Hom.Group.Defs +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Logic.Equiv.Basic +import Mathlib.Tactic.Relation.Symm set_option autoImplicit true -- testing that the attribute is recognized diff --git a/test/trace.lean b/test/trace.lean index 09a9353088abc..3ff3bd4be2429 100644 --- a/test/trace.lean +++ b/test/trace.lean @@ -1,9 +1,18 @@ import Mathlib.Tactic.Trace +import Std.Tactic.GuardMsgs +/-- +info: 7 +-/ +#guard_msgs in example : True := by trace 2 + 2 + 3 trivial +/-- +info: hello world +-/ +#guard_msgs in example : True := by trace "hello" ++ " world" trivial diff --git a/test/vec_notation.lean b/test/vec_notation.lean index 83bf6594cad96..198b8322f0d4a 100644 --- a/test/vec_notation.lean +++ b/test/vec_notation.lean @@ -11,20 +11,20 @@ open Lean open Lean.Meta open Qq -#eval do +run_cmd Elab.Command.liftTermElabM do let x : Fin 0 → ℕ := ![] - let .true ← isDefEq (toExpr x) q((![] : Fin 0 → ℕ)) | failure + guard (← isDefEq (toExpr x) q((![] : Fin 0 → ℕ))) -#eval do +run_cmd Elab.Command.liftTermElabM do let x := ![1, 2, 3] - let .true ← isDefEq (toExpr x) q(![1, 2, 3]) | failure + guard (← isDefEq (toExpr x) q(![1, 2, 3])) -#eval do +run_cmd Elab.Command.liftTermElabM do let x := ![![1, 2], ![3, 4]] - let .true ← isDefEq (toExpr x) q(![![1, 2], ![3, 4]]) | failure + guard (← isDefEq (toExpr x) q(![![1, 2], ![3, 4]])) /-! These tests are testing `PiFin.repr` -/ -#eval show MetaM Unit from guard (toString (repr (![] : _ → ℕ)) = "![]") -#eval show MetaM Unit from guard (toString (repr ![1, 2, 3]) = "![1, 2, 3]") -#eval show MetaM Unit from guard (toString (repr ![![1, 2], ![3, 4]]) = "![![1, 2], ![3, 4]]") +#guard (toString (repr (![] : _ → ℕ)) = "![]") +#guard (toString (repr ![1, 2, 3]) = "![1, 2, 3]") +#guard (toString (repr ![![1, 2], ![3, 4]]) = "![![1, 2], ![3, 4]]")