Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - chore: scope open Classical #11199

Closed
wants to merge 16 commits into from
  •  
  •  
  •  
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Algebra/Operations.lean
Original file line number Diff line number Diff line change
Expand Up @@ -346,7 +346,7 @@ end

section DecidableEq

open Classical
open scoped Classical

theorem mem_span_mul_finite_of_mem_span_mul {R A} [Semiring R] [AddCommMonoid A] [Mul A]
[Module R A] {S : Set A} {S' : Set A} {x : A} (hx : x ∈ span R (S * S')) :
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/BigOperators/Finprod.lean
Original file line number Diff line number Diff line change
Expand Up @@ -89,7 +89,7 @@ section

/- Note: we use classical logic only for these definitions, to ensure that we do not write lemmas
with `Classical.dec` in their statement. -/
open Classical
open scoped Classical

/-- Sum of `f x` as `x` ranges over the elements of the support of `f`, if it's finite. Zero
otherwise. -/
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/BigOperators/Finsupp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -688,7 +688,7 @@ theorem Finsupp.sum_apply' : g.sum k x = g.sum fun i b => k i b x :=

section

open Classical
open scoped Classical

theorem Finsupp.sum_sum_index' : (∑ x in s, f x).sum t = ∑ x in s, (f x).sum t :=
Finset.induction_on s rfl fun a s has ih => by
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Category/GroupCat/FilteredColimits.lean
Original file line number Diff line number Diff line change
Expand Up @@ -29,7 +29,7 @@ universe v u

noncomputable section

open Classical
open scoped Classical

open CategoryTheory

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean
Original file line number Diff line number Diff line change
Expand Up @@ -27,7 +27,7 @@ namespace ModuleCat

universe u

open Classical
open scoped Classical

variable (R : Type u)

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Category/MonCat/FilteredColimits.lean
Original file line number Diff line number Diff line change
Expand Up @@ -29,7 +29,7 @@ universe v u

noncomputable section

open Classical
open scoped Classical

open CategoryTheory

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Category/Ring/Adjunctions.lean
Original file line number Diff line number Diff line change
Expand Up @@ -25,7 +25,7 @@ open CategoryTheory
namespace CommRingCat
set_option linter.uppercaseLean3 false -- `CommRing`

open Classical
open scoped Classical

/-- The free functor `Type u ⥤ CommRingCat` sending a type `X` to the multivariable (commutative)
polynomials with variables `x : X`.
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Category/Ring/FilteredColimits.lean
Original file line number Diff line number Diff line change
Expand Up @@ -27,7 +27,7 @@ universe v u

noncomputable section

open Classical
open scoped Classical

open CategoryTheory

Expand Down
10 changes: 5 additions & 5 deletions Mathlib/Algebra/DirectLimit.lean
Original file line number Diff line number Diff line change
Expand Up @@ -264,7 +264,7 @@ end functorial

section Totalize

open Classical
open scoped Classical

variable (G f)

Expand All @@ -289,7 +289,7 @@ end Totalize

variable [DirectedSystem G fun i j h => f i j h]

open Classical
open scoped Classical

theorem toModule_totalize_of_le {x : DirectSum ι G} {i j : ι} (hij : i ≤ j)
(hx : ∀ k ∈ x.support, k ≤ i) :
Expand Down Expand Up @@ -633,7 +633,7 @@ theorem exists_of [Nonempty ι] [IsDirected ι (· ≤ ·)] (z : DirectLimit G f

section

open Classical
open scoped Classical

open Polynomial

Expand Down Expand Up @@ -667,7 +667,7 @@ theorem induction_on [Nonempty ι] [IsDirected ι (· ≤ ·)] {C : DirectLimit

section OfZeroExact

open Classical
open scoped Classical

variable (f' : ∀ i j, i ≤ j → G i →+* G j)

Expand Down Expand Up @@ -1003,7 +1003,7 @@ theorem exists_inv {p : Ring.DirectLimit G f} : p ≠ 0 → ∃ y, p * y = 1 :=

section

open Classical
open scoped Classical

/-- Noncomputable multiplicative inverse in a direct limit of fields. -/
noncomputable def inv (p : Ring.DirectLimit G f) : Ring.DirectLimit G f :=
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/EuclideanDomain/Defs.lean
Original file line number Diff line number Diff line change
Expand Up @@ -173,7 +173,7 @@ theorem div_zero (a : R) : a / 0 = 0 :=

section

open Classical
open scoped Classical

@[elab_as_elim]
theorem GCD.induction {P : R → R → Prop} (a b : R) (H0 : ∀ x, P 0 x)
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Field/IsField.lean
Original file line number Diff line number Diff line change
Expand Up @@ -60,7 +60,7 @@ theorem not_isField_of_subsingleton (R : Type u) [Semiring R] [Subsingleton R] :
h (Subsingleton.elim _ _)
#align not_is_field_of_subsingleton not_isField_of_subsingleton

open Classical
open scoped Classical

/-- Transferring from `IsField` to `Semifield`. -/
noncomputable def IsField.toSemifield {R : Type u} [Semiring R] (h : IsField R) : Semifield R :=
Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Algebra/GCDMonoid/Div.lean
Original file line number Diff line number Diff line change
Expand Up @@ -70,7 +70,8 @@ end Int

namespace Polynomial

open Polynomial Classical
open scoped Classical
open Polynomial

noncomputable section

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/GroupWithZero/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -36,7 +36,7 @@ and require `0⁻¹ = 0`.
-/


open Classical
open scoped Classical

open Function

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/GroupWithZero/Commute.lean
Original file line number Diff line number Diff line change
Expand Up @@ -21,7 +21,7 @@ variable [MonoidWithZero M₀]

namespace Ring

open Classical
open scoped Classical

theorem mul_inverse_rev' {a b : M₀} (h : Commute a b) :
inverse (a * b) = inverse b * inverse a := by
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Algebra/GroupWithZero/Units/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -78,7 +78,7 @@ theorem not_isUnit_zero [Nontrivial M₀] : ¬IsUnit (0 : M₀) :=

namespace Ring

open Classical
open scoped Classical

/-- Introduce a function `inverse` on a monoid with zero `M₀`, which sends `x` to `x⁻¹` if `x` is
invertible and to `0` otherwise. This definition is somewhat ad hoc, but one needs a fully (rather
Expand Down Expand Up @@ -328,7 +328,7 @@ end CommGroupWithZero

section NoncomputableDefs

open Classical
open scoped Classical

variable {M : Type*} [Nontrivial M]

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Homology/ComplexShape.lean
Original file line number Diff line number Diff line change
Expand Up @@ -43,7 +43,7 @@ so `d : X i ⟶ X j` is nonzero only when `i = j + 1`.

noncomputable section

open Classical
open scoped Classical

/-- A `c : ComplexShape ι` describes the shape of a chain complex,
with chain groups indexed by `ι`.
Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Algebra/Homology/Homology.lean
Original file line number Diff line number Diff line change
Expand Up @@ -40,7 +40,8 @@ variable {V : Type u} [Category.{v} V] [HasZeroMorphisms V]

variable {c : ComplexShape ι} (C : HomologicalComplex V c)

open Classical ZeroObject
open scoped Classical
open ZeroObject

noncomputable section

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Homology/Homotopy.lean
Original file line number Diff line number Diff line change
Expand Up @@ -18,7 +18,7 @@ We define chain homotopies, and prove that homotopic chain maps induce the same

universe v u

open Classical
open scoped Classical

noncomputable section

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Homology/HomotopyCategory.lean
Original file line number Diff line number Diff line change
Expand Up @@ -22,7 +22,7 @@ set_option autoImplicit true

universe v u

open Classical
open scoped Classical

noncomputable section

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Homology/ImageToKernel.lean
Original file line number Diff line number Diff line change
Expand Up @@ -35,7 +35,7 @@ variable {ι : Type*}

variable {V : Type u} [Category.{v} V] [HasZeroMorphisms V]

open Classical
open scoped Classical

noncomputable section

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Algebra/MonoidAlgebra/Degree.lean
Original file line number Diff line number Diff line change
Expand Up @@ -22,7 +22,8 @@ variable {R R' A T B ι : Type*}

namespace AddMonoidAlgebra

open Classical BigOperators
open scoped Classical
open BigOperators

/-!

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Order/CauSeq/Completion.lean
Original file line number Diff line number Diff line change
Expand Up @@ -198,7 +198,7 @@ instance Cauchy.commRing : CommRing (Cauchy abv) :=

end

open Classical
open scoped Classical

section

Expand Down
4 changes: 2 additions & 2 deletions Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean
Original file line number Diff line number Diff line change
Expand Up @@ -480,7 +480,7 @@ section Field

/-! ### Group operation polynomials over a field -/

open Classical
open scoped Classical

/-- The slope of the line through two affine points $(x_1, y_1)$ and $(x_2, y_2)$ in `W`.
If $x_1 \ne x_2$, then this line is the secant of `W` through $(x_1, y_1)$ and $(x_2, y_2)$,
Expand Down Expand Up @@ -691,7 +691,7 @@ lemma neg_some {x y : R} (h : W.nonsingular x y) : -some h = some (nonsingular_n
instance : InvolutiveNeg W.Point :=
⟨by rintro (_ | _) <;> simp [zero_def]; ring1⟩

open Classical
open scoped Classical

variable {F : Type u} [Field F] {W : Affine F}

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -47,7 +47,7 @@ and Chris Hughes (on an earlier repository).

noncomputable section

open Classical
open scoped Classical

universe u v

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/AlgebraicGeometry/PrimeSpectrum/Maximal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -28,7 +28,7 @@ natural inclusion into the prime spectrum to avoid API duplication for zero loci

noncomputable section

open Classical
open scoped Classical

universe u v

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/Analytic/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -74,7 +74,8 @@ noncomputable section

variable {𝕜 E F G : Type*}

open Topology Classical BigOperators NNReal Filter ENNReal
open scoped Classical
open Topology BigOperators NNReal Filter ENNReal

open Set Filter Asymptotics

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/Analytic/CPolynomial.lean
Original file line number Diff line number Diff line change
Expand Up @@ -42,7 +42,8 @@ We develop the basic properties of these notions, notably:
variable {𝕜 E F G : Type*} [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E]
[NormedAddCommGroup F] [NormedSpace 𝕜 F] [NormedAddCommGroup G] [NormedSpace 𝕜 G]

open Topology Classical BigOperators NNReal Filter ENNReal
open scoped Classical
open Topology BigOperators NNReal Filter ENNReal

open Set Filter Asymptotics

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/Analytic/Constructions.lean
Original file line number Diff line number Diff line change
Expand Up @@ -19,7 +19,8 @@ We show that the following are analytic:

noncomputable section

open Topology Classical BigOperators NNReal Filter ENNReal
open scoped Classical
open Topology BigOperators NNReal Filter ENNReal

open Set Filter Asymptotics

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/Asymptotics/Asymptotics.lean
Original file line number Diff line number Diff line change
Expand Up @@ -49,7 +49,8 @@ the Fréchet derivative.)

open Filter Set

open Topology BigOperators Classical Filter NNReal
open scoped Classical
open Topology BigOperators Filter NNReal

namespace Asymptotics

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/BoxIntegral/Box/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -56,7 +56,8 @@ open Set Function Metric Filter

noncomputable section

open NNReal Classical Topology
open scoped Classical
open NNReal Topology

namespace BoxIntegral

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/BoxIntegral/Partition/Additive.lean
Original file line number Diff line number Diff line change
Expand Up @@ -31,7 +31,8 @@ rectangular box, additive function

noncomputable section

open Classical BigOperators Function Set
open scoped Classical
open BigOperators Function Set

namespace BoxIntegral

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/BoxIntegral/Partition/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -41,7 +41,8 @@ rectangular box, partition

open Set Finset Function

open Classical NNReal BigOperators
open scoped Classical
open NNReal BigOperators

noncomputable section

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/BoxIntegral/Partition/Filter.lean
Original file line number Diff line number Diff line change
Expand Up @@ -168,7 +168,8 @@ integral, rectangular box, partition, filter

open Set Function Filter Metric Finset Bool

open Classical Topology Filter NNReal
open scoped Classical
open Topology Filter NNReal

noncomputable section

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/BoxIntegral/Partition/Split.lean
Original file line number Diff line number Diff line change
Expand Up @@ -40,7 +40,8 @@ rectangular box, partition, hyperplane

noncomputable section

open Classical BigOperators Filter
open scoped Classical
open BigOperators Filter

open Function Set Filter

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/BoxIntegral/Partition/SubboxInduction.lean
Original file line number Diff line number Diff line change
Expand Up @@ -33,7 +33,8 @@ namespace BoxIntegral

open Set Metric

open Classical Topology
open scoped Classical
open Topology

noncomputable section

Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean
Original file line number Diff line number Diff line change
Expand Up @@ -27,7 +27,8 @@ rectangular box, box partition

noncomputable section

open Classical ENNReal NNReal
open scoped Classical
open ENNReal NNReal

open Set Function

Expand Down