(Pre)images of intervals
In this file we prove a bunch of trivial lemmas like “if we add a
to all points of [b, c]
,
then we get [a + b, a + c]
”. For the functions x ↦ x ± a
, x ↦ a ± x
, and x ↦ -x
we prove
lemmas about preimages and images of all intervals. We also prove a few lemmas about images under
x ↦ a * x
, x ↦ x * a
and x ↦ x⁻¹
.
Preimages under x ↦ a + x
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
Preimages under x ↦ x + a
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
Preimages under x ↦ -x
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
Preimages under x ↦ x - a
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
Preimages under x ↦ a - x
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
Images under x ↦ a + x
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
Images under x ↦ x + a
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
Images under x ↦ -x
Images under x ↦ a - x
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
Images under x ↦ x - a
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
Multiplication and inverse in a field
@[simp]
theorem
set.preimage_mul_const_Iio
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_mul_const_Ioi
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_mul_const_Iic
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_mul_const_Ici
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_mul_const_Iio_of_neg
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_mul_const_Ioi_of_neg
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_mul_const_Iic_of_neg
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_mul_const_Ici_of_neg
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_const_mul_Iio
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_const_mul_Ioi
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_const_mul_Iic
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_const_mul_Ici
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_const_mul_Ioo
{k : Type u}
[linear_ordered_field k]
(a b : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_const_mul_Ioc
{k : Type u}
[linear_ordered_field k]
(a b : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_const_mul_Ico
{k : Type u}
[linear_ordered_field k]
(a b : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_const_mul_Icc
{k : Type u}
[linear_ordered_field k]
(a b : k)
{c : k}
(h : 0 < c) :
@[simp]
theorem
set.preimage_const_mul_Iio_of_neg
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_const_mul_Ioi_of_neg
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_const_mul_Iic_of_neg
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_const_mul_Ici_of_neg
{k : Type u}
[linear_ordered_field k]
(a : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_const_mul_Ioo_of_neg
{k : Type u}
[linear_ordered_field k]
(a b : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_const_mul_Ioc_of_neg
{k : Type u}
[linear_ordered_field k]
(a b : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_const_mul_Ico_of_neg
{k : Type u}
[linear_ordered_field k]
(a b : k)
{c : k}
(h : c < 0) :
@[simp]
theorem
set.preimage_const_mul_Icc_of_neg
{k : Type u}
[linear_ordered_field k]
(a b : k)
{c : k}
(h : c < 0) :
theorem
set.image_mul_left_Icc'
{k : Type u}
[linear_ordered_field k]
{a : k}
(h : 0 < a)
(b c : k) :
theorem
set.image_mul_left_Icc
{k : Type u}
[linear_ordered_field k]
{a b c : k}
(ha : 0 ≤ a)
(hbc : b ≤ c) :
The image under inv
of Ioo 0 a
is Ioi a⁻¹
.