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

feat(Algebra/Ring): generalise and extend material about sums of squares and semireal rings #16094

Open
wants to merge 150 commits into
base: master
Choose a base branch
from
Open
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
150 commits
Select commit Hold shift + click to select a range
b2b1bdd
Rewrote Order.group.Cone to be based on the substructure pattern.
artie2000 Aug 19, 2024
ec01c52
Add inherited instances
artie2000 Aug 19, 2024
910bc2f
Update documentation + add reverse constructions
artie2000 Aug 19, 2024
e0fdbf8
Update Mathlib/Algebra/Order/Group/Cone.lean
artie2000 Aug 19, 2024
55d74d9
more work
artie2000 Aug 19, 2024
8ae9a21
Merge branch 'real_closed_field' of https://github.com/leanprover-com…
artie2000 Aug 19, 2024
919a93f
update
artie2000 Aug 19, 2024
bbb7865
missing simp lemmas
YaelDillies Aug 19, 2024
844e8e6
more work on rings
artie2000 Aug 19, 2024
347e8c5
first nontrivial proof!
artie2000 Aug 20, 2024
e78cbd2
fixing mkOf
artie2000 Aug 20, 2024
860bd1d
positivity test
artie2000 Aug 20, 2024
b0dac72
hotfix definition
artie2000 Aug 20, 2024
d6afdda
zerzer
YaelDillies Aug 20, 2024
9b95b24
for pull
artie2000 Aug 20, 2024
ca32cb2
Merge branch 'real_closed_field' of https://github.com/leanprover-com…
artie2000 Aug 20, 2024
6e9cf21
finish fixing ring cone
artie2000 Aug 20, 2024
4eaaf0f
sum of squares stuff
artie2000 Aug 21, 2024
20e6e65
final content amount
artie2000 Aug 21, 2024
9092291
fix build
YaelDillies Aug 21, 2024
490f196
changes based on feedback
artie2000 Aug 21, 2024
f0dae6b
add multiplicative versions
artie2000 Aug 21, 2024
b38ee18
docstrings
artie2000 Aug 21, 2024
a53848e
fix namespaces
artie2000 Aug 21, 2024
5fa2878
Generalise definition
artie2000 Aug 21, 2024
c0238c4
improve documentation
artie2000 Aug 21, 2024
4ccca47
docs
artie2000 Aug 21, 2024
a5bf7ba
change definition according to feedback
artie2000 Aug 22, 2024
e4d5e9a
fix to_additive style
artie2000 Aug 22, 2024
93f78cc
typo
artie2000 Aug 22, 2024
d0b2adb
Merge branch 'real_closed_field' of https://github.com/leanprover-com…
artie2000 Aug 22, 2024
0278ee9
experimental work on sums of squares
artie2000 Aug 22, 2024
7e7133b
more work
artie2000 Aug 22, 2024
5ab1c8d
Add ring facts to SumsOfSquares
artie2000 Aug 23, 2024
46f42e1
documentation
artie2000 Aug 23, 2024
bea8c72
revert irrelevant changes from other PR
artie2000 Aug 23, 2024
700ad2d
Update Mathlib/Algebra/Ring/Semireal/Defs.lean
artie2000 Aug 23, 2024
f78cc20
Update Defs.lean
artie2000 Aug 23, 2024
6448e24
fix
artie2000 Aug 23, 2024
a70e489
Update SumsOfSquares.lean
artie2000 Aug 23, 2024
25a9734
lint fix
artie2000 Aug 23, 2024
ddca474
remove deprecated names
artie2000 Aug 23, 2024
90157db
Update Mathlib/Algebra/Ring/SumsOfSquares.lean
artie2000 Aug 24, 2024
3796e0c
Update Mathlib/Algebra/Ring/Semireal/Defs.lean
artie2000 Aug 24, 2024
5329768
revert a file
artie2000 Aug 24, 2024
fa9cf8a
changes according to comments
artie2000 Aug 24, 2024
98f9645
add requested lemmas
artie2000 Aug 26, 2024
4b8810f
Yael suggestion
artie2000 Aug 26, 2024
2e7b169
style
artie2000 Aug 26, 2024
519ce62
change proof to avoid import loop
artie2000 Aug 26, 2024
0490796
change IsSumSq.mul to use sums
artie2000 Aug 26, 2024
8140e7e
improve/tidy proofs
artie2000 Aug 27, 2024
4f7ec01
fix docstring
artie2000 Aug 27, 2024
1005c72
small fix
artie2000 Aug 27, 2024
3526514
Update Mathlib/Algebra/Group/Submonoid/Membership.lean
artie2000 Aug 27, 2024
39bc35a
Merge branch 'real_closed_field_dev' of https://github.com/leanprover…
artie2000 Aug 27, 2024
21d074b
index type convention
artie2000 Aug 27, 2024
68e067c
optimise imports
artie2000 Aug 27, 2024
d3208a1
Update Mathlib/Algebra/Ring/SumsOfSquares.lean
artie2000 Aug 27, 2024
930cd88
make some arguments implicit
artie2000 Aug 27, 2024
5817e3e
Merge branch 'real_closed_field_dev' of https://github.com/leanprover…
artie2000 Aug 27, 2024
6a93923
revert removal of deprecated names, fix names
artie2000 Aug 27, 2024
5537c09
fix name
artie2000 Aug 27, 2024
8535b28
add universe polymorphism to existential; rewrite some proofs to avoi…
artie2000 Aug 27, 2024
61e4c81
fix
artie2000 Aug 27, 2024
d93a74b
rewrite some proofs, rename some theorems
artie2000 Aug 27, 2024
1f864fd
delete redundant documentation
artie2000 Aug 27, 2024
126dc58
optimise imports
artie2000 Aug 27, 2024
696b6a7
fix
artie2000 Aug 28, 2024
702123b
fix
artie2000 Aug 28, 2024
1f3220e
fix fix
artie2000 Aug 28, 2024
edcc955
add universe polymorphism, make a bunch of lemmas simp
artie2000 Aug 28, 2024
d043402
fix simp normal form
artie2000 Aug 29, 2024
ec05973
fix deprecation
artie2000 Aug 29, 2024
7bd45db
simplify proofs
artie2000 Aug 29, 2024
8fb2c43
simplify proof
artie2000 Aug 29, 2024
a2a582c
simplify proof
artie2000 Aug 30, 2024
f3adf51
Merge branch 'master' into real_closed_field_dev
artie2000 Sep 5, 2024
4cbef8b
Merge branch 'master' into real_closed_field_dev
artie2000 Dec 23, 2024
b53eac9
update
artie2000 Dec 23, 2024
62a174d
make var implicit
artie2000 Dec 23, 2024
830dbab
reorder slightly for cleaner diff
artie2000 Dec 23, 2024
502a76f
fix
artie2000 Dec 23, 2024
166c7df
lint style
artie2000 Dec 23, 2024
0389030
min imports
artie2000 Dec 23, 2024
6ca1d33
formatting
artie2000 Dec 23, 2024
5d50c3e
golf
artie2000 Dec 23, 2024
2449a15
Update Mathlib/Algebra/Group/Submonoid/Membership.lean
artie2000 Dec 24, 2024
6aff12b
Merge branch 'real_closed_field_dev' of https://github.com/leanprover…
artie2000 Dec 24, 2024
b3cae94
fix
artie2000 Dec 24, 2024
c550ecf
version without explicit sums
artie2000 Dec 24, 2024
5e522ea
lint
artie2000 Dec 24, 2024
032fa41
lint style
artie2000 Dec 24, 2024
57adeca
min imports
artie2000 Dec 24, 2024
e3b9a38
make use of existing machinery
artie2000 Dec 25, 2024
49be1dd
rewrite Group.Even, rename, generalise
artie2000 Dec 25, 2024
6ebb7a5
fix
artie2000 Dec 25, 2024
dc7c941
move lemma
artie2000 Dec 25, 2024
17f70f4
work on missing lemma
artie2000 Dec 26, 2024
1eb4be5
fix
artie2000 Dec 26, 2024
682167b
Update Mathlib.lean
artie2000 Dec 26, 2024
189ed74
Update Mathlib.lean
artie2000 Dec 26, 2024
41e638f
Merge branch 'real_closed_field_dev' of https://github.com/leanprover…
artie2000 Dec 26, 2024
dbb5741
fill in sorry
artie2000 Dec 26, 2024
56aa9ca
fix
artie2000 Dec 26, 2024
10147cb
simplify and move proof
artie2000 Dec 26, 2024
a9de515
formatting and docs
artie2000 Dec 26, 2024
e4bd192
docs; remove defeq abuse
artie2000 Dec 26, 2024
c18ce12
fix SNF error
artie2000 Dec 26, 2024
2f0060c
revert changes to Submonoid.Membership this time around
artie2000 Dec 26, 2024
da596b5
Update Mathlib/Algebra/Group/Even/Advanced.lean
artie2000 Dec 27, 2024
102d5ec
Update Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean
artie2000 Dec 27, 2024
24e9b35
syntax
artie2000 Dec 27, 2024
0193259
revert anti-defeq-abuse lemma
artie2000 Dec 27, 2024
0a85bad
Merge branch 'master' into real_closed_field_dev
artie2000 Jan 4, 2025
91b935e
Merge branch 'master' into real_closed_field_dev
artie2000 Jan 4, 2025
1ec0056
part revert
artie2000 Jan 4, 2025
1622d34
revert cont
artie2000 Jan 4, 2025
dc7a193
fix import
artie2000 Jan 4, 2025
be6d385
fix name, improve docs
artie2000 Jan 4, 2025
d2259b5
#20449 rename
artie2000 Jan 4, 2025
99f539e
Merge branch 'master' into real_closed_field_dev
artie2000 Jan 6, 2025
d792c2f
fix merge issue + style
artie2000 Jan 6, 2025
6444fb1
docs indentation
artie2000 Jan 6, 2025
20e1887
sync
artie2000 Jan 6, 2025
1190ebc
improve docs
artie2000 Jan 6, 2025
f101491
sync
artie2000 Jan 6, 2025
a9a7dce
fix
artie2000 Jan 6, 2025
ad2ba3a
style
artie2000 Jan 6, 2025
29d9631
remove spurious deprecation
artie2000 Jan 6, 2025
d0a4b98
sync
artie2000 Jan 6, 2025
7cc36d2
style
artie2000 Jan 6, 2025
bf2a090
fix
artie2000 Jan 6, 2025
56c64c5
docs
artie2000 Jan 6, 2025
72f6efa
fix typo
artie2000 Jan 6, 2025
2da9105
fix typo
artie2000 Jan 6, 2025
03c8b03
remove extra import
artie2000 Jan 6, 2025
7e6890c
sync style
artie2000 Jan 7, 2025
e7c2293
Merge branch 'master' into real_closed_field_dev
artie2000 Jan 7, 2025
e268231
docs
artie2000 Jan 7, 2025
ba72430
fix naming
artie2000 Jan 7, 2025
7561ff8
fix naming
artie2000 Jan 7, 2025
4ad9460
fix merge issues
artie2000 Jan 7, 2025
ead992b
fix merge issue
artie2000 Jan 7, 2025
9df4ae0
fix explicit
artie2000 Jan 7, 2025
a60d2a1
restore deprecation
artie2000 Jan 7, 2025
9611f41
sync
artie2000 Jan 7, 2025
62fbdde
optimise names
artie2000 Jan 7, 2025
c0a6ebe
make sure `IsSumSq 0` is provable by `aesop`
artie2000 Jan 7, 2025
529f9c4
revert test
artie2000 Jan 7, 2025
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
34 changes: 21 additions & 13 deletions Mathlib/Algebra/Ring/Semireal/Defs.lean
Original file line number Diff line number Diff line change
@@ -1,15 +1,14 @@
/-
Copyright (c) 2024 Florent Schaffhauser. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Florent Schaffhauser
Authors: Florent Schaffhauser, Artie Khovanov
-/
import Mathlib.Algebra.Ring.SumsOfSquares

/-!
# Semireal rings

A semireal ring is a non-trivial commutative ring (with unit) in which `-1` is *not* a
sum of squares.
A semireal ring is a commutative ring (with unit) in which `-1` is *not* a sum of squares.

For instance, linearly ordered rings are semireal, because sums of squares are positive and `-1` is
not.
Expand All @@ -24,22 +23,31 @@ not.
[lam_1984](https://doi.org/10.1216/RMJ-1984-14-4-767)
-/

variable (R : Type*)
variable {R : Type*}

variable (R) in
/--
A semireal ring is a non-trivial commutative ring (with unit) in which `-1` is *not* a sum of
squares. We define the class `IsSemireal R`
for all additive monoids `R` equipped with a multiplication, a multiplicative unit and a negation.
A semireal ring is a commutative ring (with unit) in which `-1` is *not* a sum of
squares. We define the predicate `IsSemireal R` for structures `R` equipped with
a multiplication, an addition, a multiplicative unit and an additive unit.
-/
@[mk_iff]
class IsSemireal [AddMonoid R] [Mul R] [One R] [Neg R] : Prop where
non_trivial : (0 : R) ≠ 1
not_isSumSq_neg_one : ¬IsSumSq (-1 : R)
class IsSemireal [Add R] [Mul R] [One R] [Zero R] : Prop where
one_add_ne_zero {s : R} (hs : IsSumSq s) : 1 + s ≠ 0

@[deprecated (since := "2024-08-09")] alias isSemireal := IsSemireal

/-- In a semireal ring, `-1` is not a sum of squares. -/
theorem IsSemireal.not_isSumSq_neg_one [AddGroup R] [One R] [Mul R] [IsSemireal R]:
¬ IsSumSq (-1 : R) := (by simpa using one_add_ne_zero ·)

@[deprecated (since := "2024-08-09")] alias isSemireal.neg_one_not_SumSq :=
IsSemireal.not_isSumSq_neg_one

instance [LinearOrderedRing R] : IsSemireal R where
non_trivial := zero_ne_one
not_isSumSq_neg_one := fun h ↦ (not_le (α := R)).2 neg_one_lt_zero h.nonneg
/--
Linearly ordered semirings with the property `a ≤ b → ∃ c, a + c = b` (e.g. `ℕ`)
are semireal.
-/
instance [LinearOrderedSemiring R] [ExistsAddOfLE R] : IsSemireal R where
one_add_ne_zero hs amo := zero_ne_one' R (le_antisymm zero_le_one
(le_of_le_of_eq (le_add_of_nonneg_right hs.nonneg) amo))
141 changes: 120 additions & 21 deletions Mathlib/Algebra/Ring/SumsOfSquares.lean
artie2000 marked this conversation as resolved.
Show resolved Hide resolved
Original file line number Diff line number Diff line change
@@ -1,11 +1,12 @@
/-
Copyright (c) 2024 Florent Schaffhauser. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Florent Schaffhauser
Authors: Florent Schaffhauser, Artie Khovanov
-/
import Mathlib.Algebra.BigOperators.Group.Finset
import Mathlib.Algebra.Group.Submonoid.Basic
import Mathlib.Algebra.Order.Ring.Defs
import Mathlib.Algebra.Ring.Subsemiring.Basic
import Mathlib.Algebra.Group.Subgroup.Even
import Mathlib.Algebra.Ring.Parity

/-!
# Sums of squares
Expand All @@ -18,7 +19,8 @@ We introduce a predicate for sums of squares in a ring.
an inductive predicate defining the property of being a sum of squares in `R`.
`0 : R` is a sum of squares and if `S` is a sum of squares, then, for all `a : R`,
`a * a + s` is a sum of squares.
- `AddMonoid.sumSq R`: the submonoid of sums of squares in an additive monoid `R`
- `AddMonoid.sumSq R` and `Subsemiring.sumSq R`: respectively
the submonoid or subsemiring of sums of squares in an additive monoid or semiring `R`
with multiplication.

-/
Expand All @@ -37,10 +39,22 @@ inductive IsSumSq [Mul R] [Add R] [Zero R] : R → Prop

@[deprecated (since := "2024-08-09")] alias isSumSq := IsSumSq

/-- Alternative induction scheme for `IsSumSq` which uses `IsSquare`. -/
theorem IsSumSq.rec' [Mul R] [Add R] [Zero R]
{motive : (s : R) → (h : IsSumSq s) → Prop}
(zero : motive 0 zero)
(sq_add : ∀ {x s}, (hx : IsSquare x) → (hs : IsSumSq s) → motive s hs →
motive (x + s) (by rcases hx with ⟨_, rfl⟩; exact sq_add _ hs))
{s : R} (h : IsSumSq s) : motive s h :=
match h with
| .zero => zero
| .sq_add _ hs => sq_add (.mul_self _) hs (rec' zero sq_add _)

/--
In an additive monoid with multiplication,
if `s₁` and `s₂` are sums of squares, then `s₁ + s₂` is a sum of squares.
-/
@[aesop unsafe 90% apply]
theorem IsSumSq.add [AddMonoid R] [Mul R] {s₁ s₂ : R}
(h₁ : IsSumSq s₁) (h₂ : IsSumSq s₂) : IsSumSq (s₁ + s₂) := by
induction h₁ <;> simp_all [add_assoc, sq_add]
Expand Down Expand Up @@ -71,50 +85,135 @@ end AddSubmonoid
@[deprecated (since := "2025-01-03")] alias sumSqIn := AddSubmonoid.sumSq
@[deprecated (since := "2025-01-06")] alias sumSq := AddSubmonoid.sumSq

/-- In an additive unital magma with multiplication, `x * x` is a sum of squares for all `x`. -/
theorem IsSumSq.mul_self [AddZeroClass R] [Mul R] (a : R) : IsSumSq (a * a) := by
rw [← add_zero (a * a)]; exact sq_add _ zero

/--
In an additive monoid with multiplication, squares are sums of squares
In an additive unital magma with multiplication, squares are sums of squares
(see Mathlib.Algebra.Group.Even).
-/
theorem mem_sumSq_of_isSquare [AddMonoid R] [Mul R] {x : R} (hx : IsSquare x) :
x ∈ AddSubmonoid.sumSq R := by
rcases hx with ⟨y, hy⟩
rw [hy, ← AddMonoid.add_zero (y * y)]
exact IsSumSq.sq_add _ IsSumSq.zero

@[deprecated (since := "2024-08-09")] alias SquaresInSumSq := mem_sumSq_of_isSquare
@[deprecated (since := "2025-01-03")] alias mem_sumSqIn_of_isSquare := mem_sumSq_of_isSquare
@[aesop unsafe 50% apply]
theorem IsSquare.isSumSq [AddZeroClass R] [Mul R] {x : R} (hx : IsSquare x) :
IsSumSq x := by rcases hx with ⟨_, rfl⟩; apply IsSumSq.mul_self

/--
In an additive monoid with multiplication `R`, the submonoid generated by the squares is the set of
sums of squares in `R`.
-/
@[simp]
theorem AddSubmonoid.closure_isSquare [AddMonoid R] [Mul R] :
closure {x : R | IsSquare x} = sumSq R := by
refine le_antisymm (closure_le.2 (fun x hx ↦ mem_sumSq_of_isSquare hx)) (fun x hx ↦ ?_)
refine closure_eq_of_le (fun x hx ↦ IsSquare.isSumSq hx) (fun x hx ↦ ?_)
induction hx <;> aesop

@[deprecated (since := "2024-08-09")] alias SquaresAddClosure := AddSubmonoid.closure_isSquare

/--
In an additive commutative monoid with multiplication, a finite sum of sums of squares
is a sum of squares.
-/
@[aesop unsafe 90% apply]
theorem IsSumSq.sum [AddCommMonoid R] [Mul R] {ι : Type*} {I : Finset ι} {s : ι → R}
(hs : ∀ i ∈ I, IsSumSq <| s i) : IsSumSq (∑ i ∈ I, s i) := by
simpa using sum_mem (S := AddSubmonoid.sumSq R) hs

/--
In an additive commutative monoid with multiplication,
`∑ i ∈ I, x i`, where each `x i` is a square, is a sum of squares.
-/
theorem IsSumSq.sum_isSquare [AddCommMonoid R] [Mul R] {ι : Type*} (I : Finset ι) {x : ι → R}
(ha : ∀ i ∈ I, IsSquare <| x i) : IsSumSq (∑ i ∈ I, x i) := by aesop

/--
In an additive commutative monoid with multiplication,
`∑ i ∈ I, a i * a i` is a sum of squares.
-/
theorem IsSumSq.sum_mul_self [AddCommMonoid R] [Mul R] {ι : Type*} (I : Finset ι) (a : ι → R) :
IsSumSq (∑ i ∈ I, a i * a i) := by
induction I using Finset.cons_induction with
| empty => simpa using IsSumSq.zero
| cons i _ hi h => exact (Finset.sum_cons (β := R) hi) ▸ IsSumSq.sq_add (a i) h
IsSumSq (∑ i ∈ I, a i * a i) := by aesop

@[deprecated (since := "2024-12-27")] alias isSumSq_sum_mul_self := IsSumSq.sum_mul_self

namespace NonUnitalSubsemiring
variable {T : Type*} [NonUnitalCommSemiring T]

variable (T) in
/--
In a commutative (possibly non-unital) semiring `R`, `NonUnitalSubsemiring.sumSq R` is
the (possibly non-unital) subsemiring of sums of squares in `R`.
-/
def sumSq : NonUnitalSubsemiring T := (Subsemigroup.square T).nonUnitalSubsemiringClosure

@[simp] theorem sumSq_toAddSubmonoid : (sumSq T).toAddSubmonoid = .sumSq T := by
rw [sumSq, ← AddSubmonoid.closure_isSquare,
Subsemigroup.nonUnitalSubsemiringClosure_toAddSubmonoid]
simp

@[simp]
theorem mem_sumSq {s : T} : s ∈ sumSq T ↔ IsSumSq s := by
rw [← NonUnitalSubsemiring.mem_toAddSubmonoid]; simp

@[simp, norm_cast] theorem coe_sumSq : sumSq T = {s : T | IsSumSq s} := by ext; simp

@[simp] theorem closure_isSquare : closure {x : T | IsSquare x} = sumSq T := by
rw [sumSq, Subsemigroup.nonUnitalSubsemiringClosure_eq_closure]
simp

end NonUnitalSubsemiring

/--
In a commutative (possibly non-unital) semiring,
if `s₁` and `s₂` are sums of squares, then `s₁ * s₂` is a sum of squares.
-/
@[aesop unsafe 50% apply]
theorem IsSumSq.mul [NonUnitalCommSemiring R] {s₁ s₂ : R}
(h₁ : IsSumSq s₁) (h₂ : IsSumSq s₂) : IsSumSq (s₁ * s₂) := by
simpa using mul_mem (by simpa : _ ∈ NonUnitalSubsemiring.sumSq R) (by simpa)

private theorem Submonoid.square_subsemiringClosure {T : Type*} [CommSemiring T] :
(Submonoid.square T).subsemiringClosure = .closure {x : T | IsSquare x} := by
simp [Submonoid.subsemiringClosure_eq_closure]

namespace Subsemiring
variable {T : Type*} [CommSemiring T]

variable (T) in
/--
In a commutative semiring `R`, `Subsemiring.sumSq R` is the subsemiring of sums of squares in `R`.
-/
def sumSq : Subsemiring T where
__ := NonUnitalSubsemiring.sumSq T
one_mem' := by aesop

@[simp] theorem sumSq_toNonUnitalSubsemiring :
(sumSq T).toNonUnitalSubsemiring = .sumSq T := rfl

@[simp]
theorem mem_sumSq {s : T} : s ∈ sumSq T ↔ IsSumSq s := by
rw [← Subsemiring.mem_toNonUnitalSubsemiring]; simp

@[simp, norm_cast] theorem coe_sumSq : sumSq T = {s : T | IsSumSq s} := by ext; simp

@[simp] theorem closure_isSquare : closure {x : T | IsSquare x} = sumSq T := by
apply_fun toNonUnitalSubsemiring using toNonUnitalSubsemiring_injective
simp [← Submonoid.square_subsemiringClosure]

end Subsemiring

/-- In a commutative semiring, a finite product of sums of squares is a sum of squares. -/
@[aesop unsafe 50% apply]
theorem IsSumSq.prod [CommSemiring R] {ι : Type*} {I : Finset ι} {x : ι → R}
(hx : ∀ i ∈ I, IsSumSq <| x i) : IsSumSq (∏ i ∈ I, x i) := by
simpa using prod_mem (S := Subsemiring.sumSq R) (by simpa)

/--
In a linearly ordered semiring with the property `a ≤ b → ∃ c, a + c = b` (e.g. `ℕ`),
sums of squares are non-negative.
-/
theorem IsSumSq.nonneg {R : Type*} [LinearOrderedSemiring R] [ExistsAddOfLE R] {s : R}
(hs : IsSumSq s) : 0 ≤ s := by
induction hs with
| zero => simp only [le_refl]
| sq_add x _ ih => apply add_nonneg ?_ ih; simp only [← pow_two x, sq_nonneg]
induction hs using IsSumSq.rec' with
| zero => aesop
| sq_add hx hs h => exact add_nonneg (IsSquare.nonneg hx) h

@[deprecated (since := "2024-08-09")] alias isSumSq.nonneg := IsSumSq.nonneg
Loading