Skip to content

Commit

Permalink
chore: cleanup in Mathlib.Init (#6977)
Browse files Browse the repository at this point in the history
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
  • Loading branch information
semorrison and semorrison committed Sep 6, 2023
1 parent 9072536 commit 8f46576
Show file tree
Hide file tree
Showing 32 changed files with 59 additions and 31 deletions.
3 changes: 2 additions & 1 deletion Archive/Arithcc.lean
Expand Up @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Xi Wang
-/
import Mathlib.Order.Basic
import Mathlib.Tactic.Basic
import Mathlib.Data.Nat.Basic
import Mathlib.Tactic.Common

#align_import arithcc from "leanprover-community/mathlib"@"eb3595ed8610db8107b75b75ab64ab6390684155"

Expand Down
4 changes: 2 additions & 2 deletions Mathlib.lean
Expand Up @@ -2085,8 +2085,6 @@ import Mathlib.GroupTheory.Torsion
import Mathlib.GroupTheory.Transfer
import Mathlib.InformationTheory.Hamming
import Mathlib.Init.Algebra.Classes
import Mathlib.Init.Algebra.Functions
import Mathlib.Init.Algebra.Order
import Mathlib.Init.Align
import Mathlib.Init.CCLemmas
import Mathlib.Init.Classes.Order
Expand Down Expand Up @@ -2128,6 +2126,8 @@ import Mathlib.Init.Function
import Mathlib.Init.IteSimp
import Mathlib.Init.Logic
import Mathlib.Init.Meta.WellFoundedTactics
import Mathlib.Init.Order.Defs
import Mathlib.Init.Order.LinearOrder
import Mathlib.Init.Propext
import Mathlib.Init.Quot
import Mathlib.Init.Set
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Algebra/CharZero/Defs.lean
Expand Up @@ -3,6 +3,7 @@ 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.Init.Data.Nat.Lemmas
import Mathlib.Data.Int.Cast.Defs
import Mathlib.Tactic.NormCast.Tactic
import Mathlib.Algebra.NeZero
Expand Down
5 changes: 5 additions & 0 deletions Mathlib/Algebra/Field/Defs.lean
Expand Up @@ -179,3 +179,8 @@ end Field
`NeZero` should not be needed in the basic algebraic hierarchy.
-/
assert_not_exists NeZero

/-
Check that we have not imported `Mathlib.Tactic.Common` yet.
-/
assert_not_exists Mathlib.Tactic.LibrarySearch.librarySearch
1 change: 1 addition & 0 deletions Mathlib/Algebra/Field/IsField.lean
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Robert Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro
-/
import Mathlib.Algebra.Field.Defs
import Mathlib.Tactic.Common

#align_import algebra.field.defs from "leanprover-community/mathlib"@"2651125b48fc5c170ab1111afd0817c903b1fc6c"

Expand Down
1 change: 1 addition & 0 deletions Mathlib/Algebra/Group/Basic.lean
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro
-/
import Mathlib.Algebra.Group.Defs
import Mathlib.Tactic.Common

#align_import algebra.group.basic from "leanprover-community/mathlib"@"a07d750983b94c530ab69a726862c2ab6802b38c"

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Algebra/Group/Defs.lean
Expand Up @@ -5,7 +5,8 @@ Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro
-/
import Mathlib.Init.ZeroOne
import Mathlib.Init.Data.Int.Basic
import Mathlib.Tactic.Common
import Mathlib.Logic.Function.Basic
import Mathlib.Tactic.Basic

#align_import algebra.group.defs from "leanprover-community/mathlib"@"48fb5b5280e7c81672afc9524185ae994553ebf4"

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/NeZero.lean
Expand Up @@ -5,7 +5,7 @@ Authors: Eric Rodriguez
-/
import Mathlib.Logic.Basic
import Mathlib.Init.ZeroOne
import Mathlib.Init.Algebra.Order
import Mathlib.Init.Order.Defs

#align_import algebra.ne_zero from "leanprover-community/mathlib"@"f340f229b1f461aa1c8ee11e0a172d0a3b301a4a"

Expand Down
1 change: 1 addition & 0 deletions Mathlib/Algebra/Ring/Defs.lean
Expand Up @@ -6,6 +6,7 @@ Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Ne
import Mathlib.Algebra.Group.Defs
import Mathlib.Algebra.GroupWithZero.Defs
import Mathlib.Data.Int.Cast.Defs
import Mathlib.Tactic.Spread
import Mathlib.Util.AssertExists

#align_import algebra.ring.defs from "leanprover-community/mathlib"@"76de8ae01554c3b37d66544866659ff174e66e1f"
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Control/Random.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/

import Mathlib.Init.Algebra.Order
import Mathlib.Init.Order.Defs
import Mathlib.Init.Data.Nat.Lemmas
import Mathlib.Init.Data.Int.Order
import Mathlib.Control.ULiftable
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/Char.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro
-/
import Mathlib.Data.UInt
import Mathlib.Init.Algebra.Order
import Mathlib.Init.Order.Defs

#align_import data.char from "leanprover-community/mathlib"@"c4658a649d216f57e99621708b09dcb3dcccbd23"

Expand Down
1 change: 1 addition & 0 deletions Mathlib/Data/FunLike/Basic.lean
Expand Up @@ -5,6 +5,7 @@ Authors: Anne Baanen
-/
import Mathlib.Logic.Function.Basic
import Mathlib.Tactic.NormCast
import Mathlib.Util.CompileInductive

#align_import data.fun_like.basic from "leanprover-community/mathlib"@"a148d797a1094ab554ad4183a4ad6f130358ef64"

Expand Down
1 change: 1 addition & 0 deletions Mathlib/Data/Int/Cast/Basic.lean
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro, Gabriel Ebner
-/
import Mathlib.Init.Data.Nat.Lemmas
import Mathlib.Data.Int.Cast.Defs
import Mathlib.Algebra.Group.Basic

Expand Down
1 change: 1 addition & 0 deletions Mathlib/Data/Nat/Basic.lean
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights r
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro
-/
import Mathlib.Init.Data.Nat.Lemmas
import Mathlib.Order.Basic
import Mathlib.Algebra.GroupWithZero.Defs
import Mathlib.Algebra.Ring.Defs
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Data/Nat/Cast/Defs.lean
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro, Gabriel Ebner
-/
import Mathlib.Algebra.Group.Defs
import Mathlib.Tactic.SplitIfs

#align_import data.nat.cast.defs from "leanprover-community/mathlib"@"a148d797a1094ab554ad4183a4ad6f130358ef64"

Expand Down
1 change: 1 addition & 0 deletions Mathlib/Data/Option/Basic.lean
Expand Up @@ -8,6 +8,7 @@ import Mathlib.Data.Option.Defs
import Mathlib.Logic.IsEmpty
import Mathlib.Logic.Relator
import Mathlib.Tactic.Common
import Mathlib.Util.CompileInductive

#align_import data.option.basic from "leanprover-community/mathlib"@"f340f229b1f461aa1c8ee11e0a172d0a3b301a4a"

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/PNat/Defs.lean
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2017 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro, Neil Strickland
-/

import Mathlib.Init.Data.Nat.Lemmas
import Mathlib.Algebra.NeZero
import Mathlib.Data.Nat.Cast.Defs
import Mathlib.Order.Basic
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Init/Data/Int/Order.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad
-/

import Mathlib.Init.Algebra.Order
import Mathlib.Init.Order.Defs
import Mathlib.Init.Data.Int.Basic

/-! # The order relation on the integers -/
Expand Down
13 changes: 2 additions & 11 deletions Mathlib/Init/Data/Nat/Lemmas.lean
Expand Up @@ -6,7 +6,8 @@ Authors: Leonardo de Moura, Jeremy Avigad
import Std.Data.Nat.Lemmas
import Mathlib.Init.Data.Nat.Basic
import Mathlib.Init.Data.Nat.Div
import Mathlib.Init.Algebra.Functions
import Mathlib.Init.Order.LinearOrder
import Mathlib.Tactic.Cases

#align_import init.data.nat.lemmas from "leanprover-community/lean"@"38b59111b2b4e6c572582b27e8937e92fc70ac02"

Expand Down Expand Up @@ -673,16 +674,6 @@ instance decidableDvd : @DecidableRel ℕ (· ∣ ·) := fun _m _n =>

#align nat.dvd_of_mul_dvd_mul_right Nat.dvd_of_mul_dvd_mul_rightₓ

/-! iterate -/


def iterate {α : Sort u} (op : α → α) : ℕ → α → α
| 0, a => a
| succ k, a => iterate op k (op a)
#align nat.iterate Nat.iterate

notation:max f "^["n"]" => iterate f n

/-! find -/


Expand Down
File renamed without changes.
Expand Up @@ -3,10 +3,17 @@ Copyright (c) 2016 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
-/
import Mathlib.Init.Algebra.Order
import Mathlib.Init.Order.Defs

#align_import init.algebra.functions from "leanprover-community/lean"@"c2bcdbcbe741ed37c361a30d38e179182b989f76"

/-!
# Basic lemmas about linear orders.
The contents of this file came from `init.algebra.functions` in Lean 3,
and it would be good to find everything a better home.
-/

universe u

section
Expand Down
6 changes: 1 addition & 5 deletions Mathlib/Logic/Function/Basic.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl, Mario Carneiro
-/
import Mathlib.Logic.Nonempty
import Mathlib.Init.Data.Nat.Lemmas
import Mathlib.Tactic.Cases
import Mathlib.Init.Set

#align_import logic.function.basic from "leanprover-community/mathlib"@"29cb56a7b35f72758b05a30490e1f10bd62c35c1"
Expand Down Expand Up @@ -879,10 +879,6 @@ def Involutive {α} (f : α → α) : Prop :=
∀ x, f (f x) = x
#align function.involutive Function.Involutive

theorem involutive_iff_iter_2_eq_id {α} {f : α → α} : Involutive f ↔ f^[2] = id :=
funext_iff.symm
#align function.involutive_iff_iter_2_eq_id Function.involutive_iff_iter_2_eq_id

theorem _root_.Bool.involutive_not : Involutive not :=
Bool.not_not

Expand Down
15 changes: 14 additions & 1 deletion Mathlib/Logic/Function/Iterate.lean
Expand Up @@ -34,6 +34,15 @@ universe u v

variable {α : Type u} {β : Type v}

/-- Iterate a function. -/
def Nat.iterate {α : Sort u} (op : α → α) : ℕ → α → α
| 0, a => a
| succ k, a => iterate op k (op a)
#align nat.iterate Nat.iterate

@[inherit_doc Nat.iterate]
notation:max f "^["n"]" => Nat.iterate f n

namespace Function

open Function (Commute)
Expand Down Expand Up @@ -231,11 +240,15 @@ alias ⟨iterate_cancel_of_add, _⟩ := iterate_add_eq_iterate
#align function.iterate_cancel_of_add Function.iterate_cancel_of_add

lemma iterate_cancel (hf : Injective f) (ha : f^[m] a = f^[n] a) : f^[m - n] a = a := by
obtain h | h := le_total m n
obtain h | h := Nat.le_total m n
{ simp [Nat.sub_eq_zero_of_le h] }
{ exact iterate_cancel_of_add hf (by rwa [Nat.sub_add_cancel h]) }
#align function.iterate_cancel Function.iterate_cancel

theorem involutive_iff_iter_2_eq_id {α} {f : α → α} : Involutive f ↔ f^[2] = id :=
funext_iff.symm
#align function.involutive_iff_iter_2_eq_id Function.involutive_iff_iter_2_eq_id

end Function

namespace List
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Order/Basic.lean
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2014 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Mario Carneiro
-/
import Mathlib.Init.Order.LinearOrder
import Mathlib.Data.Prod.Basic
import Mathlib.Data.Subtype

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Order/Lattice.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl
-/
import Mathlib.Data.Bool.Basic
import Mathlib.Init.Algebra.Order
import Mathlib.Init.Order.Defs
import Mathlib.Order.Monotone.Basic
import Mathlib.Order.ULift
import Mathlib.Tactic.GCongr.Core
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Order/Monotone/Basic.lean
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2014 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Mario Carneiro, Yaël Dillies
-/
import Mathlib.Logic.Function.Iterate
import Mathlib.Init.Data.Int.Order
import Mathlib.Order.Compare
import Mathlib.Order.Max
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Order/RelClasses.lean
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2020 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Mario Carneiro, Yury G. Kudryashov
-/
import Mathlib.Init.Data.Nat.Lemmas
import Mathlib.Logic.IsEmpty
import Mathlib.Logic.Relation
import Mathlib.Order.Basic
Expand Down
1 change: 0 additions & 1 deletion Mathlib/Tactic/Basic.lean
Expand Up @@ -5,7 +5,6 @@ Authors: Mario Carneiro, Kyle Miller
-/
import Lean
import Std
import Mathlib.Tactic.Cases

set_option autoImplicit true

Expand Down
2 changes: 2 additions & 0 deletions Mathlib/Tactic/GCongr/Core.lean
Expand Up @@ -3,7 +3,9 @@ Copyright (c) 2023 Mario Carneiro, Heather Macbeth. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro, Heather Macbeth
-/
import Mathlib.Init.Order.Defs
import Mathlib.Tactic.Backtracking
import Mathlib.Tactic.Core
import Mathlib.Tactic.GCongr.ForwardAttr

/-!
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Tactic/PushNeg.lean
Expand Up @@ -7,7 +7,7 @@ Authors: Patrick Massot, Simon Hudon, Alice Laroche, Frédéric Dupuis, Jireh Lo
import Lean
import Mathlib.Lean.Expr
import Mathlib.Logic.Basic
import Mathlib.Init.Algebra.Order
import Mathlib.Init.Order.Defs
import Mathlib.Tactic.Conv

set_option autoImplicit true
Expand Down
1 change: 1 addition & 0 deletions test/Tauto.lean
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Simon Hudon, David Renshaw
-/
import Mathlib.Tactic.Tauto
import Mathlib.Tactic.SplitIfs

set_option autoImplicit true
section tauto₀
Expand Down
2 changes: 1 addition & 1 deletion test/push_neg.lean
Expand Up @@ -5,7 +5,7 @@ Author: Alice Laroche, Frédéric Dupuis, Jireh Loreaux
-/

import Mathlib.Tactic.PushNeg
import Mathlib.Init.Algebra.Order
import Mathlib.Init.Order.Defs

set_option autoImplicit true
variable {α β : Type} [LinearOrder β] {p q : Prop} {p' q' : α → Prop}
Expand Down

0 comments on commit 8f46576

Please sign in to comment.