Algebraic operations on upper/lower sets #
Upper/lower sets are preserved under pointwise algebraic operations in ordered groups.
theorem
IsUpperSet.smul_subset
{α : Type u_1}
[OrderedCommMonoid α]
{s : Set α}
{x : α}
(hs : IsUpperSet s)
(hx : 1 ≤ x)
:
theorem
IsUpperSet.vadd_subset
{α : Type u_1}
[OrderedAddCommMonoid α]
{s : Set α}
{x : α}
(hs : IsUpperSet s)
(hx : 0 ≤ x)
:
theorem
IsLowerSet.smul_subset
{α : Type u_1}
[OrderedCommMonoid α]
{s : Set α}
{x : α}
(hs : IsLowerSet s)
(hx : x ≤ 1)
:
theorem
IsLowerSet.vadd_subset
{α : Type u_1}
[OrderedAddCommMonoid α]
{s : Set α}
{x : α}
(hs : IsLowerSet s)
(hx : x ≤ 0)
:
theorem
IsUpperSet.smul
{α : Type u_1}
[OrderedCommGroup α]
{s : Set α}
{a : α}
(hs : IsUpperSet s)
:
IsUpperSet (a • s)
theorem
IsUpperSet.vadd
{α : Type u_1}
[OrderedAddCommGroup α]
{s : Set α}
{a : α}
(hs : IsUpperSet s)
:
IsUpperSet (a +ᵥ s)
theorem
IsLowerSet.smul
{α : Type u_1}
[OrderedCommGroup α]
{s : Set α}
{a : α}
(hs : IsLowerSet s)
:
IsLowerSet (a • s)
theorem
IsLowerSet.vadd
{α : Type u_1}
[OrderedAddCommGroup α]
{s : Set α}
{a : α}
(hs : IsLowerSet s)
:
IsLowerSet (a +ᵥ s)
theorem
Set.OrdConnected.smul
{α : Type u_1}
[OrderedCommGroup α]
{s : Set α}
{a : α}
(hs : s.OrdConnected)
:
(a • s).OrdConnected
theorem
Set.OrdConnected.vadd
{α : Type u_1}
[OrderedAddCommGroup α]
{s : Set α}
{a : α}
(hs : s.OrdConnected)
:
(a +ᵥ s).OrdConnected
theorem
IsUpperSet.mul_left
{α : Type u_1}
[OrderedCommGroup α]
{s t : Set α}
(ht : IsUpperSet t)
:
IsUpperSet (s * t)
theorem
IsUpperSet.add_left
{α : Type u_1}
[OrderedAddCommGroup α]
{s t : Set α}
(ht : IsUpperSet t)
:
IsUpperSet (s + t)
theorem
IsUpperSet.mul_right
{α : Type u_1}
[OrderedCommGroup α]
{s t : Set α}
(hs : IsUpperSet s)
:
IsUpperSet (s * t)
theorem
IsUpperSet.add_right
{α : Type u_1}
[OrderedAddCommGroup α]
{s t : Set α}
(hs : IsUpperSet s)
:
IsUpperSet (s + t)
theorem
IsLowerSet.mul_left
{α : Type u_1}
[OrderedCommGroup α]
{s t : Set α}
(ht : IsLowerSet t)
:
IsLowerSet (s * t)
theorem
IsLowerSet.add_left
{α : Type u_1}
[OrderedAddCommGroup α]
{s t : Set α}
(ht : IsLowerSet t)
:
IsLowerSet (s + t)
theorem
IsLowerSet.mul_right
{α : Type u_1}
[OrderedCommGroup α]
{s t : Set α}
(hs : IsLowerSet s)
:
IsLowerSet (s * t)
theorem
IsLowerSet.add_right
{α : Type u_1}
[OrderedAddCommGroup α]
{s t : Set α}
(hs : IsLowerSet s)
:
IsLowerSet (s + t)
theorem
IsUpperSet.neg
{α : Type u_1}
[OrderedAddCommGroup α]
{s : Set α}
(hs : IsUpperSet s)
:
IsLowerSet (-s)
theorem
IsLowerSet.neg
{α : Type u_1}
[OrderedAddCommGroup α]
{s : Set α}
(hs : IsLowerSet s)
:
IsUpperSet (-s)
theorem
IsUpperSet.div_left
{α : Type u_1}
[OrderedCommGroup α]
{s t : Set α}
(ht : IsUpperSet t)
:
IsLowerSet (s / t)
theorem
IsUpperSet.sub_left
{α : Type u_1}
[OrderedAddCommGroup α]
{s t : Set α}
(ht : IsUpperSet t)
:
IsLowerSet (s - t)
theorem
IsUpperSet.div_right
{α : Type u_1}
[OrderedCommGroup α]
{s t : Set α}
(hs : IsUpperSet s)
:
IsUpperSet (s / t)
theorem
IsUpperSet.sub_right
{α : Type u_1}
[OrderedAddCommGroup α]
{s t : Set α}
(hs : IsUpperSet s)
:
IsUpperSet (s - t)
theorem
IsLowerSet.div_left
{α : Type u_1}
[OrderedCommGroup α]
{s t : Set α}
(ht : IsLowerSet t)
:
IsUpperSet (s / t)
theorem
IsLowerSet.sub_left
{α : Type u_1}
[OrderedAddCommGroup α]
{s t : Set α}
(ht : IsLowerSet t)
:
IsUpperSet (s - t)
theorem
IsLowerSet.div_right
{α : Type u_1}
[OrderedCommGroup α]
{s t : Set α}
(hs : IsLowerSet s)
:
IsLowerSet (s / t)
theorem
IsLowerSet.sub_right
{α : Type u_1}
[OrderedAddCommGroup α]
{s t : Set α}
(hs : IsLowerSet s)
:
IsLowerSet (s - t)
Equations
- UpperSet.instOne = { one := UpperSet.Ici 1 }
Equations
- UpperSet.instZero = { zero := UpperSet.Ici 0 }
Equations
- UpperSet.instMul = { mul := fun (s t : UpperSet α) => { carrier := Set.image2 (fun (x1 x2 : α) => x1 * x2) ↑s ↑t, upper' := ⋯ } }
Equations
- UpperSet.instAdd = { add := fun (s t : UpperSet α) => { carrier := Set.image2 (fun (x1 x2 : α) => x1 + x2) ↑s ↑t, upper' := ⋯ } }
Equations
- UpperSet.instDiv = { div := fun (s t : UpperSet α) => { carrier := Set.image2 (fun (x1 x2 : α) => x1 / x2) ↑s ↑t, upper' := ⋯ } }
Equations
- UpperSet.instSub = { sub := fun (s t : UpperSet α) => { carrier := Set.image2 (fun (x1 x2 : α) => x1 - x2) ↑s ↑t, upper' := ⋯ } }
Equations
- UpperSet.instSMul = { smul := fun (a : α) (s : UpperSet α) => { carrier := (fun (x : α) => a • x) '' ↑s, upper' := ⋯ } }
Equations
- UpperSet.instVAdd = { vadd := fun (a : α) (s : UpperSet α) => { carrier := (fun (x : α) => a +ᵥ x) '' ↑s, upper' := ⋯ } }
@[simp]
@[simp]
@[simp]
@[simp]
Equations
Equations
Equations
Equations
Equations
Equations
Equations
- LowerSet.instOne = { one := LowerSet.Iic 1 }
Equations
- LowerSet.instZero = { zero := LowerSet.Iic 0 }
Equations
- LowerSet.instMul = { mul := fun (s t : LowerSet α) => { carrier := Set.image2 (fun (x1 x2 : α) => x1 * x2) ↑s ↑t, lower' := ⋯ } }
Equations
- LowerSet.instAdd = { add := fun (s t : LowerSet α) => { carrier := Set.image2 (fun (x1 x2 : α) => x1 + x2) ↑s ↑t, lower' := ⋯ } }
Equations
- LowerSet.instDiv = { div := fun (s t : LowerSet α) => { carrier := Set.image2 (fun (x1 x2 : α) => x1 / x2) ↑s ↑t, lower' := ⋯ } }
Equations
- LowerSet.instSub = { sub := fun (s t : LowerSet α) => { carrier := Set.image2 (fun (x1 x2 : α) => x1 - x2) ↑s ↑t, lower' := ⋯ } }
Equations
- LowerSet.instSMul = { smul := fun (a : α) (s : LowerSet α) => { carrier := (fun (x : α) => a • x) '' ↑s, lower' := ⋯ } }
Equations
- LowerSet.instVAdd = { vadd := fun (a : α) (s : LowerSet α) => { carrier := (fun (x : α) => a +ᵥ x) '' ↑s, lower' := ⋯ } }
@[simp]
@[simp]
@[simp]
@[simp]
Equations
Equations
Equations
Equations
Equations
Equations
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]