Definition: regular-int-seq
The factor ⌜2⌝ on the right of the inequality is essential to prove accelerate_wf

k-regular-seq(f) ==  ∀n,m:ℕ+.  (|(m (f n)) (f m)| ≤ ((2 k) (n m)))

Lemma: regular-int-seq_wf
[k:ℤ]. ∀[f:ℕ+ ⟶ ℤ].  (k-regular-seq(f) ∈ ℙ)

Definition: strong-regular-int-seq
strong-regular-int-seq(a;b;f) ==  ∀n,m:ℕ+.  ((a |(m (f n)) (f m)|) ≤ (b (n m)))

Lemma: strong-regular-int-seq_wf
[a,b:ℤ]. ∀[f:ℕ+ ⟶ ℤ].  (strong-regular-int-seq(a;b;f) ∈ ℙ)

Lemma: regular-int-seq-weakening
[n,k:ℤ]. ∀[f:ℕ+ ⟶ ℤ].  (k-regular-seq(f)) supposing (n-regular-seq(f) and (n ≤ k))

Lemma: sq_stable__regular-int-seq
[k:ℕ+]. ∀[f:ℕ+ ⟶ ℤ].  SqStable(k-regular-seq(f))

Definition: bdd-diff
bdd-diff(f;g) ==  ∃B:ℕ. ∀n:ℕ+(|(f n) n| ≤ B)

Lemma: bdd-diff_wf
[f,g:ℕ+ ⟶ ℤ].  (bdd-diff(f;g) ∈ ℙ)

Lemma: trivial-bdd-diff
[f,g:ℕ+ ⟶ ℤ].  bdd-diff(f;g) supposing ∀n:ℕ+((f n) (g n) ∈ ℤ)

Lemma: bdd-diff-equiv
EquivRel(ℕ+ ⟶ ℤ;f,g.bdd-diff(f;g))

Lemma: bdd-diff_weakening
a,b:ℕ+ ⟶ ℤ.  ((a b ∈ (ℕ+ ⟶ ℤ))  bdd-diff(a;b))

Lemma: bdd-diff_inversion
a,b:ℕ+ ⟶ ℤ.  (bdd-diff(a;b)  bdd-diff(b;a))

Lemma: bdd-diff_transitivity
a,b,c:ℕ+ ⟶ ℤ.  (bdd-diff(a;b)  bdd-diff(b;c)  bdd-diff(a;c))

Lemma: bdd-diff_functionality
x1,x2,y1,y2:ℕ+ ⟶ ℤ.  (bdd-diff(x1;x2)  bdd-diff(y1;y2)  (bdd-diff(x1;y1) ⇐⇒ bdd-diff(x2;y2)))

Lemma: bdd-diff-add
[f1,f2,g1,g2:ℕ+ ⟶ ℤ].  (bdd-diff(f1;f2)  bdd-diff(g1;g2)  bdd-diff(λn.(f1[n] g1[n]);λn.(f2[n] g2[n])))

Lemma: bdd-diff-regular-int-seq
[k,b:ℕ]. ∀[f:{f:ℕ+ ⟶ ℤk-regular-seq(f)} ]. ∀[g:ℕ+ ⟶ ℤ].
  b-regular-seq(g) supposing ∀n:ℕ+(|(f n) n| ≤ (2 b))

Lemma: bdd-diff-iff-eventual
f,g:ℕ+ ⟶ ℤ.  (∃m:ℕ+. ∃B:ℕ. ∀n:{m...}. (|(f n) n| ≤ B) ⇐⇒ bdd-diff(f;g))

Lemma: eventually-equal-implies-bdd-diff
f,g:ℕ+ ⟶ ℤ.  ((∃m:ℕ+. ∀n:{m...}. ((f n) (g n) ∈ ℤ))  bdd-diff(f;g))

Definition: real
ℝ ==  {f:ℕ+ ⟶ ℤregular-seq(f)} 

Lemma: real_wf
ℝ ∈ Type

Lemma: real-regular
[x:ℝ]. ∀[k:ℕ+].  k-regular-seq(x)

Lemma: implies-equal-real
[x,y:ℝ].  y ∈ ℝ supposing ∀n:ℕ+((x n) (y n) ∈ ℤ)

Lemma: real-has-valueall
[x:ℝ]. has-valueall(x)

Lemma: real-has-value
[x:ℝ]. (x)↓

Definition: accelerate
accelerate(k;f) ==  eval k2 in λn.eval k2 in (f m) ÷ k2

Lemma: accelerate_wf
[k:ℕ+]. ∀[f:{f:ℕ+ ⟶ ℤk-regular-seq(f)} ].  (accelerate(k;f) ∈ ℝ)

Lemma: accelerate-bdd-diff
k:ℕ+. ∀[f:{f:ℕ+ ⟶ ℤk-regular-seq(f)} ]. bdd-diff(accelerate(k;f);f)

Lemma: accelerate-real-strong-regular
k:ℕ+. ∀x:ℝ.  strong-regular-int-seq(2 k;(2 k) 1;accelerate(k;x))

Lemma: accelerate1-real-strong-regular
x:ℝstrong-regular-int-seq(2;3;accelerate(1;x))

Lemma: regular-consistency
[x,y:ℝ]. ∀[n,m:ℕ+].  ((m |(x n) n|) ≤ ((n |(x m) m|) (4 n) (4 m)))

Lemma: bdd-diff-regular
[x,y:ℕ+ ⟶ ℤ]. ∀[k,l:ℕ+].
  (∀n:ℕ+(|(x n) n| ≤ ((2 k) (2 l)))) supposing (bdd-diff(x;y) and k-regular-seq(x) and l-regular-seq(y))

Lemma: regular-less
[x,y:ℝ].  ∀n:ℕ+((x n) 4 <  (∀m:ℕ+((x m) ≤ ((y m) 4))))

Lemma: regular-less-iff
[x,y:ℝ].  (∃n:ℕ+ [(x n) 4 < n] ⇐⇒ ∀b:{4...}. ∃n:ℕ+. ∀m:{n...}. (x m) b < m)

Definition: reg-less
reg-less(x;y) ==  eval x' in eval y' in   find-ge(λn.(x n) 4 <n;1)

Lemma: reg-less_wf
[x:ℝ]. ∀[y:{y:ℝ| ∃n:ℕ+ [(x n) 4 < n]} ].  (reg-less(x;y) ∈ {n:ℕ+(x n) 4 < n} )

Lemma: fun-not-int
[f:ℕ+ ⟶ ℤ]. (isint(f) ff)

Definition: int-to-real
r(n) ==  λk.(2 n)

Lemma: int-to-real_wf
[n:ℤ]. (r(n) ∈ ℝ)

Lemma: real-list-has-valueall
[x:ℝ List]. has-valueall(x)

Lemma: real-valueall-type
valueall-type(ℝ)

Lemma: valueall-type-real-list
valueall-type(ℝ List)

Definition: reg-seq-list-add
reg-seq-list-add(L) ==  λn.cbv_list_accum(s,a.s (a n);0;L)

Lemma: reg-seq-list-add-as-l_sum
[L:(ℕ+ ⟶ ℤList]. (reg-seq-list-add(L) n.l_sum(map(λx.(x n);L))) ∈ (ℕ+ ⟶ ℤ))

Lemma: reg-seq-list-add_wf
[L:ℝ List]. (reg-seq-list-add(L) ∈ {f:ℕ+ ⟶ ℤ||L||-regular-seq(f)} )

Lemma: reg-seq-list-add_functionality_wrt_permutation
[L,L':ℝ List].
  reg-seq-list-add(L) reg-seq-list-add(L') ∈ {f:ℕ+ ⟶ ℤ||L||-regular-seq(f)}  supposing permutation(ℝ;L;L')

Lemma: reg-seq-list-add_functionality_wrt_bdd-diff
L,L':(ℕ+ ⟶ ℤList.
  (((||L|| ||L'|| ∈ ℤ) ∧ (∀i:ℕ||L||. bdd-diff(L[i];L'[i])))  bdd-diff(reg-seq-list-add(L);reg-seq-list-add(L')))

Definition: reg-seq-add
reg-seq-add(x;y) ==  λn.((x n) (y n))

Lemma: reg-seq-add_wf
[x,y:ℕ+ ⟶ ℤ].  (reg-seq-add(x;y) ∈ ℕ+ ⟶ ℤ)

Lemma: reg-seq-add_functionality_wrt_bdd-diff
x,x',y,y':ℕ+ ⟶ ℤ.  (bdd-diff(x;x')  bdd-diff(y;y')  bdd-diff(reg-seq-add(x;y);reg-seq-add(x';y')))

Definition: radd-list
radd-list(L) ==  let L' ⟵ in eval ||L'|| in if (k =z 0) then r0 else accelerate(k;reg-seq-list-add(L')) fi 

Lemma: radd-list_wf
[L:ℝ List]. (radd-list(L) ∈ ℝ)

Lemma: radd-list_functionality_wrt_permutation
[L1,L2:ℝ List].  radd-list(L1) radd-list(L2) ∈ ℝ supposing permutation(ℝ;L1;L2)

Lemma: radd-list_wf-bag
[L:bag(ℝ)]. (radd-list(L) ∈ ℝ)

Definition: radd
==  accelerate(2;reg-seq-list-add([a; b]))

Lemma: radd_wf
[a,b:ℝ].  (a b ∈ ℝ)

Lemma: radd-bdd-diff
a,b:ℝ.  bdd-diff(a b;λn.((a n) (b n)))

Lemma: radd-approx
[a,b:ℝ]. ∀[n:ℕ+].  ((a b) ((a (4 n)) (b (4 n))) ÷ 4)

Definition: rminus
-(x) ==  λn.(-(x n))

Lemma: rminus_wf
[x:ℝ]. (-(x) ∈ ℝ)

Lemma: rminus-int
[n:ℤ]. (-(r(n)) r(-n) ∈ ℝ)

Definition: reg_seq_mul
reg_seq_mul(x;y) ==  λn.[(x n) (y n) ÷ n]

Lemma: reg_seq_mul_wf
[x,y:ℕ+ ⟶ ℤ].  (reg_seq_mul(x;y) ∈ ℕ+ ⟶ ℤ)

Lemma: reg_seq_mul-regular-eventually
[x,y:ℝ].
  ∀B,b:ℕ+.
    ∀n,m:{b...}.  (|(m (reg_seq_mul(x;y) n)) (reg_seq_mul(x;y) m)| ≤ ((2 B) (n m))) 
    supposing ∀n,m:{b...}.  ((2 ((m |x n|) (n |y m|))) ≤ ((n m) ((4 B) 1)))

Lemma: reg_seq_mul-regular
[x,y:ℝ].
  ∀B:ℕ+
    B-regular-seq(reg_seq_mul(x;y)) supposing ∀n,m:ℕ+.  ((2 ((m |x n|) (n |y m|))) ≤ ((n m) ((4 B) 1)))

Definition: reg-seq-mul
reg-seq-mul(x;y) ==  λn.(((x n) (y n)) ÷ n)

Lemma: reg-seq-mul_wf
[x,y:ℕ+ ⟶ ℤ].  (reg-seq-mul(x;y) ∈ ℕ+ ⟶ ℤ)

Lemma: reg-seq-mul-regular-eventually
[x,y:ℝ].
  ∀B,b:ℕ+.
    ∀n,m:{b...}.  (|(m (reg-seq-mul(x;y) n)) (reg-seq-mul(x;y) m)| ≤ ((2 B) (n m))) 
    supposing ∀n,m:{b...}.  ((2 ((m |x n|) (n |y m|))) ≤ ((n m) ((4 B) 2)))

Lemma: reg-seq-mul-regular
[x,y:ℝ].  ∀k:ℕ+1-regular-seq(reg-seq-mul(x;y)) supposing ∀n:ℕ+((|x n| ≤ (n k)) ∧ (|y n| ≤ (n k)))

Definition: canon-bnd
canon-bnd(x) ==  |x 1| 3

Lemma: canon-bnd_wf
[x:ℝ]. (canon-bnd(x) ∈ {k:{3...}| ∀n:ℕ+(|x n| ≤ (n k))} )

Definition: canonical-bound
canonical-bound(r) ==  (|r 1| 4) ÷ 2

Lemma: canonical-bound_wf
[r:ℝ]. (canonical-bound(r) ∈ {k:{2...}| ∀n:ℕ+(|r n| ≤ ((2 n) k))} )

Lemma: reg-seq-mul_wf2
[x,y:ℝ].  (reg-seq-mul(x;y) ∈ {f:ℕ+ ⟶ ℤimax(|x 1|;|y 1|) 4-regular-seq(f)} )

Lemma: reg-seq-mul-comm
[x,y:ℕ+ ⟶ ℤ].  (reg-seq-mul(x;y) reg-seq-mul(y;x) ∈ (ℕ+ ⟶ ℤ))

Lemma: reg-seq-mul-assoc
x,y,z:ℝ.  bdd-diff(reg-seq-mul(reg-seq-mul(x;y);z);reg-seq-mul(x;reg-seq-mul(y;z)))

Lemma: reg-seq-mul_functionality_wrt_bdd-diff
x1:ℝ. ∀[x2,y1:ℕ+ ⟶ ℤ].  ∀y2:ℝ(bdd-diff(y1;y2)  bdd-diff(x1;x2)  bdd-diff(reg-seq-mul(x1;y1);reg-seq-mul(x2;y2)))

Definition: rmul
==  eval a1 in eval b1 in   accelerate(imax(|a1 1|;|b1 1|) 4;reg-seq-mul(a1;b1))

Lemma: rmul_wf
[a,b:ℝ].  (a b ∈ ℝ)

Lemma: rmul-bdd-diff-reg-seq-mul
a,b:ℝ.  bdd-diff(a b;reg-seq-mul(a;b))

Definition: req
==  ∀n:ℕ+(|(x n) n| ≤ 4)

Lemma: req_wf
[x,y:ℝ].  (x y ∈ ℙ)

Lemma: stable_req
[x,y:ℝ].  Stable{x y}

Lemma: req-iff-bdd-diff
[x,y:ℝ].  uiff(x y;bdd-diff(x;y))

Lemma: req_witness
[x,y:ℝ].  ((x y)  n.<λ_.Ax, Ax, Ax> ∈ y))

Lemma: sq_stable__req
[x,y:ℝ].  SqStable(x y)

Lemma: accelerate-req
[k:ℕ+]. ∀[x:ℝ].  ((accelerate(k;x) ∈ ℝ) ∧ (accelerate(k;x) x))

Lemma: req-equiv
EquivRel(ℝ;x,y.x y)

Lemma: req_weakening
[a,b:ℝ].  supposing b ∈ ℝ

Lemma: req-same
[a:ℝ]. (a a)

Lemma: req_inversion
[a,b:ℝ].  supposing b

Lemma: req_transitivity
[a,b,c:ℝ].  (a c) supposing ((b c) and (a b))

Lemma: req_functionality
[x1,x2,y1,y2:ℝ].  (uiff(x1 y1;x2 y2)) supposing ((y1 y2) and (x1 x2))

Lemma: req-int
[a,b:ℤ].  uiff(r(a) r(b);a b ∈ ℤ)

Lemma: radd-list_functionality
[L1,L2:ℝ List].  radd-list(L1) radd-list(L2) supposing (||L1|| ||L2|| ∈ ℤ) ∧ (∀i:ℕ||L1||. (L1[i] L2[i]))

Lemma: radd_list_nil_lemma
radd-list([]) r0

Lemma: radd-list-cons
[L:ℝ List]. ∀[x:ℝ].  (radd-list([x L]) (x radd-list(L)))

Lemma: radd-as-radd-list
[a,b:ℝ].  ((a b) radd-list([a; b]) ∈ ℝ)

Lemma: radd_functionality
[a1,a2,b1,b2:ℝ].  ((a1 b1) (a2 b2)) supposing ((a1 a2) and (b1 b2))

Lemma: rminus_functionality
[x,y:ℝ].  -(x) -(y) supposing y

Lemma: rminus_functionality_wrt_bdd-diff
x,y:ℕ+ ⟶ ℤ.  (bdd-diff(x;y)  bdd-diff(-(x);-(y)))

Lemma: rminus-rminus
[x:ℝ]. (-(-(x)) x)

Lemma: rminus-rminus-eq
[x:ℝ]. (-(-(x)) x ∈ ℝ)

Lemma: radd_comm_eq
[a,b:ℝ].  ((a b) (b a) ∈ ℝ)

Lemma: radd_comm
[a,b:ℝ].  ((a b) (b a))

Lemma: radd_assoc
[a,b,c:ℝ].  (((a b) c) (a c))

Lemma: radd-assoc
[x,y,z:ℝ].  ((x z) ((x y) z))

Lemma: radd-ac
[a,b,c:ℝ].  ((a c) (b c))

Lemma: rmul_assoc
[a,b,c:ℝ].  (((a b) c) (a c))

Lemma: rmul-assoc
[a,b,c:ℝ].  ((a c) ((a b) c))

Lemma: rmul_comm
[a,b:ℝ].  ((a b) (b a))

Lemma: rmul_functionality
[r1,r2,s1,s2:ℝ].  ((r1 s1) (r2 s2)) supposing ((s1 s2) and (r1 r2))

Lemma: rmul-ac
[a,b,c:ℝ].  ((a c) (b c))

Lemma: radd-int
[a,b:ℤ].  ((r(a) r(b)) r(a b))

Lemma: radd-zero
[x:ℝ]. ((x r0) x)

Lemma: radd-zero-both
[x:ℝ]. (((x r0) x) ∧ ((r0 x) x))

Lemma: radd-rminus
[x:ℝ]. ((x -(x)) r0)

Lemma: radd-rminus-both
[x:ℝ]. (((x -(x)) r0) ∧ ((-(x) x) r0))

Lemma: radd-rminus-assoc
[x,y:ℝ].  (((x -(x) y) y) ∧ ((-(x) y) y))

Lemma: radd-list-append
[L1,L2:ℝ List].  (radd-list(L1 L2) (radd-list(L1) radd-list(L2)))

Lemma: rmul-int
[a,b:ℤ].  ((r(a) r(b)) r(a b))

Lemma: rmul-zero
[x:ℝ]. ((x r0) r0)

Lemma: rmul-zero-both
[x:ℝ]. (((x r0) r0) ∧ ((r0 x) r0))

Lemma: rmul-one
[x:ℝ]. ((x r1) x)

Lemma: rmul-identity1
[x:ℝ]. ((r1 x) x)

Lemma: rmul-one-both
[x:ℝ]. (((x r1) x) ∧ ((r1 x) x))

Lemma: rmul-distrib1
[x,y,z:ℝ].  ((x (y z)) ((x y) (x z)))

Lemma: rmul-distrib2
[x,y,z:ℝ].  (((y z) x) ((y x) (z x)))

Lemma: rmul-distrib
[a,b,c:ℝ].  (((a (b c)) ((a b) (a c))) ∧ (((b c) a) ((b a) (c a))))

Lemma: rminus-as-rmul
[r:ℝ]. (-(r) (r(-1) r))

Lemma: rmul-minus
[r:ℝ]. ∀[n:ℤ].  ((r r(-n)) (-(r) r(n)))

Lemma: rminus-zero
-(r0) r0

Lemma: rmul_over_rminus
[a,b:ℝ].  (((-(a) b) -(a b)) ∧ ((a -(b)) -(a b)))

Lemma: rminus-radd
[r,s:ℝ].  (-(r s) ((r(-1) s) (r(-1) r)))

Lemma: radd-list-linearity
[T:Type]. ∀[x,y:T ⟶ ℝ]. ∀[a,b:ℝ]. ∀[L:T List].
  (radd-list(map(λk.((a x[k]) (b y[k]));L)) ((a radd-list(map(λk.x[k];L))) (b radd-list(map(λk.y[k];L)))))

Lemma: radd-list-linearity1
[T:Type]. ∀[x,y:T ⟶ ℝ]. ∀[L:T List].
  (radd-list(map(λk.(x[k] y[k]);L)) (radd-list(map(λk.x[k];L)) radd-list(map(λk.y[k];L))))

Lemma: radd-list-linearity2
[T:Type]. ∀[x:T ⟶ ℝ]. ∀[a:ℝ]. ∀[L:T List].  (radd-list(map(λk.(a x[k]);L)) (a radd-list(map(λk.x[k];L))))

Lemma: radd-list-linearity3
[T:Type]. ∀[x:T ⟶ ℝ]. ∀[a:ℝ]. ∀[L:T List].  (radd-list(map(λk.(x[k] a);L)) (radd-list(map(λk.x[k];L)) a))

Lemma: radd-list-one
[L:Top List]. (radd-list(map(λk.r1;L)) r(||L||))

Lemma: rnonzero-lemma1
[x:ℝ]. ∀[n:ℕ+].  ∀m:ℕ+((n ≤ m)  (m ≤ (n |x m|))) supposing 5 ≤ |x n|

Definition: reg-seq-inv
reg-seq-inv(x) ==  λm.((4 m) ÷ m)

Lemma: reg-seq-inv_wf
[b:ℕ+]. ∀[x:{f:ℕ+ ⟶ ℤb-regular-seq(f)} ]. ∀[k:ℕ+].
  reg-seq-inv(x) ∈ {f:ℕ+ ⟶ ℤ((k k) 1)-regular-seq(f)}  supposing ∀m:ℕ+((2 m) ≤ (k |x m|))

Definition: reg-seq-adjust
reg-seq-adjust(n;x) ==  λi.if (i) < (n)  then 2  else (x i)

Lemma: reg-seq-adjust_wf
[n:ℕ+]. ∀[x:ℝ].  reg-seq-adjust(n;x) ∈ {f:ℕ+ ⟶ ℤif (n =z 1) then else fi -regular-seq(f)}  supposing ∀i:ℕ+(i <\000C  (|x i| ≤ 4))

Lemma: reg-seq-adjust-property
[n:ℕ+]. ∀[x:ℝ].  ∀m:ℕ+(m ≤ (n |reg-seq-adjust(n;x) m|)) supposing 5 ≤ |x n|

Definition: rnonzero
rnonzero(x) ==  ∃n:ℕ+4 < |x n|

Lemma: rnonzero_wf
[x:ℕ+ ⟶ ℤ]. (rnonzero(x) ∈ ℙ)

Lemma: rnonzero_functionality
x,y:ℝ.  ((x y)  (rnonzero(x) ⇐⇒ rnonzero(y)))

Definition: rinv
rinv(x) ==
  eval mu-ge(λn.eval in
                    4 <|k|;1) in
  if (n =z 1)
  then accelerate(5;reg-seq-inv(x))
  else accelerate(4 ((4 n) 1);reg-seq-inv(reg-seq-adjust(n;x)))
  fi 

Lemma: rinv_wf
[x:ℝ]. (rnonzero(x)  (rinv(x) ∈ ℝ))

Lemma: rinv-functionality-lemma
x,y:ℤ. ∀a,b,n:ℕ+.
  ((n ≤ (a |x|))  (n ≤ (b |y|))  (|x y| ≤ 4)  (|((4 n) ÷ x) (4 n) ÷ y| ≤ (2 (16 b))))

Lemma: rinv_functionality
[x,y:ℝ].  (rnonzero(x)  (x y)  (rinv(x) rinv(y)))

Lemma: rinv_functionality-tst
[x,y:ℝ].  (rnonzero(x)  (x y)  (rinv(x) rinv(y)))

Lemma: rmul-rinv1
[x:ℝ]. (rnonzero(x)  ((x rinv(x)) r1))

Lemma: rinv1
rinv(r1) r1

Definition: rsub
==  -(y)

Lemma: rsub_wf
[x,y:ℝ].  (x y ∈ ℝ)

Lemma: rsub-int
[n:ℤ]. ∀m:ℤ((r(n) r(m)) r(n m))

Lemma: rsub_functionality
[x1,x2,y1,y2:ℝ].  ((x1 y1) (x2 y2)) supposing ((y1 y2) and (x1 x2))

Lemma: rmul-rsub-distrib
[a,b,c:ℝ].  (((a (b c)) ((a b) c)) ∧ (((b c) a) ((b a) a)))

Lemma: rsub-approx
[a,b:ℝ]. ∀[n:ℕ+].  ((a b) ((a (4 n)) (4 n)) ÷ 4)

Lemma: radd-preserves-req
[x,y,z:ℝ].  uiff(x y;(z x) (z y))

Definition: real_term_value
real_term_value(f;t) ==
  int_term_ind(t;
               itermConstant(n) r(n);
               itermVar(v) v;
               itermAdd(l,r) rl,rr.rl rr;
               itermSubtract(l,r) rl,rr.rl rr;
               itermMultiply(l,r) rl,rr.rl rr;
               itermMinus(x) r.-(r)) 

Lemma: real_term_value_wf
[f:ℤ ⟶ ℝ]. ∀[t:int_term()].  (real_term_value(f;t) ∈ ℝ)

Lemma: real_term_value_add_lemma
b,a,f:Top.  (real_term_value(f;a (+) b) real_term_value(f;a) real_term_value(f;b))

Lemma: real_term_value_sub_lemma
b,a,f:Top.  (real_term_value(f;a (-) b) real_term_value(f;a) real_term_value(f;b))

Lemma: real_term_value_minus_lemma
a,f:Top.  (real_term_value(f;"-"a) -(real_term_value(f;a)))

Lemma: real_term_value_mul_lemma
b,a,f:Top.  (real_term_value(f;a (*) b) real_term_value(f;a) real_term_value(f;b))

Lemma: real_term_value_const_lemma
c,f:Top.  (real_term_value(f;"c") r(c))

Lemma: real_term_value_var_lemma
c,f:Top.  (real_term_value(f;vc) c)

Definition: req_int_terms
t1 ≡ t2 ==  ∀f:ℤ ⟶ ℝ(real_term_value(f;t1) real_term_value(f;t2))

Lemma: req_int_terms_wf
[t1,t2:int_term()].  (t1 ≡ t2 ∈ ℙ)

Lemma: req_int_terms_functionality
[x1,x2,y1,y2:int_term()].  (uiff(x1 ≡ y1;x2 ≡ y2)) supposing (y1 ≡ y2 and x1 ≡ x2)

Lemma: req_int_terms_weakening
[t1,t2:int_term()].  t1 ≡ t2 supposing t1 t2 ∈ int_term()

Lemma: req_int_terms_transitivity
[t1,t2,t3:int_term()].  (t1 ≡ t3) supposing (t2 ≡ t3 and t1 ≡ t2)

Lemma: req_int_terms_inversion
[t1,t2:int_term()].  t1 ≡ t2 supposing t2 ≡ t1

Lemma: itermAdd_functionality_wrt_req
[a,b,c,d:int_term()].  (a (+) c ≡ (+) d) supposing (a ≡ and c ≡ d)

Lemma: itermSubtract_functionality_wrt_req
[a,b,c,d:int_term()].  (a (-) c ≡ (-) d) supposing (a ≡ and c ≡ d)

Lemma: itermMultiply_functionality_wrt_req
[a,b,c,d:int_term()].  (a (*) c ≡ (*) d) supposing (a ≡ and c ≡ d)

Lemma: itermMinus_functionality_wrt_req
[a,b:int_term()].  "-"a ≡ "-"b supposing a ≡ b

Lemma: imonomial-req-lemma
f:ℤ ⟶ ℝ. ∀ws:ℤ List. ∀t:int_term().
  (real_term_value(f;accumulate (with value and list item v):
                      (*) vv
                     over list:
                       ws
                     with starting value:
                      t))
  (real_term_value(f;t)
    real_term_value(f;accumulate (with value and list item v):
                         (*) vv
                        over list:
                          ws
                        with starting value:
                         "1"))))

Lemma: imonomial-term-linear-req
f:ℤ ⟶ ℝ. ∀ws:ℤ List. ∀c:ℤ.  (real_term_value(f;imonomial-term(<c, ws>)) (r(c) real_term_value(f;imonomial-term(<1,\000C ws>))))

Lemma: imonomial-term-add-req
vs:ℤ List. ∀a,b:ℤ-o. ∀f:ℤ ⟶ ℝ.  ((real_term_value(f;imonomial-term(<a, vs>)) real_term_value(f;imonomial-term(<b, vs\000C>))) real_term_value(f;imonomial-term(<b, vs>)))

Lemma: ipolynomial-term-cons-req
[m:iMonomial()]. ∀[p:iMonomial() List].  ipolynomial-term([m p]) ≡ imonomial-term(m) (+) ipolynomial-term(p)

Lemma: add-ipoly-req
p,q:iMonomial() List.  ipolynomial-term(add-ipoly(p;q)) ≡ ipolynomial-term(p) (+) ipolynomial-term(q)

Lemma: minus-poly-req
p:iPolynomial(). ipolynomial-term(minus-poly(p)) ≡ "-"ipolynomial-term(p)

Lemma: imonomial-cons-req
v:ℤ List. ∀u,a:ℤ. ∀f:ℤ ⟶ ℝ.  (real_term_value(f;imonomial-term(<a, [u v]>)) ((f u) real_term_value(f;imonomial-t\000Cerm(<a, v>))))

Lemma: mul-monomials-req
[m1,m2:iMonomial()].  imonomial-term(mul-monomials(m1;m2)) ≡ imonomial-term(m1) (*) imonomial-term(m2)

Lemma: mul-mono-poly-req
[m:iMonomial()]. ∀[p:iMonomial() List].
  ipolynomial-term(mul-mono-poly(m;p)) ≡ imonomial-term(m) (*) ipolynomial-term(p)

Lemma: mul-ipoly-req
p,q:iMonomial() List.  ipolynomial-term(mul-ipoly(p;q)) ≡ ipolynomial-term(p) (*) ipolynomial-term(q)

Lemma: real_term_polynomial
t:int_term(). ipolynomial-term(int_term_to_ipoly(t)) ≡ t

Lemma: real_polynomial_null
t:int_term(). t ≡ "0" supposing inl Ax ≤ null(int_term_to_ipoly(t))

Lemma: real_polynomial_null_orig
t:int_term(). t ≡ "0" supposing null(int_term_to_ipoly(t)) tt

Lemma: req-iff-rsub-is-0
[a,b:ℝ].  uiff(a b;(a b) r0)

Lemma: test-ring-req
v,v1,v2,v3,v4,v5,v6,v7,v8,v9:ℝ.
  ((((((v6 v3) (v2 v5) (v4 v7)) (v7 v2) (v3 v4) (v5 v6))
    (((v6 v9) (v8 v1) (v v7)) (v7 v8) (v9 v) (v1 v6)))
  ((((v v7) (v6 v5) (v4 v1)) (v1 v6) (v7 v4) (v5 v))
    (((v6 v9) (v8 v3) (v2 v7)) (v7 v8) (v9 v2) (v3 v6)))
  ((((v v3) (v2 v7) (v6 v1)) (v1 v2) (v3 v6) (v7 v))
    (((v6 v9) (v8 v5) (v4 v7)) (v7 v8) (v9 v4) (v5 v6))))
  r0)

Definition: rat_termco
rat_termco() ==
  corec(X.lbl:Atom × if lbl =a "Constant" then ℤ
                     if lbl =a "Var" then ℤ
                     if lbl =a "Add" then left:X × X
                     if lbl =a "Subtract" then left:X × X
                     if lbl =a "Multiply" then left:X × X
                     if lbl =a "Divide" then num:X × X
                     if lbl =a "Minus" then X
                     else Void
                     fi )

Lemma: rat_termco_wf
rat_termco() ∈ Type

Lemma: rat_termco-ext
rat_termco() ≡ lbl:Atom × if lbl =a "Constant" then ℤ
                          if lbl =a "Var" then ℤ
                          if lbl =a "Add" then left:rat_termco() × rat_termco()
                          if lbl =a "Subtract" then left:rat_termco() × rat_termco()
                          if lbl =a "Multiply" then left:rat_termco() × rat_termco()
                          if lbl =a "Divide" then num:rat_termco() × rat_termco()
                          if lbl =a "Minus" then rat_termco()
                          else Void
                          fi 

Definition: rat_termco_size
rat_termco_size(p) ==
  fix((λsize,p. let lbl,x 
                in if lbl =a "Constant" then 0
                   if lbl =a "Var" then 0
                   if lbl =a "Add" then (size (fst(x))) (size (snd(x)))
                   if lbl =a "Subtract" then (size (fst(x))) (size (snd(x)))
                   if lbl =a "Multiply" then (size (fst(x))) (size (snd(x)))
                   if lbl =a "Divide" then (size (fst(x))) (size (snd(x)))
                   if lbl =a "Minus" then (size x)
                   else 0
                   fi )) 
  p

Lemma: rat_termco_size_wf
[p:rat_termco()]. (rat_termco_size(p) ∈ partial(ℕ))

Definition: rat_term
rat_term() ==  {p:rat_termco()| (rat_termco_size(p))↓

Lemma: rat_term_wf
rat_term() ∈ Type

Definition: rat_term_size
rat_term_size(p) ==
  fix((λsize,p. let lbl,x 
                in if lbl =a "Constant" then 0
                   if lbl =a "Var" then 0
                   if lbl =a "Add" then (size (fst(x))) (size (snd(x)))
                   if lbl =a "Subtract" then (size (fst(x))) (size (snd(x)))
                   if lbl =a "Multiply" then (size (fst(x))) (size (snd(x)))
                   if lbl =a "Divide" then (size (fst(x))) (size (snd(x)))
                   if lbl =a "Minus" then (size x)
                   else 0
                   fi )) 
  p

Lemma: rat_term_size_wf
[p:rat_term()]. (rat_term_size(p) ∈ ℕ)

Lemma: rat_term-ext
rat_term() ≡ lbl:Atom × if lbl =a "Constant" then ℤ
                        if lbl =a "Var" then ℤ
                        if lbl =a "Add" then left:rat_term() × rat_term()
                        if lbl =a "Subtract" then left:rat_term() × rat_term()
                        if lbl =a "Multiply" then left:rat_term() × rat_term()
                        if lbl =a "Divide" then num:rat_term() × rat_term()
                        if lbl =a "Minus" then rat_term()
                        else Void
                        fi 

Definition: rtermConstant
"const" ==  <"Constant", const>

Lemma: rtermConstant_wf
[const:ℤ]. ("const" ∈ rat_term())

Definition: rtermVar
rtermVar(var) ==  <"Var", var>

Lemma: rtermVar_wf
[var:ℤ]. (rtermVar(var) ∈ rat_term())

Definition: rtermAdd
left "+" right ==  <"Add", left, right>

Lemma: rtermAdd_wf
[left,right:rat_term()].  (left "+" right ∈ rat_term())

Definition: rtermSubtract
left "-" right ==  <"Subtract", left, right>

Lemma: rtermSubtract_wf
[left,right:rat_term()].  (left "-" right ∈ rat_term())

Definition: rtermMultiply
left "*" right ==  <"Multiply", left, right>

Lemma: rtermMultiply_wf
[left,right:rat_term()].  (left "*" right ∈ rat_term())

Definition: rtermDivide
num "/" denom ==  <"Divide", num, denom>

Lemma: rtermDivide_wf
[num,denom:rat_term()].  (num "/" denom ∈ rat_term())

Definition: rtermMinus
rtermMinus(num) ==  <"Minus", num>

Lemma: rtermMinus_wf
[num:rat_term()]. (rtermMinus(num) ∈ rat_term())

Definition: rtermConstant?
rtermConstant?(v) ==  fst(v) =a "Constant"

Lemma: rtermConstant?_wf
[v:rat_term()]. (rtermConstant?(v) ∈ 𝔹)

Definition: rtermConstant-const
rtermConstant-const(v) ==  snd(v)

Lemma: rtermConstant-const_wf
[v:rat_term()]. rtermConstant-const(v) ∈ ℤ supposing ↑rtermConstant?(v)

Definition: rtermVar?
rtermVar?(v) ==  fst(v) =a "Var"

Lemma: rtermVar?_wf
[v:rat_term()]. (rtermVar?(v) ∈ 𝔹)

Definition: rtermVar-var
rtermVar-var(v) ==  snd(v)

Lemma: rtermVar-var_wf
[v:rat_term()]. rtermVar-var(v) ∈ ℤ supposing ↑rtermVar?(v)

Definition: rtermAdd?
rtermAdd?(v) ==  fst(v) =a "Add"

Lemma: rtermAdd?_wf
[v:rat_term()]. (rtermAdd?(v) ∈ 𝔹)

Definition: rtermAdd-left
rtermAdd-left(v) ==  fst(snd(v))

Lemma: rtermAdd-left_wf
[v:rat_term()]. rtermAdd-left(v) ∈ rat_term() supposing ↑rtermAdd?(v)

Definition: rtermAdd-right
rtermAdd-right(v) ==  snd(snd(v))

Lemma: rtermAdd-right_wf
[v:rat_term()]. rtermAdd-right(v) ∈ rat_term() supposing ↑rtermAdd?(v)

Definition: rtermSubtract?
rtermSubtract?(v) ==  fst(v) =a "Subtract"

Lemma: rtermSubtract?_wf
[v:rat_term()]. (rtermSubtract?(v) ∈ 𝔹)

Definition: rtermSubtract-left
rtermSubtract-left(v) ==  fst(snd(v))

Lemma: rtermSubtract-left_wf
[v:rat_term()]. rtermSubtract-left(v) ∈ rat_term() supposing ↑rtermSubtract?(v)

Definition: rtermSubtract-right
rtermSubtract-right(v) ==  snd(snd(v))

Lemma: rtermSubtract-right_wf
[v:rat_term()]. rtermSubtract-right(v) ∈ rat_term() supposing ↑rtermSubtract?(v)

Definition: rtermMultiply?
rtermMultiply?(v) ==  fst(v) =a "Multiply"

Lemma: rtermMultiply?_wf
[v:rat_term()]. (rtermMultiply?(v) ∈ 𝔹)

Definition: rtermMultiply-left
rtermMultiply-left(v) ==  fst(snd(v))

Lemma: rtermMultiply-left_wf
[v:rat_term()]. rtermMultiply-left(v) ∈ rat_term() supposing ↑rtermMultiply?(v)

Definition: rtermMultiply-right
rtermMultiply-right(v) ==  snd(snd(v))

Lemma: rtermMultiply-right_wf
[v:rat_term()]. rtermMultiply-right(v) ∈ rat_term() supposing ↑rtermMultiply?(v)

Definition: rtermDivide?
rtermDivide?(v) ==  fst(v) =a "Divide"

Lemma: rtermDivide?_wf
[v:rat_term()]. (rtermDivide?(v) ∈ 𝔹)

Definition: rtermDivide-num
rtermDivide-num(v) ==  fst(snd(v))

Lemma: rtermDivide-num_wf
[v:rat_term()]. rtermDivide-num(v) ∈ rat_term() supposing ↑rtermDivide?(v)

Definition: rtermDivide-denom
rtermDivide-denom(v) ==  snd(snd(v))

Lemma: rtermDivide-denom_wf
[v:rat_term()]. rtermDivide-denom(v) ∈ rat_term() supposing ↑rtermDivide?(v)

Definition: rtermMinus?
rtermMinus?(v) ==  fst(v) =a "Minus"

Lemma: rtermMinus?_wf
[v:rat_term()]. (rtermMinus?(v) ∈ 𝔹)

Definition: rtermMinus-num
rtermMinus-num(v) ==  snd(v)

Lemma: rtermMinus-num_wf
[v:rat_term()]. rtermMinus-num(v) ∈ rat_term() supposing ↑rtermMinus?(v)

Lemma: rat_term-induction
[P:rat_term() ⟶ ℙ]
  ((∀const:ℤP["const"])
   (∀var:ℤP[rtermVar(var)])
   (∀left,right:rat_term().  (P[left]  P[right]  P[left "+" right]))
   (∀left,right:rat_term().  (P[left]  P[right]  P[left "-" right]))
   (∀left,right:rat_term().  (P[left]  P[right]  P[left "*" right]))
   (∀num,denom:rat_term().  (P[num]  P[denom]  P[num "/" denom]))
   (∀num:rat_term(). (P[num]  P[rtermMinus(num)]))
   {∀v:rat_term(). P[v]})

Lemma: rat_term-definition
[A:Type]. ∀[R:A ⟶ rat_term() ⟶ ℙ].
  ((∀const:ℤ{x:A| R[x;"const"]} )
   (∀var:ℤ{x:A| R[x;rtermVar(var)]} )
   (∀left,right:rat_term().  ({x:A| R[x;left]}   {x:A| R[x;right]}   {x:A| R[x;left "+" right]} ))
   (∀left,right:rat_term().  ({x:A| R[x;left]}   {x:A| R[x;right]}   {x:A| R[x;left "-" right]} ))
   (∀left,right:rat_term().  ({x:A| R[x;left]}   {x:A| R[x;right]}   {x:A| R[x;left "*" right]} ))
   (∀num,denom:rat_term().  ({x:A| R[x;num]}   {x:A| R[x;denom]}   {x:A| R[x;num "/" denom]} ))
   (∀num:rat_term(). ({x:A| R[x;num]}   {x:A| R[x;rtermMinus(num)]} ))
   {∀v:rat_term(). {x:A| R[x;v]} })

Definition: rat_term_ind
rat_term_ind(v;
             rtermConstant(const) Constant[const];
             rtermVar(var) Var[var];
             rtermAdd(left,right) rec1,rec2.Add[left;
                                                  right;
                                                  rec1;
                                                  rec2];
             rtermSubtract(left,right) rec3,rec4.Subtract[left;
                                                            right;
                                                            rec3;
                                                            rec4];
             rtermMultiply(left,right) rec5,rec6.Multiply[left;
                                                            right;
                                                            rec5;
                                                            rec6];
             rtermDivide(num,denom) rec7,rec8.Divide[num;
                                                       denom;
                                                       rec7;
                                                       rec8];
             rtermMinus(num) rec9.Minus[num; rec9])  ==
  fix((λrat_term_ind,v. let lbl,v1 
                        in if lbl="Constant" then Constant[v1]
                           if lbl="Var" then Var[v1]
                           if lbl="Add"
                             then let left,v2 v1 
                                  in Add[left;
                                         v2;
                                         rat_term_ind left;
                                         rat_term_ind v2]
                           if lbl="Subtract"
                             then let left,v2 v1 
                                  in Subtract[left;
                                              v2;
                                              rat_term_ind left;
                                              rat_term_ind v2]
                           if lbl="Multiply"
                             then let left,v2 v1 
                                  in Multiply[left;
                                              v2;
                                              rat_term_ind left;
                                              rat_term_ind v2]
                           if lbl="Divide"
                             then let num,v2 v1 
                                  in Divide[num;
                                            v2;
                                            rat_term_ind num;
                                            rat_term_ind v2]
                           if lbl="Minus" then Minus[v1; rat_term_ind v1]
                           else Ax
                           fi )) 
  v

Lemma: rat_term_ind_wf
[A:Type]. ∀[R:A ⟶ rat_term() ⟶ ℙ]. ∀[v:rat_term()]. ∀[Constant:const:ℤ ⟶ {x:A| R[x;"const"]} ].
[Var:var:ℤ ⟶ {x:A| R[x;rtermVar(var)]} ]. ∀[Add:left:rat_term()
                                                 ⟶ right:rat_term()
                                                 ⟶ {x:A| R[x;left]} 
                                                 ⟶ {x:A| R[x;right]} 
                                                 ⟶ {x:A| R[x;left "+" right]} ]. ∀[Subtract:left:rat_term()
                                                                                            ⟶ right:rat_term()
                                                                                            ⟶ {x:A| R[x;left]} 
                                                                                            ⟶ {x:A| R[x;right]} 
                                                                                            ⟶ {x:A| 
                                                                                                R[x;left "-" right]} ].
[Multiply:left:rat_term() ⟶ right:rat_term() ⟶ {x:A| R[x;left]}  ⟶ {x:A| R[x;right]}  ⟶ {x:A| R[x;left "*" right]} \000C].
[Divide:num:rat_term() ⟶ denom:rat_term() ⟶ {x:A| R[x;num]}  ⟶ {x:A| R[x;denom]}  ⟶ {x:A| R[x;num "/" denom]} ].
[Minus:num:rat_term() ⟶ {x:A| R[x;num]}  ⟶ {x:A| R[x;rtermMinus(num)]} ].
  (rat_term_ind(v;
                rtermConstant(const) Constant[const];
                rtermVar(var) Var[var];
                rtermAdd(left,right) rec1,rec2.Add[left;right;rec1;rec2];
                rtermSubtract(left,right) rec3,rec4.Subtract[left;right;rec3;rec4];
                rtermMultiply(left,right) rec5,rec6.Multiply[left;right;rec5;rec6];
                rtermDivide(num,denom) rec7,rec8.Divide[num;denom;rec7;rec8];
                rtermMinus(num) rec9.Minus[num;rec9])  ∈ {x:A| R[x;v]} )

Lemma: rat_term_ind_wf_simple
[A:Type]. ∀[v:rat_term()]. ∀[Constant,Var:var:ℤ ⟶ A]. ∀[Add,Subtract,Multiply,Divide:num:rat_term()
                                                                                       ⟶ denom:rat_term()
                                                                                       ⟶ A
                                                                                       ⟶ A
                                                                                       ⟶ A]. ∀[Minus:num:rat_term()
                                                                                                      ⟶ A
                                                                                                      ⟶ A].
  (rat_term_ind(v;
                rtermConstant(const) Constant[const];
                rtermVar(var) Var[var];
                rtermAdd(left,right) rec1,rec2.Add[left;right;rec1;rec2];
                rtermSubtract(left,right) rec3,rec4.Subtract[left;right;rec3;rec4];
                rtermMultiply(left,right) rec5,rec6.Multiply[left;right;rec5;rec6];
                rtermDivide(num,denom) rec7,rec8.Divide[num;denom;rec7;rec8];
                rtermMinus(num) rec9.Minus[num;rec9])  ∈ A)

Definition: rat_term_to_ipolys
rat_term_to_ipolys(t) ==
  rat_term_ind(t;
               rtermConstant(c) <if c=0 then [] else [<c, []>], [<1, []>]>;
               rtermVar(v) <[<1, [v]>], [<1, []>]>;
               rtermAdd(left,right) rl,rr.let p1,q1 rl 
                                            in let p2,q2 rr 
                                               in <add_ipoly(mul_ipoly(p1;q2);mul_ipoly(p2;q1)), mul_ipoly(q1;q2)>;
               rtermSubtract(left,right) rl,rr.let p1,q1 rl 
                                                 in let p2,q2 rr 
                                                    in <add_ipoly(mul_ipoly(p1;q2);mul_ipoly(minus-poly(p2);q1))
                                                       mul_ipoly(q1;q2)
                                                       >;
               rtermMultiply(left,right) rl,rr.let p1,q1 rl 
                                                 in let p2,q2 rr 
                                                    in <mul_ipoly(p1;p2), mul_ipoly(q1;q2)>;
               rtermDivide(num,denom) rl,rr.let p1,q1 rl 
                                              in let p2,q2 rr 
                                                 in <mul_ipoly(p1;q2), mul_ipoly(q1;p2)>;
               rtermMinus(num) rx.let p1,q1 rx 
                                    in <minus-poly(p1), q1>

Lemma: rat_term_to_ipolys_wf
[t:rat_term()]. (rat_term_to_ipolys(t) ∈ iPolynomial() × iPolynomial())

Definition: int-radd
==  λn.(((2 k) n) (x n))

Lemma: int-radd_wf
[k:ℤ]. ∀[x:ℝ].  (k x ∈ ℝ)

Lemma: int-radd-req
[k:ℤ]. ∀[x:ℝ].  (k (r(k) x))

Definition: int-rsub
==  λn.(((2 k) n) n)

Lemma: int-rsub_wf
[k:ℤ]. ∀[x:ℝ].  (k x ∈ ℝ)

Lemma: int-rsub-req
[k:ℤ]. ∀[x:ℝ].  (k (r(k) x))

Definition: int-rmul
k1 ==  eval k1 in λn.if (k) < (0)  then -(a ((-k) n))  else if (0) < (k)  then (k n)  else 0

Lemma: int-rmul_wf
[k:ℤ]. ∀[a:ℝ].  (k a ∈ ℝ)

Lemma: int-rmul-req
[k:ℤ]. ∀[a:ℝ].  (k (r(k) a))

Lemma: int-rmul_functionality
[k1,k2:ℤ]. ∀[a,b:ℝ].  (k1 k2 b) supposing ((k1 k2 ∈ ℤand (a b))

Lemma: int-rmul-one
[a:ℝ]. (1 a)

Definition: int-rdiv

(a)/k1 ==  eval k1 in λn.((a n) ÷ k)

Lemma: int-rdiv_wf
[k:ℤ-o]. ∀[a:ℝ].  ((a)/k ∈ ℝ)

Lemma: int-rdiv-int-rmul
[k:ℤ-o]. ∀[a:ℝ].  (k (a)/k a)

Lemma: int-rdiv-one
[a:ℝ]. ((a)/1 a)

Definition: alt-int-rdiv
alt-int-rdiv(x;k) ==  if k=1 then else n.[x n ÷ k])

Lemma: alt-int-rdiv_wf
[x:ℝ]. ∀[k:ℕ+].  (alt-int-rdiv(x;k) ∈ {y:ℝx} )

Definition: rat-to-real
r(a/b) ==  (r(a))/b

Lemma: rat-to-real_wf
[a:ℤ]. ∀[b:ℤ-o].  (r(a/b) ∈ ℝ)

Definition: reg-seq-nexp
reg-seq-nexp(x;k) ==  λn.(x n^k ÷ n^k 1)

Lemma: reg-seq-nexp_wf
[x:ℝ]. ∀[k:ℕ+].  (reg-seq-nexp(x;k) ∈ {f:ℕ+ ⟶ ℤ(k ((canon-bnd(x)^k 1 ÷ 2^k 1) 1)) 1-regular-seq(f)} )

Definition: rnexp
x^k1 ==
  eval k1 in
  if (k =z 0)
  then r1
  else eval canon-bnd(x) in
       accelerate((k ((B^k 1 ÷ 2^k 1) 1)) 1;λn.(x n^k ÷ n^k 1))
  fi 

Lemma: rnexp_wf
[k:ℕ]. ∀[x:ℝ].  (x^k ∈ ℝ)

Lemma: rnexp-req
[k:ℕ]. ∀[x:ℝ].  (x^k if (k =z 0) then r1 else x^k fi )

Lemma: rnexp_zero_lemma
x:Top. (x^0 r1)

Lemma: rnexp-int
[k:ℕ]. ∀[z:ℤ].  (r(z)^k r(z^k))

Lemma: rnexp0
[k:ℕ+]. (r0^k r0)

Definition: rmax
rmax(x;y) ==  λn.imax(x n;y n)

Lemma: rmax_wf
[x,y:ℝ].  (rmax(x;y) ∈ ℝ)

Lemma: rmax_functionality_wrt_bdd-diff
[x1,x2,y1,y2:ℕ+ ⟶ ℤ].  (bdd-diff(y1;y2)  bdd-diff(x1;x2)  bdd-diff(rmax(x1;y1);rmax(x2;y2)))

Lemma: rmax_functionality
[x1,x2,y1,y2:ℝ].  (rmax(x1;y1) rmax(x2;y2)) supposing ((x1 x2) and (y1 y2))

Lemma: rmax-com
[x,y:ℝ].  (rmax(x;y) rmax(y;x))

Lemma: rmax-assoc
[x,y,z:ℝ].  (rmax(rmax(x;y);z) rmax(x;rmax(y;z)))

Definition: rmin
rmin(x;y) ==  λn.imin(x n;y n)

Lemma: rmin_wf
[x,y:ℝ].  (rmin(x;y) ∈ ℝ)

Lemma: rmin_functionality_wrt_bdd-diff
[x1,x2,y1,y2:ℕ+ ⟶ ℤ].  (bdd-diff(y1;y2)  bdd-diff(x1;x2)  bdd-diff(rmin(x1;y1);rmin(x2;y2)))

Lemma: rmin_functionality
[x1,x2,y1,y2:ℝ].  (rmin(x1;y1) rmin(x2;y2)) supposing ((x1 x2) and (y1 y2))

Lemma: rmin-com
[x,y:ℝ].  (rmin(x;y) rmin(y;x))

Lemma: rmin-assoc
[x,y,z:ℝ].  (rmin(rmin(x;y);z) rmin(x;rmin(y;z)))

Lemma: rminus-rmax
[x,y:ℝ].  (-(rmax(x;y)) rmin(-(x);-(y)))

Lemma: rminus-rmin
[x,y:ℝ].  (-(rmin(x;y)) rmax(-(x);-(y)))

Lemma: rmin-req-rminus-rmax
[x,y:ℝ].  (rmin(x;y) -(rmax(-(x);-(y))))

Lemma: rmin-rmax-distrib-strong
a,b,c:ℝ.  (rmin(a;rmax(b;c)) rmax(rmin(a;b);rmin(a;c)) ∈ ℝ)

Lemma: rmin-rmax-distrib
a,b,c:ℝ.  (rmin(a;rmax(b;c)) rmax(rmin(a;b);rmin(a;c)))

Lemma: rmin-rmax-absorption-strong
b,a:ℝ.  (rmin(b;rmax(b;a)) b ∈ ℝ)

Lemma: rmax-rmin-absorption-strong
b,a:ℝ.  (rmax(b;rmin(b;a)) b ∈ ℝ)

Lemma: rmin-rmax-absorption
b,a:ℝ.  (rmin(b;rmax(b;a)) b)

Lemma: rmax-rmin-absorption
b,a:ℝ.  (rmax(b;rmin(b;a)) b)

Definition: rabs
|x| ==  λn.|x n|

Lemma: rabs-as-rmax
[x:Top]. (|x| rmax(x;-(x)))

Lemma: rabs_wf
[x:ℝ]. (|x| ∈ ℝ)

Lemma: rabs-approx
[x,n:Top].  (|x| |x n|)

Lemma: rabs_functionality
[x,y:ℝ].  |x| |y| supposing y

Lemma: rabs_functionality_wrt_bdd-diff
x,y:ℕ+ ⟶ ℤ.  (bdd-diff(x;y)  bdd-diff(|x|;|y|))

Lemma: rabs-rminus
[x:ℝ]. (|-(x)| |x| ∈ ℝ)

Lemma: rabs-difference-symmetry
[x,y:ℝ].  (|x y| |y x|)

Lemma: rabs-rmul
[x,y:ℝ].  (|x y| (|x| |y|))

Lemma: rmin-rmax-real-decomp
[r:ℝ]. ((rmin(|r|;rmax(r0;r)) rmax(-(|r|);rmin(r0;r))) r)

Definition: rpositive
rpositive(x) ==  ∃n:ℕ+ [4 < n]

Lemma: rpositive_wf
[x:ℕ+ ⟶ ℤ]. (rpositive(x) ∈ ℙ)

Definition: rpositive2
rpositive2(x) ==  ∃n:ℕ+. ∀m:ℕ+((n ≤ m)  (m ≤ (n (x m))))

Lemma: rpositive2_wf
[x:ℕ+ ⟶ ℤ]. (rpositive2(x) ∈ ℙ)

Lemma: rpositive2_functionality
x,y:ℕ+ ⟶ ℤ.  (bdd-diff(x;y)  (rpositive2(x) ⇐⇒ rpositive2(y)))

Lemma: rpositive-iff
[x:ℝ]. (rpositive(x) ⇐⇒ rpositive2(x))

Lemma: rpositive_functionality
x,y:ℝ.  rpositive(x) ⇐⇒ rpositive(y) supposing y

Lemma: rpositive-radd
x,y:ℝ.  (rpositive(x)  rpositive(y)  rpositive(x y))

Lemma: rpositive-rmax
x,y:ℝ.  (rpositive(rmax(x;y)) ⇐⇒ rpositive(x) ∨ rpositive(y))

Lemma: rpositive-rmul
x,y:ℝ.  (rpositive(x)  rpositive(y)  rpositive(x y))

Definition: rnonneg
rnonneg(x) ==  ∀n:ℕ+((-2) ≤ (x n))

Lemma: rnonneg_wf
[x:ℕ+ ⟶ ℤ]. (rnonneg(x) ∈ ℙ)

Definition: rnonneg2
rnonneg2(x) ==  ∀n:ℕ+. ∃N:ℕ+. ∀m:{N...}. (((-2) m) ≤ (n (x m)))

Lemma: rnonneg2_wf
[x:ℕ+ ⟶ ℤ]. (rnonneg2(x) ∈ ℙ)

Lemma: rnonneg2_functionality
x,y:ℕ+ ⟶ ℤ.  (bdd-diff(x;y)  (rnonneg2(x) ⇐⇒ rnonneg2(y)))

Lemma: rnonneg-iff
[x:ℝ]. (rnonneg(x) ⇐⇒ rnonneg2(x))

Lemma: rnonneg_functionality
x,y:ℝ.  rnonneg(x) ⇐⇒ rnonneg(y) supposing y

Lemma: rpositive-implies-rnonneg
[x:ℝ]. (rpositive(x)  rnonneg(x))

Lemma: rnonneg-radd
x,y:ℝ.  (rnonneg(x)  rnonneg(y)  rnonneg(x y))

Lemma: rnonneg-int
[n:ℤ]. uiff(rnonneg(r(n));0 ≤ n)

Lemma: rpositive-int
[n:ℤ]. uiff(rpositive(r(n));0 < n)

Lemma: rpositive-radd2
x,y:ℝ.  (rpositive(x)  rnonneg(y)  rpositive(x y))

Lemma: rnonneg-rmul
x,y:ℝ.  (rnonneg(x)  rnonneg(y)  rnonneg(x y))

Lemma: sq_stable__rnonneg
[r:ℝ]. SqStable(rnonneg(r))

Lemma: rminus-nonneg
[x:ℝ]. (x r0) supposing (rnonneg(-(x)) and rnonneg(x))

Lemma: rabs-nonneg
[x:ℝ]. rnonneg(|x|)

Lemma: rmax-positive
x,y:ℝ.  ((rpositive(x) ∨ rpositive(y))  rpositive(rmax(x;y)))

Lemma: rmax-nonneg
[x,y:ℝ].  rnonneg(rmax(x;y)) supposing rnonneg(x) ∨ rnonneg(y)

Lemma: rmin-positive
x,y:ℝ.  ((rpositive(x) ∧ rpositive(y))  rpositive(rmin(x;y)))

Lemma: rmin-nonneg
[x,y:ℝ].  rnonneg(rmin(x;y)) supposing rnonneg(x) ∧ rnonneg(y)

Definition: rless
x < ==  ∃n:ℕ+ [(x n) 4 < n]

Lemma: rless_wf
[x,y:ℝ].  (x < y ∈ ℙ)

Lemma: rless-by-computation1
[x,y:ℝ].  x < supposing (x 1000) 4 < 1000

Lemma: rpositive-rless
[x:ℝ]. (r0 < ⇐⇒ rpositive(x))

Lemma: rless-iff
x,y:ℝ.  (x < ⇐⇒ ∃n:ℕ+. ∀m:ℕ+((n ≤ m)  (m ≤ (n ((y m) m)))))

Lemma: rless-iff-large-diff
x,y:ℝ.  (x < ⇐⇒ ∀b:ℕ+. ∃n:ℕ+. ∀m:ℕ+((n ≤ m)  (b ≤ ((y m) m))))

Lemma: rless-iff-rpositive
x,y:ℝ.  (x < ⇐⇒ rpositive(y x))

Lemma: rless-iff4
x,y:ℝ.  (x < ⇐⇒ ∃n:ℕ+. ∀m:{n...}. (x m) 4 < m)

Lemma: rless-iff8
x,y:ℝ.  (x < ⇐⇒ ∃m:ℕ+ [(x m) 8 < m])

Lemma: rless-iff2
x,y:ℝ.  (x < ⇐⇒ ∃n:ℕ+(x n) 4 < n)

Lemma: rless-iff2-ext
x,y:ℝ.  (x < ⇐⇒ ∃n:ℕ+(x n) 4 < n)

Definition: rlessw
rlessw(x;y) ==  quick-find(λn.(x n) 4 <n;1)

Lemma: rlessw_wf
x:ℝ. ∀y:{y:ℝx < y} .  (rlessw(x;y) ∈ x < y)

Lemma: int-rdiv-rless-witness
k:ℕ+(k ∈ (r1)/k < (r(4))/k)

Lemma: int-rdiv-rless-witness2
k:ℕ+(3 k ∈ (r1)/k < (r(2))/k)

Lemma: rless-content
[x:ℝ]. ∀[y:{y:ℝx < y} ].  (rlessw(x;y) ∈ x < y)

Lemma: sq_stable__rless
x,y:ℝ.  SqStable(x < y)

Lemma: sq_stable__rless-or2
x,y,a,b:ℝ.  SqStable((x < y) ∨ (a < b))

Lemma: sq_stable__rless-or3
x,y,a,b,c,d:ℝ.  SqStable(((x < y) ∨ (a < b)) ∨ (c < d))

Definition: rleq
x ≤ ==  rnonneg(y x)

Lemma: rleq_wf
[x,y:ℝ].  (x ≤ y ∈ ℙ)

Lemma: rleq-implies
[x,y:ℝ].  ∀n:ℕ+((x (4 n)) ≤ ((y (4 n)) 11)) supposing x ≤ y

Lemma: rleq-iff4
[x,y:ℝ].  (x ≤ ⇐⇒ ∀n:ℕ+((x n) ≤ ((y n) 4)))

Lemma: rleq-iff-not-rless
[x,y:ℝ].  uiff(y ≤ x;¬(x < y))

Lemma: rless_complement
x,y:ℝ.  (x < y) ⇐⇒ y ≤ x)

Lemma: canonical-bound-property
x:ℝ((r(-canonical-bound(x)) ≤ x) ∧ (x ≤ r(canonical-bound(x))))

Lemma: rless_functionality
x1,x2,y1,y2:ℝ.  (x1 < y1 ⇐⇒ x2 < y2) supposing ((y1 y2) and (x1 x2))

Lemma: rless-int
n,m:ℤ.  (r(n) < r(m) ⇐⇒ n < m)

Lemma: rleq-int
n,m:ℤ.  (r(n) ≤ r(m) ⇐⇒ n ≤ m)

Lemma: stable__rleq
[x,y:ℝ].  Stable{x ≤ y}

Lemma: sq_stable__rleq
[x,y:ℝ].  SqStable(x ≤ y)

Lemma: rleq_functionality
[x1,x2,y1,y2:ℝ].  (uiff(x1 ≤ y1;x2 ≤ y2)) supposing ((y1 y2) and (x1 x2))

Definition: rgt
x > ==  y < x

Lemma: rgt_wf
[x,y:ℝ].  (x > y ∈ ℙ)

Definition: rge
x ≥ ==  y ≤ x

Lemma: rge_wf
[x,y:ℝ].  (x ≥ y ∈ ℙ)

Lemma: rless_transitivity
x,y,z:ℝ.  ((x < y)  (y < z)  (x < z))

Lemma: rless_transitivity1
x,y,z:ℝ.  ((x < y)  x < supposing y ≤ z)

Lemma: rless_transitivity2
x,y,z:ℝ.  (y < z)  (x < z) supposing x ≤ y

Lemma: rleq_transitivity
[x,y,z:ℝ].  (x ≤ z) supposing ((y ≤ z) and (x ≤ y))

Lemma: rleq_antisymmetry
[x,y:ℝ].  (x y) supposing ((y ≤ x) and (x ≤ y))

Lemma: req_fake_le_antisymmetry
[x,y:ℝ].  (x y) supposing ((y ≤ x) and (x ≤ y))

Lemma: rleq_weakening
[x,y:ℝ].  x ≤ supposing y

Lemma: rleq_weakening_equal
[x,y:ℝ].  x ≤ supposing y ∈ ℝ

Lemma: rleq_functionality_wrt_implies
[a,b,c,d:ℝ].  ({a ≤ supposing b ≤ c}) supposing ((c ≤ d) and (b ≥ a))

Lemma: rless_functionality_wrt_implies
a,b,c,d:ℝ.  ({(b < c)  (a < d)}) supposing ((c ≤ d) and (b ≥ a))

Lemma: rleq_weakening_rless
[x,y:ℝ].  x ≤ supposing x < y

Definition: rbetween
x≤y≤==  (x ≤ y) ∧ (y ≤ z)

Lemma: rbetween_wf
[x,y,z:ℝ].  (x≤y≤z ∈ ℙ)

Lemma: radd_functionality_wrt_rleq
[x,y,z,t:ℝ].  ((x y) ≤ (z t)) supposing ((y ≤ t) and (x ≤ z))

Lemma: radd_functionality_wrt_rless1
x,y,z,t:ℝ.  (y < t)  ((x y) < (z t)) supposing x ≤ z

Lemma: radd_functionality_wrt_rless2
x,y,z,t:ℝ.  ((x < z)  (x y) < (z t) supposing y ≤ t)

Lemma: radd-preserves-rless
x,y,z:ℝ.  (x < ⇐⇒ (z x) < (z y))

Lemma: radd-preserves-rleq
[x,y,z:ℝ].  uiff(x ≤ y;(z x) ≤ (z y))

Lemma: rless-implies-rless
a,b:ℝ.  ∀[c,d:ℝ].  (a < b) supposing ((c < d) and ((d c) (b a)))

Lemma: rleq-implies-rleq
[a,b,c,d:ℝ].  (a ≤ b) supposing ((c ≤ d) and ((d c) (b a)))

Lemma: req-implies-req
[a,b,c,d:ℝ].  (a b) supposing ((c d) and ((d c) (b a)))

Lemma: radd-of-nonneg-is-zero
[a,b:{x:ℝr0 ≤ x} ].  uiff((a b) r0;(a r0) ∧ (b r0))

Lemma: trivial-rless-radd
a,d:ℝ.  (uiff(a < (a d);r0 < d) ∧ uiff(a < (d a);r0 < d))

Lemma: trivial-rsub-rless
a,d:ℝ.  uiff((a d) < a;r0 < d)

Lemma: trivial-rleq-radd
[a,d:ℝ].  (uiff(a ≤ (a d);r0 ≤ d) ∧ uiff(a ≤ (d a);r0 ≤ d))

Lemma: trivial-rsub-rleq
[a,d:ℝ].  uiff((a d) ≤ a;r0 ≤ d)

Lemma: radd-list_functionality_wrt_rleq
[L1,L2:ℝ List].  radd-list(L1) ≤ radd-list(L2) supposing (||L1|| ||L2|| ∈ ℤ) ∧ (∀i:ℕ||L1||. (L1[i] ≤ L2[i]))

Lemma: rmul_functionality_wrt_rleq
[x,y,z:ℝ].  ((x y) ≤ (z y)) supposing ((r0 ≤ y) and (x ≤ z))

Lemma: rmul_functionality_wrt_rleq2
[x,y,z,w:ℝ].  ((x w) ≤ (z y)) supposing ((w ≤ y) and (x ≤ z) and (((r0 ≤ x) ∧ (r0 ≤ y)) ∨ ((r0 ≤ w) ∧ (r0 ≤ z))))

Lemma: rminus-reverses-rless
x,y:ℝ.  ((x < y)  (-(y) < -(x)))

Lemma: rminus-reverses-rleq
[x,y:ℝ].  -(y) ≤ -(x) supposing x ≤ y

Lemma: rsub_functionality_wrt_rleq
[x,y,z,t:ℝ].  ((x y) ≤ (z t)) supposing ((y ≥ t) and (x ≤ z))

Lemma: rsub_functionality_wrt_rless
x,y,z,t:ℝ.  ((x < z)  (x y) < (z t) supposing t ≤ y)

Lemma: rmul_functionality_wrt_rless
x,y,z:ℝ.  ((x < z)  (r0 < y)  ((x y) < (z y)))

Lemma: rmul_functionality_wrt_rless2
x,y,z,w:ℝ.  ((r0 < w)  (x < z)  (x w) < (z y) supposing w ≤ supposing r0 ≤ z)

Lemma: rless_irreflexivity
[e:ℝ]. False supposing e < e

Lemma: radd-non-neg
x,y:ℝ.  ((r0 ≤ x)  (r0 ≤ y)  (r0 ≤ (x y)))

Lemma: rsub-rmin-rleq-rabs
[a,b:ℝ].  ((b rmin(a;b)) ≤ |a b|)

Definition: rneq
x ≠ ==  (x < y) ∨ (y < x)

Lemma: rneq_wf
[x,y:ℝ].  (x ≠ y ∈ ℙ)

Lemma: rneq-iff
x,y:ℝ.  (x ≠ ⇐⇒ ∃n:ℕ+4 < |(x n) n|)

Lemma: radd-preserves-rneq
x,y,z:ℝ.  (x ≠ ⇐⇒ x ≠ y)

Lemma: sq_stable__rneq
x,y:ℝ.  SqStable(x ≠ y)

Lemma: sqs-rneq-or
x,y:ℝ.  SqStable(x ≠ r0 ∨ y ≠ r0)

Definition: rneq-zero-or
rneq-zero-or(x;y) ==
  eval find-ge-val(λp.let a,b 
                          in 4 <|a| ∨b4 <|b|;λn.eval in
                                                    eval in
                                                      <a, b>;1) in
  let v,m1 
  in let v1,v2 
     in if (4) < (v1)
           then inl (inr m1 )
           else if (v1) < (-4)
                   then inl inl m1
                   else if (4) < (v2)  then inr inr m1    else if (v2) < (-4)  then inr (inl m1)   else Ax

Lemma: sq_stable__rneq-or
x,y:ℝ.  SqStable(x ≠ r0 ∨ y ≠ r0)

Lemma: rneq-zero-or_wf
[x,y:ℝ].  rneq-zero-or(x;y) ∈ x ≠ r0 ∨ y ≠ r0 supposing x ≠ r0 ∨ y ≠ r0

Lemma: rneq-int
n,m:ℤ.  (r(n) ≠ r(m) ⇐⇒ ¬(n m ∈ ℤ))

Lemma: rneq_functionality
x1,x2,y1,y2:ℝ.  (x1 ≠ y1 ⇐⇒ x2 ≠ y2) supposing ((y1 y2) and (x1 x2))

Lemma: rneq-zero
x:ℝ(x ≠ r0 ⇐⇒ rpositive(x) ∨ rpositive(-(x)))

Lemma: rnonzero-iff
x:ℝ(rnonzero(x) ⇐⇒ x ≠ r0)

Lemma: rneq_irreflexivity
[e:ℝ]. False supposing e ≠ e

Lemma: rneq_irrefl
[e:ℝ]. e ≠ e)

Lemma: absval-implies-rneq
x,y:ℝ.  ((∃n:ℕ+4 < |(x n) n|)  x ≠ y)

Definition: case-real
case-real(a;b;f) ==  accelerate(3;λn.if 4 <|(a n) n| ∧b (f n) then else fi )

Lemma: case-real_wf
[P:ℙ]. ∀[a,b:ℝ]. ∀[f:{n:ℕ+4 < |(a n) n|}  ⟶ ((↓P) ∨ (↓¬P))].
  (case-real(a;b;f) ∈ {z:ℝ(P  (z a)) ∧ ((¬P)  (z b))} )

Definition: case-real2
case-real2(a;b;f) ==  case-real(a;b;λn.if (a n) 4 <then (inl n) else (inr fi )

Lemma: case-real2_wf
[P:ℙ]. ∀[a,b:ℝ]. ∀[f:a ≠ b ⟶ ((↓P) ∨ (↓¬P))].  (case-real2(a;b;f) ∈ {z:ℝ(P  (z a)) ∧ ((¬P)  (z b))} )

Lemma: case-real2_wf2
[d,a,b:ℝ]. ∀[f:a ≠ b ⟶ ((↓r0 < d) ∨ (↓¬(r0 < d)))].
  (case-real2(a;b;f) ∈ {z:ℝ((r0 < d)  (z a)) ∧ ((d ≤ r0)  (z b))} )

Definition: case-real3-seq
case-real3-seq(a;b;f) ==  λn.if then else fi 

Lemma: case-real3-seq_wf
[f:ℕ+ ⟶ 𝔹]. ∀[b:ℝ]. ∀[a:ℝ supposing ∃n:ℕ+(↑(f n))].
  case-real3-seq(a;b;f) ∈ {s:ℕ+ ⟶ ℤ3-regular-seq(s)}  supposing ∀n,m:ℕ+.  ((↑(f n))  (¬↑(f m))  (|(a m) m| ≤ \000C4))

Definition: case-real3
case-real3(a;b;f) ==  accelerate(3;case-real3-seq(a;b;f))

Lemma: case-real3_wf
[f:ℕ+ ⟶ 𝔹]. ∀[b:ℝ]. ∀[a:ℝ supposing ∃n:ℕ+(↑(f n))].
  case-real3(a;b;f) ∈ ℝ supposing ∀n,m:ℕ+.  ((↑(f n))  (¬↑(f m))  (|(a m) m| ≤ 4))

Lemma: case-real3-req1
[f:ℕ+ ⟶ 𝔹]
  ∀[b,a:ℝ].  case-real3(a;b;f) supposing ∀n,m:ℕ+.  ((↑(f n))  (¬↑(f m))  (|(a m) m| ≤ 4)) 
  supposing ∃n:ℕ+(↑(f n))

Lemma: case-real3-req2
[f:ℕ+ ⟶ 𝔹]. ∀[b:ℝ]. ∀[a:Top].  (case-real3(a;b;f) b) supposing ∀n:ℕ+(¬↑(f n))

Lemma: rinv_wf2
[x:ℝ]. (x ≠ r0  (rinv(x) ∈ ℝ))

Lemma: rinv_functionality2
[x,y:ℝ].  (x ≠ r0  (x y)  (rinv(x) rinv(y)))

Lemma: rmul-rinv
[x:ℝ]. (x rinv(x)) r1 supposing x ≠ r0

Lemma: rmul-rinv2
[x:ℝ]. (rinv(x) x) r1 supposing x ≠ r0

Lemma: rmul-rinv3
[x,a:ℝ].  (x rinv(x) a) supposing x ≠ r0

Lemma: rmul_reverses_rless
x,y,z:ℝ.  ((x < z)  (y < r0)  ((z y) < (x y)))

Lemma: rinv-positive
x:ℝ((r0 < x)  (r0 < rinv(x)))

Lemma: rmul-inverse-is-rinv
[x:ℝ]. ∀[t:ℝ]. rinv(x) supposing (x t) r1 supposing x ≠ r0

Lemma: rmul-neq-zero
x,y:ℝ.  (x ≠ r0  y ≠ r0  y ≠ r0)

Lemma: rminus-neq-zero
x:ℝ(x ≠ r0  -(x) ≠ r0)

Lemma: rinv-rminus
[x:ℝ]. -(rinv(x)) rinv(-(x)) supposing x ≠ r0

Lemma: rinv-negative
x:ℝ((x < r0)  (rinv(x) < r0))

Lemma: rinv-neq-zero
x:ℝ(x ≠ r0  rinv(x) ≠ r0)

Lemma: rinv-of-rmul
[x,y:ℝ].  (rinv(x y) (rinv(x) rinv(y))) supposing (y ≠ r0 and x ≠ r0)

Lemma: rinv-of-rinv
[x:ℝ]. rinv(rinv(x)) supposing x ≠ r0

Lemma: rmul_preserves_rless
x,y,z:ℝ.  ((r0 < y)  (x < ⇐⇒ (x y) < (z y)))

Lemma: rmul_preserves_rleq
[x,y,z:ℝ].  uiff(x ≤ z;(x y) ≤ (z y)) supposing r0 < y

Lemma: rmul_preserves_rleq2
[x,y,z:ℝ].  ((x y) ≤ (z y)) supposing ((x ≤ z) and (r0 ≤ y))

Lemma: rmul_preserves_rleq3
[x,y,a,b:ℝ].  ((x y) ≤ (a b)) supposing ((x ≤ a) and (y ≤ b) and ((r0 ≤ x) ∧ (r0 ≤ y)))

Lemma: rmul_reverses_rleq
[x,y,z:ℝ].  ((z y) ≤ (x y)) supposing ((y ≤ r0) and (x ≤ z))

Lemma: rminus_functionality_wrt_rleq
[x,y:ℝ].  -(x) ≤ -(y) supposing x ≥ y

Lemma: rabs-positive
x:ℝ(x ≠ r0  rpositive(|x|))

Lemma: rabs-neq-zero
x:ℝ(x ≠ r0  (r0 < |x|))

Lemma: not-rpositive
[x:ℝ]. rnonneg(-(x)) supposing ¬rpositive(x)

Lemma: not-rless
[x,y:ℝ].  y ≤ supposing ¬(x < y)

Lemma: rless_complement-RelRST-test
[x,y,z:ℝ].  ((x ≤ y)  (z < y))  (x ≤ z))

Lemma: rleq-iff-bdd
[x,y:ℝ].  (x ≤ ⇐⇒ ∃B:ℕ. ∀n:ℕ+((x n) ≤ ((y n) B)))

Lemma: not-rneq
[x,y:ℝ].  supposing ¬x ≠ y

Lemma: rneq-symmetry
x,y:ℝ.  (x ≠  y ≠ x)

Lemma: req-iff-not-rneq
[x,y:ℝ].  uiff(x y;¬x ≠ y)

Lemma: rmul_preserves_rneq
a,b,x:ℝ.  (x ≠ r0  a ≠  a ≠ b)

Lemma: radd-rmax
[x,y,z:ℝ].  ((x rmax(y;z)) rmax(x y;x z))

Lemma: rmax-int
[a,b:ℤ].  (rmax(r(a);r(b)) r(imax(a;b)))

Lemma: rmin-int
[a,b:ℤ].  (rmin(r(a);r(b)) r(imin(a;b)))

Lemma: rabs-int
[a:ℤ]. (|r(a)| r(|a|) ∈ ℝ)

Lemma: rabs-int-rmul
[k:ℤ]. ∀[x:ℝ].  (|k x| |k| |x|)

Lemma: rabs-int-rmul-unit
[k:ℕ]. ∀[x:ℝ].  (|-1^k x| |x|)

Lemma: radd-rmin
[x,y,z:ℝ].  ((x rmin(y;z)) rmin(x y;x z))

Lemma: rabs-abs
[a:ℤ]. (|r(a)| r(|a|))

Lemma: zero-rleq-rabs
[x:ℝ]. (r0 ≤ |x|)

Lemma: square-nonneg
[x:ℝ]. (r0 ≤ (x x))

Lemma: r-triangle-inequality
[x,y:ℝ].  (|x y| ≤ (|x| |y|))

Lemma: r-triangle-inequality2
[x,y,z:ℝ].  (|x z| ≤ (|x y| |y z|))

Lemma: r-triangle-inequality-rsub
[x,y:ℝ].  (|x y| ≤ (|x| |y|))

Lemma: radd-list-rabs
L:ℝ List. (|radd-list(L)| ≤ radd-list(map(λx.|x|;L)))

Lemma: rabs-rmul-rleq
[x,y,a,b:ℝ].  (|x y| ≤ (a b)) supposing ((|y| ≤ b) and (|x| ≤ a))

Lemma: rabs-rmul-rleq-rabs
[x,y,a,b:ℝ].  (|x y| ≤ |a b|) supposing ((|y| ≤ |b|) and (|x| ≤ |a|))

Lemma: rleq-rmax
[x,y:ℝ].  ((x ≤ rmax(x;y)) ∧ (y ≤ rmax(x;y)))

Lemma: rleq-iff
x,y:ℝ.  (x ≤ ⇐⇒ ∀n:ℕ+. ∃N:ℕ+. ∀m:{N...}. (((-2) m) ≤ (n ((y m) m))))

Definition: rleq2
rleq2(x;y) ==  ∀n:ℕ+. ∃N:ℕ+. ∀m:{N...}. (((-2) m) ≤ (n ((y m) m)))

Lemma: rleq2_wf
[x,y:ℕ+ ⟶ ℤ].  (rleq2(x;y) ∈ ℙ)

Lemma: rleq2-iff-rnonneg2
[x,y:ℕ+ ⟶ ℤ].  (rleq2(x;y) ⇐⇒ rnonneg2(reg-seq-add(y;-(x))))

Lemma: rleq2_functionality
x1,y1,x2,y2:ℕ+ ⟶ ℤ.  (bdd-diff(x1;x2)  bdd-diff(y1;y2)  (rleq2(x1;y1) ⇐⇒ rleq2(x2;y2)))

Lemma: rleq-iff-rleq2
x,y:ℝ.  (x ≤ ⇐⇒ rleq2(x;y))

Lemma: rless-implies
x,y:ℝ.  ((x < y)  (∃n:ℕ+. ∀m:ℕ+((n ≤ m)  (5 ≤ ((y m) m)))))

Lemma: rmax-req
[x,y:ℝ].  rmax(x;y) supposing x ≤ y

Lemma: rmin-req
[x,y:ℝ].  rmin(x;y) supposing y ≤ x

Lemma: rmax-req2
[x,y:ℝ].  rmax(x;y) supposing y ≤ x

Lemma: rmin-req2
[x,y:ℝ].  rmin(x;y) supposing x ≤ y

Lemma: rmin-max-cases
a,b:ℝ.  (a ≠  (((rmin(a;b) a) ∧ (rmax(a;b) b)) ∨ ((rmin(a;b) b) ∧ (rmax(a;b) a))))

Lemma: rmin-classical-cases
a,b:ℝ.  (¬¬((rmin(a;b) a) ∨ (rmin(a;b) b)))

Lemma: rmin-stable-cases
a,b:ℝ. ∀P:Type.  (Stable{P}  (((rmin(a;b) a)  P) ∧ ((rmin(a;b) b)  P))  P)

Lemma: rmax_ub
[x,y,z:ℝ].  z ≤ rmax(x;y) supposing (z ≤ x) ∨ (z ≤ y)

Lemma: rmax_strict_lb
x,y,z:ℝ.  ((x < z) ∧ (y < z) ⇐⇒ rmax(x;y) < z)

Lemma: rmax_strict_ub
x,y,z:ℝ.  ((z < x) ∨ (z < y) ⇐⇒ z < rmax(x;y))

Lemma: rmax-rneq
x,y,z,w:ℝ.  (rmax(x;y) ≠ rmax(z;w)  (x ≠ z ∨ y ≠ w))

Lemma: rmax_lb
[x,y,z:ℝ].  uiff((x ≤ z) ∧ (y ≤ z);rmax(x;y) ≤ z)

Lemma: rmax_functionality_wrt_rleq
[x1,x2,y1,y2:ℝ].  (rmax(x1;y1) ≤ rmax(x2;y2)) supposing ((y1 ≤ y2) and (x1 ≤ x2))

Lemma: rmin-rleq
[x,y:ℝ].  ((rmin(x;y) ≤ x) ∧ (rmin(x;y) ≤ y))

Lemma: rmin_strict_ub
x,y,z:ℝ.  ((z < x) ∧ (z < y) ⇐⇒ z < rmin(x;y))

Lemma: rmin_ub
x,y,z:ℝ.  ((z ≤ x) ∧ (z ≤ y) ⇐⇒ z ≤ rmin(x;y))

Lemma: rmin_lb
[x,y,z:ℝ].  rmin(x;y) ≤ supposing (x ≤ z) ∨ (y ≤ z)

Lemma: rmin_functionality_wrt_rleq
[x1,x2,y1,y2:ℝ].  (rmin(x1;y1) ≤ rmin(x2;y2)) supposing ((y1 ≤ y2) and (x1 ≤ x2))

Lemma: rmin_strict_lb
x,y,z:ℝ.  ((x < z) ∨ (y < z) ⇐⇒ rmin(x;y) < z)

Lemma: rmax-minus-rmin
a,b:ℝ.  (|a b| (rmax(a;b) rmin(a;b)))

Lemma: rmin-rleq-rmax
a,b:ℝ.  (rmin(a;b) ≤ rmax(a;b))

Lemma: rmin-lb-convex
a,b,t:ℝ.  ((r0 ≤ t)  (t ≤ r1)  (rmin(a;b) ≤ ((t a) ((r1 t) b))))

Lemma: rmax-ub-convex
a,b,t:ℝ.  ((r0 ≤ t)  (t ≤ r1)  (((t a) ((r1 t) b)) ≤ rmax(a;b)))

Lemma: rabs-difference-rmax
[a,b,x,y:ℝ].  (|rmax(x;y) rmax(a;b)| ≤ rmax(|x a|;|y b|))

Lemma: rabs-of-nonneg
[x:ℝ]. |x| supposing r0 ≤ x

Lemma: rmul-nonneg-rabs
[x,y:ℝ].  (x |y|) |x y| supposing r0 ≤ x

Lemma: rabs-rabs
[x:ℝ]. (||x|| |x|)

Lemma: rabs-of-nonpos
[x:ℝ]. |x| -(x) supposing x ≤ r0

Lemma: rneq-if-rabs
x,y:ℝ.  x ≠ supposing r0 < |x y|

Lemma: rneq-if-rabs2
x,y:ℝ.  ((r0 < |x y|)  x ≠ y)

Lemma: rneq-iff-rabs
x,y:ℝ.  (x ≠ ⇐⇒ r0 < |x y|)

Lemma: sq_stable_rneq
x,y:ℝ.  SqStable(x ≠ y)

Lemma: rabs-positive-iff
x:ℝ(x ≠ r0 ⇐⇒ r0 < |x|)

Lemma: rnexp_unroll
[x:ℝ]. ∀[n:ℕ].  (x^n if (n =z 0) then r1 if (n =z 1) then else x^n fi )

Lemma: rnexp_step
[x:ℝ]. ∀[n:ℕ+].  (x^n (x^n x))

Lemma: rnexp_functionality
[n:ℕ]. ∀[x,y:ℝ].  x^n y^n supposing y

Lemma: rnexp-add
[n,m:ℕ]. ∀[x:ℝ].  ((x^n x^m) x^n m)

Lemma: rnexp1
[x:ℝ]. (x^1 x)

Lemma: rnexp2
[x:ℝ]. (x^2 (x x))

Lemma: rnexp3
[x:ℝ]. (x^3 (x x))

Lemma: rnexp-one
n:ℕ(r1^n r1)

Lemma: rnexp-add1
[n:ℕ]. ∀[x:ℝ].  (x^n (x x^n))

Lemma: rnexp-minus-one
n:ℕ(r(-1)^n if (n rem =z 0) then r1 else r(-1) fi )

Lemma: rnexp-mul
[n,m:ℕ]. ∀[x:ℝ].  (x^m^n x^m n)

Lemma: rnexp-nonneg
x:ℝ((r0 ≤ x)  (∀n:ℕ(r0 ≤ x^n)))

Lemma: rnexp-even-nonneg
n:ℕ(((n rem 2) 0 ∈ ℤ (∀x:ℝ(r0 ≤ x^n)))

Lemma: rnexp2-nonneg
x:ℝ(r0 ≤ x^2)

Lemma: rnexp-positive
x:ℝ((r0 < x)  (∀n:ℕ(r0 < x^n)))

Lemma: rnexp2-positive
x:ℝ(x ≠ r0  (r0 < x^2))

Lemma: rnexp-rless
x,y:ℝ.  ((r0 ≤ x)  (x < y)  (∀n:ℕ+(x^n < y^n)))

Lemma: rnexp-rleq
x,y:ℝ.  ((r0 ≤ x)  (x ≤ y)  (∀n:ℕ(x^n ≤ y^n)))

Lemma: rnexp_functionality_wrt_rleq
x,y:ℝ.  ((r0 ≤ x)  (x ≤ y)  (∀n:ℕ(x^n ≤ y^n)))

Lemma: rnexp-rleq-iff
x,y:ℝ.  ((r0 ≤ x)  (r0 ≤ y)  (∀n:ℕ+(x ≤ ⇐⇒ x^n ≤ y^n)))

Lemma: rnexp-req-iff
n:ℕ+. ∀x,y:ℝ.  ((r0 ≤ x)  (r0 ≤ y)  (x ⇐⇒ x^n y^n))

Lemma: rabs-rnexp
[n:ℕ]. ∀[x:ℝ].  (|x^n| |x|^n)

Lemma: rabs-rnexp2
[x:ℝ]. (|x|^2 x^2)

Lemma: square-rleq-implies
x,y:ℝ.  ((r0 ≤ y)  (x^2 ≤ y^2)  (x ≤ y))

Lemma: rmax-idempotent
[x:ℝ]. (rmax(x;x) x)

Lemma: rmin-idempotent
[x:ℝ]. (rmin(x;x) x)

Lemma: rmin-idempotent-eq
[x:ℝ]. (rmin(x;x) x ∈ ℝ)

Lemma: rpower-one
[x:ℝ]. (x^1 x)

Lemma: rpower-two
[x:ℝ]. (x^2 (x x))

Lemma: rpower-nonzero
x:ℝ(x ≠ r0  (∀n:ℕx^n ≠ r0))

Lemma: rpower-greater-one
x,q:ℝ.  ((r0 < x)  ((r1 x) < q)  (∀n:ℕ(r1 (r(n) x)) < q^n supposing 1 < n))

Definition: rdiv
(x/y) ==  rinv(y)

Lemma: rdiv_wf
[x,y:ℝ].  (x/y) ∈ ℝ supposing y ≠ r0

Lemma: rinv-as-rdiv
[y:ℝ]. rinv(y) (r1/y) supposing y ≠ r0

Lemma: rinv-mul-as-rdiv
[y,a:ℝ].  (rinv(y) a) (a/y) supposing y ≠ r0

Lemma: mul-rinv-as-rdiv
[y,a:ℝ].  (a rinv(y)) (a/y) supposing y ≠ r0

Lemma: int-rinv-cancel
[a:ℤ]. ∀[b:ℤ-o]. ∀[x:ℝ].  ((r(a b) rinv(r(b)) x) (r(a) x))

Lemma: int-rinv-cancel2
[a:ℤ]. ∀[b:ℤ-o].  ((r(a b) rinv(r(b))) r(a))

Lemma: rdiv_functionality
[x1,x2,y1,y2:ℝ].  ((x1/y1) (x2/y2)) supposing ((y1 y2) and (x1 x2) and y1 ≠ r0)

Lemma: rdiv_functionality_wrt_rleq2
[x,y,z,w:ℝ].  ((x/w) ≤ (z/y)) supposing ((x ≤ z) and (y ≤ w) and ((r0 < y) ∧ ((r0 ≤ x) ∨ (r0 ≤ z))))

Lemma: rdiv-self
[x:ℝ]. (x/x) r1 supposing x ≠ r0

Lemma: rdiv-zero
[x:ℝ]. (r0/x) r0 supposing x ≠ r0

Lemma: rmul-ident-div
[r,s:ℝ].  ((r/r) s) supposing r ≠ r0

Lemma: rmul-zero-div
[x,y:ℝ].  ((r0/y) x) r0 supposing y ≠ r0

Lemma: rmul_preserves_req
[x,y,z:ℝ].  uiff(x z;(x y) (z y)) supposing y ≠ r0

Definition: rat_term_to_real
rat_term_to_real(f;t) ==
  rat_term_ind(t;
               rtermConstant(n) <True, r(n)>;
               rtermVar(v) <True, v>;
               rtermAdd(left,right) rl,rr.let P,x rl 
                                            in let Q,y rr 
                                               in <P ∧ Q, y>;
               rtermSubtract(left,right) rl,rr.let P,x rl 
                                                 in let Q,y rr 
                                                    in <P ∧ Q, y>;
               rtermMultiply(left,right) rl,rr.let P,x rl 
                                                 in let Q,y rr 
                                                    in <P ∧ Q, y>;
               rtermDivide(num,denom) rl,rr.let P,x rl 
                                              in let Q,y rr 
                                                 in <(P ∧ Q) ∧ y ≠ r0, (x/y)>;
               rtermMinus(num) rx.let P,x rx 
                                    in <P, -(x)>

Lemma: rat_term_to_real_wf
[t:rat_term()]. ∀[f:ℤ ⟶ ℝ].  (rat_term_to_real(f;t) ∈ P:ℙ × ℝ supposing P)

Definition: req_rat_term
r ≡ p/q ==
  ∀f:ℤ ⟶ ℝ
    let P,x rat_term_to_real(f;r) 
    in  (real_term_value(f;q) ≠ r0 ∧ (x (real_term_value(f;p)/real_term_value(f;q))))

Lemma: req_rat_term_wf
[r:rat_term()]. ∀[p,q:int_term()].  (r ≡ p/q ∈ ℙ)

Lemma: rat_term_polynomial
r:rat_term(). let p,q rat_term_to_ipolys(r) in r ≡ ipolynomial-term(p)/ipolynomial-term(q)

Definition: rat-term-eq
rat-term-eq(r1;r2) ==
  let p1,q1 rat_term_to_ipolys(r1) 
  in let p2,q2 rat_term_to_ipolys(r2) 
     in null(add-ipoly(mul-ipoly(p1;q2);mul-ipoly(minus-poly(p2);q1)))

Lemma: rat-term-eq_wf
[r1,r2:rat_term()].  (rat-term-eq(r1;r2) ∈ 𝔹)

Lemma: assert-rat-term-eq
r1,r2:rat_term().
  ((↑rat-term-eq(r1;r2))
   (∀f:ℤ ⟶ ℝlet p,x rat_term_to_real(f;r1) in let q,y rat_term_to_real(f;r2) in   (x y)))

Lemma: assert-rat-term-eq2
[r1,r2:rat_term()]. ∀[f:ℤ ⟶ ℝ].
  ((snd(rat_term_to_real(f;r1))) (snd(rat_term_to_real(f;r2)))) supposing 
     ((fst(rat_term_to_real(f;r1))) and 
     (fst(rat_term_to_real(f;r2))) and 
     (inl Ax ≤ rat-term-eq(r1;r2)))

Lemma: rmul-rdiv-cancel
[a,b:ℝ].  (a (b/a)) supposing a ≠ r0

Lemma: rmul-rdiv-cancel2
[a,b:ℝ].  ((b/a) a) supposing a ≠ r0

Lemma: rmul-rdiv-cancel3
[a,b,c:ℝ].  (a (b/a) c) (b c) supposing a ≠ r0

Lemma: rmul-rdiv-cancel4
[a,b,c:ℝ].  ((b/a) c) (b c) supposing a ≠ r0

Lemma: another-test-ring-req
a,b,c,d,e,x:ℝ.  (b ≠ r0  d ≠ r0  x ≠ r0  (((a/b) (c/d) (b e/x)) ((a c/d) e/x)))

Lemma: rmul_reverses_rleq_iff
[x,y,z:ℝ].  uiff(x ≤ z;(z y) ≤ (x y)) supposing y < r0

Lemma: rmul_reverses_rless_iff
x,y,z:ℝ.  ((y < r0)  (x < ⇐⇒ (z y) < (x y)))

Lemma: rmul_preserves_rneq_iff
a,b,x:ℝ.  (x ≠ r0  (a ≠ ⇐⇒ a ≠ b))

Lemma: rmul_preserves_rneq_iff2
a,b,x:ℝ.  (x ≠ r0  (a ≠ ⇐⇒ x ≠ x))

Lemma: one-rdiv-rmul
[x,y:ℝ].  ((r1/y) x) (x/y) supposing y ≠ r0

Lemma: rmul-rdiv-cancel5
[a,b,c:ℝ].  ((b/a) (a/c)) (b/c) supposing a ≠ r0 ∧ c ≠ r0

Lemma: rmul-rdiv-cancel6
[a,b,c:ℝ].  (a (b/a c)) (b/c) supposing a ≠ r0 ∧ c ≠ r0

Lemma: rmul-rdiv-cancel7
[a,b:ℝ].  (a b/b) supposing b ≠ r0

Lemma: rmul-rdiv-cancel8
[a,b,c:ℝ].  (a b/b c) (a/c) supposing b ≠ r0 ∧ c ≠ r0

Lemma: rmul-rdiv-cancel9
[a,b,c:ℝ].  (a b/c b) (a/c) supposing b ≠ r0 ∧ c ≠ r0

Lemma: rmul-rdiv-cancel10
[a,b,c:ℝ].  (b a/b c) (a/c) supposing b ≠ r0 ∧ c ≠ r0

Lemma: req-rdiv
x,y,z:ℝ.  (z ≠ r0  (x (y/z) ⇐⇒ (x z) y))

Lemma: req-rdiv2
x,y,z,u:ℝ.  (z ≠ r0  u ≠ r0  ((x/u) (y/z) ⇐⇒ (x z) (y u)))

Lemma: radd-rdiv
[a,b,c:ℝ].  ((b/a) (c/a)) (b c/a) supposing a ≠ r0

Lemma: rsub-rdiv
[a,b,c:ℝ].  ((b/a) (c/a)) (b c/a) supposing a ≠ r0

Lemma: rmul-rdiv
[x,y,a,b:ℝ].  (((x/a) (y/b)) (x y/a b)) supposing (b ≠ r0 and a ≠ r0)

Lemma: rmul-rdiv-one
[a,b:ℝ].  (((r1/a) (r1/b)) (r1/a b)) supposing (b ≠ r0 and a ≠ r0)

Lemma: rmul-rdiv2
[x,a,b:ℝ].  ((x/a b) ((x/a) (r1/b))) supposing (b ≠ r0 and a ≠ r0)

Lemma: rdiv-rdiv
[x,a,b:ℝ].  (((x/a)/b) (x/a b)) supposing (b ≠ r0 and a ≠ r0)

Lemma: rminus-rdiv
[x,a:ℝ].  -((x/a)) (-(x)/a) supposing a ≠ r0

Lemma: rminus-rdiv2
[x,a:ℝ].  -((x/a)) (x/-(a)) supposing a ≠ r0

Lemma: radd-rneq0
x,y:ℝ.  (x y ≠ r0 ⇐⇒ x ≠ -(y))

Lemma: int-rdiv-req
[k:ℤ-o]. ∀[a:ℝ].  ((a)/k (a/r(k)))

Lemma: int-rdiv_functionality
[k1,k2:ℤ-o]. ∀[a,b:ℝ].  ((a)/k1 (b)/k2) supposing ((k1 k2 ∈ ℤand (a b))

Lemma: rsub-rmin
[x,y,z:ℝ].  ((x rmin(y;z)) rmax(x y;x z))

Lemma: rleq-iff-all-rless
[x,y:ℝ].  uiff(x ≤ y;∀e:{e:ℝr0 < e} (x ≤ (y e)))

Lemma: rat-to-real-req
[a:ℤ]. ∀[b:ℤ-o].  (r(a/b) (r(a)/r(b)))

Lemma: rinv_preserves_rless
a,b:ℝ.  ((r0 < a)  (a < b)  ((r1/b) < (r1/a)))

Lemma: rinv_preserves_rneq
a,b:ℝ.  (a ≠ r0  b ≠ r0  a ≠  (r1/a) ≠ (r1/b))

Lemma: rinverse-nonzero
x:ℝ(x ≠ r0  (r1/x) ≠ r0)

Lemma: inverse-rpower
[x:ℝ]. ∀[n:ℕ]. ((r1/x^n) (r1/x)^n) supposing x ≠ r0

Lemma: rnexp-rmul
[n:ℕ]. ∀[x,y:ℝ].  (x y^n (x^n y^n))

Lemma: rnexp-rdiv
[y,x:ℝ].  ∀[n:ℕ]. ((y^n/x^n) (y/x)^n) supposing x ≠ r0

Lemma: rabs-bounds
[x:ℝ]. ((-(|x|) ≤ x) ∧ (x ≤ |x|))

Lemma: rabs-rinv
y:ℝ(y ≠ r0  (|rinv(y)| rinv(|y|)))

Lemma: rabs-rdiv
x,y:ℝ.  (y ≠ r0  (|(x/y)| (|x|/|y|)))

Lemma: rleq-int-fractions
[a,b:ℤ]. ∀[c,d:ℕ+].  uiff((r(a)/r(c)) ≤ (r(b)/r(d));(a d) ≤ (b c))

Lemma: req-int-fractions
[a,b:ℤ]. ∀[c,d:ℤ-o].  uiff((r(a)/r(c)) (r(b)/r(d));(a d) (b c) ∈ ℤ)

Lemma: req-int-fractions2
[a,b:ℤ]. ∀[c:ℤ-o].  uiff((r(a)/r(c)) r(b);a (b c) ∈ ℤ)

Lemma: rleq-int-fractions2
[a,b:ℤ]. ∀[d:ℕ+].  uiff(r(a) ≤ (r(b)/r(d));(a d) ≤ b)

Lemma: rleq-int-fractions3
[a,b:ℤ]. ∀[d:ℕ+].  uiff((r(b)/r(d)) ≤ r(a);b ≤ (a d))

Lemma: rless-int-fractions
a,b:ℤ. ∀c,d:ℕ+.  ((r(a)/r(c)) < (r(b)/r(d)) ⇐⇒ d < c)

Lemma: decidable__rless-int-fractions
a,b:ℤ. ∀c,d:ℕ+.  Dec((r(a)/r(c)) < (r(b)/r(d)))

Lemma: rless-int-fractions2
a,b:ℤ. ∀d:ℕ+.  (r(a) < (r(b)/r(d)) ⇐⇒ d < b)

Lemma: rless-int-fractions3
a,b:ℤ. ∀d:ℕ+.  ((r(b)/r(d)) < r(a) ⇐⇒ b < d)

Lemma: rmul-int-rdiv
[x:ℝ]. ∀[a,b:ℤ].  ((r(a) (r(b)/x)) (r(a b)/x)) supposing x ≠ r0

Lemma: rmul-int-rdiv2
[x:ℝ]. ∀[a,b:ℤ].  (((r(b)/x) r(a)) (r(a b)/x)) supposing x ≠ r0

Lemma: radd-int-fractions
[a,b:ℤ]. ∀[c,d:ℕ+].  (((r(a)/r(c)) (r(b)/r(d))) (r((a d) (b c))/r(c d)))

Lemma: rsub-int-fractions
[a,b:ℤ]. ∀[c,d:ℕ+].  (((r(a)/r(c)) (r(b)/r(d))) (r((a d) c)/r(c d)))

Lemma: rmul-int-fractions
[a,b:ℤ]. ∀[c,d:ℕ+].  (((r(a)/r(c)) (r(b)/r(d))) (r(a b)/r(c d)))

Lemma: rdiv-int-fractions
a,b:ℤ. ∀c,d:ℕ+.  ((r(a)/r(c))/(r(b)/r(d))) (r(a d)/r(c b)) supposing ¬(b 0 ∈ ℤ)

Lemma: int-rdiv-int-rdiv
[k,j:ℤ-o]. ∀[x:ℝ].  (((x)/k)/j (x)/j k)

Lemma: clear-denominator1
[a,b,c,d:ℝ].  uiff(((a/b) c) d;(c a) (d b)) supposing b ≠ r0

Lemma: clear-denominator2
[a,b,c,d,e:ℝ].  uiff((((a/b) c) e) d;(c a) (d b)) supposing b ≠ r0

Lemma: fractions-req
[a,b,c,d:ℝ].  (c ≠ r0  d ≠ r0  uiff((a/c) (b/d);(a d) (b c)))

Lemma: fractions-rless
a,b,c,d:ℝ.  ((r0 < c)  (r0 < d)  ((a/c) < (b/d) ⇐⇒ (a d) < (b c)))

Lemma: fractions-rleq
a,b,c,d:ℝ.  ((r0 < c)  (r0 < d)  ((a/c) ≤ (b/d) ⇐⇒ (a d) ≤ (b c)))

Definition: rat2real
rat2real(q) ==  if isint(q) then r(q) else let a,b in (r(a))/b fi 

Lemma: rat2real_wf
[q:ℚ]. (rat2real(q) ∈ ℝ)

Lemma: rleq-rat2real
[q1,q2:ℚ].  uiff(rat2real(q1) ≤ rat2real(q2);q1 ≤ q2)

Lemma: rless-rat2real
q1,q2:ℚ.  uiff(rat2real(q1) < rat2real(q2);q1 < q2)

Lemma: rneq-rat2real
q1,q2:ℚ.  uiff(rat2real(q1) ≠ rat2real(q2);¬(q1 q2 ∈ ℚ))

Lemma: req-rat2real
q1,q2:ℚ.  uiff(rat2real(q1) rat2real(q2);q1 q2 ∈ ℚ)

Lemma: rat2real-qdiv
a:ℚ. ∀b:ℤ-o.  (rat2real((a/b)) (rat2real(a)/r(b)))

Lemma: rat2real-qdiv2
a:ℤ. ∀b:ℤ-o.  (rat2real((a/b)) (r(a)/r(b)))

Lemma: rat2real-qadd
[a,b:ℚ].  (rat2real(a b) (rat2real(a) rat2real(b)))

Lemma: rat2real-qmul
[a,b:ℚ].  (rat2real(a b) (rat2real(a) rat2real(b)))

Lemma: rat2real-qsub
[a,b:ℚ].  (rat2real(a b) (rat2real(a) rat2real(b)))

Lemma: rat2real-qavg
[a,b:ℚ].  (rat2real(qavg(a;b)) (rat2real(a) rat2real(b)/r(2)))

Lemma: rat2real-qavg-2
[a,b:ℚ].  ((r(2) rat2real(qavg(a;b))) (rat2real(a) rat2real(b)))

Lemma: rat2real-qmax
[a,b:ℚ].  (rat2real(qmax(a;b)) rmax(rat2real(a);rat2real(b)))

Lemma: rat2real-qmin
[a,b:ℚ].  (rat2real(qmin(a;b)) rmin(rat2real(a);rat2real(b)))

Lemma: radd-positive-implies
x,y:ℝ.  ((r0 < (x y))  ((r0 < x) ∨ (r0 < y)))

Lemma: rabs-difference-bound-iff
x,y,z:ℝ.  (|x y| < ⇐⇒ ((y z) < x) ∧ (x < (y z)))

Lemma: rabs-difference-bound-rleq
x,y,z:ℝ.  (|x y| ≤ ⇐⇒ ((y z) ≤ x) ∧ (x ≤ (y z)))

Lemma: rabs-rleq
x,z:ℝ.  (|x| ≤ ⇐⇒ (-(z) ≤ x) ∧ (x ≤ z))

Lemma: rabs-difference-is-zero
x,y:ℝ.  (|x y| r0 ⇐⇒ y)

Lemma: rabs-rleq-iff
x,z:ℝ.  (|x| ≤ ⇐⇒ (-(z) ≤ x) ∧ (x ≤ z))

Lemma: rabs-is-zero
x:ℝ(|x| r0 ⇐⇒ r0)

Lemma: rabs-rless-iff
x,z:ℝ.  (|x| < ⇐⇒ (-(z) < x) ∧ (x < z))

Lemma: rabs-difference-lower-bound
x,y,z:ℝ.  (z < |x y| ⇐⇒ ((z y) < x) ∨ ((z x) < y))

Lemma: rneq-rabs
a,v:ℝ.  (v ≠  -(v) ≠  v ≠ |a|)

Lemma: rless-cases1
x,y:ℝ.  ((x < y)  (∀z:ℝ((x < z) ∨ (z < y))))

Lemma: rnexp-rminus
x:ℝ. ∀n:ℕ.  (-(x)^n if isOdd(n) then -(x^n) else x^n fi )

Lemma: rnexp-rless2
x,y:ℝ.  ((x < y)  (r0 < y)  (∀n:ℕ+((↑isOdd(n))  (x^n < y^n))))

Lemma: rnexp-rless-odd
n:ℕ+((↑isOdd(n))  (∀x,y:ℝ.  ((x < y)  (x^n < y^n))))

Lemma: rnexp-req-iff-odd
n:ℕ+. ∀x,y:ℝ.  ((↑isOdd(n))  (x ⇐⇒ x^n y^n))

Lemma: rnexp-req-iff-even
n:ℕ+. ∀x,y:ℝ.  ((↑isEven(n))  (|x| |y| ⇐⇒ x^n y^n))

Lemma: square-req-iff
x,y:ℝ.  ((r0 ≤ x)  (r0 ≤ y)  (x ⇐⇒ x^2 y^2))

Lemma: squares-req
x,y:ℝ.  (y ≠ r0  (x^2 y^2 ⇐⇒ (x y) ∨ (x -(y))))

Lemma: rabs-ub
a:ℝ((r0 < a)  (∀x:ℝ(a ≤ |x| ⇐⇒ (a ≤ x) ∨ (a ≤ -(x)))))

Lemma: rabs-strict-ub
a:ℝ((r0 ≤ a)  (∀x:ℝ(a < |x| ⇐⇒ (a < x) ∨ (a < -(x)))))

Lemma: rabs-rneq
a,b:ℝ.  (|a| ≠ |b|  a ≠ b)

Lemma: real-approx
[x:ℝ]. ∀[n:ℕ+].  (|(r(2 n) x) r(x n)| ≤ r(2))

Definition: rational-approx
(x within 1/n) ==  (r(x n))/2 n

Lemma: rational-approx_wf
[x:ℕ+ ⟶ ℤ]. ∀[n:ℕ+].  ((x within 1/n) ∈ ℝ)

Lemma: rational-approx-property
x:ℝ. ∀n:ℕ+.  (|x (x within 1/n)| ≤ (r1/r(n)))

Lemma: rational-approx-property-alt
x:ℝ. ∀n:ℕ+.  (|(r(2 n) x) r(x n)| ≤ r(2))

Lemma: rational-approx-property-alt2
x:ℝ. ∀n:ℕ+.  (|r(x n)| ≤ ((r(2 n) |x|) r(2)))

Lemma: rational-approx-property-ext
x:ℝ. ∀n:ℕ+.  (|x (x within 1/n)| ≤ (r1/r(n)))

Lemma: rational-approx-property1
x:ℝ. ∀n:ℕ+.  (x ≤ ((x within 1/n) (r1/r(n))))

Lemma: rational-approx-property2
x:ℝ. ∀n:ℕ+.  (((x within 1/n) (r1/r(n))) ≤ x)

Definition: rational-lower-approx
(below within 1/n) ==  eval in eval (x m) in   (r(a))/2 m

Lemma: rational-lower-approx_wf
[x:ℕ+ ⟶ ℤ]. ∀[n:ℕ+].  ((below within 1/n) ∈ ℝ)

Lemma: rational-lower-approx-property
x:ℝ. ∀n:ℕ+.  (((below within 1/n) ≤ x) ∧ (x ≤ ((below within 1/n) (r1/r(n)))))

Definition: rational-upper-approx
above within 1/n ==  eval in eval (x m) in   (r(a))/2 m

Lemma: rational-upper-approx_wf
[x:ℕ+ ⟶ ℤ]. ∀[n:ℕ+].  (above within 1/n ∈ ℝ)

Lemma: rational-upper-approx-as-rat2real
x:ℝ. ∀n:ℕ+.  ∃q:ℚ(above within 1/n rat2real(q))

Lemma: rational-lower-approx-as-rat2real
x:ℝ. ∀n:ℕ+.  ∃q:ℚ((below within 1/n) rat2real(q))

Lemma: rational-upper-approx-property
x:ℝ. ∀n:ℕ+.  (((above within 1/n (r1/r(n))) ≤ x) ∧ (x ≤ above within 1/n))

Definition: rational-inner-approx
rational-inner-approx(x;n) ==
  eval in
  eval in
  eval if 4 <then 2
           if z <-4 then 2
           else 0
           fi  in
    (r(a))/2 m

Lemma: rational-inner-approx_wf
[x:ℕ+ ⟶ ℤ]. ∀[n:ℕ+].  (rational-inner-approx(x;n) ∈ ℝ)

Lemma: rational-inner-approx-property
x:ℝ. ∀n:ℕ+.  ((|rational-inner-approx(x;n)| ≤ |x|) ∧ (|x rational-inner-approx(x;n)| ≤ (r(2)/r(n))))

Lemma: rational-inner-approx-int
x:ℝ. ∀n:ℕ+.  ∃z:ℤ((|(r(z)/r(4 n))| ≤ |x|) ∧ (|x (r(z)/r(4 n))| ≤ (r(2)/r(n))))

Lemma: integer-between-reals
a,b:ℝ.  ((r(2) ≤ (b a))  (∃k:ℤ((a < r(k)) ∧ (r(k) < b))))

Lemma: close-reals-iff
[x,y:ℝ]. ∀[k:ℕ+].  uiff(|x y| ≤ (r1/r(k));∀m:ℕ+((|(x m) m| k) ≤ ((4 k) (2 m))))

Lemma: implies-close-reals
[x,y:ℝ]. ∀[m:ℕ+]. ∀[k:ℕ].  ((|(x m) m| ≤ (2 k))  (|x y| ≤ (r(2 k)/r(m))))

Lemma: close-reals-implies
[x,y:ℝ]. ∀[m:ℕ+].  |(x m) m| ≤ supposing |x y| ≤ (r1/r(3 m))

Definition: ravg
ravg(x;y) ==  (x y/r(2))

Lemma: ravg_wf
[x,y:ℝ].  (ravg(x;y) ∈ ℝ)

Lemma: ravg_comm
[x,y:ℝ].  (ravg(x;y) ravg(y;x))

Lemma: ravg-between
x,y:ℝ.  ((x < y)  ((x < ravg(x;y)) ∧ (ravg(x;y) < y)))

Lemma: ravg-weak-between
x,y:ℝ.  ((x ≤ y)  ((x ≤ ravg(x;y)) ∧ (ravg(x;y) ≤ y)))

Lemma: ravg-dist
x,y:ℝ.  ((|ravg(x;y) x| ((r1/r(2)) |y x|)) ∧ (|ravg(x;y) y| ((r1/r(2)) |y x|)))

Lemma: ravg-dist-when-rleq
[x,y:ℝ].  ((ravg(x;y) x) ((r1/r(2)) (y x))) ∧ ((y ravg(x;y)) ((r1/r(2)) (y x))) supposing x ≤ y

Definition: regular-upto
regular-upto(k;n;f) ==
  bdd-all(n;i.bdd-all(n;j.|((i 1) (f (j 1))) (j 1) (f (i 1))| ≤(2 k) ((i 1) 1)))

Lemma: regular-upto_wf
[k,n:ℕ]. ∀[f:ℕ+ ⟶ ℤ].  (regular-upto(k;n;f) ∈ 𝔹)

Lemma: assert-regular-upto
[k,n:ℕ]. ∀[f:ℕ+ ⟶ ℤ].  (↑regular-upto(k;n;f) ⇐⇒ ∀i,j:ℕ+1.  (|(i (f j)) (f i)| ≤ ((2 k) (i j))))

Definition: strong-regular-upto
strong-regular-upto(a;b;n;f) ==
  bdd-all(n;i.bdd-all(n;j.a |((i 1) (f (j 1))) (j 1) (f (i 1))| ≤((i 1) 1)))

Lemma: strong-regular-upto_wf
[a,b,n:ℕ]. ∀[f:ℕ+ ⟶ ℤ].  (strong-regular-upto(a;b;n;f) ∈ 𝔹)

Lemma: assert-strong-regular-upto
[a,b,n:ℕ]. ∀[f:ℕ+ ⟶ ℤ].
  (↑strong-regular-upto(a;b;n;f) ⇐⇒ ∀i,j:ℕ+1.  ((a |(i (f j)) (f i)|) ≤ (b (i j))))

Definition: seq-min-upper
seq-min-upper(k;n;f) ==
  primrec(n;1;λi,r. if ((i 1) (f r)) ((2 k) (i 1)) ≤(r (f (i 1))) ((2 k) r)
                   then r
                   else 1
                   fi )

Lemma: seq-min-upper_wf
[k,n:ℕ]. ∀[f:ℕ+ ⟶ ℤ].  (seq-min-upper(k;n;f) ∈ ℕ+)

Lemma: seq-min-upper-le
[k:ℕ]. ∀[n:ℕ+]. ∀[f:ℕ+ ⟶ ℤ].  (seq-min-upper(k;n;f) ≤ n)

Lemma: seq-min-upper-property
[k,n:ℕ]. ∀[f:ℕ+ ⟶ ℤ].
  ∀i:ℕ+1. (((i (f seq-min-upper(k;n;f))) seq-min-upper(k;n;f) (f i)) ≤ ((2 k) (seq-min-upper(k;n;f) i)))

Definition: seq-max-lower
seq-max-lower(k;n;f) ==
  primrec(n;1;λi,r. if ((i 1) (f r)) (2 k) (i 1) ≤(r (f (i 1))) (2 k) then else fi )

Lemma: seq-max-lower_wf
[k,n:ℕ]. ∀[f:ℕ+ ⟶ ℤ].  (seq-max-lower(k;n;f) ∈ ℕ+)

Lemma: seq-max-lower-le
[k:ℕ]. ∀[n:ℕ+]. ∀[f:ℕ+ ⟶ ℤ].  (seq-max-lower(k;n;f) ≤ n)

Lemma: seq-max-lower-property
[k,n:ℕ]. ∀[f:ℕ+ ⟶ ℤ].
  ∀i:ℕ+1. (((seq-max-lower(k;n;f) (f i)) (f seq-max-lower(k;n;f))) ≤ ((2 k) (seq-max-lower(k;n;f) i)))

Lemma: regular-upto-iff
k,b:ℕ+. ∀x:ℕ+ ⟶ ℤ.
  (↑regular-upto(k;b;x)
  ⇐⇒ ∀n,m:ℕ+1.
        let seq-min-upper(k;b;x) in
         let (r((x j) (2 k))/r((2 k) j)) in
         (((r((x n) k)/r((2 k) n)) ≤ z) ∧ (z ≤ (r((x n) (2 k))/r((2 k) n))))
         ∧ ((r((x m) k)/r((2 k) m)) ≤ z)
         ∧ (z ≤ (r((x m) (2 k))/r((2 k) m))))

Lemma: regular-upto-iff2
k,b:ℕ+. ∀x:ℕ+ ⟶ ℤ.
  (↑regular-upto(k;b;x)
  ⇐⇒ ∀n,m:ℕ+1.
        let seq-max-lower(k;b;x) in
         let (r((x j) k)/r((2 k) j)) in
         (((r((x n) k)/r((2 k) n)) ≤ z) ∧ (z ≤ (r((x n) (2 k))/r((2 k) n))))
         ∧ ((r((x m) k)/r((2 k) m)) ≤ z)
         ∧ (z ≤ (r((x m) (2 k))/r((2 k) m))))

Lemma: regular-iff-all-regular-upto
k:ℕ+. ∀x:ℕ+ ⟶ ℤ.  (k-regular-seq(x) ⇐⇒ ∀b:ℕ+(↑regular-upto(k;b;x)))

Lemma: regular-int-seq-iff
k:ℕ+. ∀x:ℕ+ ⟶ ℤ.
  (k-regular-seq(x)
  ⇐⇒ ∀n,m:ℕ+.
        ∃z:ℝ
         ((((r((x n) k)/r((2 k) n)) ≤ z) ∧ (z ≤ (r((x n) (2 k))/r((2 k) n))))
         ∧ ((r((x m) k)/r((2 k) m)) ≤ z)
         ∧ (z ≤ (r((x m) (2 k))/r((2 k) m)))))

Lemma: regular-iff
x:ℕ+ ⟶ ℤ
  (regular-seq(x)
  ⇐⇒ ∀n,m:ℕ+.
        ∃z:ℝ
         ((((r((x n) 2)/r(2 n)) ≤ z) ∧ (z ≤ (r((x n) 2)/r(2 n))))
         ∧ ((r((x m) 2)/r(2 m)) ≤ z)
         ∧ (z ≤ (r((x m) 2)/r(2 m)))))

Definition: ddr
ddr(x;n) ==  eval 10^n in eval N2 N ÷ in eval N2 in   r(a/N)

Lemma: ddr_wf
x:ℝ. ∀n:ℕ+.  (ddr(x;n) ∈ {y:ℝ|x y| ≤ (r1/r(5 10^(n 1)))} )

Lemma: neg-approx-of-nonneg-real
x:ℝ((r0 ≤ x)  (∀n:ℕ+(((x n) ≤ 0)  (|x n| ≤ 2))))

Lemma: rmul-rmax
[x,y,z:ℝ].  ((r0 ≤ z)  ((z rmax(x;y)) rmax(z x;z y)))

Lemma: rmul-rmin
[x,y,z:ℝ].  ((r0 ≤ z)  ((z rmin(x;y)) rmin(z x;z y)))

Lemma: rmax-rnexp
[n:ℕ]. ∀[x,y:ℝ].  ((r0 ≤ x)  (r0 ≤ y)  (rmax(x^n;y^n) rmax(x;y)^n))

Lemma: rmin-rnexp
[n:ℕ]. ∀[x,y:ℝ].  ((r0 ≤ x)  (r0 ≤ y)  (rmin(x^n;y^n) rmin(x;y)^n))

Lemma: rless-property
x,y:ℝ. ∀n:x < y.  (x n) 4 < n

Lemma: rless-cases-proof
x,y:ℝ.  ((x < y)  (∀z:ℝ((x < z) ∨ (z < y))))

Definition: rless-case
rless-case(x;y;n;z) ==  eval (12 n) in if ((x m) 4) < (z m)  then inl m  else (inr )

Lemma: rless-cases
x,y:ℝ.  ((x < y)  (∀z:ℝ((x < z) ∨ (z < y))))

Lemma: rless-case_wf
[x,y:ℝ]. ∀[n:x < y]. ∀[z:ℝ].  (rless-case(x;y;n;z) ∈ (x < z) ∨ (z < y))

Lemma: rless-cases-sq
x:ℝ. ∀y:{y:ℝx < y} . ∀z:ℝ.  ((x < z) ∨ (z < y))

Lemma: reals-close-or-rneq
K:ℕ+
  (∃g:ℝ ⟶ ℝ ⟶ ℕ[(∀x,y:ℝ.
                       ((((g y) 1 ∈ ℤ ((r1/r(K)) < (y x)))
                       ∧ (((g y) 2 ∈ ℤ ((r1/r(K)) < (x y)))
                       ∧ (((g y) 0 ∈ ℤ (|x y| < (r(2)/r(K))))))])

Definition: rclose-or-sep
rclose-or-sep(K;x;y) ==
  eval x1 (36 K) in
  if (((r1)/K x1) 4) < (|(x y) x1|)
     then if (-((x y) ((x1 20) 1))) < ((x y) ((x1 20) 1))  then 2  else 1
     else 0

Lemma: reals-close-or-rneq-ext
K:ℕ+
  (∃g:ℝ ⟶ ℝ ⟶ ℕ[(∀x,y:ℝ.
                       ((((g y) 1 ∈ ℤ ((r1/r(K)) < (y x)))
                       ∧ (((g y) 2 ∈ ℤ ((r1/r(K)) < (x y)))
                       ∧ (((g y) 0 ∈ ℤ (|x y| < (r(2)/r(K))))))])

Lemma: rclose-or-sep_wf
[K:ℕ+]. ∀[x,y:ℝ].
  (rclose-or-sep(K;x;y) ∈ {i:ℕ3| 
                           ((i 1 ∈ ℤ ((r1/r(K)) < (y x)))
                           ∧ ((i 2 ∈ ℤ ((r1/r(K)) < (x y)))
                           ∧ ((i 0 ∈ ℤ (|x y| < (r(2)/r(K))))} )

Definition: small-real-test
small-real-test(k;z) ==  rless-case((r1)/k;(r(4))/k;k;z)

Lemma: small-real-test_wf
[k:ℕ+]. ∀[z:ℝ].  (small-real-test(k;z) ∈ ((r1)/k < z) ∨ (z < (r(4))/k))

Lemma: rneq-cases
x,y:ℝ.  (x ≠  (∀z:ℝ(x ≠ z ∨ y ≠ z)))

Lemma: rneq-cotrans
x,y,z:ℝ.  (x ≠  (x ≠ z ∨ y ≠ z))

Lemma: rneq-radd
x,y,a,b:ℝ.  (x y ≠  (x ≠ a ∨ y ≠ b))

Lemma: rmul-is-negative1
x,y:ℝ.  (((x y) < r0)  (x ≠ r0 ∨ y ≠ r0))

Lemma: rmul-is-negative
x,y:ℝ.  (((x y) < r0)  ((x < r0) ∨ (y < r0)))

Lemma: rmul-is-positive
x,y:ℝ.  (r0 < (x y) ⇐⇒ ((r0 < x) ∧ (r0 < y)) ∨ ((x < r0) ∧ (y < r0)))

Lemma: combine-rless
a,b,c,d:ℝ.  ((a < b)  (c < d)  (((b c) (a d)) < ((b d) (a c))))

Lemma: rmul-negative-iff
x,y:ℝ.  ((x y) < r0 ⇐⇒ ((r0 < x) ∧ (y < r0)) ∨ ((x < r0) ∧ (r0 < y)))

Lemma: rneq-rmul
x,y,a,b:ℝ.  (x y ≠  (x ≠ a ∨ y ≠ b))

Lemma: square-req-self-iff
x:ℝ((x x) ⇐⇒ (x r1) ∨ (x r0))

Lemma: int-rmul-is-positive
k:ℤ. ∀y:ℝ.  (r0 < ⇐⇒ (0 < k ∧ (r0 < y)) ∨ (k < 0 ∧ (y < r0)))

Lemma: rdiv-is-positive
x,y:ℝ.  (y ≠ r0  (r0 < (x/y) ⇐⇒ ((r0 < x) ∧ (r0 < y)) ∨ ((x < r0) ∧ (y < r0))))

Lemma: int-rdiv-is-positive
x:ℝ. ∀k:ℤ-o.  (r0 < (x)/k ⇐⇒ ((r0 < x) ∧ 0 < k) ∨ ((x < r0) ∧ k < 0))

Lemma: epsilon/2-lemma
[k:ℕ+]. (((r1/r(2 k)) (r1/r(2 k))) ≤ (r1/r(k)))

Lemma: epsilon/n-lemma
[k,n:ℕ+].  ((r(n) (r1/r(n k))) ≤ (r1/r(k)))

Definition: real-ratio-bound
real-ratio-bound(M;x;y;a;b) ==
  eval rclose-or-sep(M;x;y) in
  if (c =z 1) then (a/y x)
  if (c =z 2) then (b/x y)
  else (r(M)/r(2)) rmin(a;b)
  fi 

Lemma: real-ratio-bound_wf
[M:ℕ+]. ∀[x,y:ℝ]. ∀[a,b:{r:ℝr0 < r} ].
  (real-ratio-bound(M;x;y;a;b) ∈ {r:ℝ((x < y)  (r ≤ (a/y x))) ∧ ((y < x)  (r ≤ (b/x y))) ∧ (r0 < r)} )

Lemma: real-ratio-bound-cases
M:ℕ+. ∀x,y:ℝ.
  ∀[a,b:{r:ℝr0 < r} ].
    ((x < y) ∧ (real-ratio-bound(M;x;y;a;b) (a/y x))) ∨ ((y < x) ∧ (real-ratio-bound(M;x;y;a;b) (b/x y))) 
    supposing (r(2)/r(M)) ≤ |x y|

Lemma: nearby-cases
n:ℕ+. ∀x,y:ℝ.  ((x < y) ∨ (y < x) ∨ (|x y| ≤ (r1/r(n))))

Lemma: nearby-cases-ext
n:ℕ+. ∀x,y:ℝ.  ((x < y) ∨ (y < x) ∨ (|x y| ≤ (r1/r(n))))

Lemma: by-nearby-cases
[P:ℝ ⟶ ℝ ⟶ ℙ]
  ∀n:ℕ+. ∀x:ℝ.
    ((∀y:{y:ℝx < y} P[x;y])
     (∀y:{y:ℝy < x} P[x;y])
     (∀y:{y:ℝ|x y| ≤ (r1/r(n))} P[x;y])
     (∀y:ℝP[x;y]))

Lemma: by-nearby-cases-ext
[P:ℝ ⟶ ℝ ⟶ ℙ]
  ∀n:ℕ+. ∀x:ℝ.
    ((∀y:{y:ℝx < y} P[x;y])
     (∀y:{y:ℝy < x} P[x;y])
     (∀y:{y:ℝ|x y| ≤ (r1/r(n))} P[x;y])
     (∀y:ℝP[x;y]))

Lemma: int-rdiv-cancel
a:ℤ. ∀n,m:ℕ+.  ((r(m a))/m (r(a))/n ∈ ℝ)

Lemma: near-inverse-of-increasing-function
f:ℝ ⟶ ℝ. ∀n,M:ℕ+. ∀z:ℝ. ∀a,b:ℤ.
  ∀k:ℕ+
    (∃c:ℤ(∃j:ℕ+ [((|f[(r(c))/j] z| ≤ (r1/r(n))) ∧ ((r(a))/k ≤ (r(c))/j) ∧ ((r(c))/j ≤ (r(b))/k))])) supposing 
       ((z ≤ f[(r(b))/k]) and 
       (f[(r(a))/k] ≤ z) and 
       (∀x,y:ℝ.
          (((r(a))/k ≤ x)
           (x < y)
           (y ≤ (r(b))/k)
           ((f[x] ≤ f[y]) ∧ (((y x) ≤ (r1/r(M)))  ((f[y] f[x]) ≤ (r1/r(n)))))))) 
  supposing a < b

Lemma: near-inverse-of-increasing-function-ext
f:ℝ ⟶ ℝ. ∀n,M:ℕ+. ∀z:ℝ. ∀a,b:ℤ.
  ∀k:ℕ+
    (∃c:ℤ(∃j:ℕ+ [((|f[(r(c))/j] z| ≤ (r1/r(n))) ∧ ((r(a))/k ≤ (r(c))/j) ∧ ((r(c))/j ≤ (r(b))/k))])) supposing 
       ((z ≤ f[(r(b))/k]) and 
       (f[(r(a))/k] ≤ z) and 
       (∀x,y:ℝ.
          (((r(a))/k ≤ x)
           (x < y)
           (y ≤ (r(b))/k)
           ((f[x] ≤ f[y]) ∧ (((y x) ≤ (r1/r(M)))  ((f[y] f[x]) ≤ (r1/r(n)))))))) 
  supposing a < b

Lemma: rbetween-convex
x,a,b:ℝ.  ((a < b)  a≤x≤ (∃t:ℝ((r0 ≤ t) ∧ (t ≤ r1) ∧ (x ((t a) ((r1 t) b))))))

Lemma: rnexp-is-positive
i:ℕ+. ∀x:ℝ.  ((r0 < |x^i|)  (r0 < |x|))

Lemma: square-rless-implies
x,y:ℝ.  ((r0 ≤ y)  (x^2 < y^2)  (x < y))

Lemma: rnexp2-positive-iff
x:ℝ(r0 < x^2 ⇐⇒ x ≠ r0)

Lemma: square-positive-iff
x:ℝ(r0 < (x x) ⇐⇒ x ≠ r0)

Lemma: rmul-nonneg
[x,y:ℝ].  r0 ≤ (x y) supposing ((r0 ≤ x) ∧ (r0 ≤ y)) ∨ ((x ≤ r0) ∧ (y ≤ r0))

Lemma: rmul-nonneg-case1
[x,y:ℝ].  r0 ≤ (x y) supposing (r0 ≤ x) ∧ (r0 ≤ y)

Lemma: rmul-nonzero
x,y:ℝ.  (x y ≠ r0 ⇐⇒ x ≠ r0 ∧ y ≠ r0)

Lemma: rdiv-nonzero
x,y:ℝ.  (y ≠ r0  ((x/y) ≠ r0 ⇐⇒ x ≠ r0))

Lemma: square-nonzero
x:ℝ(x x ≠ r0 ⇐⇒ x ≠ r0)

Lemma: square-is-zero
x:ℝ((x x) r0 ⇐⇒ r0)

Lemma: square-rless-1-iff
x:ℝ(x^2 < r1 ⇐⇒ |x| < r1)

Lemma: square-rleq-1-iff
x:ℝ(x^2 ≤ r1 ⇐⇒ |x| ≤ r1)

Lemma: square-rge-1-iff
x:ℝ(r1 ≤ x^2 ⇐⇒ r1 ≤ |x|)

Lemma: square-is-one
x:ℝ(x^2 r1 ⇐⇒ (x r1) ∨ (x -(r1)))

Lemma: square-req-1-iff
x:ℝ(x ≠ -(r1)  (x^2 r1 ⇐⇒ r1))

Lemma: imonomial-nonneg-lemma
[k:ℕ+]. ∀[m,m':iMonomial()].
  ∀f:ℤ ⟶ ℝ(r0 ≤ real_term_value(f;imonomial-term(m))) supposing mul-monomials(m';m') mul-monomials(m;<k, []>) ∈ iMo\000Cnomial()

Lemma: imonomial-nonneg
[m:iMonomial()]. ∀f:ℤ ⟶ ℝ(r0 ≤ real_term_value(f;imonomial-term(m))) supposing ↑nonneg-monomial(m)

Lemma: ipolynomial-nonneg
[p:iPolynomial()]. ∀f:ℤ ⟶ ℝ(r0 ≤ real_term_value(f;ipolynomial-term(p))) supposing ↑nonneg-poly(p)

Lemma: real-term-nonneg
[t:int_term()]. ∀f:ℤ ⟶ ℝ(r0 ≤ real_term_value(f;t)) supposing ↑nonneg-poly(int_term_to_ipoly(t))

Lemma: rationals-dense
x:ℝ. ∀y:{y:ℝx < y} .  ∃n:ℕ+. ∃m:ℤ((x < (r(m)/r(n))) ∧ ((r(m)/r(n)) < y))

Lemma: rationals-dense-ext
x:ℝ. ∀y:{y:ℝx < y} .  ∃n:ℕ+. ∃m:ℤ((x < (r(m)/r(n))) ∧ ((r(m)/r(n)) < y))

Lemma: small-reciprocal-real
x:{x:ℝr0 < x} . ∃k:ℕ+((r1/r(k)) < x)

Lemma: small-reciprocal-real-ext
x:{x:ℝr0 < x} . ∃k:ℕ+((r1/r(k)) < x)

Lemma: small-reciprocal-rneq-zero
x:ℝ(x ≠ r0  (∃k:ℕ+((r1/r(k)) < |x|)))

Definition: positive-lower-bound
positive-lower-bound(x) ==  rlessw(r0;x) 1

Lemma: positive-lower-bound_wf
[x:{x:ℝr0 < x} ]. (positive-lower-bound(x) ∈ {k:ℕ+(r1/r(k)) < x} )

Definition: integer-approx
integer-approx(x;k) ==  (x k) ÷ k

Lemma: integer-approx_wf
[x:ℝ]. ∀[k:ℕ+].  (integer-approx(x;k) ∈ {n:ℤ|x r(n)| ≤ (r1 (r1/r(k)))} )

Definition: reduce-real
reduce-real(x;b;k) ==  integer-approx((x/b);k)

Lemma: reduce-real_wf
[k:ℕ+]. ∀[x:ℝ]. ∀[b:{b:ℝr0 < b} ].  (reduce-real(x;b;k) ∈ {n:ℤ|x r(n) b| ≤ (b (b/r(k)))} )

Lemma: implies-regular
[k:ℕ+]. ∀[x:ℕ+ ⟶ ℤ].
  k-regular-seq(x) supposing ∀n,m:ℕ+.  (|(x within 1/n) (x within 1/m)| ≤ ((r(k)/r(n)) (r(k)/r(m))))

Lemma: implies-real
[x:ℕ+ ⟶ ℤ]. x ∈ ℝ supposing ∀n,m:ℕ+.  (|(x within 1/n) (x within 1/m)| ≤ ((r1/r(n)) (r1/r(m))))

Definition: converges-to
lim n→∞.x[n] ==  ∀k:ℕ+(∃N:ℕ [(∀n:ℕ((N ≤ n)  (|x[n] y| ≤ (r1/r(k)))))])

Lemma: converges-to_wf
[x:ℕ ⟶ ℝ]. ∀[y:ℝ].  (lim n→∞.x[n] y ∈ ℙ)

Lemma: rational-approx-converges-to
[x:ℝ]. lim n→∞.(x within 1/n 1) x

Lemma: rinv-converges-to-0
lim n→∞.(r1/r(n 1)) r0

Lemma: converges-to_functionality
x1,x2:ℕ ⟶ ℝ. ∀y1,y2:ℝ.  ({lim n→∞.x1[n] y1  lim n→∞.x2[n] y2}) supposing ((y1 y2) and (∀n:ℕ(x1[n] x2[n])))

Lemma: converges-to_functionality2
x1,x2:ℕ ⟶ ℝ. ∀y1,y2:ℝ.  (lim n→∞.x1[n] y1  lim n→∞.x2[n] y2) supposing ((y1 y2) and (∀n:ℕ(x1[n] x2[n])))

Definition: converges
x[n]↓ as n→∞ ==  ∃y:ℝlim n→∞.x[n] y

Lemma: converges_wf
[x:ℕ ⟶ ℝ]. (x[n]↓ as n→∞ ∈ ℙ)

Definition: diverges
n.x[n]↑ ==  ∃e:ℝ((r0 < e) ∧ (∀k:ℕ. ∃m,n:ℕ((k ≤ m) ∧ (k ≤ n) ∧ (e ≤ |x[m] x[n]|))))

Lemma: diverges_wf
[x:ℕ ⟶ ℝ]. (n.x[n]↑ ∈ ℙ)

Lemma: infinitesmal-zero
[x:ℝ]. uiff(x r0;∀[k:ℕ+]. (|x| ≤ (r1/r(k))))

Lemma: infinitesmal-difference
[x,y:ℝ].  uiff(x y;∀[k:ℕ+]. (|x y| ≤ (r1/r(k))))

Lemma: unique-limit
[x:ℕ ⟶ ℝ]. ∀[y1,y2:ℝ].  (y1 y2) supposing (lim n→∞.x[n] y2 and lim n→∞.x[n] y1)

Definition: bounded-sequence
bounded-sequence(n.x[n]) ==  ∃b:ℝ. ∀n:ℕ(|x[n]| ≤ b)

Lemma: bounded-sequence_wf
[x:ℕ ⟶ ℝ]. (bounded-sequence(n.x[n]) ∈ ℙ)

Lemma: converges-implies-bounded
x:ℕ ⟶ ℝ(x[n]↓ as n→∞  bounded-sequence(n.x[n]))

Lemma: integer-bound
x:ℝ. ∃n:ℕ+(|x| ≤ r(n))

Definition: r-bound
r-bound(x) ==  fst((TERMOF{integer-bound:o, 1:l} x))

Lemma: r-bound_wf
[x:ℝ]. (r-bound(x) ∈ ℕ+)

Lemma: r-bound-property
[x:ℝ]. ((r(-r-bound(x)) ≤ x) ∧ (x ≤ r(r-bound(x))))

Lemma: r-strict-bound-property
x:ℝ((r(-(r-bound(x) 1)) < x) ∧ (x < r(r-bound(x) 1)))

Definition: cauchy
cauchy(n.x[n]) ==  ∀k:ℕ+(∃N:ℕ [(∀n,m:ℕ.  ((N ≤ n)  (N ≤ m)  (|x[n] x[m]| ≤ (r1/r(k)))))])

Lemma: cauchy_wf
[x:ℕ ⟶ ℝ]. (cauchy(n.x[n]) ∈ ℙ)

Lemma: converges-iff-cauchy
x:ℕ ⟶ ℝ(x[n]↓ as n→∞ ⇐⇒ cauchy(n.x[n]))

Lemma: converges-iff-cauchy-ext
x:ℕ ⟶ ℝ(x[n]↓ as n→∞ ⇐⇒ cauchy(n.x[n]))

Definition: cauchy-limit
cauchy-limit(n.x[n];c) ==  accelerate(2;λn.(x[c n] n))

Lemma: cauchy-limit_wf
[x:ℕ ⟶ ℝ]. ∀[c:cauchy(n.x[n])].  (cauchy-limit(n.x[n];c) ∈ ℝ)

Lemma: converges-to-cauchy-limit
x:ℕ ⟶ ℝ. ∀c:cauchy(n.x[n]).  lim n→∞.x[n] cauchy-limit(n.x[n];c)

Lemma: converges-cauchy-witness
[x:ℕ ⟶ ℝ]. ∀[y:ℝ]. ∀[cvg:lim n→∞.x[n] y].  k.(cvg (2 k)) ∈ cauchy(n.x[n]))

Lemma: req-from-converges
[x:ℕ ⟶ ℝ]. ∀[y:ℝ]. ∀[cvg:lim n→∞.x[n] y].  (y cauchy-limit(n.x[n];λk.(cvg (2 k))))

Definition: real-from-approx
real-from-approx(n.x[n]) ==  cauchy-limit(n.x[n 1];λk.(2 k))

Lemma: real-from-approx_wf
[a:ℝ]. ∀[x:k:ℕ+ ⟶ {v:ℝ|v a| ≤ (r1/r(k))} ].  (real-from-approx(n.x[n]) ∈ {b:ℝa} )

Definition: almost-positive
almost-positive(x) ==  ¬¬(r0 < x)

Lemma: almost-positive_wf
[x:ℝ]. (almost-positive(x) ∈ ℙ)

Definition: pseudo-positive
pseudo-positive(x) ==  ∀y:ℝ((¬¬(r0 < y)) ∨ (¬¬(y < x)))

Lemma: pseudo-positive_wf
[x:ℝ]. (pseudo-positive(x) ∈ ℙ)

Lemma: pseudo-positive-almost-positive
[x:ℝ]. almost-positive(x) supposing pseudo-positive(x)

Lemma: pseudo-positive-iff
x:ℝ((r0 ≤ x)  (pseudo-positive(x) ⇐⇒ ∀y:ℝ((¬(x y)) ∨ (y r0)))))

Lemma: limit-shift
m:ℕ. ∀X:ℕ ⟶ ℝ. ∀a:ℝ.  (lim n→∞.X[n]  lim n→∞.X[n m] a)

Lemma: limit-shift-iff
m:ℕ. ∀X:ℕ ⟶ ℝ. ∀a:ℝ.  (lim n→∞.X[n] ⇐⇒ lim n→∞.X[n m] a)

Lemma: not-diverges-converges
[x:ℕ ⟶ ℝ]. (x[n]↓ as n→∞ ∧ n.x[n]↑))

Lemma: subsequence-converges
a:ℝ. ∀x,y:ℕ ⟶ ℝ.
  ((∃N:ℕ. ∀n:ℕ. ∃m:ℕ((n ≤ m) ∧ (y[n] x[m])) supposing N ≤ n)  lim n→∞.x[n]  lim n→∞.y[n] a)

Definition: converges-to-infinity
lim n →∞.x[n] = ∞ ==  ∀k:ℕ+. ∀large(n).r(k) ≤ x[n]

Lemma: converges-to-infinity_wf
[x:ℕ ⟶ ℝ]. (lim n →∞.x[n] = ∞ ∈ ℙ)

Lemma: rnexp-converges
x:ℝ((|x| < r1)  lim n→∞.x^n r0)

Lemma: rinv-exp-converges
M:ℕ+. ∀N:{2...}.  lim n→∞.(r1/r(M N^n)) r0

Lemma: rnexp-converges-ext
x:ℝ((|x| < r1)  lim n→∞.x^n r0)

Lemma: rinv-exp-converges-ext
M:ℕ+. ∀N:{2...}.  lim n→∞.(r1/r(M N^n)) r0

Lemma: rpowers-converge
x:ℝ(((|x| < r1)  lim n→∞.x^n r0) ∧ ((r1 < x)  lim n →∞.x^n = ∞))

Lemma: rpowers-converge-ext
x:ℝ(((|x| < r1)  lim n→∞.x^n r0) ∧ ((r1 < x)  lim n →∞.x^n = ∞))

Lemma: r-archimedean
x:ℝ. ∃n:ℕ((r(-n) ≤ x) ∧ (x ≤ r(n)))

Lemma: simple-converges-to
x:ℕ ⟶ ℝ. ∀a,c:ℝ.  ((∀n:ℕ(|(x n) a| ≤ ((r1/r(2^n)) c)))  lim n→∞.x a)

Lemma: common-limit-squeeze
a,b,c:ℕ ⟶ ℝ.
  ((∀n:ℕ((a[n] ≤ a[n 1]) ∧ (a[n 1] ≤ b[n 1]) ∧ (b[n 1] ≤ b[n])))
   lim n→∞.c[n] r0
   (∀n:ℕr0≤b[n] a[n]≤c[n])
   (∃y:ℝ(lim n→∞.a[n] y ∧ lim n→∞.b[n] y)))

Lemma: common-limit-squeeze-ext
a,b,c:ℕ ⟶ ℝ.
  ((∀n:ℕ((a[n] ≤ a[n 1]) ∧ (a[n 1] ≤ b[n 1]) ∧ (b[n 1] ≤ b[n])))
   lim n→∞.c[n] r0
   (∀n:ℕr0≤b[n] a[n]≤c[n])
   (∃y:ℝ(lim n→∞.a[n] y ∧ lim n→∞.b[n] y)))

Lemma: common-limit-midpoints
a,b:ℕ ⟶ ℝ.
  ((∀n:ℕ
      (((a[n 1] a[n]) ∧ (b[n 1] (a[n] b[n]/r(2)))) ∨ ((a[n 1] (a[n] b[n]/r(2))) ∧ (b[n 1] b[n]))))
   (∃y:ℝ(lim n→∞.a[n] y ∧ lim n→∞.b[n] y)))

Definition: rsum'
rsum'(n;m;k.x[k]) ==
  eval n' in
  eval m' in
  eval (m' n') in
    if (k) < (1)  then r0  else eval k2 in λj.eval k2 in Σ(x[n i] i < k) ÷ k2

Definition: rsum
Σ{x[k] n≤k≤m} ==  eval n' in eval m' in   let xs ⟵ map(λk.x[k];[n', m' 1)) in radd-list(xs)

Lemma: rsum_wf
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  {x[k] n≤k≤m} ∈ ℝ)

Lemma: rsum'-eq-rsum
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  (rsum'(n;m;k.x[k]) = Σ{x[k] n≤k≤m} ∈ (ℕ+ ⟶ ℤ))

Lemma: rsum'_wf
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  (rsum'(n;m;k.x[k]) ∈ ℝ)

Lemma: rsum'-rsum
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  {x[k] n≤k≤m} rsum'(n;m;k.x[k]))

Lemma: rsum_unroll
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].
  {x[k] n≤k≤m} if m <then r0 if (m =z n) then x[n] else Σ{x[k] n≤k≤1} x[m] fi )

Lemma: rsum_single
[n:ℤ]. ∀[x:{n..n 1-} ⟶ ℝ].  {x[k] n≤k≤n} x[n])

Lemma: rsum-telescopes
[n:ℤ]. ∀[m:{n...}]. ∀[x,y:{n..m 1-} ⟶ ℝ].
  Σ{x[k] y[k] n≤k≤m} (x[m] y[n]) supposing ∀i:{n..m-}. (y[i 1] x[i])

Definition: rpolynomial
i≤n. a_i x^i) ==  Σ{(a i) x^i 0≤i≤n}

Lemma: rpolynomial_wf
[n:ℕ]. ∀[a:ℕ1 ⟶ ℝ]. ∀[x:ℝ].  ((Σi≤n. a_i x^i) ∈ ℝ)

Lemma: rpolynomial_unroll
[n:ℕ]. ∀[a:ℕ1 ⟶ ℝ]. ∀[x:ℝ].
  ((Σi≤n. a_i x^i) if (n =z 0) then else ((a n) x^n) i≤1. a_i x^i) fi )

Definition: pointwise-req
x[k] y[k] for k ∈ [n,m] ==  ∀k:ℤ((n ≤ k)  (k ≤ m)  (x[k] y[k]))

Lemma: pointwise-req_wf
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].  (x[k] y[k] for k ∈ [n,m] ∈ ℙ)

Definition: pointwise-rleq
x[k] ≤ y[k] for k ∈ [n,m] ==  ∀k:ℤ((n ≤ k)  (k ≤ m)  (x[k] ≤ y[k]))

Lemma: pointwise-rleq_wf
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].  (x[k] ≤ y[k] for k ∈ [n,m] ∈ ℙ)

Lemma: rsum_functionality
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].  Σ{x[k] n≤k≤m} = Σ{y[k] n≤k≤m} supposing x[k] y[k] for k ∈ [n,m]

Lemma: rsum_functionality2
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].
  Σ{x[k] n≤k≤m} = Σ{y[k] n≤k≤m} supposing ∀k:ℤ((n ≤ k)  (k ≤ m)  (x[k] y[k]))

Lemma: rsum_functionality_wrt_rleq
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].  Σ{x[k] n≤k≤m} ≤ Σ{y[k] n≤k≤m} supposing x[k] ≤ y[k] for k ∈ [n,m]

Lemma: rsum_functionality_wrt_rleq2
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].
  Σ{x[k] n≤k≤m} ≤ Σ{y[k] n≤k≤m} supposing ∀k:ℤ((n ≤ k)  (k ≤ m)  (x[k] ≤ y[k]))

Lemma: rsum_functionality_wrt_rleq3
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].
  Σ{x[k] n≤k≤m} ≤ Σ{y[k] n≤k≤m} supposing ∀k:ℤ((n ≤ k)  (k ≤ m)  (y[k] ≥ x[k]))

Lemma: rpolynomial_functionality
[n:ℕ]. ∀[a,b:ℕ1 ⟶ ℝ]. ∀[x,y:ℝ].
  ((Σi≤n. a_i x^i) i≤n. b_i y^i)) supposing ((x y) and a[k] b[k] for k ∈ [0,n])

Lemma: rsum-telescopes2
[n:ℤ]. ∀[m:{n...}]. ∀[x,y:{n..m 1-} ⟶ ℝ].
  Σ{x[k] y[k] n≤k≤m} (x[n] y[m]) supposing ∀i:{n..m-}. (x[i 1] y[i])

Lemma: rsum-rewrite-test
Σ{r(i) 1≤i≤10} ≤ Σ{r(1 i) 1≤i≤10}

Lemma: rsum-split
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ]. ∀[k:ℤ].
  {x[i] n≤i≤m} {x[i] n≤i≤k} + Σ{x[i] 1≤i≤m})) supposing ((k ≤ m) and (n ≤ k))

Lemma: rsum-split2
[n,m:ℤ]. ∀[k:{n..m 1-}]. ∀[x:{n..m 1-} ⟶ ℝ].  {x[i] n≤i≤m} {x[i] n≤i≤k} + Σ{x[i] 1≤i≤m}))

Lemma: rsum-shift
[k,n,m:ℤ]. ∀[x:Top].  {x[i] n≤i≤m} ~ Σ{x[i k] k≤i≤k})

Lemma: rsum-split-shift
[k,n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].
  {x[i] n≤i≤m} {x[i] n≤i≤k} + Σ{x[k 1] 0≤i≤1})) supposing ((k ≤ m) and (n ≤ k))

Lemma: rsum-single
[n:ℤ]. ∀[x:{i:ℤn ∈ ℤ}  ⟶ ℝ].  {x[i] n≤i≤n} x[n])

Lemma: rsum-empty
[n,m:ℤ]. ∀[x:Top].  Σ{x[i] n≤i≤m} r0 supposing m < n

Lemma: rsum-split-first
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  Σ{x[i] n≤i≤m} (x[n] + Σ{x[i] 1≤i≤m}) supposing n ≤ m

Lemma: rsum-split-first-shift
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  Σ{x[i] n≤i≤m} (x[n] + Σ{x[i 1] n≤i≤1}) supposing n ≤ m

Lemma: rsum-split-last
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  Σ{x[i] n≤i≤m} {x[i] n≤i≤1} x[m]) supposing n ≤ m

Lemma: rsum-difference
[k,n,m:ℤ]. ∀[x:{k..m 1-} ⟶ ℝ].
  ((Σ{x[i] k≤i≤m} - Σ{x[i] k≤i≤n}) = Σ{x[i] 1≤i≤m}) supposing ((n ≤ m) and (k ≤ n))

Lemma: rsum-difference2
[k,n,m:ℤ]. ∀[x,y:{k..m 1-} ⟶ ℝ].
  ((Σ{x[i] k≤i≤m} - Σ{y[i] k≤i≤n}) = Σ{x[i] 1≤i≤m}) supposing 
     ((∀i:{k..n 1-}. (x[i] y[i])) and 
     (n ≤ m) and 
     (k ≤ n))

Lemma: rsum_linearity1
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].  {x[k] y[k] n≤k≤m} {x[k] n≤k≤m} + Σ{y[k] n≤k≤m}))

Lemma: rsum_linearity2
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ]. ∀[y:ℝ].  {y x[k] n≤k≤m} (y * Σ{x[k] n≤k≤m}))

Lemma: rsum_linearity3
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ]. ∀[y:ℝ].  {x[k] n≤k≤m} {x[k] n≤k≤m} y))

Lemma: rsum_linearity-rsub
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].  {x[k] y[k] n≤k≤m} {x[k] n≤k≤m} - Σ{y[k] n≤k≤m}))

Lemma: rsum-rminus
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  {-(x[k]) n≤k≤m} -(Σ{x[k] n≤k≤m}))

Lemma: rsum-constant
[n,m:ℤ]. ∀[a:ℝ].  {a n≤k≤m} (a * Σ{r1 n≤k≤m}))

Lemma: rsum-constant2
[n,m:ℤ]. ∀[a:ℝ].  {a n≤k≤m} (a if m <then r0 else r((m n) 1) fi ))

Lemma: rsum-zero
[n,m:ℤ].  {r0 n≤k≤m} r0)

Lemma: rsum-zero-req
[n,m:ℤ]. ∀[f:{n..m 1-} ⟶ ℝ].  Σ{f[k] n≤k≤m} r0 supposing ∀k:{n..m 1-}. (f[k] r0)

Lemma: rsum-one
[n,m:ℤ].  {r1 n≤k≤m} if m <then r0 else r((m n) 1) fi )

Lemma: rsum_product
[a,b,c,d:ℤ]. ∀[x:{a..b 1-} ⟶ ℝ]. ∀[y:{c..d 1-} ⟶ ℝ].
  ((Σ{x[i] a≤i≤b} * Σ{y[j] c≤j≤d}) = Σ{x[i] y[j] c≤j≤d} a≤i≤b})

Lemma: rsum_nonneg
[n,m:ℤ]. ∀[y:{n..m 1-} ⟶ ℝ].  r0 ≤ Σ{y[k] n≤k≤m} supposing r0 ≤ y[k] for k ∈ [n,m]

Lemma: rsum_int
[n:ℕ]. ∀[y:ℕn ⟶ ℤ].  {r(y[k]) 0≤k≤1} r(Σ(y[k] k < n)))

Lemma: rabs-rsum
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  (|Σ{x[i] n≤i≤m}| ≤ Σ{|x[i]| n≤i≤m})

Lemma: rsum-positive-implies
n,m:ℤ. ∀x:{n..m 1-} ⟶ ℝ.  ((r0 < Σ{x[i] n≤i≤m})  (∃i:{n..m 1-}. (r0 < |x[i]|)))

Lemma: rsum-of-nonneg-positive-iff
n,m:ℤ. ∀x:{n..m 1-} ⟶ ℝ.  ((∀i:{n..m 1-}. (r0 ≤ x[i]))  (r0 < Σ{x[i] n≤i≤m} ⇐⇒ ∃i:{n..m 1-}. (r0 < x[i])))

Lemma: rsum-of-nonneg-zero-iff
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].
  uiff(Σ{x[i] n≤i≤m} r0;∀i:{n..m 1-}. (x[i] r0)) supposing ∀i:{n..m 1-}. (r0 ≤ x[i])

Lemma: item-rleq-rsum-of-nonneg
n,m:ℤ. ∀x:{n..m 1-} ⟶ ℝ.  ((∀i:{n..m 1-}. (r0 ≤ x[i]))  (∀i:{n..m 1-}. (x[i] ≤ Σ{x[i] n≤i≤m})))

Lemma: rmul-rpolynomial
[n:ℕ]. ∀[a:ℕ1 ⟶ ℝ]. ∀[x,b:ℝ].  ((b i≤n. a_i x^i)) i≤n. λi.(b (a i))_i x^i))

Lemma: rpolynomial-rmul
[n:ℕ]. ∀[a:ℕ1 ⟶ ℝ]. ∀[x,b:ℝ].  (((Σi≤n. a_i x^i) b) i≤n. λi.((a i) b)_i x^i))

Lemma: rdiv-rpolynomial
[n:ℕ]. ∀[a:ℕ1 ⟶ ℝ]. ∀[x,b:ℝ].  ((Σi≤n. a_i x^i)/b) i≤n. λi.(a i/b)_i x^i) supposing b ≠ r0

Lemma: shift-rpolynomial
[n:ℕ]. ∀[a:ℕ1 ⟶ ℝ]. ∀[x:ℝ].
  ((x i≤n. a_i x^i)) i≤1. λi.if (i =z 0) then r0 else (i 1) fi _i x^i))

Lemma: add-rpolynomials
[n,m:ℕ]. ∀[a:ℕ1 ⟶ ℝ]. ∀[b:ℕ1 ⟶ ℝ]. ∀[x:ℝ].
  ((Σi≤n. a_i x^i) i≤m. b_i x^i)) i≤n. λi.if i ≤then (a i) (b i) else fi _i x^i) supposing m ≤ n

Lemma: subtract-rpolynomials
[n,m:ℕ]. ∀[a:ℕ1 ⟶ ℝ]. ∀[b:ℕ1 ⟶ ℝ]. ∀[x:ℝ].
  ((Σi≤n. a_i x^i) i≤m. b_i x^i)) i≤n. λi.if i ≤then (a i) else fi _i x^i) supposing m ≤ n

Lemma: add-rpolynomials-same-degree
[n:ℕ]. ∀[a,b:ℕ1 ⟶ ℝ]. ∀[x:ℝ].  (((Σi≤n. a_i x^i) i≤n. b_i x^i)) i≤n. λi.((a i) (b i))_i x^i))

Lemma: subtract-rpolynomials-same-degree
[n:ℕ]. ∀[a,b:ℕ1 ⟶ ℝ]. ∀[x:ℝ].  (((Σi≤n. a_i x^i) i≤n. b_i x^i)) i≤n. λi.((a i) i)_i x^i))

Definition: rpolydiv
rpolydiv(n;a;z) ==  λi.primrec(n i;a n;λj,r. ((a (n 1)) (z r)))

Lemma: rpolydiv_wf
[n:ℤ]. ∀[a:ℕ1 ⟶ ℝ]. ∀[z:ℝ].  (rpolydiv(n;a;z) ∈ ℕn ⟶ ℝ)

Lemma: rpolydiv-rec
[n:ℤ]. ∀[a,z:Top].
  ((rpolydiv(n;a;z) (n 1) n) ∧ (∀i:ℕ1. (rpolydiv(n;a;z) (a (i 1)) (z (rpolydiv(n;a;z) (i 1))))))

Lemma: rpolydiv-property
[n:ℕ+]. ∀[a:ℕ1 ⟶ ℝ]. ∀[z,x:ℝ].
  ((Σi≤n. a_i x^i) (((x z) i≤1. rpolydiv(n;a;z)_i x^i)) i≤n. a_i z^i)))

Lemma: rpolynomial-linear-factor
n:ℕ+. ∀a:ℕ1 ⟶ ℝ. ∀z:ℝ.
  ∃b:ℕn ⟶ ℝ((∀[x:ℝ]. ((Σi≤n. a_i x^i) ((x z) i≤1. b_i x^i)))) ∧ ((b (n 1)) (a n))) 
  supposing i≤n. a_i z^i) r0

Definition: rprod
rprod(n;m;k.x[k]) ==  if m <then r1 else rprod(n;m 1;k.x[k]) x[m] fi 

Lemma: rprod_wf
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  (rprod(n;m;k.x[k]) ∈ ℝ)

Lemma: rprod-empty
[n,m:ℤ]. ∀[x:Top].  rprod(n;m;k.x[k]) r1 supposing m < n

Lemma: rprod_functionality
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].  rprod(n;m;k.x[k]) rprod(n;m;k.y[k]) supposing x[k] y[k] for k ∈ [n,m]

Lemma: rprod-single
[n:ℤ]. ∀[x:{n..n 1-} ⟶ ℝ].  (rprod(n;n;k.x[k]) x[n])

Lemma: rprod-of-positive
n,m:ℤ. ∀x:{n..m 1-} ⟶ ℝ.  r0 < rprod(n;m;k.x[k]) supposing ∀k:{n..m 1-}. (r0 < x[k])

Lemma: rprod-rminus
n,m:ℤ. ∀x:{n..m 1-} ⟶ ℝ.  rprod(n;m;k.-(x[k])) (r(-1)^(m n) rprod(n;m;k.x[k])) supposing n ≤ m

Lemma: rprod-of-negative
n,m:ℤ. ∀x:{n..m 1-} ⟶ ℝ.
  (((m rem 2) 1 ∈ ℤ (r0 < rprod(n;m;k.x[k]))) ∧ (((m rem 2) 0 ∈ ℤ (rprod(n;m;k.x[k]) < r0)) 
  supposing (∀k:{n..m 1-}. (x[k] < r0)) ∧ (n ≤ m)

Lemma: rprod-rsub-symmetry
n,m:ℤ. ∀x,y:{n..m 1-} ⟶ ℝ.
  rprod(n;m;k.x[k] y[k]) (r(-1)^(m n) rprod(n;m;k.y[k] x[k])) supposing n ≤ m

Lemma: rprod-split
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ]. ∀[i:ℤ].
  rprod(n;m;k.x[k]) (rprod(n;i;k.x[k]) rprod(i 1;m;k.x[k])) supposing (i ≤ m) ∧ (n ≤ (i 1))

Lemma: rprod-split-last
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  rprod(n;m;k.x[k]) (rprod(n;m 1;k.x[k]) x[m]) supposing n ≤ m

Lemma: rprod-split-first
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  rprod(n;m;k.x[k]) (x[n] rprod(n 1;m;k.x[k])) supposing n ≤ m

Lemma: rprod-is-zero
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  rprod(n;m;k.x[k]) r0 supposing ∃k:{n..m 1-}. (x[k] r0)

Lemma: rprod-as-itop
[n,m:ℤ]. ∀[x:Top].  x,y. (x y),r1) n ≤ k < m. x[k] rprod(n;m 1;k.x[k]))

Lemma: rsum-as-itop
[n,m:ℤ]. ∀[x:{n..m-} ⟶ ℝ].  x,y. (x y),r0) n ≤ k < m. x[k] = Σ{x[k] n≤k≤1})

Definition: r-list-sum
r-list-sum(L) ==  reduce(λx,y. (x y);r0;L)

Lemma: r-list-sum_wf
[L:ℝ List]. (r-list-sum(L) ∈ ℝ)

Lemma: r-list-sum_functionality
[L1,L2:ℝ List].  r-list-sum(L1) r-list-sum(L2) supposing (||L1|| ||L2|| ∈ ℤ) ∧ (∀i:ℕ||L1||. (L1[i] L2[i]))

Lemma: rpolynomial-complete-factors
n:ℕ+. ∀a:ℕ1 ⟶ ℝ. ∀z:ℕn ⟶ ℝ.
  ((∀i,j:ℕn.  ((¬(i j ∈ ℤ))  i ≠ j))
   ∀[x:ℝ]. ((Σi≤n. a_i x^i) ((a n) rprod(0;n 1;j.x j))) supposing ∀j:ℕn. ((Σi≤n. a_i j^i) r0))

Lemma: rpolynomial-complete-factors-ordered
n:ℕ+. ∀a:ℕ1 ⟶ ℝ. ∀z:ℕn ⟶ ℝ.
  ((∀j:ℕ1. ((z j) < (z (j 1))))
   ∀[x:ℝ]. ((Σi≤n. a_i x^i) ((a n) rprod(0;n 1;j.x j))) supposing ∀j:ℕn. ((Σi≤n. a_i j^i) r0))

Lemma: rpolynomial-complete-roots-unique
[n:ℕ+]. ∀[a:ℕ1 ⟶ ℝ].
  ∀[z,y:ℕn ⟶ ℝ].
    (∀[j:ℕn]. ((z j) (y j))) supposing 
       ((∀j:ℕn. ((Σi≤n. a_i j^i) r0)) and 
       (∀j:ℕn. ((Σi≤n. a_i j^i) r0)) and 
       (∀j:ℕ1. ((y j) < (y (j 1)))) and 
       (∀j:ℕ1. ((z j) < (z (j 1))))) 
  supposing n ≠ r0

Lemma: exists-rneq-iff
n:ℕ. ∀a,b:ℕn ⟶ ℝ.  (∃i:ℕn. a[i] ≠ b[i] ⇐⇒ r0 < Σ{|a[i] b[i]| 0≤i≤1})

Lemma: sq_stable_ex_rneq
n:ℕ. ∀a,b:ℕn ⟶ ℝ.  SqStable(∃i:ℕn. a[i] ≠ b[i])

Lemma: sq_stable_double_ex_rneq
m,n:ℕ. ∀a,b:ℕm ⟶ ℕn ⟶ ℝ.  SqStable(∃i:ℕm. ∃j:ℕn. a[i;j] ≠ b[i;j])

Lemma: sq_stable_ex_nonzero
n:ℕ. ∀a:ℕn ⟶ ℝ.  SqStable(∃i:ℕn. a[i] ≠ r0)

Lemma: rsum-triangle-inequality1
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].  ((Σ{|x[i]| n≤i≤m} - Σ{|y[i]| n≤i≤m}) ≤ Σ{|x[i] y[i]| n≤i≤m})

Lemma: rsum-triangle-inequality2
[n,m:ℤ]. ∀[x,y:{n..m 1-} ⟶ ℝ].  ((Σ{|y[i]| n≤i≤m} - Σ{|x[i]| n≤i≤m}) ≤ Σ{|x[i] y[i]| n≤i≤m})

Lemma: real-binomial
[n:ℕ]. ∀[a,b:ℝ].  (a b^n = Σ{r(choose(n;i)) a^n b^i 0≤i≤n})

Lemma: rpolynomial-composition1
n:ℕ. ∀a:ℕ1 ⟶ ℝ. ∀b,c,d:ℝ.  ∃a':ℕ1 ⟶ ℝ. ∀x:ℝ((Σi≤n. a'_i x^i) ((Σi≤n. a_i ((c b) x) b^i) d))

Definition: rmaximum
rmaximum(n;m;k.x[k]) ==  primrec(m n;x[n];λi,s. rmax(s;x[n 1]))

Lemma: rmaximum_wf
[n,m:ℤ].  ∀[x:{n..m 1-} ⟶ ℝ]. (rmaximum(n;m;k.x[k]) ∈ ℝsupposing n ≤ m

Lemma: rmaximum_functionality
[n,m:ℤ].
  ∀[x,y:{n..m 1-} ⟶ ℝ].
    rmaximum(n;m;k.x[k]) rmaximum(n;m;k.y[k]) supposing ∀k:ℤ((n ≤ k)  (k ≤ m)  (x[k] y[k])) 
  supposing n ≤ m

Lemma: rmaximum_functionality_wrt_rleq
[n,m:ℤ].
  ∀[x,y:{n..m 1-} ⟶ ℝ].
    rmaximum(n;m;k.x[k]) ≤ rmaximum(n;m;k.y[k]) supposing ∀k:ℤ((n ≤ k)  (k ≤ m)  (x[k] ≤ y[k])) 
  supposing n ≤ m

Lemma: rmaximum-split
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ]. ∀[k:ℤ].
  (rmaximum(n;m;i.x[i]) rmax(rmaximum(n;k;i.x[i]);rmaximum(k 1;m;i.x[i]))) supposing (k < and (n ≤ k))

Lemma: rmaximum-shift
[k,n,m:ℤ]. ∀[x:Top].  rmaximum(n;m;i.x[i]) rmaximum(n k;m k;i.x[i k]) supposing n ≤ m

Lemma: rmaximum_ub
[k,n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ].  (x[k] ≤ rmaximum(n;m;i.x[i])) supposing ((k ≤ m) and (n ≤ k))

Lemma: rmaximum-lub
n,m:ℤ. ∀x:{n..m 1-} ⟶ ℝ. ∀r:ℝ.  (∀k:{n..m 1-}. (x[k] ≤ r))  (rmaximum(n;m;i.x[i]) ≤ r) supposing n ≤ m

Lemma: rmaximum-select
n,m:ℤ.  ∀x:{n..m 1-} ⟶ ℝ. ∀e:ℝ.  ((r0 < e)  (∃i:{n..m 1-}. ((rmaximum(n;m;i.x[i]) e) < x[i]))) supposing n ≤ m

Lemma: rmaximum-constant
[n,m:ℤ].  ∀[x:{n..m 1-} ⟶ ℝ]. ∀[r:ℝ].  rmaximum(n;m;i.x[i]) supposing ∀i:{n..m 1-}. (x[i] r) supposing n ≤ m

Definition: rminimum
rminimum(n;m;k.x[k]) ==  primrec(m n;x[n];λi,s. rmin(s;x[n 1]))

Lemma: rminimum_wf
[n,m:ℤ].  ∀[x:{n..m 1-} ⟶ ℝ]. (rminimum(n;m;k.x[k]) ∈ ℝsupposing n ≤ m

Lemma: rminimum_functionality
[n,m:ℤ].
  ∀[x,y:{n..m 1-} ⟶ ℝ].
    rminimum(n;m;k.x[k]) rminimum(n;m;k.y[k]) supposing ∀k:ℤ((n ≤ k)  (k ≤ m)  (x[k] y[k])) 
  supposing n ≤ m

Lemma: rminimum_functionality_wrt_rleq
[n,m:ℤ].
  ∀[x,y:{n..m 1-} ⟶ ℝ].
    rminimum(n;m;k.x[k]) ≤ rminimum(n;m;k.y[k]) supposing ∀k:ℤ((n ≤ k)  (k ≤ m)  (x[k] ≤ y[k])) 
  supposing n ≤ m

Lemma: rminimum-split
[n,m:ℤ]. ∀[x:{n..m 1-} ⟶ ℝ]. ∀[k:ℤ].
  (rminimum(n;m;i.x[i]) rmin(rminimum(n;k;i.x[i]);rminimum(k 1;m;i.x[i]))) supposing (k < and (n ≤ k))

Lemma: rminimum-shift
[k,n,m:ℤ]. ∀[x:Top].  rminimum(n;m;i.x[i]) rminimum(n k;m k;i.x[i k]) supposing n ≤ m

Lemma: rminimum_lb
[k,n,m:ℤ].  (∀[x:{n..m 1-} ⟶ ℝ]. (rminimum(n;m;i.x[i]) ≤ x[k])) supposing ((k ≤ m) and (n ≤ k))

Lemma: rminimum-glb
n,m:ℤ. ∀x:{n..m 1-} ⟶ ℝ. ∀r:ℝ.  (∀k:{n..m 1-}. (r ≤ x[k]))  (r ≤ rminimum(n;m;i.x[i])) supposing n ≤ m

Lemma: rminimum-cases
n,m:ℤ.
  ∀x:{n..m 1-} ⟶ ℝ(¬¬(∃a:{n..m 1-}. ((rminimum(n;m;i.x[i]) x[a]) ∧ (∀j:{n..m 1-}. (x[a] ≤ x[j]))))) 
  supposing n ≤ m

Lemma: rminimum-select
n,m:ℤ.  ∀x:{n..m 1-} ⟶ ℝ. ∀e:ℝ.  ((r0 < e)  (∃i:{n..m 1-}. (x[i] < (rminimum(n;m;i.x[i]) e)))) supposing n ≤ m

Lemma: rminimum-positive
n,m:ℤ.  ∀x:{n..m 1-} ⟶ ℝ(r0 < rminimum(n;m;i.x[i]) ⇐⇒ ∀i:{n..m 1-}. (r0 < x[i])) supposing n ≤ m

Lemma: rminimum-constant
[n,m:ℤ].  ∀[x:{n..m 1-} ⟶ ℝ]. ∀[r:ℝ].  rminimum(n;m;i.x[i]) supposing ∀i:{n..m 1-}. (x[i] r) supposing n ≤ m

Lemma: Cauchy-Schwarz1
[n:ℕ]. ∀[x,y:ℕ1 ⟶ ℝ].
  ((Σ{x[i] y[i] 0≤i≤n} * Σ{x[i] y[i] 0≤i≤n}) ≤ {x[i] x[i] 0≤i≤n} * Σ{y[i] y[i] 0≤i≤n}))

Lemma: Cauchy-Schwarz1-strict
n:ℕ. ∀x,y:ℕ1 ⟶ ℝ.
  ((∃i,j:ℕ1. x[j] y[i] ≠ x[i] y[j])
   ((Σ{x[i] y[i] 0≤i≤n} * Σ{x[i] y[i] 0≤i≤n}) < {x[i] x[i] 0≤i≤n} * Σ{y[i] y[i] 0≤i≤n})))

Lemma: Cauchy-Schwarz1-strict-iff
n:ℕ. ∀x,y:ℕ1 ⟶ ℝ.
  (∃i,j:ℕ1. x[j] y[i] ≠ x[i] y[j]
  ⇐⇒ {x[i] y[i] 0≤i≤n} * Σ{x[i] y[i] 0≤i≤n}) < {x[i] x[i] 0≤i≤n} * Σ{y[i] y[i] 0≤i≤n}))

Lemma: Cauchy-Schwarz2
[n:ℕ]. ∀[x,y:ℕn ⟶ ℝ].
  ((Σ{x[i] y[i] 0≤i≤1} * Σ{x[i] y[i] 0≤i≤1}) ≤ {x[i] x[i] 0≤i≤1}
  * Σ{y[i] y[i] 0≤i≤1}))

Lemma: Cauchy-Schwarz2-strict
n:ℕ. ∀x,y:ℕn ⟶ ℝ.
  (∃i,j:ℕn. x[j] y[i] ≠ x[i] y[j]
  ⇐⇒ {x[i] y[i] 0≤i≤1} * Σ{x[i] y[i] 0≤i≤1}) < {x[i] x[i] 0≤i≤1}
      * Σ{y[i] y[i] 0≤i≤1}))

Lemma: radd-limit
x,y:ℕ ⟶ ℝ. ∀a,b:ℝ.  (lim n→∞.x[n]  lim n→∞.y[n]  lim n→∞.x[n] y[n] b)

Lemma: constant-limit
a,b:ℝ.  (lim n→∞.a ⇐⇒ b)

Lemma: rminus-limit
x:ℕ ⟶ ℝ. ∀a:ℝ.  (lim n→∞.x[n]  lim n→∞.-(x[n]) -(a))

Lemma: rsub-limit
x,y:ℕ ⟶ ℝ. ∀a,b:ℝ.  (lim n→∞.x[n]  lim n→∞.y[n]  lim n→∞.x[n] y[n] b)

Lemma: rmax-limit
x,y:ℕ ⟶ ℝ. ∀a,b:ℝ.  (lim n→∞.x[n]  lim n→∞.y[n]  lim n→∞.rmax(x[n];y[n]) rmax(a;b))

Lemma: const-rmul-limit-with-bound
x:ℕ ⟶ ℝ. ∀a,c:ℝ. ∀m:ℕ+.  ((|c| ≤ r(m))  lim n→∞.x[n]  lim n→∞.c x[n] a)

Lemma: rmul-limit
x,y:ℕ ⟶ ℝ. ∀a,b:ℝ.  (lim n→∞.x[n]  lim n→∞.y[n]  lim n→∞.x[n] y[n] b)

Lemma: rleq-limit
[x,y:ℕ ⟶ ℝ]. ∀[a,b:ℝ].  (a ≤ b) supposing ((∀n:ℕ(x[n] ≤ y[n])) and lim n→∞.y[n] and lim n→∞.x[n] a)

Lemma: constant-rleq-limit
[x:ℝ]. ∀[y:ℕ ⟶ ℝ]. ∀[a:ℝ].  (x ≤ a) supposing ((∀n:ℕ(x ≤ y[n])) and lim n→∞.y[n] a)

Lemma: rleq-limit-constant
[x:ℝ]. ∀[y:ℕ ⟶ ℝ]. ∀[a:ℝ].  (a ≤ x) supposing ((∀n:ℕ(y[n] ≤ x)) and lim n→∞.y[n] a)

Lemma: rless-iff-rleq
x,y:ℝ.  (x < ⇐⇒ ∃m:ℕ+(x ≤ (y (r1/r(m)))))

Lemma: rless-implies-rleq
x,y:ℝ.  ((x < y)  (∃m:ℕ+(x ≤ (y (r1/r(m))))))

Lemma: req-iff-rabs-rleq
x,y:ℝ.  (x ⇐⇒ ∀m:ℕ+(|x y| ≤ (r1/r(m))))

Lemma: req-iff-rabs-rleq-bound
x,y:ℝ.  (x ⇐⇒ ∃B:ℕ+. ∀m:ℕ+(|x y| ≤ (r(B)/r(m))))

Lemma: near-real-implies-real
[x:ℕ+ ⟶ ℤ]. ∀[y:ℝ].  x ∈ {x:ℝy}  supposing ∀n:ℕ+(|(x within 1/n) y| ≤ (r1/r(n)))

Lemma: rational-approx-implies-req
[k:ℕ+]. ∀[x:ℝ]. ∀[a:ℕ+ ⟶ ℤ].
  ((∀n:ℕ+(|x (r(a n)/r(2 n))| ≤ (r(k)/r(n))))  {k-regular-seq(a) ∧ (accelerate(k;a) x)})

Lemma: accelerate-rational-approx
[k:ℕ+]. ∀[x:ℝ]. ∀[a:ℕ+ ⟶ ℤ].  ((∀n:ℕ+(|x (r(a n)/r(2 n))| ≤ (r(k)/r(n))))  (accelerate(k;a) ∈ {y:ℝx} ))

Lemma: req-iff-rational-approx
[x:ℝ]. ∀[a:ℕ+ ⟶ ℤ].  (regular-seq(a) ∧ (a x) ⇐⇒ ∀n:ℕ+(|x (r(a n)/r(2 n))| ≤ (r1/r(n))))

Lemma: rabs-diff-rmul
[a,b,c,d,x,y:ℝ].  ((|a b| ≤ x)  (|c d| ≤ y)  (|(a c) d| ≤ ((|a| y) (|d| x))))

Lemma: rabs-diff-rdiv
[a,b,c,d,x,y:ℝ].
  (c ≠ r0  d ≠ r0  (|a b| ≤ x)  (|(r1/c) (r1/d)| ≤ y)  (|(a/c) (b/d)| ≤ ((|a| y) (|(r1/d)| x))))

Definition: ireal-approx
j-approx(x;M;z) ==  |x (r(z)/r(2 M))| ≤ (r(j)/r(M))

Lemma: ireal-approx_wf
[j:ℕ]. ∀[x:ℝ]. ∀[M:ℕ+]. ∀[z:ℤ].  (j-approx(x;M;z) ∈ ℙ)

Lemma: ireal-approx_functionality
[j:ℕ]. ∀[M:ℕ+]. ∀[z:ℤ]. ∀[x,y:ℝ].  j-approx(x;M;z) ⇐⇒ j-approx(y;M;z) supposing y

Lemma: ireal-approx-1
[x:ℝ]. ∀[M:ℕ+].  1-approx(x;M;x M)

Lemma: ireal-approx-radd
[x,y:ℝ]. ∀[j,i:ℕ]. ∀[M:ℕ+]. ∀[a,b:ℤ].  (j-approx(x;M;a)  i-approx(y;M;b)  i-approx(x y;M;a b))

Lemma: ireal-approx-radd-int
[x,y:ℝ]. ∀[j:ℕ]. ∀[M:ℕ+]. ∀[a,n:ℤ].  (j-approx(x;M;a)  j-approx(x r(n);M;a (2 M)))

Lemma: ireal-approx-rmul
[x,y:ℝ]. ∀[j,M:ℕ+]. ∀[a,b:ℤ].
  ∀k:ℕ+
    ((|x| ≤ (r1/r(4)))
     ((2 |b|) ≤ (k M))
     j-approx(x;k M;a)
     j-approx(y;M;b)
     j-approx(x y;M;(a b) ÷ M))

Lemma: ireal-approx-rmul2
[x,y:ℝ]. ∀[j,M:ℕ+]. ∀[a,b:ℤ].
  ∀k:ℕ+((|x| ≤ (r1/r(2)))  (|b| ≤ k)  j-approx(x;k;a)  j-approx(y;2 M;b)  j-approx(x y;M;(a b) ÷ k))

Definition: poly-approx-aux
poly-approx-aux(a;x;xM;M;n;k)
==r if (k =z 0)
    then M
    else eval poly-approx-aux(a;x;xM;M;n 1;k 1) in
         eval ((2 |u|) ÷ M) in
         eval if (b =z 1) then xM else (b M) fi  in
         eval (u z) ÷ in
         eval in
           t
    fi 

Lemma: poly-approx-aux_wf
[k:ℕ]. ∀[a:ℕ ⟶ ℝ]. ∀[x:ℝ]. ∀[xM:ℤ]. ∀[M:ℕ+]. ∀[n:ℕ].  (poly-approx-aux(a;x;xM;M;n;k) ∈ ℤ)

Lemma: poly-approx-aux-property
[k:ℕ]. ∀[a:ℕ ⟶ ℝ]. ∀[x:ℝ]. ∀[xM:ℤ]. ∀[M:ℕ+]. ∀[n:ℕ].
  ((|x| ≤ (r1/r(4)))  1-approx(x;M;xM)  1-approx(Σ{(a (n i)) x^i 0≤i≤k};M;poly-approx-aux(a;x;xM;M;n;k)))

Definition: poly-approx
poly-approx(a;x;k;N) ==
  eval (k 1) in
  eval in
  eval xM in
  eval poly-approx-aux(a;x;xM;M;0;k) in
    z ÷ B

Lemma: poly-approx_wf
[k:ℕ]. ∀[a:ℕ ⟶ ℝ]. ∀[x:ℝ]. ∀[N:ℕ+].  (poly-approx(a;x;k;N) ∈ ℤ)

Lemma: poly-approx-property
[k:ℕ]. ∀[a:ℕ ⟶ ℝ]. ∀[x:ℝ]. ∀[N:ℕ+].  ((|x| ≤ (r1/r(4)))  1-approx(Σ{(a i) x^i 0≤i≤k};N;poly-approx(a;x;k;N)))

Definition: rless-witness
rless-witness(x;y;p) ==  fst((TERMOF{rless-implies-rleq:o, 1:l} p))

Lemma: rless-witness_wf
[x,y:ℝ]. ∀[p:x < y].  (rless-witness(x;y;p) ∈ ℕ+)

Lemma: rless-witness-property
[x,y:ℝ]. ∀[p:x < y].  ((x ≤ (y (r1/r(rless-witness(x;y;p))))) ∧ ((x (r1/r(rless-witness(x;y;p)))) ≤ y))

Lemma: rinv-limit
x:ℕ ⟶ ℝ. ∀a:ℝ.  (lim n→∞.x[n]  (∀n:ℕx[n] ≠ r0)  a ≠ r0  lim n→∞.(r1/x[n]) (r1/a))

Definition: regularize
regularize(k;f) ==
  λn.if regular-upto(k;n;f)
     then n
     else eval mu(λn.(¬bregular-upto(k;n;f))) in
          eval seq-min-upper(k;m;f) in
            ((n ((f j) (2 k))) ÷ k)
     fi 

Lemma: regularize_wf
[k:ℕ+]. ∀[f:ℕ+ ⟶ ℤ].  (regularize(k;f) ∈ ℕ+ ⟶ ℤ)

Lemma: regular-upto-regularize
f:ℕ+ ⟶ ℤ. ∀k,m:ℕ.  ((↑regular-upto(k;m 1;f))  {∀n:ℕ((n ≤ m)  (regularize(k;f) n))})

Lemma: regularize-k-regular
k:ℕ+. ∀f:ℕ+ ⟶ ℤ.  k-regular-seq(regularize(k;f))

Lemma: regularize-regular
k:ℕ+. ∀x:{f:ℕ+ ⟶ ℤk-regular-seq(f)} .  (regularize(k;x) x ∈ (ℕ+ ⟶ ℤ))

Lemma: regularize-real
k:ℕ+. ∀x:ℝ.  (regularize(k;x) x ∈ ℝ)

Lemma: weak-continuity-principle-real
x:ℝ. ∀F:ℝ ⟶ 𝔹. ∀G:n:ℕ+ ⟶ {y:ℝy ∈ (ℕ+n ⟶ ℤ)} .  (∃n:ℕ+ [F (G n)])

Lemma: weak-continuity-principle-real-double
x:ℝ. ∀F,H:ℝ ⟶ 𝔹. ∀G:n:ℕ+ ⟶ {y:ℝy ∈ (ℕ+n ⟶ ℤ)} .  (∃n:ℕ+ [(F (G n) ∧ (G n))])

Lemma: weak-continuity-principle-real-nat
x:ℝ. ∀F:ℝ ⟶ ℕ. ∀G:n:ℕ+ ⟶ {y:ℝy ∈ (ℕ+n ⟶ ℤ)} .  (∃n:ℕ+ [((F x) (F (G n)) ∈ ℕ)])

Lemma: pseudo-positive-is-positive
x:ℝr0 < supposing pseudo-positive(x)

Lemma: real-weak-Markov
x,y:ℝ.  x ≠ supposing ∀z:ℝ((¬(z x)) ∨ (z y)))

Lemma: rneq-by-function
x,y,a,b:ℝ. ∀f:ℝ ⟶ ℝ.  (a ≠  (f[x] a)  (f[y] b)  (∀x,y:ℝ.  ((x y)  (f[x] f[y])))  x ≠ y)

Lemma: test-rat-term-poly
[x,y:ℝ].  (r1/r1 (y/x x)) (x x/y (x x)) supposing (y < -(x x)) ∧ x ≠ r0

Lemma: real-weak-Markov-ext
x,y:ℝ.  x ≠ supposing ∀z:ℝ((¬(z x)) ∨ (z y)))

Lemma: pseudo-positive-is-positive-proof2
x:ℝr0 < supposing pseudo-positive(x)

Lemma: rneq-function
f:ℝ ⟶ ℝ((∀x,y:ℝ.  ((x y)  (f[x] f[y])))  (∀x,y:ℝ.  (f[x] ≠ f[y]  x ≠ y)))

Lemma: cantor-lemma
x,y,e,z:ℝ.
  (∃x',y':ℝ((x ≤ x') ∧ (x' < y') ∧ (y' ≤ y) ∧ ((z < x') ∨ (y' < z)) ∧ ((y' x') < e))) supposing 
     ((x < y) and 
     (r0 < e))

Lemma: cantor-lemma2
z:ℕ ⟶ ℝ
  ∃f:ℕ ⟶ {p:ℝ × ℝ(fst(p)) < (snd(p))}  ⟶ {p:ℝ × ℝ(fst(p)) < (snd(p))} 
   ∀n:ℕ. ∀p:{p:ℝ × ℝ(fst(p)) < (snd(p))} .
     let x,y 
     in let x',y' 
        in (x ≤ x') ∧ (x' < y') ∧ (y' ≤ y) ∧ (((z n) < x') ∨ (y' < (z n))) ∧ ((y' x') < (r1/r(n 1)))

Lemma: reals-uncountable

z:ℕ ⟶ ℝ. ∀x,y:ℝ.  ((x < y)  (∃u:ℝ((x ≤ u) ∧ (u ≤ y) ∧ (∀n:ℕu ≠ n))))

Lemma: reals-uncountable-simple
f:ℕ ⟶ ℝSurj(ℕ;ℝ;f))

Definition: series-sum
Note that in Σn.x[n] a, the index is in ℕ  (i.e, 0,1,2,...) rather
than ℕ+  (1,2,3,...) as in Bishop Bridges.⋅

Σn.x[n] ==  lim n→∞{x[i] 0≤i≤n} a

Lemma: series-sum_wf
[x:ℕ ⟶ ℝ]. ∀[a:ℝ].  n.x[n] a ∈ ℙ)

Lemma: series-sum_functionality
x,y:ℕ ⟶ ℝ. ∀a,b:ℝ.  ({Σn.x[n]  Σn.y[n] b}) supposing ((a b) and (∀n:ℕ(x[n] y[n])))

Lemma: series-sum-unique
[x:ℕ ⟶ ℝ]. ∀[a,b:ℝ].  (a b) supposing n.x[n] and Σn.x[n] a)

Lemma: series-sum-0
Σi.r0 r0

Lemma: series-sum-constant
x:ℝ. Σi.if (i =z 0) then else r0 fi  x

Lemma: series-sum-linear1
x,y:ℕ ⟶ ℝ. ∀a,b:ℝ.  n.x[n]  Σn.y[n]  Σn.x[n] y[n] b)

Lemma: series-sum-linear2
x:ℕ ⟶ ℝ. ∀a,c:ℝ.  n.x[n]  Σn.c x[n] a)

Lemma: series-sum-linear3
x:ℕ ⟶ ℝ. ∀a,c:ℝ.  n.x[n]  Σn.x[n] c)

Lemma: triangular-reciprocal-series-sum
Σn.(r1/r(t(n 1))) r(2)

Definition: series-converges
Σn.x[n]↓ ==  ∃a:ℝ. Σn.x[n] a

Lemma: series-converges_wf
[x:ℕ ⟶ ℝ]. n.x[n]↓ ∈ ℙ)

Lemma: series-converges-tail
x:ℕ ⟶ ℝn.x[n]↓  (∀y:ℕ ⟶ ℝ((∃N:ℕ. ∀n:{N...}. (y[n] x[n]))  Σn.y[n]↓)))

Lemma: series-converges-tail2
N:ℕ. ∀x:ℕ ⟶ ℝ.  n.x[n N]↓  Σn.x[n]↓)

Lemma: series-converges-tail2-ext
N:ℕ. ∀x:ℕ ⟶ ℝ.  n.x[n N]↓  Σn.x[n]↓)

Lemma: series-converges-rmul
c:ℝ. ∀x:ℕ ⟶ ℝ.  n.x[n]↓  Σn.x[n] c↓)

Lemma: series-converges-rmul-ext
c:ℝ. ∀x:ℕ ⟶ ℝ.  n.x[n]↓  Σn.x[n] c↓)

Lemma: series-converges-limit-zero
x:ℕ ⟶ ℝn.x[n]↓  lim n→∞.x[n] r0)

Definition: converges-absolutely
converges-absolutely(n.x[n]) ==  Σn.|x[n]|↓

Lemma: converges-absolutely_wf
[x:ℕ ⟶ ℝ]. (converges-absolutely(n.x[n]) ∈ ℙ)

Lemma: comparison-test
y:ℕ ⟶ ℝn.y[n]↓  (∀x:ℕ ⟶ ℝ. Σn.x[n]↓ supposing ∀n:ℕ(|x[n]| ≤ y[n])))

Lemma: comparison-test-ext
y:ℕ ⟶ ℝn.y[n]↓  (∀x:ℕ ⟶ ℝ. Σn.x[n]↓ supposing ∀n:ℕ(|x[n]| ≤ y[n])))

Lemma: converges-absolutely-converges
x:ℕ ⟶ ℝ(converges-absolutely(n.x[n])  Σn.x[n]↓)

Definition: series-diverges
Σn.x[n]↑ ==  n.Σ{x[i] 0≤i≤n}↑

Lemma: series-diverges_wf
[x:ℕ ⟶ ℝ]. n.x[n]↑ ∈ ℙ)

Lemma: series-diverges_functionality
[x,y:ℕ ⟶ ℝ].  n.x[n]↑  Σn.y[n]↑supposing ∀n:ℕ(x[n] y[n])

Lemma: series-diverges-rmul
[x:ℕ ⟶ ℝ]. n.x[n]↑  (∀c:ℝ(c ≠ r0  Σn.c x[n]↑)))

Lemma: series-diverges-trivially
z:ℝ((r0 < z)  (∀x:ℕ ⟶ ℝ((∀k:ℕ. ∃n:ℕ((k ≤ n) ∧ (z ≤ |x[n]|)))  Σn.x[n]↑)))

Lemma: series-diverges-tail
x:ℕ ⟶ ℝ. ∀N:ℕ.  n.x[N n]↑  Σn.x[n]↑)

Lemma: series-diverges-tail-iff
x:ℕ ⟶ ℝ. ∀N:ℕ.  n.x[N n]↑ ⇐⇒ Σn.x[n]↑)

Lemma: harmonic-series-diverges
Σn.(r1/r(n 1))↑

Lemma: harmonic-series-diverges-to-infinity
n:ℕ((r1 (r(n)/r(2))) ≤ Σ{(r1/r(i)) 1≤i≤2^n})

Lemma: comparison-test-for-divergence
x,y:ℕ ⟶ ℝ.  n.y[n]↑  (∃N:ℕ. ∀n:{N...}. ((r0 ≤ y[n]) ∧ (y[n] ≤ x[n])))  Σn.x[n]↑)

Lemma: partial-geometric-series
n:ℕ. ∀c:ℝ.  (c ≠ r1  {c^i 0≤i≤n} (r1 c^n 1/r1 c)))

Lemma: geometric-series-converges
c:{c:ℝ(r0 ≤ c) ∧ (c < r1)} . Σn.c^n (r1/r1 c)

Lemma: geometric-series-one-half
We have  Σn.(r1/r(2)^n) r(2) because 0 ∈ ℤ is included in the series.⋅

Σn.(r1/r(2)^n) r(2)

Lemma: geometric-series-converges-ext
c:{c:ℝ(r0 ≤ c) ∧ (c < r1)} . Σn.c^n (r1/r1 c)

Lemma: ratio-test
x:ℕ ⟶ ℝ. ∀N:ℕ.
  ((∀c:{c:ℝ(r0 ≤ c) ∧ (c < r1)} ((∀n:{N...}. (|x[n 1]| ≤ (c |x[n]|)))  Σn.x[n]↓))
  ∧ (∀c:{c:ℝr1 < c} ((∀n:{N...}. ((c |x[n]|) < |x[n 1]|))  Σn.x[n]↑)))

Lemma: ratio-test-ext
x:ℕ ⟶ ℝ. ∀N:ℕ.
  ((∀c:{c:ℝ(r0 ≤ c) ∧ (c < r1)} ((∀n:{N...}. (|x[n 1]| ≤ (c |x[n]|)))  Σn.x[n]↓))
  ∧ (∀c:{c:ℝr1 < c} ((∀n:{N...}. ((c |x[n]|) < |x[n 1]|))  Σn.x[n]↑)))

Lemma: ratio-test-corollary
x:ℕ ⟶ ℝ
  ((∀n:ℕx[n] ≠ r0)  (∀L:ℝ(lim n→∞.|(x[n 1]/x[n])|  (((L < r1)  Σn.x[n]↓) ∧ ((r1 < L)  Σn.x[n]↑)))))

Lemma: Kummer-criterion
a,x:ℕ ⟶ ℝ.
  ((lim n→∞.a[n] x[n] r0
   (∃c:{c:ℝr0 < c} 
       ∃N:ℕ
        ((∀n:{N...}. ((r0 < a[n]) ∧ (r0 < x[n])))
        ∧ (∀n:{N...}. ((r0 < a[n]) ∧ (c ≤ ((a[n] x[n]/x[n 1]) a[n 1]))))))
   Σn.x[n]↓)
  ∧ ((∃N:ℕ
       ((∀n:{N...}. ((r0 < a[n]) ∧ (r0 < x[n])))
       ∧ (∀n:{N...}. (((a[n] x[n]/x[n 1]) a[n 1]) ≤ r0))
       ∧ Σn.(r1/a[N n])↑))
     Σn.x[n]↑))

Lemma: r-archimedean-implies
x:ℝ. ∀m:ℕ+.  ∃M:ℕ+(((r1/r(M)) x) ≤ (r1/r(m)))

Lemma: r-archimedean-implies2
x:ℝ. ∀d:{d:ℝr0 < d} .  ∃M:ℕ+((x/r(M)) ≤ d)

Lemma: r-archimedean-rabs
x:ℝ. ∃n:ℕ(|x| ≤ r(n))

Lemma: r-archimedean-rabs-ext
x:ℝ. ∃n:ℕ(|x| ≤ r(n))

Lemma: r-archimedean2
x:ℝ. ∃N:ℕ. ∀n:{N...}. (|(x/r(n 1))| ≤ (r1/r(2)))

Lemma: Raabe-lemma
y:ℕ ⟶ ℝ. ∀c:ℝ.
  ((r0 < c)
   (∀N:ℕ+((∀n:{N...}. (r0 < y[n]))  (∀n:{N...}. (c ≤ (r(n) ((y[n]/y[n 1]) r1))))  lim n→∞.y[n] r0)))

Lemma: Raabe-test
x:ℕ ⟶ ℝ. ∀L:ℝ.
  ((∀n:ℕ(r0 < x[n]))
   lim n→∞.r(n) ((x[n]/x[n 1]) r1) L
   (((r1 < L)  Σn.x[n]↓) ∧ ((L < r1)  Σn.x[n]↑)))

Lemma: alternating-series-tail-bound
x:ℕ ⟶ ℝ. ∀M:ℕ.
  ((∀n:ℕ((M ≤ n)  ((r0 ≤ x[n]) ∧ (x[n 1] ≤ x[n]))))
   lim n→∞.x[n] r0
   (∀a:{M...}. ∀b:ℕ.  (|Σ{-1^i x[i] a≤i≤b}| ≤ x[a])))

Lemma: alternating-series-converges
x:ℕ ⟶ ℝ((∃M:ℕ. ∀n:ℕ(M <  ((r0 ≤ x[n]) ∧ (x[n 1] ≤ x[n]))))  lim n→∞.x[n] r0  Σn.-1^n x[n]↓)

Lemma: alternating-series-converges-ext
x:ℕ ⟶ ℝ((∃M:ℕ. ∀n:ℕ(M <  ((r0 ≤ x[n]) ∧ (x[n 1] ≤ x[n]))))  lim n→∞.x[n] r0  Σn.-1^n x[n]↓)

Lemma: rdiv-factorial-lemma1
n:ℕ((r((2 (n 1))!)/r((2 n)!)) r(((n 2) 2) ((n 2) 1)))

Lemma: rdiv-factorial-lemma2
x:ℝ. ∀b:ℕ.  (((x x) ≤ r(b b))  (∀n:ℕ(((b ÷ 2) ≤ n)  ((x x) ≤ (r((2 (n 1))!)/r((2 n)!))))))

Lemma: rdiv-factorial-lemma3
x:ℝ. ∀b:ℕ.
  (((x x) ≤ r(b b))  (∀n:ℕ(((b ÷ 2) ≤ n)  ((x x) ≤ (r(((2 (n 1)) 1)!)/r(((2 n) 1)!))))))

Lemma: rdiv-factorial-limit-zero-from-bound
x:ℝ. ∀n:ℕ.  ((|x| ≤ r(n))  lim n→∞.(|x|^n/r((n)!)) r0)

Lemma: rdiv-factorial-limit-zero-from-bound2
x:ℝ. ∀n:ℕ.  ((|x| ≤ r(n))  lim n→∞.(|x|^n/r((n 1)!)) r0)

Lemma: rdiv-factorial-limit-zero
x:ℝlim n→∞.(|x|^n/r((n)!)) r0

Lemma: sine-exists
x:ℝ. ∃a:ℝ. Σi.-1^i (x^(2 i) 1)/((2 i) 1)! a

Lemma: sine-exists-ext
x:ℝ. ∃a:ℝ. Σi.-1^i (x^(2 i) 1)/((2 i) 1)! a

Definition: sine
sine(x) ==  fst((TERMOF{sine-exists-ext:o, 1:l} x))

Lemma: sine_wf
[x:ℝ]. (sine(x) ∈ ℝ)

Lemma: sine-is-limit
x:ℝ. Σi.-1^i (x^(2 i) 1)/((2 i) 1)! sine(x)

Lemma: sine_functionality
[x,y:ℝ].  sine(x) sine(y) supposing y

Lemma: sine0
sine(r0) r0

Lemma: sine-rminus
x:ℝ(sine(-(x)) -(sine(x)))

Definition: sine-approx
sine-approx(x;k;N) ==
  eval poly-approx(λi.(r(if (i rem =z 0) then else -1 fi ))/((2 i) 1)!;x^2;k;2 N) in
  eval |u| in
  eval in
    (u z) ÷ b

Lemma: sine-approx_wf
[x:ℝ]. ∀[k:ℕ]. ∀[N:ℕ+].  (sine-approx(x;k;N) ∈ ℤ)

Lemma: sine-approx-property
[x:ℝ]. ∀[k:ℕ]. ∀[N:ℕ+].
  ((|x| ≤ (r1/r(2)))  1-approx(Σ{-1^i (x^(2 i) 1)/((2 i) 1)! 0≤i≤k};N;sine-approx(x;k;N)))

Lemma: sine-poly-approx-1
[x:{x:ℝ(r0 ≤ x) ∧ (x ≤ r1)} ]. ∀[k:ℕ].
  (|sine(x) - Σ{-1^i (x^(2 i) 1)/((2 i) 1)! 0≤i≤k}| ≤ (x^(2 k) 3/r(((2 k) 3)!)))

Lemma: sine-poly-approx
[x:{x:ℝ|x| ≤ (r1/r(2))} ]. ∀[k:ℕ]. ∀[N:ℕ+].
  (|sine(x) (r(sine-approx(x;k;N))/r(2 N))| ≤ ((|x|^(2 k) 3/r(((2 k) 3)!)) (r1/r(N))))

Lemma: sine-approx-lemma
a:{2...}. ∀N:ℕ+.  (∃k:ℕ [(N ≤ (a^((2 k) 3) ((2 k) 3)!))])

Lemma: sine-approx-lemma-bad
a:{2...}. ∀N:ℕ.  (∃k:ℕ [(N ≤ (a^((2 k) 3) ((2 k) 3)!))])

Lemma: sine-approx-lemma-bad-ext
a:{2...}. ∀N:ℕ.  (∃k:ℕ [(N ≤ (a^((2 k) 3) ((2 k) 3)!))])

Lemma: cosine-approx-lemma
a:{2...}. ∀N:ℕ+.  (∃k:ℕ [(N ≤ (a^((2 k) 2) ((2 k) 2)!))])

Comment: extract-sine-lemma-non-uniform
Using CompNatInd' we got this extract for sine-approx-lemma :

⌜λa,N. eval aa in
       eval a^3 in
         if (b) < (N)
            then letrec rec(d)=λk,b. eval k' in
                                     eval b' aa ((2 k) 4) ((2 k) 5) in
                                       if (b') < (N)  then rec (d 1) k' b'  else k' in
                  rec  
                 
                 
                 b
            else 0⌝⋅

Lemma: sine-approx-lemma-ext
a:{2...}. ∀N:ℕ+.  (∃k:ℕ [(N ≤ (a^((2 k) 3) ((2 k) 3)!))])

Lemma: cosine-approx-lemma-ext
a:{2...}. ∀N:ℕ+.  (∃k:ℕ [(N ≤ (a^((2 k) 2) ((2 k) 2)!))])

Lemma: sine-approx-for-small
a:{2...}. ∀N:ℕ+. ∀x:{x:ℝ|x| ≤ (r1/r(a))} .  (∃z:ℤ [(|sine(x) (r(z)/r(2 N))| ≤ (r(2)/r(N)))])

Lemma: sine-approx-for-small-ext
a:{2...}. ∀N:ℕ+. ∀x:{x:ℝ|x| ≤ (r1/r(a))} .  (∃z:ℤ [(|sine(x) (r(z)/r(2 N))| ≤ (r(2)/r(N)))])

Definition: sine-small
sine-small(x) ==  accelerate(2;λN.sine-approx(x;genfact-inv(N;48;k.4 ((2 k) 3) ((2 k) 2));N))

Lemma: sine-small_wf
[x:{x:ℝ|x| ≤ (r1/r(2))} ]. (sine-small(x) ∈ {y:ℝsine(x)} )

Lemma: cosine-exists
x:ℝ. ∃a:ℝ. Σi.-1^i (x^2 i)/(2 i)! a

Lemma: cosine-exists-ext
x:ℝ. ∃a:ℝ. Σi.-1^i (x^2 i)/(2 i)! a

Definition: cosine
cosine(x) ==  fst((TERMOF{cosine-exists-ext:o, 1:l} x))

Lemma: cosine_wf
[x:ℝ]. (cosine(x) ∈ ℝ)

Lemma: cosine-is-limit
x:ℝ. Σi.-1^i (x^2 i)/(2 i)! cosine(x)

Lemma: cosine_functionality
[x,y:ℝ].  cosine(x) cosine(y) supposing y

Lemma: cosine0
cosine(r0) r1

Lemma: cosine-rminus
x:ℝ(cosine(-(x)) cosine(x))

Definition: cosine-approx
cosine-approx(x;k;N) ==  poly-approx(λi.(r(if (i rem =z 0) then else -1 fi ))/(2 i)!;x^2;k;N)

Lemma: cosine-approx_wf
[x:ℝ]. ∀[k:ℕ]. ∀[N:ℕ+].  (cosine-approx(x;k;N) ∈ ℤ)

Lemma: cosine-approx-property
[x:ℝ]. ∀[k:ℕ]. ∀[N:ℕ+].  ((|x| ≤ (r1/r(2)))  1-approx(Σ{-1^i (x^2 i)/(2 i)! 0≤i≤k};N;cosine-approx(x;k;N)))

Lemma: cosine-poly-approx-1
[x:{x:ℝ(r0 ≤ x) ∧ (x ≤ r1)} ]. ∀[k:ℕ].
  (|cosine(x) - Σ{-1^i (x^2 i)/(2 i)! 0≤i≤k}| ≤ (x^(2 k) 2/r(((2 k) 2)!)))

Lemma: cosine-poly-approx
[x:{x:ℝ|x| ≤ (r1/r(2))} ]. ∀[k:ℕ]. ∀[N:ℕ+].
  (|cosine(x) (r(cosine-approx(x;k;N))/r(2 N))| ≤ ((|x|^(2 k) 2/r(((2 k) 2)!)) (r1/r(N))))

Lemma: cosine-approx-for-small
a:{2...}. ∀N:ℕ+. ∀x:{x:ℝ|x| ≤ (r1/r(a))} .  (∃z:ℤ [(|cosine(x) (r(z)/r(2 N))| ≤ (r(2)/r(N)))])

Lemma: cosine-approx-for-small-ext
a:{2...}. ∀N:ℕ+. ∀x:{x:ℝ|x| ≤ (r1/r(a))} .  (∃z:ℤ [(|cosine(x) (r(z)/r(2 N))| ≤ (r(2)/r(N)))])

Definition: cosine-small
cosine-small(x) ==  accelerate(2;λN.cosine-approx(x;genfact-inv(N;8;k.4 ((2 k) 2) ((2 k) 1));N))

Lemma: cosine-small_wf
[x:{x:ℝ|x| ≤ (r1/r(2))} ]. (cosine-small(x) ∈ {y:ℝcosine(x)} )

Lemma: exp-series-converges
x:ℝ. Σn.(x^n)/(n)!↓

Lemma: exp-exists
x:ℝ. ∃a:ℝ. Σn.(x^n)/(n)! a

Lemma: exp-exists-ext
x:ℝ. ∃a:ℝ. Σn.(x^n)/(n)! a

Definition: rexp
e^x ==  fst((TERMOF{exp-exists-ext:o, 1:l} x))

Lemma: rexp_wf
[x:ℝ]. (e^x ∈ ℝ)

Lemma: rexp-is-limit
x:ℝ. Σn.(x^n)/(n)! e^x

Lemma: rexp_functionality
[x,y:ℝ].  e^x e^y supposing y

Lemma: rexp0
e^r0 r1

Lemma: rexp-of-nonneg-stronger
x:ℝ((r0 ≤ x)  ((r1 x) ≤ e^x))

Lemma: rexp-of-nonneg
x:ℝ((r0 ≤ x)  (r1 ≤ e^x))

Lemma: rexp-of-positive
x:ℝ((r0 < x)  (r1 < e^x))

Definition: rexp-approx
rexp-approx(x;k;N) ==  poly-approx(λi.(r1)/(i)!;x;k;N)

Lemma: rexp-approx_wf
[x:ℝ]. ∀[k:ℕ]. ∀[N:ℕ+].  (rexp-approx(x;k;N) ∈ ℤ)

Lemma: rexp-approx-property
[x:ℝ]. ∀[k:ℕ]. ∀[N:ℕ+].  ((|x| ≤ (r1/r(4)))  1-approx(Σ{(x^i)/(i)! 0≤i≤k};N;rexp-approx(x;k;N)))

Definition: rset
Set(ℝ==  {A:ℝ ⟶ ℙ| ∀x,y:ℝ.  ((x y)  (A x)  (A y))} 

Lemma: rset_wf
Set(ℝ) ∈ 𝕌'

Definition: rset-member
x ∈ ==  x

Lemma: rset-member_wf
[A:Set(ℝ)]. ∀[x:ℝ].  (x ∈ A ∈ ℙ)

Definition: rseteq
rseteq(A;B) ==  ∀x:ℝ(x ∈ ⇐⇒ x ∈ B)

Lemma: rseteq_wf
[A,B:Set(ℝ)].  (rseteq(A;B) ∈ ℙ)

Lemma: rseteq_weakening
[A,B:Set(ℝ)].  rseteq(A;B) supposing B ∈ Set(ℝ)

Lemma: rset-member_functionality
A,B:Set(ℝ). ∀x,y:ℝ.
  ((∀y:ℝSqStable(y ∈ A))  (∀y:ℝSqStable(y ∈ B))  rseteq(A;B)  {(x ∈ A)  (y ∈ B)} supposing y)

Definition: mk-rset
{x:ℝ P[x]} ==  λx.P[x]

Lemma: mk-rset_wf
[P:ℝ ⟶ ℙ]. {x:ℝ P[x]} ∈ Set(ℝsupposing ∀x,y:ℝ.  ((x y)  P[x]  P[y])

Lemma: member_mk_rset_lemma
z,P:Top.  (z ∈ {x:ℝ P[x]} P[z])

Definition: upper-bound
A ≤ ==  ∀x:ℝ((x ∈ A)  (x ≤ b))

Lemma: upper-bound_wf
[A:Set(ℝ)]. ∀[b:ℝ].  (A ≤ b ∈ ℙ)

Lemma: upper-bound_functionality
[A:Set(ℝ)]. ∀[b,c:ℝ].  {A ≤ supposing A ≤ b} supposing b ≤ c

Definition: strict-upper-bound
A < ==  ∀x:ℝ((x ∈ A)  (x < b))

Lemma: strict-upper-bound_wf
[A:Set(ℝ)]. ∀[b:ℝ].  (A < b ∈ ℙ)

Lemma: strict-upper-bound_functionality
A:Set(ℝ). ∀b,c:ℝ.  {A <  A < c} supposing b ≤ c

Definition: lower-bound
lower-bound(A;b) ==  ∀x:ℝ((x ∈ A)  (b ≤ x))

Lemma: lower-bound_wf
[A:Set(ℝ)]. ∀[b:ℝ].  (lower-bound(A;b) ∈ ℙ)

Definition: bounded-above
bounded-above(A) ==  ∃b:ℝA ≤ b

Lemma: bounded-above_wf
[A:Set(ℝ)]. (bounded-above(A) ∈ ℙ)

Lemma: bounded-above-strict
[A:Set(ℝ)]. (bounded-above(A)  (∃b:ℝA < b))

Definition: bounded-below
bounded-below(A) ==  ∃b:ℝlower-bound(A;b)

Lemma: bounded-below_wf
[A:Set(ℝ)]. (bounded-below(A) ∈ ℙ)

Definition: sup
sup(A) ==  A ≤ b ∧ (∀e:ℝ((r0 < e)  (∃x:ℝ((x ∈ A) ∧ ((b e) < x)))))

Lemma: sup_wf
[A:Set(ℝ)]. ∀[b:ℝ].  (sup(A) b ∈ ℙ)

Lemma: sup-unique
[A:Set(ℝ)]. ∀[b,c:ℝ].  (b c) supposing (sup(A) and sup(A) b)

Definition: inf
inf(A) ==  lower-bound(A;b) ∧ (∀e:ℝ((r0 < e)  (∃x:ℝ((x ∈ A) ∧ (x < (b e))))))

Lemma: inf_wf
[A:Set(ℝ)]. ∀[b:ℝ].  (inf(A) b ∈ ℙ)

Lemma: inf_functionality
[A,A':Set(ℝ)].  ∀b,b':ℝ.  ((b b')  rseteq(A;A')  (inf(A) ⇐⇒ inf(A') b'))

Lemma: sup_functionality
[A,A':Set(ℝ)].  ∀b,b':ℝ.  ((b b')  rseteq(A;A')  (sup(A) ⇐⇒ sup(A') b'))

Lemma: rleq_inf
[A:Set(ℝ)]. ∀[b,c:ℝ].  (inf(A)  (c ≤ ⇐⇒ ∀x:ℝ((x ∈ A)  (c ≤ x))))

Lemma: inf-rleq
[A:Set(ℝ)]. ∀b,c:ℝ.  (inf(A)  (b ≤ ⇐⇒ ∀e:ℝ((r0 < e)  (∃x:ℝ((x ∈ A) ∧ ((x e) ≤ c))))))

Lemma: inf-rless
[A:Set(ℝ)]. ∀b,c:ℝ.  (inf(A)  (b < ⇐⇒ ∃x:ℝ((x ∈ A) ∧ (x < c))))

Lemma: inf-unique
[A:Set(ℝ)]. ∀[b,c:ℝ].  (inf(A)  inf(A)  (b c))

Definition: member-closure
y ∈ closure(A) ==  ∃x:ℕ ⟶ ℝ(lim n→∞.x[n] y ∧ (∀n:ℕ(A x[n])))

Lemma: member-closure_wf
[A:ℝ ⟶ ℙ]. ∀[y:ℝ].  (y ∈ closure(A) ∈ ℙ)

Definition: closed-rset
closed-rset(A) ==  ∀y:ℝ(y ∈ closure(A)  (A y))

Lemma: closed-rset_wf
[A:ℝ ⟶ ℙ]. (closed-rset(A) ∈ ℙ)

Definition: upper-bounds
upper-bounds(A) ==  λb.A ≤ b

Lemma: upper-bounds_wf
[A:Set(ℝ)]. (upper-bounds(A) ∈ Set(ℝ))

Definition: strict-upper-bounds
strict-upper-bounds(A) ==  λb.A < b

Lemma: strict-upper-bounds_wf
[A:Set(ℝ)]. (strict-upper-bounds(A) ∈ Set(ℝ))

Lemma: upper-bounds-closed
[A:Set(ℝ)]. closed-rset(upper-bounds(A))

Lemma: sup-iff-closure
[A:Set(ℝ)]. ∀x:ℝ(sup(A) ⇐⇒ A ≤ x ∧ x ∈ closure(A))

Lemma: closures-meet
[P,Q:ℝ ⟶ ℙ].
  ((∃a,b:ℝ((P a) ∧ (Q b) ∧ (a ≤ b)))
   (∃c:ℝ
       (((r0 ≤ c) ∧ (c < r1))
       ∧ (∀a,b:ℝ.
            (((P a) ∧ (Q b) ∧ (a ≤ b))
             (∃a',b':ℝ((P a') ∧ (Q b') ∧ (a ≤ a') ∧ (a' ≤ b') ∧ (b' ≤ b) ∧ ((b' a') ≤ ((b a) c))))))))
   (∃y:ℝ(y ∈ closure(P) ∧ y ∈ closure(Q))))

Lemma: closures-meet-sq
[P,Q:ℝ ⟶ ℙ].
  ((∃a:{a:ℝa} (∃b:ℝ [((Q b) ∧ (a ≤ b))]))
   (∃c:{c:ℝ(r0 ≤ c) ∧ (c < r1)} 
       ∀a:{a:ℝa} . ∀b:{b:ℝ(Q b) ∧ (a ≤ b)} .
         ∃a':{a':ℝa'} (∃b':{b':ℝ(Q b') ∧ (a' ≤ b')}  [((a ≤ a') ∧ (b' ≤ b) ∧ ((b' a') ≤ ((b a) c)))]))
   (∃y:ℝ(y ∈ closure(λz.(↓z)) ∧ y ∈ closure(λz.(↓z)))))

Lemma: closures-meet-sq-ext
[P,Q:ℝ ⟶ ℙ].
  ((∃a:{a:ℝa} (∃b:ℝ [((Q b) ∧ (a ≤ b))]))
   (∃c:{c:ℝ(r0 ≤ c) ∧ (c < r1)} 
       ∀a:{a:ℝa} . ∀b:{b:ℝ(Q b) ∧ (a ≤ b)} .
         ∃a':{a':ℝa'} (∃b':{b':ℝ(Q b') ∧ (a' ≤ b')}  [((a ≤ a') ∧ (b' ≤ b) ∧ ((b' a') ≤ ((b a) c)))]))
   (∃y:ℝ(y ∈ closure(λz.(↓z)) ∧ y ∈ closure(λz.(↓z)))))

Lemma: closures-meet'
P,Q:Set(ℝ).
  ((∃a,b:ℝ((a ∈ P) ∧ (b ∈ Q) ∧ (a < b)))
   (∃c:ℝ
       (((r0 ≤ c) ∧ (c < r1))
       ∧ (∀a,b:ℝ.
            (((a ∈ P) ∧ (b ∈ Q) ∧ (a < b))
             (∃a',b':ℝ((a' ∈ P) ∧ (b' ∈ Q) ∧ (a ≤ a') ∧ (a' < b') ∧ (b' ≤ b) ∧ ((b' a') ≤ ((b a) c))))))))
   (∃y:ℝ(y ∈ closure(P) ∧ y ∈ closure(Q))))

Lemma: closures-meet-sq'
[P,Q:ℝ ⟶ ℙ].
  ((∃a:{a:ℝa} (∃b:ℝ [((Q b) ∧ (a < b))]))
   (∃c:{c:ℝ(r0 ≤ c) ∧ (c < r1)} 
       ∀a:{a:ℝa} . ∀b:{b:ℝ(Q b) ∧ (a < b)} .
         ∃a':{a':ℝa'} (∃b':{b':ℝ(Q b') ∧ (a' < b')}  [((a ≤ a') ∧ (b' ≤ b) ∧ ((b' a') ≤ ((b a) c)))]))
   (∃y:ℝ(y ∈ closure(λz.(↓z)) ∧ y ∈ closure(λz.(↓z)))))

Lemma: closures-meet-sq'-ext
[P,Q:ℝ ⟶ ℙ].
  ((∃a:{a:ℝa} (∃b:ℝ [((Q b) ∧ (a < b))]))
   (∃c:{c:ℝ(r0 ≤ c) ∧ (c < r1)} 
       ∀a:{a:ℝa} . ∀b:{b:ℝ(Q b) ∧ (a < b)} .
         ∃a':{a':ℝa'} (∃b':{b':ℝ(Q b') ∧ (a' < b')}  [((a ≤ a') ∧ (b' ≤ b) ∧ ((b' a') ≤ ((b a) c)))]))
   (∃y:ℝ(y ∈ closure(λz.(↓z)) ∧ y ∈ closure(λz.(↓z)))))

Lemma: least-upper-bound
[A:Set(ℝ)]
  ((∃x:ℝ(x ∈ A))
   bounded-above(A)
   (∃b:ℝsup(A) ⇐⇒ ∀x,y:ℝ.  ((x < y)  ((∃a:ℝ((a ∈ A) ∧ (x < a))) ∨ A ≤ y))))

Definition: totally-bounded
totally-bounded(A) ==
  ∀e:ℝ((r0 < e)  (∃n:ℕ+. ∃a:ℕn ⟶ ℝ((∀i:ℕn. (a i ∈ A)) ∧ (∀x:ℝ((x ∈ A)  (∃i:ℕn. (|x i| < e)))))))

Lemma: totally-bounded_wf
[A:Set(ℝ)]. (totally-bounded(A) ∈ ℙ)

Lemma: totally-bounded-implies-nonvoid
[A:Set(ℝ)]. (totally-bounded(A)  (∃x:ℝ(x ∈ A)))

Lemma: totally-bounded-bounded-above
[A:Set(ℝ)]. (totally-bounded(A)  bounded-above(A))

Lemma: totally-bounded-sup
[A:Set(ℝ)]. (totally-bounded(A)  (∃b:ℝsup(A) b))

Definition: rset-neg
-(A) ==  λx.(-(x) ∈ A)

Lemma: rset-neg_wf
[A:Set(ℝ)]. (-(A) ∈ Set(ℝ))

Lemma: member_rset_neg_lemma
x,A:Top.  (x ∈ -(A) -(x) ∈ A)

Lemma: inf-as-sup
[A:Set(ℝ)]. ∀b:ℝ(inf(A) ⇐⇒ sup(-(A)) -(b))

Lemma: totally-bounded-neg
[A:Set(ℝ)]. (totally-bounded(A) ⇐⇒ totally-bounded(-(A)))

Lemma: totally-bounded-inf
[A:Set(ℝ)]. (totally-bounded(A)  (∃b:ℝinf(A) b))

Definition: interval
An interval has two endpoint. There are three kinds of endpoints
that we represent with the type ⌜ℝ + ℝ Top⌝.
An endpoint ⌜inr anything ⌝ represents "infinity"  
(positive or negative, depending on whether it is the right or left endpoint).
An endpoint of the form ⌜inl x⌝  where ⌜x ∈ ℝ + ℝ⌝ is "finite" endpoint.
If ⌜(inl a) ∈ (ℝ + ℝ)⌝ then the interval is "closed" at that endpoint  (written [a,..
or ...,a] and if ⌜(inl a) ∈ (ℝ + ℝ)⌝ then the interval is "open" at 
that endpoint (written  (a,...   or ...,a)  ).

Therefore, the type ℝ + ℝ Top × (ℝ + ℝ Top) 
represents nine kinds of intervals:
[a,b] (a,b], (-infinity, b]
[a,b), (a,b), (-infintiy, b)
[a, infinity), (a,infinity), (-infinity, infinity).⋅

Interval ==  ℝ + ℝ Top × (ℝ + ℝ Top)

Lemma: interval_wf
Interval ∈ Type

Definition: i-finite
i-finite(I) ==  let l,u in (↑isl(l)) ∧ (↑isl(u))

Lemma: i-finite_wf
[J:Interval]. (i-finite(J) ∈ ℙ)

Lemma: decidable__i-finite
J:Interval. Dec(i-finite(J))

Definition: endpoints
endpoints(I) ==  let l,u in <case outl(l) of inl(a) => inr(a) => a, case outl(u) of inl(b) => inr(b) => b>

Lemma: endpoints_wf
[I:Interval]. endpoints(I) ∈ ℝ × ℝ supposing i-finite(I)

Definition: left-endpoint
left-endpoint(I) ==  fst(endpoints(I))

Lemma: left-endpoint_wf
[I:Interval]. left-endpoint(I) ∈ ℝ supposing i-finite(I)

Definition: right-endpoint
right-endpoint(I) ==  snd(endpoints(I))

Lemma: right-endpoint_wf
[I:Interval]. right-endpoint(I) ∈ ℝ supposing i-finite(I)

Definition: iproper
iproper(I) ==  i-finite(I)  (left-endpoint(I) < right-endpoint(I))

Lemma: iproper_wf
[I:Interval]. (iproper(I) ∈ ℙ)

Lemma: sq_stable__iproper
I:Interval. SqStable(iproper(I))

Definition: i-length
|I| ==  right-endpoint(I) left-endpoint(I)

Lemma: i-length_wf
[I:Interval]. |I| ∈ ℝ supposing i-finite(I)

Lemma: iproper-length
I:Interval. ((i-finite(I) ∧ iproper(I))  (r0 < |I|))

Lemma: iproper-length-iff
I:Interval. (i-finite(I)  (r0 < |I| ⇐⇒ iproper(I)))

Definition: i-closed
i-closed(I) ==  let l,u in (↑((¬bisl(l)) ∨bisl(outl(l)))) ∧ (↑((¬bisl(u)) ∨bisl(outl(u))))

Lemma: i-closed_wf
[I:Interval]. (i-closed(I) ∈ ℙ)

Lemma: decidable__i-closed
I:Interval. Dec(i-closed(I))

Definition: i-member
r ∈ ==
  let l,u 
  in case l
      of inl(x) =>
      case x
       of inl(a) =>
       case of inl(y) => case of inl(b) => (a ≤ r) ∧ (r ≤ b) inr(b) => (a ≤ r) ∧ (r < b) inr(y) => a ≤ r
       inr(a) =>
       case of inl(y) => case of inl(b) => (a < r) ∧ (r ≤ b) inr(b) => (a < r) ∧ (r < b) inr(y) => a < r
      inr(x) =>
      case of inl(y) => case of inl(b) => r ≤ inr(b) => r < inr(y) => True

Lemma: i-member_wf
[I:Interval]. ∀[x:ℝ].  (x ∈ I ∈ ℙ)

Lemma: i-member_functionality
I:Interval. ∀a,b:ℝ.  a ∈ ⇐⇒ b ∈ supposing b

Lemma: sq_stable__i-member
I:Interval. ∀r:ℝ.  SqStable(r ∈ I)

Lemma: i-member-finite
[I:Interval]. ∀[r:ℝ]. left-endpoint(I)≤r≤right-endpoint(I) supposing r ∈ supposing i-finite(I)

Lemma: i-finite-iff-bounded
I:Interval. (i-finite(I) ⇐⇒ ∃a,b:ℝ. ∀[r:ℝ]. a≤r≤supposing r ∈ I)

Definition: rccint
[l, u] ==  <inl inl l, inl inl u>

Lemma: rccint_wf
[a,b:ℝ].  ([a, b] ∈ Interval)

Lemma: member_rccint_lemma
r,u,l:Top.  (r ∈ [l, u] (l ≤ r) ∧ (r ≤ u))

Definition: rcoint
[l, u) ==  <inl inl l, inl (inr )>

Lemma: rcoint_wf
[l,u:ℝ].  ([l, u) ∈ Interval)

Lemma: member_rcoint_lemma
r,u,l:Top.  (r ∈ [l, u) (l ≤ r) ∧ (r < u))

Definition: rciint
[l, ∞==  <inl inl l, inr ⋅ >

Lemma: rciint_wf
[l:ℝ]. ([l, ∞) ∈ Interval)

Lemma: member_rciint_lemma
r,l:Top.  (r ∈ [l, ∞l ≤ r)

Definition: rocint
(l, u] ==  <inl (inr ), inl inl u>

Lemma: rocint_wf
[l,u:ℝ].  ((l, u] ∈ Interval)

Lemma: member_rocint_lemma
r,u,l:Top.  (r ∈ (l, u] (l < r) ∧ (r ≤ u))

Definition: rooint
(l, u) ==  <inl (inr ), inl (inr )>

Lemma: rooint_wf
[l,u:ℝ].  ((l, u) ∈ Interval)

Lemma: member_rooint_lemma
r,u,l:Top.  (r ∈ (l, u) (l < r) ∧ (r < u))

Definition: roiint
(l, ∞==  <inl (inr ), inr ⋅ >

Lemma: roiint_wf
[l:ℝ]. ((l, ∞) ∈ Interval)

Lemma: member_roiint_lemma
r,l:Top.  (r ∈ (l, ∞l < r)

Definition: ricint
(-∞u] ==  <inr ⋅ inl inl u>

Lemma: ricint_wf
[u:ℝ]. ((-∞u] ∈ Interval)

Lemma: member_ricint_lemma
r,u:Top.  (r ∈ (-∞u] r ≤ u)

Definition: rioint
(-∞u) ==  <inr ⋅ inl (inr )>

Lemma: rioint_wf
[u:ℝ]. ((-∞u) ∈ Interval)

Lemma: member_rioint_lemma
r,u:Top.  (r ∈ (-∞u) r < u)

Definition: riiint
(-∞, ∞==  <inr ⋅ inr ⋅ >

Lemma: riiint_wf
(-∞, ∞) ∈ Interval

Lemma: member_riiint_lemma
r:Top. (r ∈ (-∞, ∞True)

Lemma: left_endpoint_rccint_lemma
b,a:Top.  (left-endpoint([a, b]) a)

Lemma: right_endpoint_rccint_lemma
b,a:Top.  (right-endpoint([a, b]) b)

Lemma: iproper-roiint
x:ℝiproper((x, ∞))

Lemma: iproper-riiint
iproper((-∞, ∞))

Definition: closed-rational-interval
closed-rational-interval(a;b;c;d) ==  [r(a/b), r(c/d)]

Lemma: closed-rational-interval_wf
[a:ℤ]. ∀[b:ℤ-o]. ∀[c:ℤ]. ∀[d:ℤ-o].  (closed-rational-interval(a;b;c;d) ∈ Interval)

Lemma: interval-totally-bounded
a:ℝ. ∀b:{b:ℝa ≤ b} .  totally-bounded(λx.(x ∈ [a, b]))

Definition: lower-right-endpoint
lower-right-endpoint(a;b;n) ==  (a b)/n 2

Lemma: lower-right-endpoint_wf
[a,b:ℝ]. ∀[n:ℕ+].  (lower-right-endpoint(a;b;n) ∈ ℝ)

Definition: raise-left-endpoint
raise-left-endpoint(a;b;n) ==  (n b)/n 2

Lemma: raise-left-endpoint_wf
[a,b:ℝ]. ∀[n:ℕ+].  (raise-left-endpoint(a;b;n) ∈ ℝ)

Lemma: lower-right-endpoint-rless
a,b:ℝ. ∀n:ℕ+.  ((a < b)  ((a < lower-right-endpoint(a;b;n)) ∧ (lower-right-endpoint(a;b;n) < b)))

Lemma: raise-left-endpoint-rless
a,b:ℝ. ∀n:ℕ+.  ((a < b)  ((a < raise-left-endpoint(a;b;n)) ∧ (raise-left-endpoint(a;b;n) < b)))

Lemma: raise-lower-endpoints-rless
a,b:ℝ. ∀n:ℕ+.  ((a < b)  (raise-left-endpoint(a;b;n) < lower-right-endpoint(a;b;n)))

Definition: old-i-approx
old-i-approx(I;n) ==
  let l,u 
  in case l
      of inl(x) =>
      case x
       of inl(a) =>
       case of inl(y) => case of inl(b) => [a, b] inr(b) => [a, (r1/r(n))] inr(y) => [a, r(n)]
       inr(a) =>
       case u
        of inl(y) =>
        case of inl(b) => [a (r1/r(n)), b] inr(b) => [a (r1/r(n)), (r1/r(n))]
        inr(y) =>
        [a (r1/r(n)), r(n)]
      inr(x) =>
      case of inl(y) => case of inl(b) => [r(-n), b] inr(b) => [r(-n), (r1/r(n))] inr(y) => [r(-n), r(n)]

Definition: i-approx
i-approx(I;n) ==
  let l,u 
  in case l
      of inl(x) =>
      case x
       of inl(a) =>
       case of inl(y) => case of inl(b) => [a, b] inr(b) => [a, (r1/r(n))] inr(y) => [a, r(n)]
       inr(a) =>
       case u
        of inl(y) =>
        case of inl(b) => [a (r1/r(n)), b] inr(b) => [a (r1/r(n)), (r1/r(n))]
        inr(y) =>
        [a (r1/r(n)), r(n)]
      inr(x) =>
      case of inl(y) => case of inl(b) => [r(-n), b] inr(b) => [r(-n), (r1/r(n))] inr(y) => [r(-n), r(n)]

Lemma: i-approx_wf
[n:ℕ+]. ∀[I:Interval].  (i-approx(I;n) ∈ Interval)

Lemma: i-finite-approx
n,m:ℕ+. ∀I:Interval.  (i-finite(i-approx(I;n)) ⇐⇒ i-finite(i-approx(I;m)))

Lemma: i-approx-monotonic
I:Interval. ∀n,m:ℕ+.  ∀x:ℝ((x ∈ i-approx(I;n))  (x ∈ i-approx(I;m))) supposing n ≤ m

Lemma: i-member-implies
I:Interval. ∀r:ℝ.
  ((r ∈ I)
   (∃n,M:ℕ+
       ((r ∈ i-approx(I;n))
       ∧ (∀y:{y:ℝy ∈ I} ((|y r| ≤ (r1/r(M)))  (y ∈ i-approx(I;n))))
       ∧ (iproper(I)  iproper(i-approx(I;n))))))

Lemma: i-member-iff
I:Interval. ∀r:ℝ.  (r ∈ ⇐⇒ ∃n:ℕ+(r ∈ i-approx(I;n)))

Lemma: i-member-proper-iff
I:Interval. (iproper(I)  (∀r:ℝ(r ∈ ⇐⇒ ∃n:ℕ+(iproper(i-approx(I;n)) ∧ (r ∈ i-approx(I;n))))))

Lemma: i-member-approx
I:Interval. ∀x:ℝ. ∀n:ℕ+.  ((x ∈ i-approx(I;n))  (x ∈ I))

Lemma: i-member-witness
I:Interval. ∀r:ℝ.  ((r ∈ I)  (∃n:ℕ+(r ∈ i-approx(I;n))))

Lemma: i-member-finite-closed
I:Interval. (i-closed(I)  i-finite(I)  (∀r:ℝ(r ∈ ⇐⇒ left-endpoint(I)≤r≤right-endpoint(I))))

Lemma: i-member-between
I:Interval. ∀a,b:ℝ.  ((a ∈ I)  (b ∈ I)  (∀r:ℝ((a ≤ r)  (r ≤ b)  (r ∈ I))))

Definition: rfun
I ⟶ℝ ==  {x:ℝx ∈ I}  ⟶ ℝ

Lemma: rfun_wf
[I:Interval]. (I ⟶ℝ ∈ Type)

Lemma: rfun_subtype_1
[a,b,c:ℝ].  ((a ≤ c)  (c ≤ b)  ([a, b] ⟶ℝ ⊆[a, c] ⟶ℝ))

Lemma: rfun_subtype_2
[a,b,c:ℝ].  ((a ≤ c)  (c ≤ b)  ([a, b] ⟶ℝ ⊆[c, b] ⟶ℝ))

Lemma: rfun_subtype_3
[a,b,c,d:ℝ].  ((a ≤ c)  (c ≤ d)  (d ≤ b)  ([a, b] ⟶ℝ ⊆[c, d] ⟶ℝ))

Definition: r-ap
f(x) ==  x

Lemma: r-ap_wf
[I:Interval]. ∀[f:I ⟶ℝ]. ∀[x:ℝ].  f(x) ∈ ℝ supposing x ∈ I

Definition: rfun-eq
rfun-eq(I;f;g) ==  ∀x:{x:ℝx ∈ I} (f(x) g(x))

Lemma: rfun-eq_wf
[I:Interval]. ∀[f,g:I ⟶ℝ].  (rfun-eq(I;f;g) ∈ ℙ)

Lemma: rfun-eq_weakening
[I:Interval]. ∀[f,g:I ⟶ℝ].  ((f g ∈ I ⟶ℝ rfun-eq(I;f;g))

Lemma: rfun-eq_inversion
[I:Interval]. ∀[f,g:I ⟶ℝ].  rfun-eq(I;g;f) supposing rfun-eq(I;f;g)

Lemma: rfun-eq_transitivity
[I:Interval]. ∀[f,g,h:I ⟶ℝ].  (rfun-eq(I;f;h)) supposing (rfun-eq(I;g;h) and rfun-eq(I;f;g))

Lemma: r-ap_functionality
[I:Interval]. ∀[f,g:I ⟶ℝ]. ∀[x:{x:ℝx ∈ I} ].  f(x) g(x) supposing rfun-eq(I;f;g)

Definition: i-witness
i-witness(I;r;p) ==  fst((TERMOF{i-member-witness:o, 1:l} p))

Lemma: i-witness_wf
[I:Interval]. ∀[r:ℝ]. ∀[p:r ∈ I].  (i-witness(I;r;p) ∈ ℕ+)

Lemma: i-witness-property
I:Interval. ∀r:ℝ. ∀p:r ∈ I.  (r ∈ i-approx(I;i-witness(I;r;p)))

Definition: i-type
i-type(I) ==  n:ℕ+ × {r:ℝr ∈ i-approx(I;n)} 

Lemma: i-type_wf
[I:Interval]. (i-type(I) ∈ Type)

Definition: i-real
real(p) ==  snd(p)

Lemma: i-real_wf
[I:Interval]. ∀[p:i-type(I)].  (real(p) ∈ ℝ)

Definition: i-nonvoid
i-nonvoid(I) ==  ∃r:ℝ(r ∈ I)

Lemma: i-nonvoid_wf
[I:Interval]. (i-nonvoid(I) ∈ ℙ)

Lemma: iproper-nonvoid
I:Interval. (iproper(I)  i-nonvoid(I))

Lemma: i-finite-closed-is-rccint
[I:Interval]. [left-endpoint(I), right-endpoint(I)] supposing i-finite(I) ∧ i-closed(I)

Definition: icompact
icompact(I) ==  i-nonvoid(I) ∧ i-closed(I) ∧ i-finite(I)

Lemma: icompact_wf
[I:Interval]. (icompact(I) ∈ ℙ)

Lemma: icompact-is-rccint
[I:Interval]. [left-endpoint(I), right-endpoint(I)] supposing icompact(I)

Lemma: rccint-icompact
a,b:ℝ.  (a ≤ ⇐⇒ icompact([a, b]))

Lemma: i-approx-closed
I:Interval. ∀n:ℕ+.  i-closed(i-approx(I;n))

Lemma: i-approx-finite
I:Interval. ∀n:ℕ+.  i-finite(i-approx(I;n))

Lemma: i-approx-compact
I:Interval. ∀n:ℕ+. ∀r:ℝ.  ((r ∈ i-approx(I;n))  icompact(i-approx(I;n)))

Lemma: implies-approx-compact
I:Interval. ∀n:ℕ+.  ((∃r:ℝ(r ∈ i-approx(I;n)))  icompact(i-approx(I;n)))

Lemma: i-approx-approx
I:Interval. ∀n,m:Top.  (i-approx(i-approx(I;n);m) i-approx(I;n))

Lemma: i-approx-of-compact
I:Interval. (icompact(I)  (∀n:ℕ+(i-approx(I;n) I ∈ Interval)))

Lemma: icompact-length-nonneg
[I:Interval]. r0 ≤ |I| supposing icompact(I)

Lemma: icompact-endpoints-rleq
[I:Interval]. left-endpoint(I) ≤ right-endpoint(I) supposing icompact(I)

Lemma: icompact-endpoints
I:Interval. (icompact(I)  ((left-endpoint(I) ∈ I) ∧ (right-endpoint(I) ∈ I)))

Lemma: i-member-compact
I:Interval. (icompact(I)  (∀r:ℝ(r ∈ ⇐⇒ left-endpoint(I)≤r≤right-endpoint(I))))

Lemma: iproper-approx
I:Interval
  (iproper(I)  (∀n:ℕ+(icompact(i-approx(I;n))  (icompact(i-approx(I;n 1)) ∧ iproper(i-approx(I;n 1))))))

Lemma: i-approx-elim
I:Interval. ∀n:ℕ+.  (i-nonvoid(i-approx(I;n))  (∃a:ℝ. ∃b:{b:ℝa ≤ b} (i-approx(I;n) [a, b] ∈ Interval)))

Lemma: i-member-diff-bound
I:Interval. (icompact(I)  (∀x,y:{x:ℝx ∈ I} .  (|x y| ≤ |I|)))

Lemma: rmin-i-member
I:Interval. ∀a,b:ℝ.  ((a ∈ I)  (b ∈ I)  (rmin(a;b) ∈ I))

Lemma: rmax-i-member
I:Interval. ∀a,b:ℝ.  ((a ∈ I)  (b ∈ I)  (rmax(a;b) ∈ I))

Definition: real-closed-interval-lattice
For real numbers ⌜a ≤ b⌝the reals in the closed interval ⌜[a, b]⌝
form distributive lattice with meet ⌜rmin(x;y)⌝ and join ⌜rmax(x;y)⌝.
But they do not form bounded distributive lattice with and b
because we can't prove that ⌜rmin(a;x)⌝ is the same real as ⌜a⌝.
We can only prove that it is an "equivalent" real using the `req` relation
(that we display ⌜y⌝).

Thus the closed interval forms generalized-bounded-distributive-lattice.⋅

real-closed-interval-lattice(a;b) ==  mk-general-bounded-dist-lattice({r:ℝr ∈ [a, b]} 2y.rmin(x;y);λ2y.rmax(x;y)\000C;a;b;λ2y.x y)

Lemma: real-closed-interval-lattice_wf
[a,b:ℝ].  real-closed-interval-lattice(a;b) ∈ GeneralBoundedDistributiveLattice supposing a ≤ b

Definition: real-interval-lattice
real-interval-lattice(I) ==  mk-distributive-lattice({r:ℝr ∈ I} ; λ2y.rmin(x;y); λ2y.rmax(x;y))

Lemma: real-interval-lattice_wf
[I:Interval]. (real-interval-lattice(I) ∈ DistributiveLattice)

Lemma: i-member-convex
I:Interval. ∀a,b:ℝ.  ((a ∈ I)  (b ∈ I)  (∀t:ℝ((r0 ≤ t)  (t ≤ r1)  ((t a) ((r1 t) b) ∈ I))))

Lemma: i-member-convex'
I:Interval. ∀a,b:ℝ.
  ((a ∈ I)  (b ∈ I)  (∀z:{z:ℝr0 < z} . ∀x,y:{z:ℝr0 ≤ z} .  (((x y) z)  (((x a) (y b)/z) ∈ I))))

Lemma: i-member-convex2
I:Interval. ∀a,b:ℝ.  ((a ∈ I)  (b ∈ I)  (∀n:ℕ+. ∀i,j:ℕ.  (((i j) n ∈ ℤ ((i b)/n ∈ I))))

Lemma: i-member-convex3
I:Interval. ∀a,b:ℝ.  ((a ∈ I)  (b ∈ I)  (∀n:ℕ+. ∀i:ℕ.  (((i 1) n ∈ ℤ ((i b)/n ∈ I))))

Lemma: i-member-convex4
I:Interval. ∀a,b:ℝ.  ((a ∈ I)  (b ∈ I)  (∀n:ℕ+. ∀i:ℕ.  (((1 i) n ∈ ℤ ((a b)/n ∈ I))))

Lemma: sq_stable__icompact
I:Interval. SqStable(icompact(I))

Lemma: i-closed-finite-rep
I:Interval. (i-closed(I)  i-finite(I)  (∃a,b:ℝ(I [a, b] ∈ Interval)))

Lemma: i-approx-rep
I:Interval. ∀n:ℕ+. ∀r:ℝ.  ((r ∈ i-approx(I;n))  (∃a,b:ℝ((a ≤ b) ∧ (i-approx(I;n) [a, b] ∈ Interval))))

Lemma: i-approx-rep2
I:Interval. ∀n:ℕ+.  ∃a,b:ℝ(i-approx(I;n) [a, b] ∈ Interval)

Lemma: i-type-member
I:Interval. ∀p:i-type(I).  (real(p) ∈ I)

Lemma: member-i-type
[I:Interval]. ∀[r:ℝ]. ∀[p:r ∈ I].  (<i-witness(I;r;p), r> ∈ i-type(I))

Definition: subinterval
I ⊆ J  ==  ∀r:ℝ((r ∈ I)  (r ∈ J))

Lemma: subinterval_wf
[I,J:Interval].  (I ⊆ J  ∈ ℙ)

Lemma: subinterval-riiint
[I:Interval]. I ⊆ (-∞, ∞

Lemma: i-finite-subinterval
I,J:Interval.  (I ⊆ J   i-finite(J)  i-finite(I))

Lemma: subinterval_transitivity
[I,J,K:Interval].  (I ⊆ J   J ⊆ K   I ⊆ )

Lemma: subinterval-subtype
[I,J:Interval].  {x:ℝx ∈ I}  ⊆{x:ℝx ∈ J}  supposing I ⊆ 

Lemma: i-approx-is-subinterval
I:Interval. ∀n:ℕ+.  i-approx(I;n) ⊆ 

Lemma: compact-subinterval
I:Interval. ∀J:{J:Interval| icompact(J)} .  (J ⊆ I   (∃n:{n:ℕ+icompact(i-approx(I;n))} J ⊆ i-approx(I;n) ))

Lemma: iproper-subinterval
I,J:Interval.  (I ⊆ J   iproper(I)  iproper(J))

Lemma: rcc-subinterval
I:Interval. ∀a,b:ℝ.  ([a, b] ⊆ I  ⇐⇒ (a ≤ b)  ((a ∈ I) ∧ (b ∈ I)))

Lemma: trivial-subinterval
I:Interval. [left-endpoint(I), right-endpoint(I)] ⊆ I  supposing icompact(I)

Lemma: subinterval-trivial
I:Interval. I ⊆ [left-endpoint(I), right-endpoint(I)]  supposing icompact(I)

Lemma: between-rmin-rmax
[x,y,t:ℝ].  uiff((rmin(x;y) ≤ t) ∧ (t ≤ rmax(x;y));(|t x| ≤ |y x|) ∧ (|t y| ≤ |y x|))

Lemma: rmin-rmax-subinterval
I:Interval. ∀a,b:ℝ.  ((a ∈ I)  (b ∈ I)  [rmin(a;b), rmax(a;b)] ⊆ )

Lemma: rmin3-rmax3-subinterval
I:Interval. ∀a,b,c:ℝ.  ((a ∈ I)  (b ∈ I)  (c ∈ I)  [rmin(a;rmin(b;c)), rmax(a;rmax(b;c))] ⊆ )

Lemma: member-rcc-min-max
[x,y,t:ℝ].  uiff(t ∈ [rmin(x;y), rmax(x;y)];(|t x| ≤ |y x|) ∧ (|t y| ≤ |y x|))

Lemma: i-approx-containing2
I:Interval. ∀a,b:ℝ.  (((a ∈ I) ∧ (b ∈ I))  (∃n:ℕ+((a ∈ i-approx(I;n)) ∧ (b ∈ i-approx(I;n)))))

Lemma: i-approx-containing
I:Interval. ∀x:ℝ.  ((x ∈ I)  (∃m:ℕ+(icompact(i-approx(I;m)) ∧ (x ∈ i-approx(I;m)))))

Lemma: compact-proper-interval-near-member
J:Interval
  (icompact(J)
   iproper(J)
   (∀x:ℝ((x ∈ J)  (∀r:ℝ((r0 < r)  (∃y:ℝ((y ∈ J) ∧ (|y x| ≤ r) ∧ (r0 < |y x|))))))))

Lemma: icompact-is-subinterval
I:Interval. (icompact(I)  (∃n:ℕ+I ⊆ i-approx((-∞, ∞);n) ))

Definition: sublevelset
sublevelset(I;f;c) ==  λx.((x ∈ I) ∧ (f(x) ≤ c))

Lemma: sublevelset_wf
[I:Interval]. ∀[f:I ⟶ℝ]. ∀[c:ℝ].  (sublevelset(I;f;c) ∈ ℝ ⟶ ℙ)

Definition: superlevelset
superlevelset(I;f;c) ==  λx.((x ∈ I) ∧ (c ≤ f(x)))

Lemma: superlevelset_wf
[I:Interval]. ∀[f:I ⟶ℝ]. ∀[c:ℝ].  (superlevelset(I;f;c) ∈ ℝ ⟶ ℙ)

Lemma: rfun_subtype
[I,J:Interval].  I ⟶ℝ ⊆J ⟶ℝ supposing J ⊆ 

Definition: frs-non-dec
frs-non-dec(L) ==  ∀i,j:ℕ||L||.  ((i ≤ j)  (L[i] ≤ L[j]))

Lemma: frs-non-dec_wf
[p:ℝ List]. (frs-non-dec(p) ∈ ℙ)

Lemma: sq_stable__frs-non-dec
[p:ℝ List]. SqStable(frs-non-dec(p))

Definition: frs-refines
frs-refines(p;q) ==  (∀x∈q.(∃y∈p. y))

Lemma: frs-refines_wf
[p,q:ℝ List].  (frs-refines(p;q) ∈ ℙ)

Lemma: frs-refines_weakening
[p,q:ℝ List].  ((p q ∈ (ℝ List))  frs-refines(p;q))

Lemma: frs-refines_transitivity
[p,q,r:ℝ List].  (frs-refines(p;q)  frs-refines(q;r)  frs-refines(p;r))

Definition: frs-increasing
frs-increasing(p) ==  ∀i,j:ℕ||p||.  (i <  (p[i] < p[j]))

Lemma: frs-increasing_wf
[p:ℝ List]. (frs-increasing(p) ∈ ℙ)

Lemma: frs-increasing-non-dec
[p:ℝ List]. (frs-increasing(p)  frs-non-dec(p))

Lemma: frs-increasing-sorted-by
[p:ℝ List]. (frs-increasing(p) ⇐⇒ sorted-by(λx,y. (x < y);p))

Lemma: frs-non-dec-sum
p:ℝ List. (1 < ||p||  frs-non-dec(p)  {p[i 1] p[i] 0≤i≤||p|| 2} (last(p) p[0])))

Lemma: frs-increasing-cons
p:ℝ List. ∀c:ℝ.  (frs-increasing([c p]) ⇐⇒ (0 < ||p||  (c < p[0])) ∧ frs-increasing(p))

Lemma: frs-non-dec-sorted-by
[p:ℝ List]. (frs-non-dec(p) ⇐⇒ sorted-by(λx,y. (x ≤ y);p))

Definition: frs-separated
frs-separated(p;q) ==  (∀x∈p.(∀y∈q.x ≠ y))

Lemma: frs-separated_wf
[p,q:ℝ List].  (frs-separated(p;q) ∈ ℙ)

Lemma: frs-increasing-separated-common-refinement
p,q:ℝ List.
  (frs-increasing(p)
   frs-increasing(q)
   frs-separated(p;q)
   (∃r:ℝ List. (frs-increasing(r) ∧ frs-refines(r;p) ∧ frs-refines(r;q) ∧ frs-refines(p q;r))))

Definition: frs-mesh
frs-mesh(p) ==  if ||p|| <then r0 else rmaximum(0;||p|| 2;i.p[i 1] p[i]) fi 

Lemma: frs-mesh_wf
[p:ℝ List]. (frs-mesh(p) ∈ ℝ)

Lemma: frs-mesh-nonneg
[p:ℝ List]. (frs-non-dec(p)  (r0 ≤ frs-mesh(p)))

Lemma: adjacent-frs-points
[p:ℝ List]. ∀[i:ℕ||p|| 1].  (frs-non-dec(p)  r0≤p[i 1] p[i]≤frs-mesh(p))

Lemma: frs-mesh-bound
[p:ℝ List]. ∀[x:ℝ].  ((r0 ≤ x)  (∀[i:ℕ||p|| 1]. ((p[i 1] p[i]) ≤ x))  (frs-mesh(p) ≤ x))

Lemma: avoid-real
a,b,x:ℝ.  ((a < b)  (∃a',b':ℝ((a ≤ a') ∧ (b' ≤ b) ∧ (a' < b') ∧ ((x < a') ∨ (b' < x)))))

Lemma: avoid-reals
L:ℝ List. ∀a,b:ℝ.  ((a < b)  (∃a',b':ℝ((a ≤ a') ∧ (b' ≤ b) ∧ (a' < b') ∧ (∀x∈L.(x < a') ∨ (b' < x)))))

Lemma: avoid-reals-simple
L:ℝ List. ∀a,b:ℝ.  ((a < b)  (∃c:ℝ((a ≤ c) ∧ (c < b) ∧ (∀x∈L.c ≠ x))))

Definition: partitions
partitions(I;p) ==  frs-non-dec(p) ∧ (0 < ||p||  ((left-endpoint(I) ≤ p[0]) ∧ (last(p) ≤ right-endpoint(I))))

Lemma: partitions_wf
[I:Interval]. ∀[p:ℝ List].  partitions(I;p) ∈ ℙ supposing icompact(I)

Lemma: sq_stable__partitions
I:Interval. ∀p:ℝ List.  (icompact(I)  SqStable(partitions(I;p)))

Definition: partition
partition(I) ==  {p:ℝ List| partitions(I;p)} 

Lemma: partition_wf
[I:Interval]. partition(I) ∈ Type supposing icompact(I)

Lemma: partition-point-member
I:Interval. (icompact(I)  (∀p:partition(I). (∀x∈p.x ∈ I)))

Definition: nearby-partitions
nearby-partitions(e;p;q) ==  (||p|| ||q|| ∈ ℤ) ∧ (∀i:ℕ||p||. (|p[i] q[i]| ≤ e))

Lemma: nearby-partitions_wf
[e:ℝ]. ∀[p,q:ℝ List].  (nearby-partitions(e;p;q) ∈ ℙ)

Lemma: nearby-partitions_functionality
p,q:ℝ List. ∀e1,e2:ℝ.  ((e1 ≤ e2)  {nearby-partitions(e1;p;q)  nearby-partitions(e2;p;q)})

Lemma: nearby-frs-mesh
e:{e:ℝr0 < e} . ∀p,q:ℝ List.  (nearby-partitions(e;q;p)  (frs-mesh(p) ≤ (frs-mesh(q) (r(2) e))))

Lemma: nearby-increasing-partition
I:Interval
  ((icompact(I) ∧ iproper(I))
   (∀p:partition(I). ∀e:{e:ℝr0 < e} .  ∃q:partition(I). (frs-increasing(q) ∧ nearby-partitions(e;p;q))))

Lemma: nearby-increasing-partition-avoids
I:Interval
  ((icompact(I) ∧ iproper(I))
   (∀p:partition(I)
        (frs-increasing(p)
         (∀L:ℝ List. ∀e:{e:ℝr0 < e} .
              ∃q:partition(I). (frs-increasing(q) ∧ nearby-partitions(e;p;q) ∧ frs-separated(q;L))))))

Definition: full-partition
full-partition(I;p) ==  [left-endpoint(I) (p [right-endpoint(I)])]

Lemma: full-partition_wf
[I:Interval]. ∀[p:partition(I)]. (full-partition(I;p) ∈ ℝ List) supposing icompact(I)

Lemma: last-full-partition
[I:Interval]. ∀[p:partition(I)]. (last(full-partition(I;p)) right-endpoint(I) ∈ ℝsupposing icompact(I)

Lemma: full-partition-non-dec
[I:Interval]. ∀[p:partition(I)]. frs-non-dec(full-partition(I;p)) supposing icompact(I)

Lemma: full-partition-point-member
I:Interval. (icompact(I)  (∀p:partition(I). (∀x∈full-partition(I;p).x ∈ I)))

Definition: partition-mesh
partition-mesh(I;p) ==  frs-mesh(full-partition(I;p))

Lemma: partition-mesh_wf
[I:Interval]. ∀[p:partition(I)]. (partition-mesh(I;p) ∈ ℝsupposing icompact(I)

Lemma: partition-mesh-nil
[I:Top]. (partition-mesh(I;[]) |I|)

Lemma: partition-mesh-nonneg
[I:Interval]. ∀[p:partition(I)]. (r0 ≤ partition-mesh(I;p)) supposing icompact(I)

Lemma: adjacent-partition-points
[I:Interval]
  ∀[p:partition(I)]
    (((¬0 < ||p||)  r0≤right-endpoint(I) left-endpoint(I)≤partition-mesh(I;p))
    ∧ (0 < ||p||
       (r0≤p[0] left-endpoint(I)≤partition-mesh(I;p)
         ∧ (∀i:ℕ||p|| 1. r0≤p[i 1] p[i]≤partition-mesh(I;p))
         ∧ r0≤right-endpoint(I) last(p)≤partition-mesh(I;p)))) 
  supposing icompact(I)

Lemma: adjacent-full-partition-points
[I:Interval]
  ∀[p:partition(I)]
    ∀i:ℕ||full-partition(I;p)|| 1. r0≤full-partition(I;p)[i 1] full-partition(I;p)[i]≤partition-mesh(I;p) 
  supposing icompact(I)

Lemma: partition-endpoints
[I:Interval]
  ∀[p:partition(I)]. ∀[x:ℝ].  full-partition(I;p)[0]≤x≤full-partition(I;p)[||full-partition(I;p)|| 1] supposing x ∈ 
  supposing icompact(I)

Lemma: partition-split-cons
[I:Interval]
  ∀[a:ℝ]. ∀[bs:ℝ List].
    (partitions(I;[a bs])  (partitions([left-endpoint(I), a];[]) ∧ partitions([a, right-endpoint(I)];bs))) 
  supposing icompact(I)

Lemma: partition-split-cons-mesh
[I:Interval]
  ∀[a:ℝ]. ∀[bs:ℝ List].
    (partitions(I;[a bs])
     (partitions([left-endpoint(I), a];[])
       ∧ partitions([a, right-endpoint(I)];bs)
       ∧ (left-endpoint(I) ≤ a)
       ∧ (a ≤ right-endpoint(I))
       ∧ (partition-mesh([left-endpoint(I), a];[]) ≤ partition-mesh(I;[a bs]))
       ∧ (partition-mesh([a, right-endpoint(I)];bs) ≤ partition-mesh(I;[a bs])))) 
  supposing icompact(I)

Lemma: partition-lemma
e:ℝ
  ((r0 < e)
   (∀n:ℕ+. ∀f:ℕn ⟶ ℝ.
        ∀x:ℝ. ∃i:ℕn. (|x i| ≤ e) supposing 0≤x≤(n 1) supposing ∀i:ℕ1. r0≤(f (i 1)) i≤e))

Lemma: firstn-partition
I:Interval
  (icompact(I)  (∀a:ℝ. ∀p:partition(I). ∀i:ℕ||p||.  ((a p[i])  (firstn(i;p) ∈ partition([left-endpoint(I), a])))))

Lemma: nth_tl-partition
I:Interval
  (icompact(I)
   (∀a:ℝ. ∀p:partition(I). ∀i:ℕ||p||.  ((a p[i])  (nth_tl(i 1;p) ∈ partition([a, right-endpoint(I)])))))

Lemma: mesh-property
I:Interval
  (icompact(I)
   (∀p:partition(I). ∀e:ℝ.
        ((r0 < e)
         ∀x:ℝ((x ∈ I)  (∃i:ℕ||full-partition(I;p)||. (|x full-partition(I;p)[i]| ≤ e))) 
           supposing partition-mesh(I;p) ≤ e)))

Definition: trivial-partition
trivial-partition(I) ==  []

Lemma: trivial-partition_wf
[I:Interval]. trivial-partition(I) ∈ partition(I) supposing i-finite(I)

Lemma: mesh-trivial-partition
[I:Top]. (partition-mesh(I;[]) |I|)

Definition: uniform-partition
uniform-partition(I;k) ==
  mklist(k 1;λi.(((r(k) r(i 1)) left-endpoint(I)) (r(i 1) right-endpoint(I))/r(k)))

Lemma: uniform-partition_wf
[I:Interval]. ∀[k:ℕ+]. (uniform-partition(I;k) ∈ partition(I)) supposing icompact(I)

Lemma: uniform-partition-point
[I:Interval]
  ∀[k:ℕ+]
    ∀i:ℕ1
      ((full-partition(I;uniform-partition(I;k))[i] r(k))
      (((r(k) r(i)) left-endpoint(I)) (r(i) right-endpoint(I)))) 
  supposing icompact(I)

Lemma: mesh-uniform-partition
[I:Interval]. ∀[k:ℕ+]. (partition-mesh(I;uniform-partition(I;k)) (|I|/r(k))) supposing icompact(I)

Lemma: partition-exists
I:Interval. (icompact(I)  (∀e:ℝ((r0 < e)  (∃p:partition(I). (partition-mesh(I;p) ≤ e)))))

Lemma: simple-partition-exists
a,b:ℝ.
  ((a ≤ b)
   (∀e:ℝ
        ((r0 < e)
         (∃M:ℕ+
             ∃g:ℕ1 ⟶ {x:ℝx ∈ [a, b]} 
              (((g 0) a ∈ ℝ) ∧ ((g M) b ∈ ℝ) ∧ (∀i:ℕM. (((g i) ≤ (g (i 1))) ∧ (((g (i 1)) i) ≤ e))))))))

Definition: is-partition-choice
is-partition-choice(p;x) ==  ∀i:ℕ||p|| 1. (x i ∈ [p[i], p[i 1]])

Lemma: is-partition-choice_wf
[p:ℝ List]. ∀[x:ℕ||p|| 1 ⟶ ℝ].  (is-partition-choice(p;x) ∈ ℙ)

Lemma: sq_stable__is-partition-choice
[p:ℝ List]. ∀[x:ℕ||p|| 1 ⟶ ℝ].  SqStable(is-partition-choice(p;x))

Definition: partition-choice
partition-choice(p) ==  i:ℕ||p|| 1 ⟶ {x:ℝx ∈ [p[i], p[i 1]]} 

Lemma: partition-choice_wf
[p:ℝ List]. (partition-choice(p) ∈ Type)

Lemma: partition-choice-subtype
[p:ℝ List]. ({f:ℕ||p|| 1 ⟶ ℝis-partition-choice(p;f)}  ⊆partition-choice(p))

Lemma: partition-choice-indep-funtype
[I:Interval]
  ∀[p:partition(I)]. (partition-choice(full-partition(I;p)) ⊆(ℕ||p|| 1 ⟶ {x:ℝx ∈ I} )) supposing icompact(I)

Lemma: implies-is-partition-choice
I:Interval
  (icompact(I)
   (∀p:partition(I). ∀x:partition-choice(full-partition(I;p)).  is-partition-choice(full-partition(I;p);x)))

Lemma: partition-choice-member
I:Interval
  (icompact(I)
   (∀p:partition(I). ∀x:partition-choice(full-partition(I;p)). ∀i:ℕ||full-partition(I;p)|| 1.  (x i ∈ {x:ℝx ∈ I} )\000C))

Definition: partition-choice-ap
x[i] ==  i

Lemma: partition-choice-ap_wf
I:Interval
  (icompact(I)  (∀p:partition(I). ∀x:partition-choice(full-partition(I;p)). ∀i:ℕ||p|| 1.  (x[i] ∈ {x:ℝx ∈ I} )))

Definition: partition-sum
S(f;p) ==  Σ{(f (x i)) (p[i 1] p[i]) 0≤i≤||p|| 2}

Lemma: partition-sum_wf
I:Interval
  (icompact(I)
   (∀p:partition(I). ∀f:I ⟶ℝ. ∀x:partition-choice(full-partition(I;p)).  (S(f;full-partition(I;p)) ∈ ℝ)))

Lemma: partition-sum_functionality
I:Interval
  (icompact(I)
   (∀p:partition(I). ∀q:ℝ List.
        ((||q|| ||p|| ∈ ℤ)
         (∀i:ℕ||q||. (q[i] p[i]))
         (∀f:I ⟶ℝ. ∀x:partition-choice(full-partition(I;p)).
              (S(f;full-partition(I;p)) S(f;full-partition(I;q)))))))

Lemma: constant-partition-sum
I:Interval
  (icompact(I)  (∀p:partition(I). ∀f:I ⟶ℝ. ∀z:ℝ.  ((z ∈ I)  (S(f;full-partition(I;p)) ((f z) |I|)))))

Definition: partition-refines
refines ==  frs-refines(P;Q)

Lemma: partition-refines_wf
I:Interval. ∀P,Q:partition(I).  (P refines Q ∈ ℙsupposing icompact(I)

Lemma: partition-refines_weakening
I:Interval. ∀P,Q:partition(I).  ((P Q ∈ partition(I))  refines Q) supposing icompact(I)

Lemma: partition-refines_transitivity
I:Interval. ∀P,Q,R:partition(I).  (P refines  refines  refines R) supposing icompact(I)

Lemma: partition-refines-cons
I:Interval
  (icompact(I)
   (∀a:ℝ. ∀bs:ℝ List.
        (partitions(I;[a bs])
         (0 < ||bs||  (a < hd(bs)))
         (∀p:partition(I)
              (p refines [a bs]
               (∃q:partition([left-endpoint(I), a])
                   ∃r:partition([a, right-endpoint(I)])
                    (q refines []
                    ∧ refines bs
                    ∧ (∃x:ℝ((x a) ∧ (p (q [x r]) ∈ (ℝ List))))
                    ∧ ||r|| ||q|| < ||p||
                    ∧ (∀x:partition-choice(full-partition(I;p))
                         (is-partition-choice(full-partition([left-endpoint(I), a];q);x)
                         ∧ is-partition-choice(full-partition([a, right-endpoint(I)];r);λi.(x (i ||q|| 1))))))))))))

Lemma: uniform-partition-refines
a:ℝ. ∀b:{b:ℝa ≤ b} . ∀k,m:ℕ+.  uniform-partition([a, b];m k) refines uniform-partition([a, b];k)

Lemma: uniform-partition-increasing
a,b:ℝ.  ((a < b)  (∀k:ℕ+frs-increasing(uniform-partition([a, b];k))))

Definition: rrange
f[x](x∈I) ==  λy.∃x:ℝ((x ∈ I) ∧ (f[x] y))

Lemma: rrange_wf
[I:Interval]. ∀[f:I ⟶ℝ].  (f[x](x∈I) ∈ Set(ℝ))

Lemma: rset-member-rrange
I:Interval. ∀[f:I ⟶ℝ]. ∀r:ℝ((r ∈ I)  (f[r] ∈ f[x](x∈I)))

Definition: continuous
Bishop defines f[x] to be continuous on compact interval if
for every ∈ > there is delta > such that
x,y:ℝ.  ((x ∈ J)  (y ∈ J)  (|x y| ≤ delta)  (|f[x] f[y]| ≤ ∈))

Bishop then defines f[x] to be continuous on an arbitrary interval if
if it is continuous on every compact subinterval of I.
The constructive content of this definition is called the modulus of continuity
of f.

In order to keep the modulus of continuity simpler (but equivalent),
 we modify this definition in two ways.
First, rather than quantify over all ∈ > we can quantify over n > 0
and use (r1/r(n)) in place of ∈.
Second, rather than quantify over all compact subintervals of I, we
define uniform family of compact subintervals i-approx(I;m), indexed by 
m > 0, that  "fill up" the interval I.

This modification give the modulus of continuity the type like
+ ⟶ ℕ+ ⟶ ℝ⋅

f[x] continuous for x ∈ ==
  ∀m:{m:ℕ+icompact(i-approx(I;m))} . ∀n:ℕ+.
    (∃d:ℝ [((r0 < d)
          ∧ (∀x,y:ℝ.  ((x ∈ i-approx(I;m))  (y ∈ i-approx(I;m))  (|x y| ≤ d)  (|f[x] f[y]| ≤ (r1/r(n))))))])

Lemma: continuous_wf
[I:Interval]. ∀[f:I ⟶ℝ].  (f[x] continuous for x ∈ I ∈ ℙ)

Definition: proper-continuous
f[x] (proper)continuous for x ∈ ==
  ∀m:{m:ℕ+icompact(i-approx(I;m)) ∧ iproper(i-approx(I;m))} . ∀n:ℕ+.
    (∃d:ℝ [((r0 < d)
          ∧ (∀x,y:ℝ.  ((x ∈ i-approx(I;m))  (y ∈ i-approx(I;m))  (|x y| ≤ d)  (|f[x] f[y]| ≤ (r1/r(n))))))])

Lemma: proper-continuous_wf
[I:Interval]. ∀[f:I ⟶ℝ].  (f[x] (proper)continuous for x ∈ I ∈ ℙ)

Lemma: proper-continuous-implies
[I:Interval]. ∀[f:I ⟶ℝ].
  (f[x] (proper)continuous for x ∈ I
   (∀m:ℕ+(icompact(i-approx(I;m))  iproper(i-approx(I;m))  f[x] continuous for x ∈ i-approx(I;m))))

Lemma: continuous-rneq
I:Interval. ∀f:I ⟶ℝ.  (f[x] continuous for x ∈  (∀a,b:{x:ℝx ∈ I} .  (f[a] ≠ f[b]  a ≠ b)))

Definition: dense-in-interval
dense-in-interval(I;X) ==  ∀a,b:{r:ℝr ∈ I} .  ((a < b)  (∃x:ℝ(((a < x) ∧ (x < b)) ∧ (X x))))

Lemma: dense-in-interval_wf
[I:Interval]. ∀[X:{a:ℝa ∈ I}  ⟶ ℙ].  (dense-in-interval(I;X) ∈ ℙ)

Definition: metric
metric(X) ==  {d:X ⟶ X ⟶ ℝ(∀x,y,z:X.  ((d z) ≤ ((d y) (d y)))) ∧ (∀x:X. ((d x) r0))} 

Lemma: metric_wf
[X:Type]. (metric(X) ∈ Type)

Lemma: metric-on-void
[X:Type]. Top ⊆metric(X) supposing ¬X

Lemma: same-metric
[X:Type]. ∀[d:metric(X)]. ∀[d':X ⟶ X ⟶ ℝ].  d' ∈ metric(X) supposing ∀x,y:X.  ((d' y) (d y))

Definition: metric-eq
metric-eq(X;d;d') ==  ∀x,y:X.  ((d' y) (d y))

Lemma: metric-eq_wf
[X:Type]. ∀[d,d':X ⟶ X ⟶ ℝ].  (metric-eq(X;d;d') ∈ ℙ)

Lemma: metric-symmetry
[X:Type]. ∀[d:metric(X)].  ∀x,y:X.  ((d y) (d x))

Lemma: metric-nonneg
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  (r0 ≤ (d y))

Lemma: metric-on-subtype
[X,Y:Type].  metric(X) ⊆metric(Y) supposing Y ⊆X

Definition: meq
x ≡ ==  (d y) r0

Lemma: meq_wf
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  (x ≡ y ∈ ℙ)

Lemma: stable__meq
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  Stable{x ≡ y}

Lemma: sq_stable__meq
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  SqStable(x ≡ y)

Lemma: meq-equiv
[X:Type]. ∀[d:metric(X)].  EquivRel(X;x,y.x ≡ y)

Lemma: meq-same
[X:Type]. ∀[d:metric(X)]. ∀[x:X].  x ≡ x

Lemma: meq_inversion
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  y ≡ supposing x ≡ y

Lemma: meq_weakening
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  x ≡ supposing y ∈ X

Lemma: meq_transitivity
[X:Type]. ∀[d:metric(X)]. ∀[x,y,z:X].  (x ≡ z) supposing (x ≡ and y ≡ z)

Lemma: meq_functionality
[X:Type]. ∀[d:metric(X)]. ∀[x1,x2,y1,y2:X].  (uiff(x1 ≡ y1;x2 ≡ y2)) supposing (y1 ≡ y2 and x1 ≡ x2)

Definition: mdist
mdist(d;x;y) ==  y

Lemma: mdist_wf
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  (mdist(d;x;y) ∈ ℝ)

Lemma: mdist-nonneg
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  (r0 ≤ mdist(d;x;y))

Lemma: mdist-symm
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  (mdist(d;x;y) mdist(d;y;x))

Lemma: mdist-triangle-inequality1
[X:Type]. ∀[d:metric(X)]. ∀[x,y,z:X].  (mdist(d;x;z) ≤ (mdist(d;x;y) mdist(d;z;y)))

Lemma: mdist-triangle-inequality
[X:Type]. ∀[d:metric(X)]. ∀[x,y,z:X].  (mdist(d;x;z) ≤ (mdist(d;x;y) mdist(d;y;z)))

Lemma: mdist-same
[X:Type]. ∀[d:metric(X)]. ∀[x:X].  (mdist(d;x;x) r0)

Lemma: mdist_functionality
[X:Type]. ∀[d:metric(X)]. ∀[x,y,x',y':X].  (mdist(d;x;y) mdist(d;x';y')) supposing (x ≡ x' and y ≡ y')

Lemma: meq-iff-mdist-rleq
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  (x ≡ ⇐⇒ ∀k:ℕ+(mdist(d;x;y) ≤ (r1/r(k))))

Definition: msep
==  r0 < mdist(d;x;y)

Lemma: msep_wf
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  (x y ∈ ℙ)

Lemma: msep_functionality
[X:Type]. ∀d:metric(X). ∀x,y,x',y':X.  (x ≡ x'  y ≡ y'  {x ⇐⇒ x' y'})

Lemma: sq_stable__msep
[X:Type]. ∀d:metric(X). ∀x,y:X.  SqStable(x y)

Lemma: msep-symm
[X:Type]. ∀d:metric(X). ∀x,y:X.  uiff(x y;y x)

Lemma: msep-or
[X:Type]. ∀d:metric(X). ∀x,y:X.  (x  (∀z:X. (x z ∨ y)))

Lemma: not-msep
[X:Type]. ∀[d:metric(X)]. ∀[x,y:X].  x ≡ supposing ¬y

Lemma: msep-not-meq
[X:Type]. ∀d:metric(X). ∀x,y:X.  (x  x ≡ y))

Definition: rmetric
rmetric() ==  λx,y. |x y|

Lemma: rmetric_wf
rmetric() ∈ metric(ℝ)

Lemma: rmetric-meq
[x,y:ℝ].  uiff(x ≡ y;x y)

Definition: induced-metric
induced-metric(d;f) ==  λx,y. (d (f x) (f y))

Lemma: induced-metric_wf
[X:Type]. ∀[d:metric(X)]. ∀[Y:Type]. ∀[f:Y ⟶ X].  (induced-metric(d;f) ∈ metric(Y))

Definition: induced-rmetric
induced-rmetric(f) ==  λx,y. |(f x) y|

Lemma: induced-rmetric_wf
[Y:Type]. ∀[f:Y ⟶ ℝ].  (induced-rmetric(f) ∈ metric(Y))

Definition: real-subset-metric
real-subset-metric() ==  induced-rmetric(λp.(fst(p)))

Lemma: real-subset-metric_wf
[P:ℝ ⟶ ℙ]. (real-subset-metric() ∈ metric(x:ℝ × P[x]))

Definition: scale-metric
c*d ==  λx,y. (c (d y))

Lemma: scale-metric_wf
[X:Type]. ∀[c:{c:ℝr0 ≤ c} ]. ∀[d:metric(X)].  (c*d ∈ metric(X))

Definition: metric-leq
d1 ≤ d2 ==  ∀x,y:X.  (mdist(d1;x;y) ≤ mdist(d2;x;y))

Lemma: metric-leq_wf
[X:Type]. ∀[d1,d2:metric(X)].  (d1 ≤ d2 ∈ ℙ)

Lemma: scale-metric-leq-iff
[X:Type]. ∀[d1,d2:metric(X)]. ∀[c:{c:ℝr0 < c} ].  (c*d1 ≤ d2 ⇐⇒ d1 ≤ (r1/c)*d2)

Lemma: metric-leq-meq
[X:Type]. ∀[d1,d2:metric(X)].  (d1 ≤ d2  (∀x,y:X.  (x ≡  x ≡ y)))

Definition: metric-space
MetricSpace ==  X:Type × metric(X)

Lemma: metric-space_wf
MetricSpace ∈ 𝕌'

Definition: mk-metric-space
with ==  <X, d>

Lemma: mk-metric-space_wf
[X:Type]. ∀[d:metric(X)].  (X with d ∈ MetricSpace)

Definition: prod-metric
prod-metric(k;d) ==  λv,w. Σ{mdist(d i;v i;w i) 0≤i≤1}

Lemma: prod-metric_wf
[k:ℕ]. ∀[X:ℕk ⟶ Type]. ∀[d:i:ℕk ⟶ metric(X[i])].  (prod-metric(k;d) ∈ metric(i:ℕk ⟶ X[i]))

Lemma: prod-metric-meq
[k:ℕ]. ∀[X:ℕk ⟶ Type]. ∀[d:i:ℕk ⟶ metric(X[i])]. ∀[p,q:i:ℕk ⟶ X[i]].  uiff(p ≡ q;∀i:ℕk. i ≡ i)

Definition: prod2-metric
prod2-metric(dX;dY) ==  λv,w. let x1,y1 in let x2,y2 in mdist(dX;x1;x2) mdist(dY;y1;y2)

Lemma: prod2-metric_wf
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)].  (prod2-metric(dX;dY) ∈ metric(X × Y))

Lemma: prod2-metric-meq
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)]. ∀[p,q:X × Y].  uiff(p ≡ q;fst(p) ≡ fst(q) ∧ snd(p) ≡ snd(q))

Definition: prod-metric-space
prod-metric-space(k;X) ==  <i:ℕk ⟶ (fst((X i))), prod-metric(k;λi.(snd((X i))))>

Lemma: prod-metric-space_wf
[k:ℕ]. ∀[X:ℕk ⟶ MetricSpace].  (prod-metric-space(k;X) ∈ MetricSpace)

Definition: union-metric-space
union-metric-space(T;eq;X) ==
  i:T × (fst((X i))) with λp,q. let i,v in let j,w in if eq then rmin(mdist(snd((X i));v;w);r1) else r1 fi 

Lemma: union-metric-space_wf
[T:Type]. ∀[eq:EqDecider(T)]. ∀[X:T ⟶ MetricSpace].  (union-metric-space(T;eq;X) ∈ MetricSpace)

Definition: interval-metric-space
interval-metric-space(I) ==  <{x:ℝx ∈ I} rmetric()>

Lemma: interval-metric-space_wf
[I:Interval]. (interval-metric-space(I) ∈ MetricSpace)

Definition: unit-interval-ms
==  interval-metric-space([r0, r1])

Lemma: unit-interval-ms_wf
I ∈ MetricSpace

Definition: unit-prod
(X;d) ==  prod-metric-space(2;λi.if (i =z 0) then else <X, d> fi )

Lemma: unit-prod_wf
[X:Type]. ∀[d:metric(X)].  (I (X;d) ∈ MetricSpace)

Lemma: mdist-difference
[X:Type]. ∀[d:metric(X)]. ∀[x,a,b:X].  (|mdist(d;x;a) mdist(d;x;b)| ≤ mdist(d;a;b))

Lemma: mdist-difference2
[X:Type]. ∀[d:metric(X)]. ∀[x,a,b,y:X].  (|mdist(d;x;y) mdist(d;a;b)| ≤ (mdist(d;x;a) mdist(d;y;b)))

Definition: is-mfun
f:FUN(X;Y) ==  ∀x1,x2:X.  (x1 ≡ x2  f[x1] ≡ f[x2])

Lemma: is-mfun_wf
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)]. ∀[f:X ⟶ Y].  (f:FUN(X;Y) ∈ ℙ)

Lemma: is-mfun-compose
[X,Y,Z:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)]. ∀[d'':metric(Z)]. ∀[f:X ⟶ Y]. ∀[g:Y ⟶ Z].
  (g  f:FUN(X;Z)) supposing (g:FUN(Y;Z) and f:FUN(X;Y))

Lemma: stable__is-mfun
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)]. ∀[f:X ⟶ Y].  Stable{f:FUN(X;Y)}

Lemma: sq_stable__is-mfun
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)]. ∀[f:X ⟶ Y].  SqStable(f:FUN(X;Y))

Definition: mfun
FUN(X ⟶ Y) ==  {f:X ⟶ Y| f:FUN(X;Y)} 

Lemma: mfun_wf
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)].  (FUN(X ⟶ Y) ∈ Type)

Lemma: mfun-subtype
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)]. ∀[A:Type].  FUN(X ⟶ A) ⊆FUN(X ⟶ Y) supposing A ⊆Y

Lemma: mfun-subtype2
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)]. ∀[A:Type].  FUN(X ⟶ Y) ⊆FUN(A ⟶ Y) supposing A ⊆X

Lemma: id-mfun
[X:Type]. ∀[d:metric(X)].  x.x ∈ FUN(X ⟶ X))

Lemma: compose-mfun
[X,Y,Z:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)]. ∀[dZ:metric(Z)]. ∀[f:FUN(X ⟶ Y)]. ∀[g:FUN(Y ⟶ Z)].
  (g f ∈ FUN(X ⟶ Z))

Definition: mcompose
mcompose(f;g) ==  f

Lemma: mcompose_wf
[X,Y,Z:Type].
  ∀dx:metric(X). ∀dy:metric(Y). ∀dz:metric(Z).  ∀[f:FUN(X ⟶ Y)]. ∀[g:FUN(Y ⟶ Z)].  (mcompose(f;g) ∈ FUN(X ⟶ Z))

Lemma: simple-glueing
[X:Type]. ∀[dX:metric(X)]. ∀[A,B:X ⟶ ℙ].
  ((∀x1,x2:X.  (x1 ≡ x2  A[x1]  A[x2]))
   (∀x1,x2:X.  (x1 ≡ x2  B[x1]  B[x2]))
   (∀x:X. (A[x] ∨ B[x]))
   (∀[Y:Type]. ∀[dY:metric(Y)].
        ∀f:FUN({x:X| A[x]}  ⟶ Y). ∀g:FUN({x:X| B[x]}  ⟶ Y).
          ∃h:FUN(X ⟶ Y). ∀x:X. ((A[x]  x ≡ x) ∧ (B[x]  x ≡ x)) 
          supposing ∀x:X. ((A[x] ∧ B[x])  x ≡ x)))

Definition: homeomorphic
homeomorphic(X;dX;Y;dY) ==  ∃f:FUN(X ⟶ Y). (∃g:FUN(Y ⟶ X) [((∀x:X. (f x) ≡ x) ∧ (∀y:Y. (g y) ≡ y))])

Lemma: homeomorphic_wf
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)].  (homeomorphic(X;dX;Y;dY) ∈ ℙ)

Lemma: homeomorphic_functionality
[X,Y,X',Y':Type].
  (∀[dX:metric(X)]. ∀[dY:metric(Y)].  homeomorphic(X;dX;Y;dY) ≡ homeomorphic(X';dX;Y';dY)) supposing (X ≡ X' and Y ≡ Y')

Lemma: homeomorphic_weakening
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)].  (X ≡  (dX dY ∈ metric(X))  homeomorphic(X;dX;Y;dY))

Definition: homeo-inv
homeo-inv(h) ==  let f,g in <g, f>

Lemma: homeo-inv_wf
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)]. ∀[h:homeomorphic(X;dX;Y;dY)].  (homeo-inv(h) ∈ homeomorphic(Y;dY;X;dX))

Lemma: homeomorphic_inversion
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)].  (homeomorphic(X;dX;Y;dY)  homeomorphic(Y;dY;X;dX))

Lemma: homeomorphic_transitivity
[X,Y,Z:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)]. ∀[dZ:metric(Z)].
  (homeomorphic(X;dX;Y;dY)  homeomorphic(Y;dY;Z;dZ)  homeomorphic(X;dX;Z;dZ))

Definition: homeomorphic+
homeomorphic+(X;dX;Y;dY) ==
  ∃f:FUN(X ⟶ Y)
   ((∃g:FUN(Y ⟶ X) [((∀x:X. (f x) ≡ x) ∧ (∀y:Y. (g y) ≡ y))])
   ∧ (∃B:ℕ+ [(∀x1,x2:X.  (mdist(dY;f x1;f x2) ≤ (r(B) mdist(dX;x1;x2))))]))

Lemma: homeomorphic+_wf
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)].  (homeomorphic+(X;dX;Y;dY) ∈ ℙ)

Definition: is-msfun
is-msfun(X;d;Y;d';f) ==  ∀x1,x2:X.  (f[x1] f[x2]  x1 x2)

Lemma: is-msfun_wf
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)]. ∀[f:X ⟶ Y].  (is-msfun(X;d;Y;d';f) ∈ ℙ)

Lemma: sq_stable__is-msfun
[X,Y:Type].  ∀d:metric(X). ∀[d':metric(Y)]. ∀[f:X ⟶ Y].  SqStable(is-msfun(X;d;Y;d';f))

Definition: msfun
msfun(X;d;Y;d') ==  {f:X ⟶ Y| is-msfun(X;d;Y;d';f)} 

Lemma: msfun_wf
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)].  (msfun(X;d;Y;d') ∈ Type)

Definition: m-cont-real-fun
m-cont-real-fun(X;d;x.f[x]) ==
  ∀e:{e:ℝr0 < e} . ∃delta:{e:ℝr0 < e} . ∀x,x':X.  ((mdist(d;x;x') < delta)  (|f[x] f[x']| < e))

Lemma: m-cont-real-fun_wf
[X:Type]. ∀[d:metric(X)]. ∀[f:X ⟶ ℝ].  (m-cont-real-fun(X;d;x.f[x]) ∈ ℙ)

Lemma: m-cont-real-fun-is-mfun
[X:Type]. ∀[d:metric(X)]. ∀[f:X ⟶ ℝ].  (m-cont-real-fun(X;d;x.f[x])  λx.f[x]:FUN(X;ℝ))

Lemma: mdist-m-cont
[X:Type]. ∀d:metric(X). ∀a:X.  m-cont-real-fun(X;d;x.mdist(d;a;x))

Definition: m-cont-metric-fun
m-cont-metric-fun(X;dX;Y;dY;x.f[x]) ==
  ∀e:{e:ℝr0 < e} . ∃delta:{e:ℝr0 < e} . ∀x,x':X.  ((mdist(dX;x;x') < delta)  (mdist(dY;f[x];f[x']) < e))

Lemma: m-cont-metric-fun_wf
[X:Type]. ∀[dX:metric(X)]. ∀[Y:Type]. ∀[dY:metric(Y)]. ∀[f:X ⟶ Y].  (m-cont-metric-fun(X;dX;Y;dY;x.f[x]) ∈ ℙ)

Definition: meqfun
meqfun(d;A;f;g) ==  ∀a:A. a ≡ a

Lemma: meqfun_wf
[A,X:Type]. ∀[d:metric(X)]. ∀[f,g:A ⟶ X].  (meqfun(d;A;f;g) ∈ ℙ)

Lemma: meqfun-equiv-rel
[A,X:Type]. ∀[d:metric(X)].  EquivRel(A ⟶ X;f,g.meqfun(d;A;f;g))

Lemma: meqfun-equiv-rel-mfun
[A,X:Type]. ∀[dA:metric(A)]. ∀[d:metric(X)].  EquivRel(FUN(A ⟶ X);f,g.meqfun(d;A;f;g))

Definition: metric-subspace
metric-subspace(X;d;A) ==  strong-subtype(A;X) ∧ (∀a:A. ∀x:X.  (x ≡  (x ∈ A)))

Lemma: metric-subspace_wf
[X:Type]. ∀[d:metric(X)]. ∀[A:Type].  (metric-subspace(X;d;A) ∈ ℙ)

Lemma: sq_stable__metric-subspace
[X:Type]. ∀[d:metric(X)]. ∀[A:Type].  SqStable(metric-subspace(X;d;A))

Lemma: set-metric-subspace
[X:Type]. ∀[d:metric(X)]. ∀[P:X ⟶ ℙ].  metric-subspace(X;d;{x:X| P[x]} supposing ∀x,y:X.  (P[x]  y ≡  P[y])

Lemma: set-metric-subspace2
[X:Type]. ∀[d:metric(X)]. ∀[P,Q:X ⟶ ℙ].
  (metric-subspace({x:X| Q[x]} ;d;{x:X| P[x]} )) supposing ((∀x,y:X.  (P[x]  y ≡  P[y])) and (∀x:X. (P[x]  Q[x])\000C))

Lemma: stable-union-metric-subspace
[X:Type]. ∀[d:metric(X)]. ∀[T:Type]. ∀[P:T ⟶ X ⟶ ℙ].
  metric-subspace(X;d;stable-union(X;T;i,x.P[i;x])) supposing ∀i:T. ∀x,y:X.  (P[i;x]  y ≡  P[i;y])

Definition: m-interior-point
m-interior-point(X;d;A;p) ==  ∃M:ℕ+. ∀x:X. ((mdist(d;x;p) ≤ (r1/r(M)))  (x ∈ A))

Lemma: m-interior-point_wf
[X,A:Type].  ∀[d:metric(X)]. ∀[p:A].  (m-interior-point(X;d;A;p) ∈ ℙsupposing strong-subtype(A;X)

Definition: m-interior
m-interior(X;d;A) ==  {p:A| m-interior-point(X;d;A;p)} 

Lemma: m-interior_wf
[X,A:Type].  ∀[d:metric(X)]. (m-interior(X;d;A) ∈ Type) supposing strong-subtype(A;X)

Definition: m-boundary
m-boundary(X;d;A) ==  {p:A| ¬m-interior-point(X;d;A;p)} 

Lemma: m-boundary_wf
[X,A:Type].  ∀[d:metric(X)]. (m-boundary(X;d;A) ∈ Type) supposing strong-subtype(A;X)

Definition: m-ball
m-ball(X;d;c;r) ==  {x:X| mdist(d;c;x) ≤ r} 

Lemma: m-ball_wf
[X:Type]. ∀[d:metric(X)]. ∀[c:X]. ∀[r:ℝ].  (m-ball(X;d;c;r) ∈ Type)

Definition: m-open-ball
m-open-ball(X;d;c;r) ==  {x:X| mdist(d;c;x) < r} 

Lemma: m-open-ball_wf
[X:Type]. ∀[d:metric(X)]. ∀[c:X]. ∀[r:ℝ].  (m-open-ball(X;d;c;r) ∈ Type)

Definition: m-sphere
m-sphere(X;d;c;r) ==  {x:X| mdist(d;c;x) r} 

Lemma: m-sphere_wf
[X:Type]. ∀[d:metric(X)]. ∀[c:X]. ∀[r:ℝ].  (m-sphere(X;d;c;r) ∈ Type)

Lemma: m-ball-boundary-subtype-m-sphere
[X:Type]. ∀[d:metric(X)]. ∀[c:X]. ∀[r:ℝ].  (m-boundary(X;d;m-ball(X;d;c;r)) ⊆m-sphere(X;d;c;r))

Lemma: m-sphere-subtype-m-ball-boundary
[X:Type]. ∀[d:metric(X)]. ∀[c:X]. ∀[r:ℝ].
  m-sphere(X;d;c;r) ⊆m-boundary(X;d;m-ball(X;d;c;r)) 
  supposing ∀c,x:X. ∀M:ℕ+.
              ∃y:X. ((mdist(d;c;y) (mdist(d;c;x) mdist(d;x;y))) ∧ (r0 < mdist(d;y;x)) ∧ (mdist(d;y;x) ≤ (r1/r(M))))

Lemma: m-ball-boundary
[X:Type]. ∀[d:metric(X)]. ∀[c:X]. ∀[r:ℝ].
  m-boundary(X;d;m-ball(X;d;c;r)) ≡ m-sphere(X;d;c;r) 
  supposing ∀c,x:X. ∀M:ℕ+.
              ∃y:X. ((mdist(d;c;y) (mdist(d;c;x) mdist(d;x;y))) ∧ (r0 < mdist(d;y;x)) ∧ (mdist(d;y;x) ≤ (r1/r(M))))

Definition: m-closed-subspace
m-closed-subspace(X;d;A) ==  ∀x:X. ((∀k:ℕ+. ∃a:A. (mdist(d;x;a) ≤ (r1/r(k))))  (x ∈ A))

Lemma: m-closed-subspace_wf
[X:Type]. ∀[d:metric(X)]. ∀[A:Type].  m-closed-subspace(X;d;A) ∈ ℙ supposing metric-subspace(X;d;A)

Lemma: mfun-strong-subtype
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)]. ∀[A:Type].
  strong-subtype(FUN(X ⟶ A);FUN(X ⟶ Y)) supposing metric-subspace(Y;d';A)

Definition: mfun-class
mfun-class(X;dx;Y;dy) ==  f,g:FUN(X ⟶ Y)//meqfun(dy;X;f;g)

Lemma: mfun-class_wf
[X:Type]. ∀[dx:metric(X)]. ∀[Y:Type]. ∀[dy:metric(Y)].  (mfun-class(X;dx;Y;dy) ∈ Type)

Lemma: mfun-class-strong-subtype
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)]. ∀[A:Type].
  strong-subtype(mfun-class(X;d;A;d');mfun-class(X;d;Y;d')) supposing metric-subspace(Y;d';A)

Definition: image-space
f[X] ==  y:Y × {x:X| y ≡ x} 

Lemma: image-space_wf
[X,Y:Type]. ∀[dY:metric(Y)]. ∀[f:X ⟶ Y].  (f[X] ∈ Type)

Definition: image-ap
f[x] ==  <x, x>

Lemma: image-ap_wf
[X,Y:Type]. ∀[d:metric(Y)]. ∀[f:X ⟶ Y]. ∀[x:X].  (f[x] ∈ f[X])

Definition: image-metric
image-metric(d) ==  λp,q. (d (fst(p)) (fst(q)))

Lemma: image-metric_wf
[X,Y:Type]. ∀[f:X ⟶ Y]. ∀[d:metric(Y)].  (image-metric(d) ∈ metric(f[X]))

Definition: m-unif-cont
UC(f:X ⟶ Y) ==  ∀k:ℕ+. ∃delta:{d:ℝr0 < d} . ∀x,y:X.  ((mdist(dx;x;y) ≤ delta)  (mdist(dy;f x;f y) ≤ (r1/r(k))))

Lemma: m-unif-cont_wf
[X:Type]. ∀[dx:metric(X)]. ∀[Y:Type]. ∀[dy:metric(Y)]. ∀[f:X ⟶ Y].  (UC(f:X ⟶ Y) ∈ ℙ)

Definition: m-ptwise-cont
PtwiseCONT(f:X ⟶ Y) ==
  ∀x:X. ∀k:ℕ+.  ∃delta:{d:ℝr0 < d} . ∀y:X. ((mdist(dx;x;y) ≤ delta)  (mdist(dy;f x;f y) ≤ (r1/r(k))))

Lemma: m-ptwise-cont_wf
[X:Type]. ∀[dx:metric(X)]. ∀[Y:Type]. ∀[dy:metric(Y)]. ∀[f:X ⟶ Y].  (PtwiseCONT(f:X ⟶ Y) ∈ ℙ)

Definition: m-retraction
Retract(X ⟶ A) ==  {f:X ⟶ A| f:FUN(X;A) ∧ (∀a:A. a ≡ a)} 

Lemma: m-retraction_wf
[X,A:Type]. ∀[d:metric(X)].  Retract(X ⟶ A) ∈ Type supposing A ⊆X

Lemma: m-retraction_functionality
[X,Y:Type].
  ∀[A,B:Type].  ∀[d:metric(X)]. Retract(X ⟶ A) ≡ Retract(Y ⟶ B) supposing A ⊆supposing A ≡ supposing X ≡ Y

Lemma: m-retraction-subtype1
[X,Y,A:Type].  ∀[d:metric(X)]. Retract(X ⟶ A) ⊆Retract(Y ⟶ A) supposing A ⊆supposing Y ⊆X

Lemma: m-retraction-subtype2
[X,A,B:Type].  ∀[d:metric(X)]. Retract(X ⟶ A) ⊆Retract(X ⟶ B) supposing B ⊆supposing A ≡ B

Definition: homeo-image
homeo-image(A;Y;dY;h) ==  {y:Y| ∃a:A. y ≡ (fst(h)) a} 

Lemma: homeo-image_wf
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)]. ∀[h:homeomorphic(X;dX;Y;dY)]. ∀[A:Type].
  homeo-image(A;Y;dY;h) ∈ Type supposing A ⊆X

Lemma: homeo-image-inverse
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)]. ∀[h:homeomorphic(X;dX;Y;dY)]. ∀[A:Type].
  homeo-image(homeo-image(A;Y;dY;h);X;dX;homeo-inv(h)) ≡ supposing metric-subspace(X;dX;A)

Lemma: trivial-homeo-image
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)]. ∀[h:homeomorphic(X;dX;Y;dY)].  homeo-image(X;Y;dY;h) ≡ Y

Lemma: homeo-image-homeomorphic-subtype
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)]. ∀[h:homeomorphic(X;dX;Y;dY)]. ∀[A:Type].
  h ∈ homeomorphic(A;dX;homeo-image(A;Y;dY;h);dY) supposing metric-subspace(X;dX;A)

Lemma: homeo-image-homeomorphic
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)].
  ∀h:homeomorphic(X;dX;Y;dY). ∀[A:Type]. homeomorphic(A;dX;homeo-image(A;Y;dY;h);dY) supposing metric-subspace(X;dX;A)

Lemma: homeo-image-boundary
[X,Y:Type]. ∀[dX:metric(X)]. ∀[dY:metric(Y)].
  ∀h:homeomorphic(X;dX;Y;dY)
    ∀[A:Type]
      homeo-image(m-boundary(X;dX;A);Y;dY;h) ≡ m-boundary(Y;dY;homeo-image(A;Y;dY;h)) supposing metric-subspace(X;dX;A) 
    supposing PtwiseCONT(fst(h):X ⟶ Y) ∧ PtwiseCONT(snd(h):Y ⟶ X)

Lemma: homeomorphic-retraction
[X,A:Type]. ∀[d:metric(X)].
  Retract(X ⟶ A)  (∀Y:Type. ∀d':metric(Y). ∀h:homeomorphic(X;d;Y;d').  Retract(Y ⟶ homeo-image(A;Y;d';h))) 
  supposing metric-subspace(X;d;A)

Definition: fixedpoint-property
FP(X) ==  ∀f:FUN(X ⟶ X). ∀n:ℕ+.  ∃x:X. (mdist(d;f x;x) ≤ (r1/r(n)))

Lemma: fixedpoint-property_wf
[X:Type]. ∀[d:metric(X)].  (FP(X) ∈ ℙ)

Definition: m-open
m-open(X;d;x.A[x]) ==  ∀x:X. (A[x]  (∃k:ℕ+. ∀y:X. ((mdist(d;x;y) ≤ (r1/r(k)))  A[y])))

Lemma: m-open_wf
[X:Type]. ∀[d:metric(X)]. ∀[A:X ⟶ ℙ].  (m-open(X;d;x.A[x]) ∈ ℙ)

Definition: m-set
m-set(X;d;x.A[x]) ==  ∀x,y:X.  (x ≡  (A[x] ⇐⇒ A[y]))

Lemma: m-set_wf
[X:Type]. ∀[d:metric(X)]. ∀[A:X ⟶ ℙ].  (m-set(X;d;x.A[x]) ∈ ℙ)

Lemma: m-open-set
[X:Type]. ∀[d:metric(X)]. ∀[A:X ⟶ ℙ].  (m-open(X;d;x.A[x])  m-set(X;d;x.A[x]))

Definition: m-open-cover
m-open-cover(X;d;I;i,x.A[i; x]) ==  (∀i:I. m-open(X;d;x.A[i; x])) ∧ (∀x:X. ∃i:I. A[i; x])

Lemma: m-open-cover_wf
[X:Type]. ∀[d:metric(X)]. ∀[I:Type]. ∀[A:I ⟶ X ⟶ ℙ].  (m-open-cover(X;d;I;i,x.A[i;x]) ∈ ℙ)

Lemma: m-open-cover-iff
[X:Type]. ∀[d:metric(X)]. ∀[I:Type]. ∀[A:I ⟶ X ⟶ ℙ].
  (m-open-cover(X;d;I;i,x.A[i;x])
  ⇐⇒ (∀i:I. m-open(X;d;x.A[i;x])) ∧ (∃b:X ⟶ ℕ+. ∃c:X ⟶ I. ∀x,y:X.  ((mdist(d;x;y) ≤ (r1/r(b x)))  A[c x;y])))

Definition: mcauchy
mcauchy(d;n.x[n]) ==  ∀k:ℕ+(∃N:ℕ [(∀n,m:ℕ.  ((N ≤ n)  (N ≤ m)  (mdist(d;x[n];x[m]) ≤ (r1/r(k)))))])

Lemma: mcauchy_wf
[X:Type]. ∀[d:X ⟶ X ⟶ ℝ]. ∀[x:ℕ ⟶ X].  (mcauchy(d;n.x[n]) ∈ ℙ)

Definition: mconverges-to
lim n→∞.x[n] ==  ∀k:ℕ+(∃N:ℕ [(∀n:ℕ((N ≤ n)  (mdist(d;x[n];y) ≤ (r1/r(k)))))])

Lemma: mconverges-to_wf
[X:Type]. ∀[d:X ⟶ X ⟶ ℝ]. ∀[x:ℕ ⟶ X]. ∀[y:X].  (lim n→∞.x[n] y ∈ ℙ)

Lemma: mconverges-to_functionality
[X:Type]. ∀d:metric(X). ∀x1,x2:ℕ ⟶ X.  ∀[y:X]. {lim n→∞.x1[n]  lim n→∞.x2[n] y} supposing ∀n:ℕx1[n] ≡ x2[n]

Lemma: constant-mconverges-to
[X:Type]. ∀[d:metric(X)]. ∀[y:X].  lim n→∞.y y

Lemma: subsequence-mconverges-to
[X:Type]. ∀[d:metric(X)]. ∀[a:X].
  ∀x,y:ℕ ⟶ X.  (subsequence(a,b.a ≡ b;n.x[n];n.y[n])  lim n→∞.x[n]  lim n→∞.y[n] a)

Lemma: m-unique-limit
[X:Type]. ∀d:metric(X). ∀x:ℕ ⟶ X.  ∀[y1,y2:X].  (y1 ≡ y2) supposing (lim n→∞.x[n] y1 and lim n→∞.x[n] y2)

Definition: mconverges
x[n]↓ as n→∞ ==  ∃y:X. lim n→∞.x[n] y

Lemma: mconverges_wf
[X:Type]. ∀[d:X ⟶ X ⟶ ℝ]. ∀[x:ℕ ⟶ X].  (x[n]↓ as n→∞ ∈ ℙ)

Lemma: mconverges-implies-mcauchy
[X:Type]. ∀[d:metric(X)]. ∀[x:ℕ ⟶ X].  (x[n]↓ as n→∞  mcauchy(d;n.x[n]))

Lemma: subsequence-mconverges
[X:Type]. ∀[d:metric(X)]. ∀[a:X].  ∀x,y:ℕ ⟶ X.  (subsequence(a,b.a ≡ b;n.x[n];n.y[n])  x[n]↓ as n→∞  y[n]↓ as n→∞)

Definition: mdiverges
n.x[n]↑ ==  ∃e:ℝ((r0 < e) ∧ (∀k:ℕ. ∃m,n:ℕ((k ≤ m) ∧ (k ≤ n) ∧ (e ≤ mdist(d;x[m];x[n])))))

Lemma: mdiverges_wf
[X:Type]. ∀[d:X ⟶ X ⟶ ℝ]. ∀[x:ℕ ⟶ X].  (n.x[n]↑ ∈ ℙ)

Definition: mcomplete
mcomplete(M) ==  let X,d in ∀x:ℕ ⟶ X. (mcauchy(d;n.x n)  n↓ as n→∞)

Lemma: mcomplete_wf
[M:MetricSpace]. (mcomplete(M) ∈ ℙ)

Definition: cauchy-mlimit
cauchy-mlimit(cmplt;x;c) ==  fst((cmplt c))

Lemma: cauchy-mlimit_wf
[X:Type]. ∀[d:metric(X)]. ∀[cmplt:mcomplete(X with d)]. ∀[x:ℕ ⟶ X]. ∀[c:mcauchy(d;n.x n)].
  (cauchy-mlimit(cmplt;x;c) ∈ X)

Lemma: converges-to-cauchy-mlimit
[X:Type]
  ∀d:metric(X). ∀cmplt:mcomplete(X with d). ∀x:ℕ ⟶ X. ∀c:mcauchy(d;n.x n).  lim n→∞.x cauchy-mlimit(cmplt;x;c)

Lemma: cauchy-mlimit-unique
[X:Type]. ∀[d:metric(X)]. ∀[cmplt:mcomplete(X with d)]. ∀[x:ℕ ⟶ X]. ∀[c:mcauchy(d;n.x n)]. ∀[z:X].
  cauchy-mlimit(cmplt;x;c) ≡ supposing lim n→∞.x z

Lemma: reals-complete
mcomplete(ℝ with rmetric())

Lemma: reals-complete-ext
mcomplete(ℝ with rmetric())

Lemma: real-interval-complete
a,b:ℝ.  mcomplete({x:ℝx ∈ [a, b]}  with rmetric())

Lemma: metric-leq-cauchy
[X:Type]. ∀[d1,d2:metric(X)].  ∀c:{c:ℝr0 < c} (d1 ≤ c*d2  (∀x:ℕ ⟶ X. (mcauchy(d2;n.x n)  mcauchy(d1;n.x n))))

Lemma: scale-metric-cauchy
[X:Type]. ∀[d:metric(X)].  ∀c:{c:ℝr0 < c} . ∀x:ℕ ⟶ X.  (mcauchy(d;n.x n) ⇐⇒ mcauchy(c*d;n.x n))

Lemma: scale-metric-converges
[X:Type]. ∀[d:metric(X)].  ∀c:{c:ℝr0 < c} . ∀x:ℕ ⟶ X. ∀y:X.  (lim n→∞.x ⇐⇒ lim n→∞.x y)

Lemma: scale-metric-complete
[X:Type]. ∀[d:metric(X)].  ∀c:{c:ℝr0 < c} (mcomplete(X with d) ⇐⇒ mcomplete(X with c*d))

Lemma: prod-metric-space-complete
k:ℕ. ∀X:ℕk ⟶ MetricSpace.  ((∀i:ℕk. mcomplete(X i))  mcomplete(prod-metric-space(k;X)))

Lemma: union-metric-space-complete
T:Type. ∀eq:EqDecider(T). ∀X:T ⟶ MetricSpace.  ((∀i:T. mcomplete(X i))  mcomplete(union-metric-space(T;eq;X)))

Lemma: metric-weak-Markov
[X:Type]. ∀d:metric(X). (mcomplete(X with d)  (∀x,y:X.  ((∀z:X. ((¬z ≡ x) ∨ z ≡ y)))  y)))

Lemma: metric-strong-extensionality
[X:Type]
  ∀d:metric(X)
    (mcomplete(X with d)
     (∀Y:Type. ∀d':metric(Y). ∀f:X ⟶ Y.
          ((∀x1,x2:X.  (x1 ≡ x2  f[x1] ≡ f[x2]))  (∀x1,x2:X.  (f[x1] f[x2]  x1 x2)))))

Lemma: m-strong-extensionality
[X:Type]. ∀d:metric(X). (mcomplete(X with d)  (∀Y:Type. ∀d':metric(Y). ∀f:FUN(X ⟶ Y).  is-msfun(X;d;Y;d';f)))

Lemma: msfun-ext-mfun
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)].  (mcomplete(X with d)  msfun(X;d;Y;d') ≡ FUN(X ⟶ Y))

Lemma: punctured-homeomorphism
[X,Y:Type]. ∀[d:metric(X)]. ∀[d':metric(Y)].
  (mcomplete(X with d)
   mcomplete(Y with d')
   (∀h:homeomorphic(X;d;Y;d'). ∀p:Y.  (h ∈ homeomorphic({x:X| (snd(h)) p} ;d;{y:Y| p} ;d'))))

Lemma: metric-leq-converges-to
[X:Type]. ∀[d1,d2:metric(X)].  (d2 ≤ d1  (∀x:ℕ ⟶ X. ∀y:X.  (lim n→∞.x[n]  lim n→∞.x[n] y)))

Lemma: metric-leq-complete
[X:Type]. ∀[d1,d2:metric(X)].
  (d2 ≤ d1
   (∀x:ℕ ⟶ X. (mcauchy(d2;n.x n)  (∃y:ℕ ⟶ X. (subsequence(a,b.a ≡ b;n.x n;n.y n) ∧ mcauchy(d1;n.y n)))))
   mcomplete(X with d1)
   mcomplete(X with d2))

Lemma: equiv-metrics-preserve-complete
[X:Type]. ∀[d1,d2:metric(X)].
  ((∃c1,c2:{s:ℝr0 < s} (c1*d1 ≤ d2 ∧ c2*d2 ≤ d1))  (mcomplete(X with d1) ⇐⇒ mcomplete(X with d2)))

Lemma: m-closed-iff-complete
[X:Type]
  ∀d:metric(X)
    (mcomplete(X with d)  (∀[A:Type]. (metric-subspace(X;d;A)  (m-closed-subspace(X;d;A) ⇐⇒ mcomplete(A with d)))))

Definition: m-TB
This is useful formulation of "metric total boundedness".
The lemma m-TB-iff proves that it is equivalent to
k:ℕ. ∃n:ℕ+. ∃xs:ℕn ⟶ X. ∀x:X. ∃i:ℕn. (mdist(d;x;xs i) ≤ (r1/r(k 1)))

This says that for every epsilon (1/(k+1)) there is finite collection, xs,
of points in (some might call it "subfinite" collection) such that
every point in is within distance epsilon from some point in xs.

The constructive content of this is
1) the function from in ℕ to in ℕ+
2) for given k, the xs:ℕk ⟶ X
3) the "decider" or "chooser" that for given in finds
   the index in ℕof the point in the xs for which the distance to is
   less than epsilon.⋅

m-TB(X;d) ==
  {tb:B:ℕ ⟶ ℕ+ × k:ℕ ⟶ ℕk ⟶ X × (X ⟶ k:ℕ ⟶ ℕk)| 
   let B,xs,c tb in 
   ∀p:X. ∀k:ℕ.  (mdist(d;p;xs (c k)) ≤ (r1/r(k 1)))} 

Lemma: m-TB_wf
[X:Type]. ∀[d:metric(X)].  (m-TB(X;d) ∈ Type)

Lemma: m-TB-iff
[X:Type]. ∀[d:metric(X)].  (m-TB(X;d) ⇐⇒ ∀k:ℕ. ∃n:ℕ+. ∃xs:ℕn ⟶ X. ∀x:X. ∃i:ℕn. (mdist(d;x;xs i) ≤ (r1/r(k 1))))

Lemma: real-interval-m-TB
a:ℝ. ∀b:{b:ℝa ≤ b} .  m-TB({x:ℝx ∈ [a, b]} ;rmetric())

Lemma: m-TB-product
m:ℕ. ∀[X:ℕm ⟶ Type]. ∀[d:i:ℕm ⟶ metric(X[i])].  ((∀i:ℕm. m-TB(X[i];d[i]))  m-TB(i:ℕm ⟶ X[i];prod-metric(m;d)))

Lemma: continuous-image-m-TB
[X:Type]
  ∀dX:metric(X). ∀[Y:Type]. ∀dY:metric(Y). ∀f:X ⟶ Y.  (UC(f:X ⟶ Y)  m-TB(X;dX)  m-TB(f[X];image-metric(dY)))

Lemma: m-TB-sup-and-inf
[X:Type]
  ∀dX:metric(X)
    (m-TB(X;dX)
     (∀f:X ⟶ ℝ(UC(f:X ⟶ ℝ ((∃a:ℝinf(λr.∃x:X. (r (f x))) a) ∧ (∃b:ℝsup(λr.∃x:X. (r (f x))) b)))))

Definition: m-inf
m-inf{i:l}(d;mtb;f;mc) ==  fst(fst((TERMOF{m-TB-sup-and-inf:o, 1:l, i:l} mtb mc)))

Lemma: m-inf_wf
[X:Type]. ∀[d:metric(X)]. ∀[mtb:m-TB(X;d)]. ∀[f:X ⟶ ℝ]. ∀[mc:UC(f:X ⟶ ℝ)].  (m-inf{i:l}(d;mtb;f;mc) ∈ ℝ)

Lemma: m-inf-property1
[X:Type]
  ∀d:metric(X). ∀mtb:m-TB(X;d). ∀f:X ⟶ ℝ. ∀mc:UC(f:X ⟶ ℝ).  inf(λr.∃x:X. (r (f x))) m-inf{i:l}(d;mtb;f;mc)

Lemma: m-inf-property
[X:Type]
  ∀d:metric(X). ∀mtb:m-TB(X;d). ∀f:X ⟶ ℝ. ∀mc:UC(f:X ⟶ ℝ).
    ((∀x:X. (m-inf{i:l}(d;mtb;f;mc) ≤ (f x))) ∧ (∀e:ℝ((r0 < e)  (∃x:X. ((f x) < (m-inf{i:l}(d;mtb;f;mc) e))))))

Definition: m-sup
m-sup{i:l}(d;mtb;f;mc) ==  fst(snd((TERMOF{m-TB-sup-and-inf:o, 1:l, i:l} mtb mc)))

Lemma: m-sup_wf
[X:Type]. ∀[d:metric(X)]. ∀[mtb:m-TB(X;d)]. ∀[f:X ⟶ ℝ]. ∀[mc:UC(f:X ⟶ ℝ)].  (m-sup{i:l}(d;mtb;f;mc) ∈ ℝ)

Lemma: m-sup-property1
[X:Type]
  ∀d:metric(X). ∀mtb:m-TB(X;d). ∀f:X ⟶ ℝ. ∀mc:UC(f:X ⟶ ℝ).  sup(λr.∃x:X. (r (f x))) m-sup{i:l}(d;mtb;f;mc)

Lemma: m-sup-property
[X:Type]
  ∀d:metric(X). ∀mtb:m-TB(X;d). ∀f:X ⟶ ℝ. ∀mc:UC(f:X ⟶ ℝ).
    ((∀x:X. ((f x) ≤ m-sup{i:l}(d;mtb;f;mc))) ∧ (∀e:ℝ((r0 < e)  (∃x:X. ((m-sup{i:l}(d;mtb;f;mc) e) < (f x))))))

Definition: mtb-cantor
When mtb is witness to metric space being totally bounded,
the first component of mtb is function of type ⌜ℕ ⟶ ℕ+.
This gives us "general Cantor space" k:ℕ ⟶ ℕB[k]⋅

mtb-cantor(mtb) ==  k:ℕ ⟶ ℕ(fst(mtb)) k

Lemma: mtb-cantor_wf
[X:Type]. ∀[d:metric(X)]. ∀[mtb:m-TB(X;d)].  (mtb-cantor(mtb) ∈ Type)

Definition: mtb-seq
When mtb is witness that metric space is totally bounded
and is point in the general Cantor space, mtb-cantor(mtb),
then at each natural number there are finitely many points xs(k) 
(that every other point in is within 1/(k+1) of one of them), 
and s(k) is the index of one of these points.

Thus we get sequence of points in X.
Not every such sequence will converge in X, but for every point p
in there is always one, coming from mtb-point-cantor(mtb;p),
that does converge to (see mtb-seq-mtb-point-cantor-mconverges-to).


mtb-seq(mtb;s) ==  let B,xs,c mtb in λk.(xs (s k))

Lemma: mtb-seq_wf
[X:Type]. ∀[d:metric(X)]. ∀[mtb:m-TB(X;d)]. ∀[s:mtb-cantor(mtb)].  (mtb-seq(mtb;s) ∈ ℕ ⟶ X)

Definition: mtb-point-cantor
For point in X, this gives the sequence of indexes to the points
that are better and better "approximations" to p.
So, it is point in the general Cantor space mtb-cantor(mtb)
and mtb-seq(mtb;s) is Cauchy sequence (in fact, regular sequence) that
converges to p.⋅

mtb-point-cantor(mtb;p) ==  let B,xs,c mtb in p

Lemma: mtb-point-cantor_wf
[X:Type]. ∀[d:metric(X)]. ∀[mtb:m-TB(X;d)]. ∀[p:X].  (mtb-point-cantor(mtb;p) ∈ mtb-cantor(mtb))

Lemma: mtb-seq-mtb-point-cantor-mconverges-to
[X:Type]. ∀d:metric(X). ∀mtb:m-TB(X;d). ∀x:X.  lim n→∞.mtb-seq(mtb;mtb-point-cantor(mtb;x)) x

Definition: m-k-regular
This is the "metric" version of the notion of (k)regular sequence
(that was used to define the real numbers).
Every (k)regular sequence is Cauchy sequence (m-k-regular-mcauchy)
so it will converge if the metric space is complete.⋅

m-k-regular(d;k;s) ==  ∀n,m:ℕ.  (mdist(d;s n;s m) ≤ ((r(k)/r(n 1)) (r(k)/r(m 1))))

Lemma: m-k-regular_wf
[X:Type]. ∀[d:metric(X)]. ∀[k:ℕ]. ∀[s:ℕ ⟶ X].  (m-k-regular(d;k;s) ∈ ℙ)

Lemma: m-k-regular-monotone
[n,k:ℕ].  ∀[X:Type]. ∀[d:metric(X)]. ∀[s:ℕ ⟶ X].  m-k-regular(d;k;s) supposing m-k-regular(d;n;s) supposing n ≤ k

Lemma: m-k-regular-mcauchy
[X:Type]. ∀[d:metric(X)]. ∀[s:ℕ ⟶ X].  ∀b:ℕ+(m-k-regular(d;b;s)  mcauchy(d;n.s n))

Definition: m-reg-test
For sequence of integers, s, we can decide is is regular upto some bound b
because n ≤ is decidable on integers.
For the metric version of regularity, we won't be able to decide
whether the distance mdist(d;x;y) is or is not less than or equal to given
(rational) real number.
But if a < then we can decide (a < z) ∨ (z < b).
So we set to be the bound for 2-regularity and the bound for 3-regularity.
Then to test whether we can add point onto sequence s,
we will either find that the extended sequence is not 2-regular,
 m-not-reg(d;s;n),or that it is (still) 3-regular.

2-regular sequence, s, has ∀n:ℕm-not-reg(d;s;n) ff,
(see m-regular-not-m-not-reg)
and ∀n:ℕb. m-not-reg(d;s;n) ff implies that is at least 3-regular
(see not-m-not-reg-3regular).⋅

m-reg-test(d;b;s;x) ==
  int-seg-case(0;b;λn.rless-case((r(2)/r(n 1)) (r(2)/r(b 1));(r(3)/r(n 1)) (r(3)/r(b 1));rlessw((r(2)/r(n
                      1))
                      (r(2)/r(b 1));(r(3)/r(n 1)) (r(3)/r(b 1)));mdist(d;s n;x)))

Lemma: m-reg-test_wf
[X:Type]
  ∀d:metric(X). ∀b:ℕ. ∀s:ℕb ⟶ X. ∀x:X.
    (m-reg-test(d;b;s;x) ∈ (∃n:ℕb. (((r(2)/r(n 1)) (r(2)/r(b 1))) < mdist(d;s n;x)))
     ∨ (∀n:ℕb. (mdist(d;s n;x) < ((r(3)/r(n 1)) (r(3)/r(b 1))))))

Definition: m-not-reg
m-not-reg(d;s;n) ==  isl(m-reg-test(d;n;s;s n))

Lemma: m-not-reg_wf
[X:Type]. ∀[d:metric(X)]. ∀[n:ℕ]. ∀[s:ℕ1 ⟶ X].  (m-not-reg(d;s;n) ∈ 𝔹)

Lemma: m-regular-not-m-not-reg
[X:Type]. ∀[d:metric(X)]. ∀[s:ℕ ⟶ X].  (m-k-regular(d;2;s)  (∀n:ℕm-not-reg(d;s;n) ff))

Lemma: not-m-not-reg-3regular
[X:Type]. ∀[d:metric(X)]. ∀[s:ℕ ⟶ X]. ∀[b:ℕ].
  ((∀n:ℕb. m-not-reg(d;s;n) ff)  (∀n,m:ℕb.  (mdist(d;s n;s m) ≤ ((r(3)/r(n 1)) (r(3)/r(m 1))))))

Definition: first-m-not-reg
We can search for the first place, up to bound k, 
where sequence is not 2-regular.
We either locate the first failure of 2-regularity
or we find that the sequence is at least 3-regular up to k.
(see first-m-not-reg-property). ⋅

first-m-not-reg(d;s;k) ==  search(k;λn.m-not-reg(d;s;n))

Lemma: first-m-not-reg_wf
[X:Type]. ∀[d:metric(X)]. ∀[k:ℕ]. ∀[s:ℕk ⟶ X].  (first-m-not-reg(d;s;k) ∈ ℕ1)

Lemma: first-m-not-reg-property
[X:Type]
  ∀d:metric(X). ∀k:ℕ. ∀s:ℕk ⟶ X.
    ((first-m-not-reg(d;s;k) 0 ∈ ℤ ⇐⇒ ∀n:ℕk. m-not-reg(d;s;n) ff)
    ∧ let first-m-not-reg(d;s;k) in
          (∀n:ℕi. m-not-reg(d;s;n) ff) ∧ m-not-reg(d;s;i) tt 
      supposing 0 < first-m-not-reg(d;s;k))

Definition: m-regularize
For any sequence of points s, we "regularize" by testing if there
is failure of 2-regularity. If so, we make the regularized sequence
become constant from that index on. Otherwise, the sequence will be
unchanged, and will be at least 3-regular.

We know that any 2-regular sequence has m-regularize(d;s) s
(see m-regularize-of-regular)
and the regularization of any sequence is Cauchy sequence
(with modulus of Cauchyness λk.(6 k) )
(see m-regularize-mcauchy)⋅

m-regularize(d;s) ==  λn.eval first-m-not-reg(d;s;n 1) in if 0 <then (m 2) else fi 

Lemma: m-regularize_wf_finite
[X:Type]. ∀[d:metric(X)]. ∀[b:ℕ]. ∀[s:ℕb ⟶ X].  (m-regularize(d;s) ∈ ℕb ⟶ X)

Lemma: m-regularize_wf
[X:Type]. ∀[d:metric(X)]. ∀[s:ℕ ⟶ X].  (m-regularize(d;s) ∈ ℕ ⟶ X)

Lemma: m-regularize-of-regular
[X:Type]. ∀[d:metric(X)]. ∀[s:ℕ ⟶ X].  m-regularize(d;s) s ∈ (ℕ ⟶ X) supposing m-k-regular(d;2;s)

Lemma: m-regularize-mcauchy
[X:Type]. ∀[d:metric(X)]. ∀[s:ℕ ⟶ X].  k.(6 k) ∈ mcauchy(d;n.m-regularize(d;s) n))

Lemma: mtb-point-cantor-seq-regular
[X:Type]. ∀[d:metric(X)].  ∀mtb:m-TB(X;d). ∀p:X.  m-k-regular(d;1;mtb-seq(mtb;mtb-point-cantor(mtb;p)))

Definition: mtb-cantor-map
Given point in mtb-cantor(mtb), the sequence mtb-seq(mtb;p) may not
be Cauchy, but its regularization m-regularize(d;mtb-seq(mtb;p)) is
(and λk.(6 k) is witness -- modulus of Cauchyness -- for that).
Thus (given witness cmplt for the completenss of X) it converges
to point in X.

So we get map from the general Cantor space, mtb-cantor(mtb), to X.
It is onto X  (mtb-cantor-map-ontoand is uniformly
continuous (mtb-cantor-map-continuous).
Hence, every compact metric space is the continuous image of Cantor
space.

We get an even stronger version mtb-cantor-map-onto-common
of the onto property of this map. It says that for points and in X
that are withing 1/(n+1) of each other (mdist(d;x;y) ≤ (r1/r(n 1)))
then we can find points and in mtb-cantor(mtb) such that
maps to (mtb-cantor-map(d;cmplt;mtb;p) ≡ x) and maps to y
(mtb-cantor-map(d;cmplt;mtb;q) ≡ y) and and agree on their first
values  (they are equal in the type i:ℕn ⟶ ℕ(fst(mtb)) i).

Because of the theorem general-cantor-to-int-uniform-continuity,
 we can then prove that every FUNCTION
(f in X ⟶ ℝ for which ∀x,y:X.  (x ≡  ((f x) (f y))) )
from compact metric space to the reals is uniformly continuous.
(see compact-metric-to-real-continuity).

Then, using the fact that ⌜X × X⌝ is also compact metric space
(see m-TB-product and prod-metric-space-complete)
we can easily prove that every FUNCTION from to metric space Y
is uniformly continuous. compact-metric-to-metric-continuity

mtb-cantor-map(d;cmplt;mtb;p) ==  cauchy-mlimit(cmplt;m-regularize(d;mtb-seq(mtb;p));λk.(6 k))

Lemma: mtb-cantor-map_wf
[X:Type]. ∀[d:metric(X)]. ∀[cmplt:mcomplete(X with d)]. ∀[mtb:m-TB(X;d)]. ∀[p:mtb-cantor(mtb)].
  (mtb-cantor-map(d;cmplt;mtb;p) ∈ X)

Lemma: mtb-cantor-map-onto-common
[X:Type]
  ∀d:metric(X). ∀cmplt:mcomplete(X with d). ∀mtb:m-TB(X;d). ∀n:ℕ. ∀x,y:X.
    ((mdist(d;x;y) ≤ (r1/r(n 1)))
     (∃p,q:mtb-cantor(mtb)
         ((p q ∈ (i:ℕn ⟶ ℕ(fst(mtb)) i)) ∧ mtb-cantor-map(d;cmplt;mtb;p) ≡ x ∧ mtb-cantor-map(d;cmplt;mtb;q) ≡ y)))

Definition: mcompact
mcompact(X;d) ==  mcomplete(X with d) × m-TB(X;d)

Lemma: mcompact_wf
[X:Type]. ∀[d:metric(X)].  (mcompact(X;d) ∈ Type)

Lemma: compact-metric-to-real-continuity
[X:Type]. ∀d:metric(X). ∀c:mcompact(X;d). ∀f:FUN(X ⟶ ℝ).  UC(f:X ⟶ ℝ)

Definition: compact-mc
compact-mc{i:l}(d;c;f) ==  TERMOF{compact-metric-to-real-continuity:o, 1:l, i:l} f

Lemma: compact-mc_wf
[X:Type]. ∀d:metric(X). ∀c:mcompact(X;d). ∀f:FUN(X ⟶ ℝ).  (compact-mc{i:l}(d;c;f) ∈ UC(f:X ⟶ ℝ))

Definition: compact-sup
compact-sup{i:l}(d;c;f) ==  m-sup{i:l}(d;snd(c);f;compact-mc{i:l}(d;c;f))

Lemma: compact-sup_wf
[X:Type]. ∀d:metric(X). ∀c:mcompact(X;d). ∀f:FUN(X ⟶ ℝ).  (compact-sup{i:l}(d;c;f) ∈ ℝ)

Lemma: compact-sup-property
[X:Type]
  ∀d:metric(X). ∀c:mcompact(X;d). ∀f:FUN(X ⟶ ℝ).
    ((∀x:X. ((f x) ≤ compact-sup{i:l}(d;c;f))) ∧ (∀e:ℝ((r0 < e)  (∃x:X. ((compact-sup{i:l}(d;c;f) e) < (f x))))))

Definition: compact-inf
compact-inf{i:l}(d;c;f) ==  m-inf{i:l}(d;snd(c);f;compact-mc{i:l}(d;c;f))

Lemma: compact-inf_wf
[X:Type]. ∀d:metric(X). ∀c:mcompact(X;d). ∀f:FUN(X ⟶ ℝ).  (compact-inf{i:l}(d;c;f) ∈ ℝ)

Lemma: compact-inf-property
[X:Type]
  ∀d:metric(X). ∀c:mcompact(X;d). ∀f:FUN(X ⟶ ℝ).
    ((∀x:X. (compact-inf{i:l}(d;c;f) ≤ (f x))) ∧ (∀e:ℝ((r0 < e)  (∃x:X. ((f x) < (compact-inf{i:l}(d;c;f) e))))))

Lemma: compact-metric-to-metric-continuity
X:Type. ∀d:metric(X).  (mcompact(X;d)  (∀Y:Type. ∀dY:metric(Y). ∀f:FUN(X ⟶ Y).  UC(f:X ⟶ Y)))

Lemma: fixedpoint-property_functionality
X:Type. ∀d:metric(X). ∀Y:Type. ∀d':metric(Y).  (homeomorphic(X;d;Y;d')  mcompact(X;d)  FP(X)  FP(Y))

Lemma: mtb-cantor-map-onto
[X:Type]. ∀[d:metric(X)]. ∀[cmplt:mcomplete(X with d)]. ∀[mtb:m-TB(X;d)]. ∀[x:X].
  mtb-cantor-map(d;cmplt;mtb;mtb-point-cantor(mtb;x)) ≡ x

Lemma: mtb-cantor-map-continuous
[X:Type]. ∀[d:metric(X)]. ∀[cmplt:mcomplete(X with d)]. ∀[mtb:m-TB(X;d)]. ∀[k:ℕ+]. ∀[p,q:mtb-cantor(mtb)].
  mdist(d;mtb-cantor-map(d;cmplt;mtb;p);mtb-cantor-map(d;cmplt;mtb;q)) ≤ (r1/r(k)) 
  supposing ∀i:ℕ((i ≤ (18 k))  ((p i) (q i) ∈ ℤ))

Lemma: compact-metric-to-int-bounded
[X:Type]. ∀d:metric(X). ∀cmplt:mcomplete(X with d). ∀mtb:m-TB(X;d). ∀f:X ⟶ ℤ.  ∃B:ℕ. ∀x:X. ∃y:X. (x ≡ y ∧ (|f y| ≤ B))

Definition: dist-fun
dist-fun(d;x) ==  λa.mdist(d;x;a)

Lemma: dist-fun_wf
[X:Type]. ∀[d:metric(X)]. ∀[x:X].  (dist-fun(d;x) ∈ FUN(X ⟶ ℝ))

Definition: compact-dist
dist(x;A) ==  compact-inf{i:l}(d;c;dist-fun(d;x))

Lemma: compact-dist_wf
[X:Type]. ∀[d:metric(X)]. ∀[A:Type].  ∀[c:mcompact(A;d)]. ∀[x:X].  (dist(x;A) ∈ ℝsupposing A ⊆X

Lemma: compact-dist-nonneg
[X:Type]. ∀[d:metric(X)]. ∀[A:Type].  ∀[c:mcompact(A;d)]. ∀[x:X].  (r0 ≤ dist(x;A)) supposing A ⊆X

Lemma: compact-dist-positive
[X:Type]
  ∀d:metric(X). ∀A:Type.
    ∀c:mcompact(A;d). ∀x:X.  (r0 < dist(x;A) ⇐⇒ ∃n:ℕ+. ∀a:A. ((r1/r(n)) ≤ mdist(d;x;a))) supposing A ⊆X

Lemma: compact-dist-zero
[X:Type]
  ∀d:metric(X). ∀A:Type.
    ∀c:mcompact(A;d). ∀x:X.  (dist(x;A) r0 ⇐⇒ ∀n:ℕ+. ∃a:A. (mdist(d;x;a) < (r1/r(n)))) supposing A ⊆X

Lemma: compact-dist-zero-in-complete
[X:Type]
  ∀d:metric(X)
    (mcomplete(X with d)
     (∀[A:Type]. (metric-subspace(X;d;A)  (∀c:mcompact(A;d). ∀x:X.  (dist(x;A) r0 ⇐⇒ x ∈ A)))))

Lemma: not-in-compact-separated
[X:Type]
  ∀d:metric(X)
    ∀[A:Type]
      ∀c:mcompact(A;d). ∀x:X.  (x ∈ A) ⇐⇒ ¬¬(∃n:ℕ+. ∀a:A. ((r1/r(n)) ≤ mdist(d;x;a)))) 
      supposing mcomplete(X with d) ∧ metric-subspace(X;d;A)

Lemma: mcomplete-stable-union
[X:Type]. ∀[d:metric(X)]. ∀[T:Type]. ∀[P:T ⟶ X ⟶ ℙ].
  finite(T)  mcomplete(X with d)  (∀i:T. mcompact({x:X| P[i;x]} ;d))  mcomplete(stable-union(X;T;i,x.P[i;x]) with \000Cd) 
  supposing ∀i:T. ∀x,y:X.  (P[i;x]  y ≡  P[i;y])

Lemma: mcompact-stable-union
[X:Type]
  ∀d:metric(X)
    ∀[T:Type]. ∀[P:T ⟶ X ⟶ ℙ].
      finite(T)  mcomplete(X with d)  (∀i:T. mcompact({x:X| P[i;x]} ;d))   mcompact(stable-union(X;T;i,x.P[i;x\000C]);d) 
      supposing ∀i:T. ∀x,y:X.  (P[i;x]  y ≡  P[i;y])

Lemma: mcompact_functionality_wrt_homeomorphic+
X,Y:Type. ∀d:metric(X). ∀d':metric(Y).  (homeomorphic+(Y;d';X;d)  mcompact(X;d)  mcompact(Y;d'))

Lemma: mcompact_functionality
[X,Y:Type].  ∀d:metric(X). (mcompact(X;d) ⇐⇒ mcompact(Y;d)) supposing X ≡ Y

Lemma: mcompact-product
k:ℕ. ∀X:ℕk ⟶ Type. ∀d:i:ℕk ⟶ metric(X i).  ((∀i:ℕk. mcompact(X i;d i))  mcompact(i:ℕk ⟶ (X i);prod-metric(k;d)))

Lemma: mcompact-interval
a,b:ℝ.  mcompact({x:ℝx ∈ [a, b]} ;rmetric()) supposing a ≤ b

Lemma: interval-to-int-bounded
a:ℝ. ∀b:{b:ℝa ≤ b} . ∀f:{x:ℝx ∈ [a, b]}  ⟶ ℤ.
  ∃B:ℕ. ∀x:{x:ℝx ∈ [a, b]} . ∃y:{x:ℝx ∈ [a, b]} ((x y) ∧ (|f y| ≤ B))

Lemma: fixedpoint-property-iff
X:Type. ∀d:metric(X).  (mcompact(X;d)  (FP(X) ⇐⇒ ∀f:FUN(X ⟶ X). (∀x:X. x))))

Lemma: mcompact-finite-subcover
[X:Type]
  ∀d:metric(X)
    (mcompact(X;d)
     (∀[I:Type]. ∀[A:I ⟶ X ⟶ ℙ].  (m-open-cover(X;d;I;i,x.A[i;x])  (∃n:ℕ+. ∃L:ℕn ⟶ I. ∀x:X. ∃j:ℕn. A[L j;x]))))

Lemma: finite-subcover-implies-m-TB
[X:Type]
  ∀d:metric(X)
    ((∀[I:Type]. ∀[A:I ⟶ X ⟶ ℙ].  (m-open-cover(X;d;I;i,x.A[i;x])  (∃n:ℕ+. ∃L:ℕn ⟶ I. ∀x:X. ∃j:ℕn. A[L j;x])))
     m-TB(X;d))

Definition: incr-binary-seq
IBS ==  {s:ℕ ⟶ ℕ2| ∀i:ℕ((s i) ≤ (s (1 i)))} 

Lemma: incr-binary-seq_wf
IBS ∈ Type

Lemma: ibs-property
[s:IBS]. ∀[m:ℕ].  ∀[n:ℕ]. (s n) 1 ∈ ℤ supposing m ≤ supposing (s m) 1 ∈ ℤ

Definition: mkibs
mkibs(n.p[n]) ==  λn.if p[n] then else fi 

Lemma: mkibs_wf
[p:ℕ ⟶ 𝔹]. mkibs(n.p[n]) ∈ IBS supposing ∀n:ℕ((↑p[n])  (↑p[n 1]))

Definition: rless_ibs
rless_ibs(x;y) ==  λn.if (∃m∈upto(n 1).(x (m 1)) 4 <(m 1))_b then else fi 

Lemma: rless_ibs_wf
[x,y:ℝ].  (rless_ibs(x;y) ∈ IBS)

Lemma: rless_ibs_property
x,y:ℝ.
  ((x < ⇐⇒ ∃n:ℕ((rless_ibs(x;y) n) 1 ∈ ℤ))
  ∧ (∀n:ℕ
       (((rless_ibs(x;y) n) 0 ∈ ℤ)
        (((y < x) ∧ (∀m:ℕ((rless_ibs(x;y) m) 0 ∈ ℤ))) ∨ (|x y| ≤ (r(4)/r(n 1)))))))

Definition: real-fun
real-fun(f;a;b) ==  ∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  ((f x) (f y)))

Lemma: real-fun_wf
[a,b:ℝ]. ∀[f:[a, b] ⟶ℝ].  (real-fun(f;a;b) ∈ ℙ)

Definition: real-cont
real-cont(f;a;b) ==  ∀k:ℕ+. ∃d:{d:ℝr0 < d} . ∀x,y:{x:ℝx ∈ [a, b]} .  ((|x y| ≤ d)  (|(f x) y| ≤ (r1/r(k))))

Lemma: real-cont_wf
[a,b:ℝ]. ∀[f:[a, b] ⟶ℝ].  (real-cont(f;a;b) ∈ ℙ)

Lemma: real-continuity
a,b:ℝ.  ∀f:[a, b] ⟶ℝreal-cont(f;a;b) supposing real-fun(f;a;b) supposing a ≤ b

Lemma: stable__real-fun
[a,b:ℝ]. ∀[f:[a, b] ⟶ℝ].  Stable{real-fun(f;a;b)}

Lemma: sq_stable__real-fun
[a,b:ℝ]. ∀[f:[a, b] ⟶ℝ].  SqStable(real-fun(f;a;b))

Lemma: real-fun-iff-continuous
a,b:ℝ.  ∀f:[a, b] ⟶ℝ(real-fun(f;a;b) ⇐⇒ real-cont(f;a;b)) supposing a ≤ b

Lemma: real-cont-iff-continuous
a,b:ℝ.  ((a ≤ b)  (∀f:[a, b] ⟶ℝ(real-cont(f;a;b) ⇐⇒ f[x] continuous for x ∈ [a, b])))

Definition: ifun
ifun(f;I) ==  real-fun(f;left-endpoint(I);right-endpoint(I))

Lemma: ifun_wf
[I:Interval]. ∀[f:I ⟶ℝ].  ifun(f;I) ∈ ℙ supposing icompact(I)

Lemma: ifun-alt
I:Interval. ∀[f:I ⟶ℝ]. (ifun(f;I)) supposing ((∀x,y:{x:ℝx ∈ I} .  ((x y)  ((f x) (f y)))) and icompact(I))

Lemma: sq_stable__ifun
[I:Interval]. ∀[f:I ⟶ℝ].  SqStable(ifun(f;I)) supposing icompact(I)

Lemma: ifun_subtype_1
[a,b,c:ℝ].  ((a ≤ c)  (c ≤ b)  ({f:[a, b] ⟶ℝifun(f;[a, b])}  ⊆{f:[a, c] ⟶ℝifun(f;[a, c])} ))

Lemma: ifun_subtype_2
[a,b,c:ℝ].  ((a ≤ c)  (c ≤ b)  ({f:[a, b] ⟶ℝifun(f;[a, b])}  ⊆{f:[c, b] ⟶ℝifun(f;[c, b])} ))

Lemma: ifun_subtype_3
[a,b,c,d:ℝ].  ((a ≤ c)  (c ≤ d)  (d ≤ b)  ({f:[a, b] ⟶ℝifun(f;[a, b])}  ⊆{f:[c, d] ⟶ℝifun(f;[c, d])} ))

Lemma: ifun_subtype_subinterval
[I,J:{J:Interval| icompact(J)} ].  {f:I ⟶ℝifun(f;I)}  ⊆{f:J ⟶ℝifun(f;J)}  supposing J ⊆ 

Lemma: ifun-continuous
I:Interval. (icompact(I)  (∀f:{f:I ⟶ℝifun(f;I)} f[x] continuous for x ∈ I))

Lemma: function-is-continuous
I:Interval. ∀f:I ⟶ℝ.  ((∀x,y:{t:ℝt ∈ I} .  ((x y)  (f[x] f[y])))  f[x] continuous for x ∈ I)

Lemma: ifun-iff-continuous
I:Interval. (icompact(I)  (∀f:I ⟶ℝ(ifun(λx.f[x];I) ⇐⇒ f[x] continuous for x ∈ I)))

Definition: real-sfun
real-sfun(f;a;b) ==  ∀x,y:{x:ℝx ∈ [a, b]} .  (f x ≠  x ≠ y)

Lemma: real-sfun_wf
[a,b:ℝ]. ∀[f:[a, b] ⟶ℝ].  (real-sfun(f;a;b) ∈ ℙ)

Lemma: real-fun-implies-sfun-general
[I:Interval]. ∀[f:I ⟶ℝ].
  ∀x,y:{x:ℝx ∈ I} .  (f x ≠  x ≠ y) supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  ((f x) (f y)))

Lemma: strict-increasing-implies-inv-strict-increasing
[I:Interval]. ∀[f:I ⟶ℝ].
  (∀x,y:{x:ℝx ∈ I} .  (((f x) < (f y))  (x < y))) supposing 
     ((∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f x) < (f y)))) and 
     (∀x,y:{x:ℝx ∈ I} .  ((x y)  ((f x) (f y)))))

Lemma: real-fun-implies-sfun
[a:ℝ]. ∀[b:{b:ℝa ≤ b} ]. ∀[f:[a, b] ⟶ℝ].  real-sfun(f;a;b) supposing real-fun(f;a;b)

Lemma: real-fun-implies-sfun-ext
[a:ℝ]. ∀[b:{b:ℝa ≤ b} ]. ∀[f:[a, b] ⟶ℝ].  real-sfun(f;a;b) supposing real-fun(f;a;b)

Lemma: function-on-compact
a:ℝ. ∀b:{b:ℝa ≤ b} . ∀f:[a, b] ⟶ℝ.
  ((∀x,y:{t:ℝt ∈ [a, b]} .  ((x y)  (f[x] f[y])))
   (∀n:ℕ+
        (∃d:ℝ [((r0 < d)
              ∧ (∀x,y:ℝ.  ((x ∈ [a, b])  (y ∈ [a, b])  (|x y| ≤ d)  (|f[x] f[y]| ≤ (r1/r(n))))))])))

Definition: discontinuous
discontinuous(f;x) ==
  ∃epsilon:{e:ℝr0 < e} . ∀delta:{e:ℝr0 < e} . ∃y:ℝ((|x y| < delta) ∧ (epsilon < |(f x) y|))

Lemma: discontinuous_wf
[f:ℝ ⟶ ℝ]. ∀[x:ℝ].  (discontinuous(f;x) ∈ ℙ)

Lemma: not-discontinuous
[f:ℝ ⟶ ℝ]. ∀[x:ℝ]. discontinuous(f;x)) supposing ∀x,y:ℝ.  ((x y)  ((f x) (f y)))

Lemma: extensional-discrete-real-fun-is-constant
a,b:ℝ. ∀f:{x:ℝx ∈ [a, b]}  ⟶ ℤ.
  ∀x,y:{x:ℝx ∈ [a, b]} .  ((f x) (f y) ∈ ℤsupposing ∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  ((f x) (f y) ∈ ℤ))

Lemma: extensional-real-to-bool-constant
f:ℝ ⟶ 𝔹. ∀x,y:ℝ.  supposing ∀x,y:ℝ.  ((x y)  y)

Lemma: extensional-interval-to-bool-constant
a:ℝ. ∀b:{b:ℝa ≤ b} . ∀f:{x:ℝx ∈ [a, b]}  ⟶ 𝔹.
  ∀x,y:{x:ℝx ∈ [a, b]} .  supposing ∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  y)

Definition: discrete-type
discrete-type(T) ==  ∀f:ℝ ⟶ T. ((∀x,y:ℝ.  ((x y)  ((f x) (f y) ∈ T)))  (∀x,y:ℝ.  ((f x) (f y) ∈ T)))

Lemma: discrete-type_wf
[T:Type]. (discrete-type(T) ∈ ℙ)

Lemma: strong-subtype-discrete-type
[A,B:Type].  (discrete-type(A)) supposing (discrete-type(B) and strong-subtype(A;B))

Lemma: int-discrete
discrete-type(ℤ)

Lemma: decidable-equality-implies-discrete
[T:Type]. ((∀x,y:T.  Dec(x y ∈ T))  discrete-type(T))

Lemma: product-discrete
A:Type. ∀B:A ⟶ Type.  (discrete-type(A)  (∀a:A. discrete-type(B[a]))  discrete-type(a:A × B[a]))

Lemma: prod-discrete
A,B:Type.  (discrete-type(A)  discrete-type(B)  discrete-type(A × B))

Lemma: function-discrete
A:Type. ∀B:A ⟶ Type.  ((∀a:A. discrete-type(B[a]))  discrete-type(a:A ⟶ B[a]))

Lemma: fun-discrete
A,B:Type.  (discrete-type(B)  discrete-type(A ⟶ B))

Lemma: union-discrete
A,B:Type.  (discrete-type(A)  discrete-type(B)  discrete-type(A B))

Lemma: partial-int-not-discrete
¬discrete-type(partial(ℤ))

Lemma: base-not-discrete
¬discrete-type(Base)

Definition: real-disjoint
real-disjoint(x.A[x];y.B[y]) ==  ∀x,y:ℝ.  ((x y)  (A[x] ∧ B[y])))

Lemma: real-disjoint_wf
[A,B:ℝ ⟶ ℙ].  (real-disjoint(x.A[x];y.B[y]) ∈ ℙ)

Definition: real-separation
real-separation(x.A[x];y.B[y]) ==  real-disjoint(x.A[x];y.B[y]) ∧ (∃x:ℝA[x]) ∧ (∃y:ℝB[y]) ∧ (∀r:ℝ(A[r] ∨ B[r]))

Lemma: real-separation_wf
[A,B:ℝ ⟶ ℙ].  (real-separation(x.A[x];y.B[y]) ∈ ℙ)

Lemma: no-real-separation
[A,B:ℝ ⟶ ℙ].  real-separation(x.A[x];y.B[y]))

Lemma: no-real-separation-corollary
[A,B:ℝ ⟶ ℙ].  ((∃x:ℝA[x])  (∃y:ℝB[y])  (∀r:ℝ(A[r] ∨ B[r]))  (¬¬(∃x,y:ℝ((x y) ∧ A[x] ∧ B[y]))))

Lemma: no-nontrivial-decidable-real-prop
[A:ℝ ⟶ ℙ]. ((∀x,y:ℝ.  ((x y)  (A[x] ⇐⇒ A[y])))  (∀r:ℝ(A[r] ∨ A[r])))  ((∀x:ℝA[x]) ∨ (∀x:ℝA[x]))))

Lemma: separated-decider-not-extensional
a,b:ℝ.  ((a < b)  (∀d:∀u:ℝ((a < u) ∨ (u < b)). (¬¬(∃x,y:ℝ((x y) ∧ (↑isl(d x)) ∧ (↑isr(d y)))))))

Lemma: separated-decider-not-extensional2
a,b:ℝ.  ((a < b)  (∀d:∀u:ℝ((a < u) ∨ (u < b)). (∀x,y:ℝ.  ((x y)  isl(d x) isl(d y))))))

Definition: cover-seq
cover-seq(d;a;b;n) ==
  primrec(n;<a, b>i,p. let a,b in case (a b/r(2)) of inl(z) => <(a b/r(2)), b> inr(z) => <a, (a b/r(2))>\000C)

Lemma: cover-seq_wf
[A,B:ℝ ⟶ ℙ]. ∀[d:r:ℝ ⟶ (A[r] B[r])]. ∀[a,b:ℝ]. ∀[n:ℕ].  (cover-seq(d;a;b;n) ∈ ℝ × ℝ)

Lemma: cover-seq-0
[d,a,b:Top].  (cover-seq(d;a;b;0) ~ <a, b>)

Lemma: cover-seq-property
[A,B:ℝ ⟶ ℙ].
  ∀d:r:ℝ ⟶ (A[r] B[r]). ∀a,b:ℝ.
    (A[a]
     B[b]
     (∀n:ℕ
          (A[fst(cover-seq(d;a;b;n))]
          ∧ B[snd(cover-seq(d;a;b;n))]
          ∧ ((cover-seq(d;a;b;n 1) let a,b cover-seq(d;a;b;n) in <a, (a b/r(2))> ∈ (ℝ × ℝ))
            ∨ (cover-seq(d;a;b;n 1) let a,b cover-seq(d;a;b;n) in <(a b/r(2)), b> ∈ (ℝ × ℝ))))))

Lemma: cover-seq-property-ext
[A,B:ℝ ⟶ ℙ].
  ∀d:r:ℝ ⟶ (A[r] B[r]). ∀a,b:ℝ.
    (A[a]
     B[b]
     (∀n:ℕ
          (A[fst(cover-seq(d;a;b;n))]
          ∧ B[snd(cover-seq(d;a;b;n))]
          ∧ ((cover-seq(d;a;b;n 1) let a,b cover-seq(d;a;b;n) in <a, (a b/r(2))> ∈ (ℝ × ℝ))
            ∨ (cover-seq(d;a;b;n 1) let a,b cover-seq(d;a;b;n) in <(a b/r(2)), b> ∈ (ℝ × ℝ))))))

Lemma: inhabited-covers-reals-implies
[A,B:ℝ ⟶ ℙ].
  ((∃a:ℝA[a])
   (∃b:ℝB[b])
   (∀r:ℝ(A[r] ∨ B[r]))
   (∃f,g:ℕ ⟶ ℝ. ∃x:ℝ((∀n:ℕA[f n]) ∧ (∀n:ℕB[g n]) ∧ lim n→∞.f x ∧ lim n→∞.g x)))

Lemma: inhabited-covers-real-implies
[A,B:ℝ ⟶ ℙ].
  ((∃a:ℝA[a])
   (∃b:ℝB[b])
   (∀r:ℝ(A[r] ∨ B[r]))
   (∃f,g:ℕ ⟶ ℝ. ∃x:ℝ((∀n:ℕA[f n]) ∧ (∀n:ℕB[g n]) ∧ lim n→∞.f x ∧ lim n→∞.g x)))

Definition: cover-real
cover-real(d; a; b; cb) ==  accelerate(2;λn.((fst(cover-seq(d;a;b;log(2;n (cb 1))))) n))

Lemma: inhabited-covers-real-implies-ext
[A,B:ℝ ⟶ ℙ].
  ((∃a:ℝA[a])
   (∃b:ℝB[b])
   (∀r:ℝ(A[r] ∨ B[r]))
   (∃f,g:ℕ ⟶ ℝ. ∃x:ℝ((∀n:ℕA[f n]) ∧ (∀n:ℕB[g n]) ∧ lim n→∞.f x ∧ lim n→∞.g x)))

Definition: interval-retraction
interval-retraction(u;v;r) ==  rmin(v;rmax(u;r))

Lemma: interval-retraction_wf
[u,v,x:ℝ].  interval-retraction(u;v;x) ∈ {x:ℝx ∈ [u, v]}  supposing u ≤ v

Lemma: interval-retraction_functionality
[u,v,x,u',v',x':ℝ].
  (interval-retraction(u;v;x) interval-retraction(u';v';x')) supposing ((x x') and (v v') and (u u'))

Lemma: interval-retraction-req
[u,v:ℝ]. ∀[x:{x:ℝx ∈ [u, v]} ].  (interval-retraction(u;v;x) x)

Lemma: weak-continuity-principle-interval
u,v:ℝ. ∀x:{x:ℝx ∈ [u, v]} .
  ∃x':{x':ℝx' x} 
   ∀F:{x:ℝx ∈ [u, v]}  ⟶ 𝔹. ∀G:n:ℕ+ ⟶ {y:ℝ(x y ∈ (ℕ+n ⟶ ℤ)) ∧ (y ∈ [u, v])} .
     ∃z:{x:ℝx ∈ [u, v]} (∃n:ℕ+ [(F x' z ∧ (z (G n)))])

Definition: WCPR
WCPR(F;x;G) ==  WCP(λf.(F regularize(1;f));x;G)

Lemma: WCPR_wf
[F:ℝ ⟶ 𝔹]. ∀[x:ℝ]. ∀[G:n:ℕ+ ⟶ {y:ℝy ∈ (ℕ+n ⟶ ℤ)} ].  (WCPR(F;x;G) ∈ ℕ+)

Lemma: weak-continuity-principle-real-ext
x:ℝ. ∀F:ℝ ⟶ 𝔹. ∀G:n:ℕ+ ⟶ {y:ℝy ∈ (ℕ+n ⟶ ℤ)} .  (∃n:ℕ+ [F (G n)])

Lemma: int-int-retraction-reals-1
k:{2...}. ∃r:(ℤ ⟶ ℤ) ⟶ ℝ. ∀x:ℝ(accelerate(k;x) (r i.if i <then else fi )) ∈ ℝ)

Lemma: nat-int-retraction-reals-1
k:{2...}. ∃r:(ℕ ⟶ ℤ) ⟶ ℝ. ∀x:ℝ(accelerate(k;x) (r n.(x (n 1)))) ∈ ℝ)

Lemma: int-int-retraction-reals
r:(ℤ ⟶ ℤ) ⟶ ℝ. ∀x:ℝ(x (r i.if i <then else fi )))

Lemma: nat-int-retraction-reals
r:(ℕ ⟶ ℤ) ⟶ ℝ. ∀x:ℝ(x (r n.(x (n 1)))))

Definition: blend-seq
blend-seq(k;x;y) ==  λi.if i <then else fi 

Lemma: blend-seq_wf
[k:ℕ+]. ∀[x,y:ℕ+ ⟶ ℤ].  (blend-seq(k;x;y) ∈ ℕ+ ⟶ ℤ)

Lemma: blend-close-reals
[k:ℕ+]. ∀[x,y:ℝ].  ((|x y| ≤ (r1/r(k)))  3-regular-seq(blend-seq(k;x;y)))

Definition: blended-real
blended-real(k;x;y) ==  accelerate(3;blend-seq(k;x;y))

Lemma: blended-real_wf
[k:ℕ+]. ∀[x,y:ℝ].  blended-real(k;x;y) ∈ ℝ supposing |x y| ≤ (r1/r(k))

Lemma: blended-real-req
[k:ℕ+]. ∀[x,y:ℝ].  blended-real(k;x;y) supposing |x y| ≤ (r1/r(k))

Lemma: blended-real-agrees
[k:ℕ+]. ∀[x,y:ℝ].  ∀n:ℕ+k ÷ 6. ((blended-real(k;x;y) n) (accelerate(3;x) n) ∈ ℤ)

Lemma: connectedness-main-lemma
x:ℝ. ∀g:ℕ ⟶ ℝ.  (lim n→∞.g  (∀P:ℝ ⟶ 𝔹. ∃z:{z:ℝaccelerate(3;x)} (∃n:ℕ [(z (g n))])))

Lemma: better-continuity-for-reals
x:ℝ. ∃x':{x':ℝx' x} . ∀g:ℕ ⟶ ℝ(lim n→∞.g  (∀P:ℝ ⟶ 𝔹. ∃z:{z:ℝx'} (∃n:ℕ [(z (g n))])))

Lemma: not-all-nonneg-or-nonpos
¬(∀x:ℝ((r0 ≤ x) ∨ (x ≤ r0)))

Lemma: union-of-intervals-not-interval
¬(∀t:{t:ℝt ∈ [r(-1), r1]} ((↓t ∈ [r(-1), r0]) ∨ (↓t ∈ [r0, r1])))

Lemma: dense-in-reals-iff
X:ℝ ⟶ ℙ(dense-in-interval((-∞, ∞);X) ⇐⇒ ∀x:ℝ. ∀n:ℕ+.  ∃y:ℝ((X y) ∧ (|x y| < (r1/r(n)))))

Lemma: dense-in-reals-implies-accel
X:ℝ ⟶ ℙ
  (dense-in-interval((-∞, ∞);X)
   (∀x:ℝ. ∀y:{y:ℝx} .  ((X x)  (X y)))
   (∀x:ℝ. ∀k:ℕ+.  ∃y:ℝ((y accelerate(3;x) ∈ (ℕ+k ⟶ ℤ)) ∧ (X y))))

Lemma: dense-in-reals-implies
X:ℝ ⟶ ℙ
  (dense-in-interval((-∞, ∞);X)
   (∀x:ℝ. ∀y:{y:ℝx} .  ((X x)  (X y)))
   (∀x:ℝ. ∃x':ℝ((x' x) ∧ (∀k:ℕ+. ∃y:ℝ((y x' ∈ (ℕ+k ⟶ ℤ)) ∧ (X y))))))

Lemma: connectedness-main-lemma-ext
x:ℝ. ∀g:ℕ ⟶ ℝ.  (lim n→∞.g  (∀P:ℝ ⟶ 𝔹. ∃z:{z:ℝaccelerate(3;x)} (∃n:ℕ [(z (g n))])))

Definition: connected
Connected(X) ==
  ∀[A,B:X ⟶ ℙ].
    ((∀x:X. ∀y:{y:X| y} .  (A[y]  A[x]))
     (∀x:X. ∀y:{y:X| y} .  (B[y]  B[x]))
     (∃a:X. A[a])
     (∃b:X. B[b])
     (∀r:X. (A[r] ∨ B[r]))
     (∃r:X. (A[r] ∧ B[r])))

Lemma: connected_wf
[X:Type]. Connected(X) ∈ ℙsupposing X ⊆r ℝ

Lemma: reals-connected
Connected(ℝ)

Lemma: closed-interval-connected
u,v:ℝ.  Connected({x:ℝx ∈ [u, v]} )

Lemma: IVT-from-connected
u,v:ℝ. ∀f:[u, v] ⟶ℝ.
  ((u ≤ v)
   (∀x,y:{x:ℝx ∈ [u, v]} .  ((x y)  (f(x) f(y))))
   (f(u) < r0)
   (r0 < f(v))
   (∀e:{e:ℝr0 < e} . ∃c:{c:ℝc ∈ [u, v]} (|f(c)| < e)))

Lemma: interval-connected
I:Interval. Connected({x:ℝx ∈ I} )

Lemma: real-subset-connected-lemma
X:ℝ ⟶ ℙ
  (dense-in-interval((-∞, ∞);X)
   (∀a,b:{x:ℝx}  ⟶ 𝔹.
        ((∃x:{x:ℝx} (↑(a x)))
         (∃x:{x:ℝx} (↑(b x)))
         (∀x:{x:ℝx} ((↑(a x)) ∨ (↑(b x))))
         (∃f,g:ℕ ⟶ {x:ℝx} 
             ∃x:ℝ((∀n:ℕ(↑(a (f n)))) ∧ (∀n:ℕ(↑(b (g n)))) ∧ lim n→∞.f x ∧ lim n→∞.g x)))))

Lemma: real-subset-connected
X:ℝ ⟶ ℙ
  ((∀x:ℝSqStable(X x))
   (∀x:ℝ. ∀y:{y:ℝy} .  ((X y)  (X x)))
   dense-in-interval((-∞, ∞);X)
   (∀Q:{x:ℝx}  ⟶ 𝔹. ∃Q':ℝ ⟶ 𝔹. ∀x:{x:ℝx} Q' x)
   Connected({x:ℝx} ))

Definition: VesleySchema
VesleySchema ==
  ∀P:(ℕ+ ⟶ ℤ) ⟶ ℙ
    ((∀f:ℕ+ ⟶ ℤ. ∀k:ℕ+.  ∃g:{x:ℕ+ ⟶ ℤ| ¬(P x)} (g f ∈ (ℕ+k ⟶ ℤ)))
     (∀Q:{x:ℕ+ ⟶ ℤ| ¬(P x)}  ⟶ 𝔹. ∃Q':(ℕ+ ⟶ ℤ) ⟶ 𝔹. ∀x:{x:ℕ+ ⟶ ℤ| ¬(P x)} Q' x))

Lemma: VesleySchema_wf
VesleySchema ∈ ℙ'

Definition: VesleySchema2
VesleySchema2 ==
  ∀P:(ℕ+ ⟶ ℤ) ⟶ ℙ
    ((∀f:ℕ+ ⟶ ℤ. ∀k:ℕ+.  ∃g:{x:ℕ+ ⟶ ℤ| ¬(P x)} (g regularize(1;f) ∈ (ℕ+k ⟶ ℤ)))
     (∀Q:{x:ℕ+ ⟶ ℤ| ¬(P x)}  ⟶ 𝔹. ∃Q':(ℕ+ ⟶ ℤ) ⟶ 𝔹. ∀x:{x:ℕ+ ⟶ ℤ| ¬(P x)} Q' x))

Lemma: VesleySchema2_wf
VesleySchema2 ∈ ℙ'

Definition: VesleyAxiom
VesleyAxiom ==
  ∀P:ℝ ⟶ ℙ
    (dense-in-interval((-∞, ∞);λx.(¬(P x)))
     (∀x:ℝ. ∀y:{y:ℝy} .  ((P y)  (P x)))
     (∀Q:{x:ℝ| ¬(P x)}  ⟶ 𝔹. ∃Q':ℝ ⟶ 𝔹. ∀x:{x:ℝ| ¬(P x)} Q' x))

Lemma: VesleyAxiom_wf
VesleyAxiom ∈ ℙ'

Lemma: Vesley-subset-connected
VesleyAxiom
 (∀P:ℝ ⟶ ℙ
      ((∀x:ℝ. ∀y:{y:ℝy} .  (P[y]  P[x]))  dense-in-interval((-∞, ∞);λx.(¬P[x]))  Connected({x:ℝ| ¬P[x]} )))

Lemma: Vesley-connected-rationals
VesleyAxiom  Connected({x:ℝ| ¬¬(∃a:ℤ. ∃b:ℤ-o(x (r(a)/r(b))))} )

Definition: KripkeSchema
KripkeSchema ==  ∀P:ℙ. ∃f:ℕ ⟶ ℕ(((∀n:ℕ((f n) 0 ∈ ℕ))  P)) ∧ ((∃n:ℕ0 < n)  P))

Lemma: KripkeSchema_wf
KripkeSchema ∈ ℙ'

Definition: decdr-to-bool
bool(d) ==  λx.if then inl Ax else inr Ax  fi 

Lemma: decdr-to-bool_wf
[T:Type]. ∀[A,B:T ⟶ ℙ]. ∀[d:x:T ⟶ (A[x] B[x])].  (bool(d) ∈ T ⟶ 𝔹)

Definition: cover-search-left
cover-search-left(d;a;b;x) ==
  WCPR(λx.isl(d x);accelerate(3;cover-real(bool(d); a; b; x
                                           ));λk.blended-real(6
                                                 k;cover-real(bool(d); a; b; x
                                                                );snd(cover-seq(bool(d);a;b;imax(log(2;(4 k)
                                                                                                 (x 1))
                                                                                                 1;log(2;(2 k)
                                                                                                 (x 1)))))))

Definition: cover-search-right
cover-search-right(d;a;b;x) ==
  6
  WCPR(λx.isr(d x);accelerate(3;cover-real(bool(d); a; b; x
                                             ));λk.blended-real(6 k;cover-real(bool(d); a; b; x
                                                                                 );fst(cover-seq(bool(d);a;b;log(2;(4
                                                                                                 6
                                                                                                 k)
                                                                                                 (x 1))
                                                                                                 1))))

Definition: altered-seq1
altered-seq1(d; a; b; x; n) ==  blended-real(6 cover-search-left(d;a;b;x);cover-real(bool(d); a; b; x);n)

Definition: altered-seq2
altered-seq2(d; a; b; x; n) ==  blended-real(cover-search-right(d;a;b;x);cover-real(bool(d); a; b; x);n)

Lemma: reals-connected-ext
Connected(ℝ)

Definition: non-rational
non-rational() ==  {x:ℝ| ∀a:ℤ. ∀b:ℕ+.  (x (r(a)/r(b))))} 

Lemma: non-rational_wf
non-rational() ∈ Type

Definition: real-continuity-principle
real-continuity-principle() ==  ∀I:Interval. ∀f:I ⟶ℝ.  f[x] continuous for x ∈ I

Lemma: real-continuity-principle_wf
real-continuity-principle() ∈ ℙ

Definition: proper-real-continuity
proper-real-continuity() ==
  ∀a,b:ℝ.
    ((a < b)
     (∀f:[a, b] ⟶ℝ. ∀n:ℕ+.
          (∃d:ℝ [((r0 < d)
                ∧ (∀x,y:ℝ.  ((x ∈ [a, b])  (y ∈ [a, b])  (|x y| ≤ d)  (|(f x) y| ≤ (r1/r(n))))))])))

Lemma: proper-real-continuity_wf
proper-real-continuity() ∈ ℙ

Lemma: cantor-middle-third-lemma
x,a,b:ℝ.  ((x ∈ [(2 b)/3, b]) ∨ (x ∈ [a, (a b)/3])) supposing ((x ∈ [a, b]) and (a < b))

Lemma: cantor-common-middle-third-lemma
x,y,a,b:ℝ.
  (((x ∈ [(2 b)/3, b]) ∧ (y ∈ [(2 b)/3, b]))
     ∨ ((x ∈ [a, (a b)/3]) ∧ (y ∈ [a, (a b)/3]))) supposing 
     (((x ∈ [a, b]) ∧ (y ∈ [a, b]) ∧ (|x y| ≤ (b a/r(6)))) and 
     (a < b))

Definition: unit-interval-fan
unit-interval-fan(f;n) ==
  primrec(n;<0, 1>i,p. let a,b 
                       in if i
                          then eval a' (2 a) in
                               eval b' in
                                 <a', b'>
                          else eval a' in
                               eval b' (2 b) in
                                 <a', b'>
                          fi )

Lemma: unit-interval-fan_wf
[f:ℕ ⟶ 𝔹]. ∀[n:ℕ].  (unit-interval-fan(f;n) ∈ ℤ × ℤ)

Definition: cantor_ivl
cantor_ivl(a;b;f;n) ==
  let k,j unit-interval-fan(f;n) 
  in eval 3^n in
     <eval k' in (k' b)/d, eval j' in (j' b)/d>

Lemma: cantor_ivl_wf
[a,b:ℝ]. ∀[f:ℕ ⟶ 𝔹]. ∀[n:ℕ].  (cantor_ivl(a;b;f;n) ∈ ℝ × ℝ)

Definition: cantor-interval
cantor-interval(a;b;f;n) ==
  primrec(n;<a, b>i,p. let a',b' in if then <(2 a' b')/3, b'> else <a', (a' b')/3> fi )

Lemma: cantor-interval_wf
[a,b:ℝ]. ∀[n:ℕ]. ∀[f:ℕn ⟶ 𝔹].  (cantor-interval(a;b;f;n) ∈ ℝ × ℝ)

Lemma: cantor-interval-length
[a,b:ℝ]. ∀[n:ℕ]. ∀[f:ℕn ⟶ 𝔹].  (((snd(cantor-interval(a;b;f;n))) fst(cantor-interval(a;b;f;n))) (2^n a)/3^n)

Lemma: cantor-interval-rleq
[a,b:ℝ].  ∀[n:ℕ]. ∀[f:ℕn ⟶ 𝔹].  ((fst(cantor-interval(a;b;f;n))) ≤ (snd(cantor-interval(a;b;f;n)))) supposing a ≤ b

Lemma: cantor-interval-rless
a,b:ℝ.  ((a < b)  (∀n:ℕ. ∀f:ℕn ⟶ 𝔹.  ((fst(cantor-interval(a;b;f;n))) < (snd(cantor-interval(a;b;f;n))))))

Lemma: cantor-interval-inclusion
[a,b:ℝ].
  ∀[f:ℕ ⟶ 𝔹]. ∀[n:ℕ]. ∀[m:{n...}].
    (((fst(cantor-interval(a;b;f;n))) ≤ (fst(cantor-interval(a;b;f;m))))
    ∧ ((fst(cantor-interval(a;b;f;m))) ≤ (snd(cantor-interval(a;b;f;m))))
    ∧ ((snd(cantor-interval(a;b;f;m))) ≤ (snd(cantor-interval(a;b;f;n))))) 
  supposing a ≤ b

Lemma: cantor-interval-inclusion2
[a,b:ℝ].
  ∀[m:ℕ]. ∀[n:ℕm]. ∀[f:ℕm ⟶ 𝔹].
    (((fst(cantor-interval(a;b;f;n))) ≤ (fst(cantor-interval(a;b;f;m))))
    ∧ ((fst(cantor-interval(a;b;f;m))) ≤ (snd(cantor-interval(a;b;f;m))))
    ∧ ((snd(cantor-interval(a;b;f;m))) ≤ (snd(cantor-interval(a;b;f;n))))) 
  supposing a ≤ b

Lemma: cantor-interval-cauchy
a,b:ℝ.  ∀[f:ℕ ⟶ 𝔹]. cauchy(n.fst(cantor-interval(a;b;f;n))) supposing a ≤ b

Definition: cantor_cauchy
cantor_cauchy(a;b;k) ==
  eval eval y' (b 4) in
           eval y' y' (-(a 4)) in
             y' ÷ in
  if (0) < (x)
     then if (x 4) ÷ 2=0 then else (2 log(2;((x 4) ÷ 2) k))
     else if ((-x) 4) ÷ 2=0 then else (2 log(2;(((-x) 4) ÷ 2) k))

Lemma: cantor-interval-cauchy-ext
a,b:ℝ.  ∀[f:ℕ ⟶ 𝔹]. cauchy(n.fst(cantor-interval(a;b;f;n))) supposing a ≤ b

Lemma: cantor-interval-converges
a,b:ℝ.  ∀f:ℕ ⟶ 𝔹fst(cantor-interval(a;b;f;n))↓ as n→∞ supposing a ≤ b

Lemma: cantor-interval-req
a,b:ℝ. ∀f:ℕ ⟶ 𝔹. ∀n:ℕ.
  (((fst(cantor-interval(a;b;f;n))) (fst(cantor_ivl(a;b;f;n))))
  ∧ ((snd(cantor-interval(a;b;f;n))) (snd(cantor_ivl(a;b;f;n)))))

Lemma: cantor_ivl-converges
a,b:ℝ.  ∀f:ℕ ⟶ 𝔹fst(cantor_ivl(a;b;f;n))↓ as n→∞ supposing a ≤ b

Lemma: cantor-interval-converges-ext
a,b:ℝ.  ∀f:ℕ ⟶ 𝔹fst(cantor-interval(a;b;f;n))↓ as n→∞ supposing a ≤ b

Lemma: cantor_ivl-converges-ext
a,b:ℝ.  ∀f:ℕ ⟶ 𝔹fst(cantor_ivl(a;b;f;n))↓ as n→∞ supposing a ≤ b

Definition: cantor-to-interval
cantor-to-interval(a;b;f) ==  λn.eval in let x,y cantor-interval(a;b;f;cantor_cauchy(a;b;m)) in (x m) ÷ 4

Lemma: cantor-to-interval_wf1
a,b:ℝ.  ∀f:ℕ ⟶ 𝔹(cantor-to-interval(a;b;f) ∈ {x:ℝlim n→∞.fst(cantor-interval(a;b;f;n)) x} supposing a ≤ b

Lemma: cantor-to-interval_wf
a,b:ℝ.  ∀f:ℕ ⟶ 𝔹(cantor-to-interval(a;b;f) ∈ {x:ℝ(a ≤ x) ∧ (x ≤ b)} supposing a ≤ b

Lemma: cantor-to-interval-req
a,b,x:ℝ. ∀f:ℕ ⟶ 𝔹.
  ((∀n:ℕ(x ∈ [fst(cantor-interval(a;b;f;n)), snd(cantor-interval(a;b;f;n))]))  (cantor-to-interval(a;b;f) x))

Lemma: cantor-to-interval-onto-lemma
a,b:ℝ.
  ∀x:ℝ. ∀n:ℕ. ∀f:{f:ℕn ⟶ 𝔹x ∈ [fst(cantor-interval(a;b;f;n)), snd(cantor-interval(a;b;f;n))]} .
    ∃g:{g:ℕ1 ⟶ 𝔹x ∈ [fst(cantor-interval(a;b;g;n 1)), snd(cantor-interval(a;b;g;n 1))]} (g f ∈ (ℕn ⟶ 𝔹))\000C 
  supposing a < b

Lemma: cantor-to-interval-onto-proper
a,b:ℝ.  ∀x:ℝ((a ≤ x)  (x ≤ b)  (∃f:ℕ ⟶ 𝔹(cantor-to-interval(a;b;f) x))) supposing a < b

Lemma: cantor-to-interval-onto-common
a,b:ℝ.
  ∀x,y:ℝ.
    ((x ∈ [a, b])
     (y ∈ [a, b])
     (∀n:ℕ
          ((|x y| ≤ (2^n a)/6 3^n)
           (∃f,g:ℕ ⟶ 𝔹
               (((cantor-to-interval(a;b;f) x) ∧ (cantor-to-interval(a;b;g) y)) ∧ (f g ∈ (ℕn ⟶ 𝔹))))))) 
  supposing a < b

Lemma: proper-interval-to-int-bounded
a,b:ℝ.
  ∀f:{x:ℝx ∈ [a, b]}  ⟶ ℤ. ∃B:ℕ. ∀x:{x:ℝx ∈ [a, b]} . ∃y:{x:ℝx ∈ [a, b]} ((x y) ∧ (|f y| ≤ B)) supposing a < \000Cb

Definition: cantor_to_interval
cantor_to_interval(a;b;f) ==  ((b a/r1 (b a)) (cantor-to-interval(a;b r1;f) a))

Lemma: cantor_to_interval_wf
[a:ℝ]. ∀[b:{b:ℝa ≤ b} ]. ∀[f:ℕ ⟶ 𝔹].  (cantor_to_interval(a;b;f) ∈ {x:ℝx ∈ [a, b]} )

Lemma: old-proof-of-real-continuity
a,b:ℝ.  ∀f:[a, b] ⟶ℝreal-cont(f;a;b) supposing real-fun(f;a;b) supposing a ≤ b

Definition: real-cont-ps
real-cont-ps(k;a;b;f;x;N) ==
  uniform-continuity-pi-search(
  λm@0.(primrec(x;λ%4,m,%5. ⋅i,x,%7,m,%8. case FiniteCantorDecide(λs.if 
                                                                  cantor-to-interval(a;b;λx.if (x) < ((i 1) 1)
                                                                                               then x
                                                                                               else tt) 
                                                                  (N k)=f 
                                                                          cantor-to-interval(a;b;λx.if (x) < ((i 1) 
                                                                                                      1)
                                                                                                       then x
                                                                                                       else ff) 
                                                                          (N k)
                                                               then inr _.⋅
                                                               else (inl _.⋅));(i 1) 1;λx.x)
                                    of inl(x1) =>
                                    let x1,f x1 
                                    in inr %14.⋅
                                    inr(x1) =>
                                    if m=(i 1) then inl f,g,%3. ⋅else (x f,g,%3. ⋅m ⋅)) 
        f,g,eq. ⋅
        m@0 
        ⋅);
  x;0)

Definition: real-cont-br
real-cont-br(a; b; f; k; N) ==
  FAN(λn,s. if primrec(n;eval Kleene-M(λf@0.int2nat(f cantor-to-interval(a;b;λx.if f@0 x=0 then ff else tt) 
                                                        (N k))) 
                                  
                                  x.if then else fi in
                         if is an integer then inl v
                         else inr i,r. eval Kleene-M(λf@0.int2nat(f 
                                                                        cantor-to-interval(a;b;λx.if f@0 x=0
                                                                                                  then ff
                                                                                                  else tt) 
                                                                        (N k))) 
                                                  
                                                  x.if then else fi in
                                         if is an integer then ff
                                         else r)
           then tt
           else inr x.⋅
           fi )

Lemma: function-proper-continuous
I:Interval. ∀f:I ⟶ℝ.
  ((iproper(I)  (∀x,y:{t:ℝt ∈ I} .  ((x y)  (f[x] f[y]))))  f[x] (proper)continuous for x ∈ I)

Lemma: proper-continuous-is-continuous
I:Interval. (iproper(I)  (∀f:I ⟶ℝ(f[x] (proper)continuous for x ∈  f[x] continuous for x ∈ I)))

Lemma: countable-Heine-Borel-proper
a:ℝ. ∀b:{b:ℝa < b} .
  ∀[C:ℕ ⟶ {x:ℝx ∈ [a, b]}  ⟶ ℙ]
    ((∀n:ℕ. ∀x:{x:ℝx ∈ [a, b]} . ∀y:{y:{x:ℝx ∈ [a, b]} y} .  (C[n;x]  C[n;y]))
     (∀x:{x:ℝx ∈ [a, b]} . ∃n:ℕC[n;x])
     (∃k:ℕ. ∀x:{x:ℝx ∈ [a, b]} . ∃n:ℕk. C[n;x]))

Lemma: countable-Heine-Borel-weak
a:ℝ. ∀b:{b:ℝa ≤ b} .
  ∀[C:ℕ ⟶ {x:ℝx ∈ [a, b]}  ⟶ ℙ]
    ((∀n:ℕ. ∀x:{x:ℝx ∈ [a, b]} . ∀y:{y:{x:ℝx ∈ [a, b]} y} .  (C[n;x]  C[n;y]))
     (∀x:{x:ℝx ∈ [a, b]} . ∃n:ℕC[n;x])
     (∃k:ℕ. ∀x:{x:ℝx ∈ [a, b]} (¬¬(∃n:ℕk. C[n;x]))))

Lemma: continuous-implies-functional
I:Interval. ∀f:I ⟶ℝ.  (f[x] continuous for x ∈  (∀a,b:{x:ℝx ∈ I} .  ((a b)  (f[a] f[b]))))

Lemma: continuous_functionality_wrt_subinterval
I:Interval. ∀[f:I ⟶ℝ]. ∀J:Interval. (J ⊆ I   f[x] continuous for x ∈  f[x] continuous for x ∈ J)

Lemma: continuous_functionality_wrt_rfun-eq
I:Interval. ∀[f1,f2:I ⟶ℝ].  (rfun-eq(I;λx.f1[x];λx.f2[x])  f1[x] continuous for x ∈  f2[x] continuous for x ∈ I)

Lemma: proper-continuous-implies-functional
I:Interval. ∀f:I ⟶ℝ.
  (f[x] (proper)continuous for x ∈  iproper(I)  (∀a,b:{x:ℝx ∈ I} .  ((a b)  (f[a] f[b]))))

Lemma: continuous-limit
I:Interval. ∀f:I ⟶ℝ. ∀y:ℝ. ∀x:ℕ ⟶ ℝ.
  (f(x) continuous for x ∈  lim n→∞.x[n]  (y ∈ I)  (∀n:ℕ(x[n] ∈ I))  lim n→∞.f(x[n]) f(y))

Lemma: function-limit
I:Interval. ∀f:I ⟶ℝ. ∀y:ℝ. ∀x:ℕ ⟶ ℝ.
  ((∀x,y:{x:ℝx ∈ I} .  ((x y)  (f(x) f(y))))
   lim n→∞.x[n] y
   (y ∈ I)
   (∀n:ℕ(x[n] ∈ I))
   lim n→∞.f(x[n]) f(y))

Lemma: dense-in-interval-implies
I:Interval
  ∀[X:{a:ℝa ∈ I}  ⟶ ℙ]
    (dense-in-interval(I;X)
     (∃u,v:{a:ℝa ∈ I} u ≠ v)
     (∀a:{a:ℝa ∈ I} . ∃x:ℕ ⟶ {a:ℝa ∈ I} ((∀n:ℕ(X (x n))) ∧ lim n→∞.x a)))

Lemma: functions-equal-on-dense
I:Interval
  ∀[X:{a:ℝa ∈ I}  ⟶ ℙ]
    (dense-in-interval(I;X)
     (∃u,v:{a:ℝa ∈ I} u ≠ v)
     (∀f,g:I ⟶ℝ.
          ((∀x,y:{x:ℝx ∈ I} .  ((x y)  (f(x) f(y))))
           (∀x,y:{x:ℝx ∈ I} .  ((x y)  (g(x) g(y))))
           (∀x:{x:ℝx ∈ I} ((X x)  (f(x) g(x))))
           (∀a:{x:ℝx ∈ I} (f(a) g(a))))))

Lemma: rationals-dense-in-interval
I:Interval. dense-in-interval(I;λr.∃z:ℤ. ∃d:ℕ+(r (r(z)/r(d))))

Lemma: interior-dense-in-interval
a,b:ℝ.  dense-in-interval([a, b];λr.((a < r) ∧ (r < b)))

Lemma: dyadic-rationals-dense
dense-in-interval((-∞, ∞);λr.∃n:ℤ. ∃m:ℕ+(r (r(n)/r(2^m))))

Lemma: dyadic-scaled-rationals-dense
a:ℝ((r0 < a)  dense-in-interval((-∞, ∞);λr.∃n:ℤ. ∃m:ℕ+(r (r(n) (a/r(2^m))))))

Lemma: functions-equal-on-rationals
I:Interval
  ((∃u,v:{a:ℝa ∈ I} u ≠ v)
   (∀f,g:I ⟶ℝ.
        ((∀x,y:{x:ℝx ∈ I} .  ((x y)  (f(x) f(y))))
         (∀x,y:{x:ℝx ∈ I} .  ((x y)  (g(x) g(y))))
         (∀z:ℤ. ∀d:ℕ+.  (((r(z)/r(d)) ∈ I)  (f((r(z)/r(d))) g((r(z)/r(d))))))
         (∀a:{x:ℝx ∈ I} (f(a) g(a))))))

Lemma: functions-equal-on-interior
a:ℝ. ∀b:{b:ℝa < b} . ∀f,g:[a, b] ⟶ℝ.
  ((∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  (f(x) f(y))))
   (∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  (g(x) g(y))))
   (∀x:{x:ℝx ∈ (a, b)} (f(x) g(x)))
   (∀x:{x:ℝx ∈ [a, b]} (f(x) g(x))))

Lemma: total-function-limit
f:ℝ ⟶ ℝ. ∀y:ℝ. ∀x:ℕ ⟶ ℝ.  ((∀x,y:ℝ.  ((x y)  (f[x] f[y])))  lim n→∞.x[n]  lim n→∞.f[x[n]] f[y])

Lemma: total-function-rational-approx
f:ℝ ⟶ ℝ. ∀y:ℝ.  ((∀x,y:ℝ.  ((x y)  (f[x] f[y])))  lim n→∞.f[(y within 1/n 1)] f[y])

Lemma: real-path-comp-exists
f,g:[r0, r1] ⟶ℝ.
  ((∀x,y:{x:ℝx ∈ [r0, r1]} .  ((x y)  (f(x) f(y))))
   (∀x,y:{x:ℝx ∈ [r0, r1]} .  ((x y)  (g(x) g(y))))
   (f(r1) g(r0))
   (∃h:[r0, r1] ⟶ℝ
       ((∀t:{x:ℝx ∈ [r0, (r1/r(2))]} (h(t) f(r(2) t)))
       ∧ (∀t:{x:ℝx ∈ [(r1/r(2)), r1]} (h(t) g((r(2) t) r1)))
       ∧ (∀x,y:{x:ℝx ∈ [r0, r1]} .  ((x y)  (h(x) h(y)))))))

Lemma: continuous-const
[I:Interval]. ∀[a:ℝ].  continuous for x ∈ I

Lemma: continuous-id
[I:Interval]. continuous for x ∈ I

Lemma: continuous-add
[I:Interval]. ∀[f,g:I ⟶ℝ].
  (f[x] continuous for x ∈  g[x] continuous for x ∈  f[x] g[x] continuous for x ∈ I)

Lemma: continuous-add-ext
[I:Interval]. ∀[f,g:I ⟶ℝ].
  (f[x] continuous for x ∈  g[x] continuous for x ∈  f[x] g[x] continuous for x ∈ I)

Lemma: continuous-sum
I:Interval. ∀n,m:ℤ. ∀f:{n..m 1-} ⟶ I ⟶ℝ.
  ((∀i:{n..m 1-}. f[i;x] continuous for x ∈ I)  Σ{f[i;x] n≤i≤m} continuous for x ∈ I)

Lemma: continuous-max
[I:Interval]. ∀[f,g:I ⟶ℝ].
  (f[x] continuous for x ∈  g[x] continuous for x ∈  rmax(f[x];g[x]) continuous for x ∈ I)

Lemma: continuous-minus
[I:Interval]. ∀[f:I ⟶ℝ].  (f[x] continuous for x ∈  -(f[x]) continuous for x ∈ I)

Lemma: continuous-sub
[I:Interval]. ∀[f,g:I ⟶ℝ].
  (f[x] continuous for x ∈  g[x] continuous for x ∈  f[x] g[x] continuous for x ∈ I)

Lemma: continuous-abs
[I:Interval]. ∀[f:I ⟶ℝ].  (f[x] continuous for x ∈  |f[x]| continuous for x ∈ I)

Lemma: continuous-abs-ext
[I:Interval]. ∀[f:I ⟶ℝ].  (f[x] continuous for x ∈  |f[x]| continuous for x ∈ I)

Lemma: continuous-abs-subtype
[I:Interval]. ∀[f:I ⟶ℝ].  (f[x] continuous for x ∈ I ⊆|f[x]| continuous for x ∈ I)

Lemma: continuous-range-totally-bounded
I:Interval. ∀f:I ⟶ℝ.
  (f[x] continuous for x ∈  (∀m:ℕ+(i-nonvoid(i-approx(I;m))  totally-bounded(f[x](x∈i-approx(I;m))))))

Lemma: continuous-compact-range-totally-bounded
I:Interval. ∀f:I ⟶ℝ.  (icompact(I)  f[x] continuous for x ∈  totally-bounded(f[x](x∈I)))

Lemma: sup-range
I:{I:Interval| icompact(I)} . ∀f:I ⟶ℝ.  (f[x] continuous for x ∈  (∃y:ℝsup(f(x)(x∈I)) y))

Lemma: sup-range-no-mc
sup-range is Bishop's theorem that the range of continuous function
has supremum. To apply that theorem we need modulus of continuity for f.

Since we can prove Brouwer's theorem Error :real-continuity1
 that all functions on proper interval are continuous,
we can prove this version of the theorem that does not require 
modulus of continuity.

We have to take some care to handle the general case where the interval
may not be proper.

The constructive content of this theorem will be terribly inefficient
since it will first construct modulus of continuity (using bar recursion).
But, the constructive content of the original theorem of Bishop was already
very inefficient. We should think of the supremum as the "specification" of
(well defined) real number, that one may be able to compute efficiently 
by other means.


I:{I:Interval| icompact(I)} . ∀f:{f:I ⟶ℝifun(f;I)} .  ∃y:ℝsup(f(x)(x∈I)) y

Lemma: inf-range
I:{I:Interval| icompact(I)} . ∀f:I ⟶ℝ.  (f[x] continuous for x ∈  (∃y:ℝinf(f(x)(x∈I)) y))

Lemma: inf-range-no-mc
I:{I:Interval| icompact(I)} . ∀f:{f:I ⟶ℝifun(f;I)} .  ∃y:ℝinf(f(x)(x∈I)) y

Definition: range-sup
sup{f[x]|x ∈ I} ==  fst((TERMOF{sup-range:o, 1:l} x.f[x]) mc))

Lemma: range-sup_wf
[I:{I:Interval| icompact(I)} ]. ∀[f:I ⟶ℝ]. ∀[mc:f[x] continuous for x ∈ I].  (sup{f[x]|x ∈ I} ∈ ℝ)

Lemma: range-sup-property
I:{I:Interval| icompact(I)} . ∀f:I ⟶ℝ. ∀mc:f[x] continuous for x ∈ I.  sup(f[x](x∈I)) sup{f[x]|x ∈ I}

Definition: range_sup
sup{f[x] x ∈ I} ==  fst((TERMOF{sup-range-no-mc:o, 1:l} x.f[x])))

Lemma: range_sup_wf
[I:{I:Interval| icompact(I)} ]. ∀[f:{x:ℝx ∈ I}  ⟶ ℝ].
  sup{f[x] x ∈ I} ∈ ℝ supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Definition: range_inf
inf{f[x] x ∈ I} ==  fst((TERMOF{inf-range-no-mc:o, 1:l} x.f[x])))

Lemma: range_inf_wf
[I:{I:Interval| icompact(I)} ]. ∀[f:{x:ℝx ∈ I}  ⟶ ℝ].
  inf{f[x] x ∈ I} ∈ ℝ supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: range_sup-property
I:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  sup(f[x](x∈I)) sup{f[x] x ∈ I} supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: range_sup-bound
I:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  ∀[c:ℝ]. sup{f[x] x ∈ I} ≤ supposing ∀x:ℝ((x ∈ I)  (f[x] ≤ c)) 
  supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: range_sup-const
I:{I:Interval| icompact(I)} . ∀[c:ℝ]. (sup{c x ∈ I} c)

Lemma: rleq-range_sup
I:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  ∀[c:{c:ℝc ∈ I} ]. (f[c] ≤ sup{f[x] x ∈ I}) supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: rleq-range_sup-2
I:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  ∀[c:ℝ]. c ≤ sup{f[x] x ∈ I} supposing ∃x:ℝ((x ∈ I) ∧ (c ≤ f[x])) 
  supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: range_sup_functionality_wrt_subinterval
I:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  ∀J:{J:Interval| icompact(J)} (J ⊆ I   (sup{f[x] x ∈ J} ≤ sup{f[x] x ∈ I})) 
  supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Definition: range-inf
inf{f[x]|x ∈ I} ==  fst((TERMOF{inf-range:o, 1:l} x.f[x]) mc))

Lemma: range-inf_wf
[I:{I:Interval| icompact(I)} ]. ∀[f:I ⟶ℝ]. ∀[mc:f[x] continuous for x ∈ I].  (inf{f[x]|x ∈ I} ∈ ℝ)

Lemma: range-inf-property
I:{I:Interval| icompact(I)} . ∀f:I ⟶ℝ. ∀mc:f[x] continuous for x ∈ I.  inf(f[x](x∈I)) inf{f[x]|x ∈ I}

Lemma: range_inf-property
I:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  inf(f[x](x∈I)) inf{f[x] x ∈ I} supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: range_inf-bound
I:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  ∀[c:ℝ]. c ≤ inf{f[x] x ∈ I} supposing ∀x:ℝ((x ∈ I)  (c ≤ f[x])) 
  supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: range_inf_functionality
I:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  ∀g:{x:ℝx ∈ I}  ⟶ ℝinf{f[x] x ∈ I} inf{g[x] x ∈ I} supposing ∀x:{x:ℝx ∈ I} (f[x] g[x]) 
  supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: range_sup_functionality2
I,J:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  ∀g:{x:ℝx ∈ J}  ⟶ ℝ
    (sup{f[x] x ∈ I} sup{g[x] x ∈ J}) supposing 
       (((∀x:{x:ℝx ∈ I} . ∃y:{x:ℝx ∈ J} (f[x] g[y])) ∧ (∀x:{x:ℝx ∈ J} . ∃y:{x:ℝx ∈ I} (f[y] g[x]))) and 
       (∀x,y:{x:ℝx ∈ J} .  ((x y)  (g[x] g[y])))) 
  supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: range_sup_functionality
I:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  ∀g:{x:ℝx ∈ I}  ⟶ ℝsup{f[x] x ∈ I} sup{g[x] x ∈ I} supposing ∀x:{x:ℝx ∈ I} (f[x] g[x]) 
  supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Definition: Inorm
||f[x]||_I ==  sup{|f[x]||x ∈ I}

Lemma: Inorm_wf
[I:{I:Interval| icompact(I)} ]. ∀[f:I ⟶ℝ]. ∀[mc:f[x] continuous for x ∈ I].  (||f[x]||_I ∈ ℝ)

Lemma: Inorm-bound
[I:{I:Interval| icompact(I)} ]. ∀[f:I ⟶ℝ]. ∀[mc:f[x] continuous for x ∈ I]. ∀[x:{r:ℝr ∈ I} ].  (|f[x]| ≤ ||f[x]||_I)

Lemma: Inorm-non-neg
[I:{I:Interval| icompact(I)} ]. ∀[f:I ⟶ℝ]. ∀[mc:f[x] continuous for x ∈ I].  (r0 ≤ ||f[x]||_I)

Lemma: function-diff-small-or-interval-proper
I:Interval. ∀f:I ⟶ℝ. ∀x,y:{x:ℝx ∈ I} . ∀e:{e:ℝr0 < e} .
  (f[x] continuous for x ∈  ((|f[x] f[y]| ≤ e) ∨ (r0 < |x y|)))

Definition: I-norm
||f[x]||_x:I ==  sup{|f[x]| x ∈ I}

Lemma: I-norm_wf
[I:{I:Interval| icompact(I)} ]. ∀[f:{x:ℝx ∈ I}  ⟶ ℝ].
  ||f[x]||_x:I ∈ ℝ supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: I-norm-bound
[I:{I:Interval| icompact(I)} ]. ∀[f:{x:ℝx ∈ I}  ⟶ ℝ].
  ∀[x:{r:ℝr ∈ I} ]. (|f[x]| ≤ ||f[x]||_x:I) supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: I-norm-rleq
[I:{I:Interval| icompact(I)} ]. ∀[f:{x:ℝx ∈ I}  ⟶ ℝ].
  ∀[c:ℝ]. ||f[x]||_x:I ≤ supposing ∀[x:{r:ℝr ∈ I} ]. (|f[x]| ≤ c) 
  supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: I-norm_functionality_wrt_subinterval
[I:{I:Interval| icompact(I)} ]. ∀[f:{x:ℝx ∈ I}  ⟶ ℝ].
  ∀[J:{J:Interval| icompact(J)} ]. ||f[x]||_x:J ≤ ||f[x]||_x:I supposing J ⊆ I  
  supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: I-norm_functionality
[I:{I:Interval| icompact(I)} ]. ∀[f:{x:ℝx ∈ I}  ⟶ ℝ].
  ∀[g:{x:ℝx ∈ I}  ⟶ ℝ]. ||f[x]||_x:I ||g[x]||_x:I supposing ∀x:{x:ℝx ∈ I} (f[x] g[x]) 
  supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: I-norm-non-neg
[I:{I:Interval| icompact(I)} ]. ∀[f:{x:ℝx ∈ I}  ⟶ ℝ].
  r0 ≤ ||f[x]||_x:I supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: I-norm-positive-implies
I:{I:Interval| icompact(I)} . ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  (r0 < ||f[x]||_x:I)  (∃c:{c:ℝc ∈ I} (r0 < |f[c]|)) supposing ∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y]))

Lemma: real-fun-bounded
a:ℝ. ∀b:{b:ℝa ≤ b} . ∀f:[a, b] ⟶ℝ.
  ∃bnd:ℝ((r0 ≤ bnd) ∧ (∀x:ℝ((x ∈ [a, b])  (|f x| ≤ bnd)))) 
  supposing ∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  (f[x] f[y]))

Lemma: real-fun-uniformly-positive
a:ℝ. ∀b:{b:ℝa ≤ b} . ∀f:[a, b] ⟶ℝ.
  (real-fun(f;a;b)  (∀x:{x:ℝx ∈ [a, b]} (r0 < (f x)))  (∃c:{c:ℝr0 < c} . ∀x:{x:ℝx ∈ [a, b]} (c < (f x))))

Lemma: real-fun-uniformly-greater
a:ℝ. ∀b:{b:ℝa ≤ b} . ∀f:[a, b] ⟶ℝ.
  (real-fun(f;a;b)
   (∀c:ℝ((∀x:{x:ℝx ∈ [a, b]} (c < (f x)))  (∃c':{c':ℝc < c'} . ∀x:{x:ℝx ∈ [a, b]} (c' ≤ (f x))))))

Lemma: real-fun-uniformly-less
a:ℝ. ∀b:{b:ℝa ≤ b} . ∀f:[a, b] ⟶ℝ.
  (real-fun(f;a;b)
   (∀c:ℝ((∀x:{x:ℝx ∈ [a, b]} ((f x) < c))  (∃c':{c':ℝc' < c} . ∀x:{x:ℝx ∈ [a, b]} ((f x) ≤ c')))))

Lemma: continuous-mul
I:Interval. ∀f,g:I ⟶ℝ.  (f[x] continuous for x ∈  g[x] continuous for x ∈  f[x] g[x] continuous for x ∈ I)

Lemma: continuous-rnexp
I:Interval. ∀n:ℕ.  x^n continuous for x ∈ I

Lemma: continuous-rnexp2
I:Interval. ∀f:I ⟶ℝ.  (f[x] continuous for x ∈  (∀n:ℕf[x]^n continuous for x ∈ I))

Lemma: continuous-rpolynomial
n:ℕ. ∀a:ℕ1 ⟶ ℝ. ∀I:Interval.  i≤n. a_i x^i) continuous for x ∈ I

Definition: nonzero-on
f[x]≠r0 for x ∈ ==
  ∀m:{m:ℕ+icompact(i-approx(I;m))} (∃c:ℝ [((r0 < c) ∧ (∀x:ℝ((x ∈ i-approx(I;m))  (c ≤ |f[x]|))))])

Lemma: nonzero-on_wf
[I:Interval]. ∀[f:I ⟶ℝ].  (f[x]≠r0 for x ∈ I ∈ ℙ)

Lemma: const-nonzero-on
[I:Interval]. ∀a:ℝ(a ≠ r0  a≠r0 for x ∈ I)

Lemma: nonzero-on-implies
I:Interval. ∀f:I ⟶ℝ.  (f[x]≠r0 for x ∈  (∀x:ℝ((x ∈ I)  f[x] ≠ r0)))

Definition: convex-on
convex-on(I;x.f[x]) ==
  ∀x,y,t:ℝ.  ((x ∈ I)  (y ∈ I)  (t ∈ [r0, r1])  (f[(t x) ((r1 t) y)] ≤ ((t f[x]) ((r1 t) f[y]))))

Lemma: convex-on_wf
[I:Interval]. ∀[f:I ⟶ℝ].  (convex-on(I;x.f[x]) ∈ ℙ)

Lemma: implies-convex-on
[I:Interval]. ∀[f:I ⟶ℝ].
  ((∀x,y:ℝ.  ((x ∈ I)  (y ∈ I)  (x y)  (f[x] f[y])))
   (∀x,y:ℝ.
        ((x < y)
         (∀t:ℝ
              ((x ∈ I)
               (y ∈ I)
               (t ∈ [r0, r1])
               (f[(t x) ((r1 t) y)] ≤ ((t f[x]) ((r1 t) f[y])))))))
   convex-on(I;x.f[x]))

Definition: concave-on
concave-on(I;x.f[x]) ==
  ∀x,y,t:ℝ.  ((x ∈ I)  (y ∈ I)  (t ∈ [r0, r1])  (((t f[x]) ((r1 t) f[y])) ≤ f[(t x) ((r1 t) y)]))

Lemma: concave-on_wf
[I:Interval]. ∀[f:I ⟶ℝ].  (concave-on(I;x.f[x]) ∈ ℙ)

Lemma: concave-positive-nonzero-on
I:Interval. ∀f:I ⟶ℝ.
  ((∀x,y:ℝ.  ((x ∈ I)  (y ∈ I)  (x y)  (f[x] f[y])))
   (∀x:ℝ((x ∈ I)  (r0 < f[x])))
   concave-on(I;x.f[x])
   f[x]≠r0 for x ∈ I)

Lemma: convex-negative-nonzero-on
I:Interval. ∀f:I ⟶ℝ.
  ((∀x,y:ℝ.  ((x ∈ I)  (y ∈ I)  (x y)  (f[x] f[y])))
   (∀x:ℝ((x ∈ I)  (f[x] < r0)))
   convex-on(I;x.f[x])
   f[x]≠r0 for x ∈ I)

Lemma: rmul-nonzero-on
I:Interval. ∀f,g:I ⟶ℝ.  (f[x]≠r0 for x ∈  g[x]≠r0 for x ∈  f[x] g[x]≠r0 for x ∈ I)

Lemma: continuous-rinv
I:Interval. ∀f:I ⟶ℝ.  (f[x] continuous for x ∈  f[x]≠r0 for x ∈  (r1/f[x]) continuous for x ∈ I)

Lemma: continuous-rdiv
I:Interval. ∀f,g:I ⟶ℝ.
  (f[x] continuous for x ∈  g[x] continuous for x ∈  g[x]≠r0 for x ∈  (f[x]/g[x]) continuous for x ∈ I)

Definition: maps-compact
maps-compact(I;J;x.f[x]) ==
  ∀n:{n:ℕ+icompact(i-approx(I;n))} 
    ∃m:{m:ℕ+icompact(i-approx(J;m))} . ∀x:{x:ℝx ∈ i-approx(I;n)} (f[x] ∈ i-approx(J;m))

Lemma: maps-compact_wf
[I,J:Interval]. ∀[f:I ⟶ℝ].  (maps-compact(I;J;x.f[x]) ∈ ℙ)

Definition: interval-fun
interval-fun(I;J;x.f[x]) ==  (∀x:{x:ℝx ∈ I} (f[x] ∈ J)) ∧ (∀x,y:{x:ℝx ∈ I} .  f[x] f[y] supposing y)

Lemma: interval-fun_wf
[I,J:Interval]. ∀[f:I ⟶ℝ].  (interval-fun(I;J;x.f[x]) ∈ ℙ)

Lemma: interval-fun-real-fun
[a,b:ℝ]. ∀[f:[a, b] ⟶ℝ].  real-fun(f;a;b) supposing ∃J:Interval. interval-fun([a, b];J;x.f[x])

Lemma: interval-fun-maps-compact
I,J:Interval. ∀f:I ⟶ℝ.  (interval-fun(I;J;x.f[x])  maps-compact(I;J;x.f[x]))

Definition: maps-compact-proper
maps-compact-proper(I;J;x.f[x]) ==
  ∀n:{n:ℕ+icompact(i-approx(I;n)) ∧ iproper(i-approx(I;n))} 
    ∃m:{m:ℕ+icompact(i-approx(J;m)) ∧ iproper(i-approx(J;m))} . ∀x:{x:ℝx ∈ i-approx(I;n)} (f[x] ∈ i-approx(J;m))

Lemma: maps-compact-proper_wf
[I,J:Interval]. ∀[f:I ⟶ℝ].  (maps-compact-proper(I;J;x.f[x]) ∈ ℙ)

Lemma: proper-maps-compact
I,J:Interval. ∀f:I ⟶ℝ.  (iproper(J)  maps-compact(I;J;x.f[x])  maps-compact-proper(I;J;x.f[x]))

Lemma: proper-continuous-maps-compact
I:Interval. ∀f:I ⟶ℝ.  (f[x] (proper)continuous for x ∈  maps-compact-proper(I;(-∞, ∞);x.f[x]))

Lemma: continuous-maps-compact
I:Interval. ∀f:I ⟶ℝ.  (iproper(I)  f[x] (proper)continuous for x ∈  maps-compact(I;(-∞, ∞);x.f[x]))

Lemma: function-maps-compact
I:Interval. ∀f:I ⟶ℝ.
  (iproper(I)  (∀x,y:{t:ℝt ∈ I} .  ((x y)  (f[x] f[y])))  maps-compact(I;(-∞, ∞);x.f[x]))

Lemma: monotone-maps-compact
I,J:Interval. ∀f:I ⟶ℝ.
  (((∀x,y:{t:ℝt ∈ I} .  ((x ≤ y)  (f[x] ≤ f[y]))) ∨ (∀x,y:{t:ℝt ∈ I} .  ((x ≤ y)  (f[y] ≤ f[x]))))
   (∀x:{t:ℝt ∈ I} (f[x] ∈ J))
   maps-compact(I;J;x.f[x]))

Lemma: has-minimum-maps-compact
I:Interval. ∀l:ℝ. ∀f:I ⟶ℝ.
  ((∀x,y:{t:ℝt ∈ I} .  ((x y)  (f[x] f[y])))
   (∀x:{t:ℝt ∈ I} (l < f[x]))
   (∀a:{a:ℝa ∈ I} . ∀b:{b:ℝ(b ∈ I) ∧ (a ≤ b)} .  ∃c:{t:ℝt ∈ [a, b]} . ∀x:{t:ℝt ∈ [a, b]} (f[c] ≤ f[x]))
   maps-compact(I;(l, ∞);x.f[x]))

Lemma: continuous-composition-from-maps-compact
I,J:Interval. ∀f:I ⟶ℝ. ∀g:J ⟶ℝ.
  (maps-compact(I;J;x.f[x])  f[x] continuous for x ∈  g[x] continuous for x ∈  g[f[x]] continuous for x ∈ I)

Lemma: continuous-composition
Bishop's proof continuous-composition-from-maps-compact
that composition of continuous functions is continuous requires
an extra hypothesis, maps-compact(I;J;x.f[x]).

But because we can prove Brouwer's theorem on continuity,
 Error :real-continuity1 
we can prove this theorem that does not require the extra hypothesis.⋅

I:Interval
  (iproper(I)
   (∀J:Interval. ∀f:{x:ℝx ∈ I}  ⟶ {y:ℝy ∈ J} . ∀g:J ⟶ℝ.
        (f[x] continuous for x ∈  g[x] continuous for x ∈  g[f[x]] continuous for x ∈ I)))

Lemma: function-values-near-same-sign
I:Interval. ∀f:{x:ℝx ∈ I}  ⟶ ℝ.
  (icompact(I)
   (∀x,y:{x:ℝx ∈ I} .  ((x y)  (f[x] f[y])))
   (∀x:{x:ℝx ∈ I} 
        ((r0 < |f[x]|)
         (∃d:{d:ℝr0 < d} 
             ∀y:{x:ℝx ∈ I} ((|x y| ≤ d)  ((r0 < f[x] ⇐⇒ r0 < f[y]) ∧ (f[x] < r0 ⇐⇒ f[y] < r0)))))))

Lemma: i-closed-closed
I:Interval. (i-closed(I)  closed-rset(λx.(x ∈ I)))

Lemma: superlevelset-closed
I:Interval. ∀[f:I ⟶ℝ]. ∀[c:ℝ].  (i-closed(I)  f(x) continuous for x ∈  closed-rset(superlevelset(I;f;c)))

Lemma: sublevelset-closed
I:Interval. ∀[f:I ⟶ℝ]. ∀[c:ℝ].  (i-closed(I)  f(x) continuous for x ∈  closed-rset(sublevelset(I;f;c)))

Lemma: intermediate-value-theorem
I:Interval. ∀f:I ⟶ℝ.
  (f[x] continuous for x ∈ I
   (∀a,b:{x:ℝx ∈ I} .
        ((f(a) < f(b))
         (∀y:{y:ℝy ∈ [f(a), f(b)]} . ∀e:{e:ℝr0 < e} .  ∃x:{x:ℝx ∈ [rmin(a;b), rmax(a;b)]} (|f(x) y| < e)))))

Lemma: intermediate-value-lemma
a:ℝ. ∀b:{b:ℝa < b} . ∀f:[a, b] ⟶ℝ.
  ∀c:{c:ℝ(f(a) ≤ c) ∧ (c ≤ f(b))} 
    ((∃d:{d:ℝ(r0 ≤ d) ∧ (d < r1)} 
       ∀a1:{a1:ℝ(a1 ∈ [a, b]) ∧ (f(a1) ≤ c)} . ∀b1:{b1:ℝ(b1 ∈ [a, b]) ∧ (c ≤ f(b1)) ∧ (a1 ≤ b1)} .
         ∃a2:{a2:ℝ(a2 ∈ [a, b]) ∧ (f(a2) ≤ c)} (∃b2:{b2:ℝ(b2 ∈ [a, b]) ∧ (c ≤ f(b2))}  [((a1 ≤ a2) ∧ (a2 ≤ b2) ∧ (\000Cb2 ≤ b1) ∧ ((b2 a2) ≤ ((b1 a1) d)))]))
     (∃x:ℝ [((x ∈ [a, b]) ∧ (f(x) c))])) 
  supposing ∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  (f[x] f[y]))

Lemma: intermediate-value-lemma-ext
a:ℝ. ∀b:{b:ℝa < b} . ∀f:[a, b] ⟶ℝ.
  ∀c:{c:ℝ(f(a) ≤ c) ∧ (c ≤ f(b))} 
    ((∃d:{d:ℝ(r0 ≤ d) ∧ (d < r1)} 
       ∀a1:{a1:ℝ(a1 ∈ [a, b]) ∧ (f(a1) ≤ c)} . ∀b1:{b1:ℝ(b1 ∈ [a, b]) ∧ (c ≤ f(b1)) ∧ (a1 ≤ b1)} .
         ∃a2:{a2:ℝ(a2 ∈ [a, b]) ∧ (f(a2) ≤ c)} (∃b2:{b2:ℝ(b2 ∈ [a, b]) ∧ (c ≤ f(b2))}  [((a1 ≤ a2) ∧ (a2 ≤ b2) ∧ (\000Cb2 ≤ b1) ∧ ((b2 a2) ≤ ((b1 a1) d)))]))
     (∃x:ℝ [((x ∈ [a, b]) ∧ (f(x) c))])) 
  supposing ∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  (f[x] f[y]))

Lemma: intermediate-value-lemma'
a:ℝ. ∀b:{b:ℝa < b} . ∀f:[a, b] ⟶ℝ.
  ∀c:{c:ℝ(f(a) ≤ c) ∧ (c ≤ f(b))} 
    ((∃d:{d:ℝ(r0 ≤ d) ∧ (d < r1)} 
       ∀a1:{a1:ℝ(a1 ∈ [a, b]) ∧ (f(a1) ≤ c)} . ∀b1:{b1:ℝ(b1 ∈ [a, b]) ∧ (c ≤ f(b1)) ∧ (a1 < b1)} .
         ∃a2:{a2:ℝ(a2 ∈ [a, b]) ∧ (f(a2) ≤ c)} (∃b2:{b2:ℝ(b2 ∈ [a, b]) ∧ (c ≤ f(b2))}  [((a1 ≤ a2) ∧ (a2 < b2) ∧ (\000Cb2 ≤ b1) ∧ ((b2 a2) ≤ ((b1 a1) d)))]))
     (∃x:ℝ [((x ∈ [a, b]) ∧ (f(x) c))])) 
  supposing ∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  (f[x] f[y]))

Lemma: intermediate-value-lemma'-ext
a:ℝ. ∀b:{b:ℝa < b} . ∀f:[a, b] ⟶ℝ.
  ∀c:{c:ℝ(f(a) ≤ c) ∧ (c ≤ f(b))} 
    ((∃d:{d:ℝ(r0 ≤ d) ∧ (d < r1)} 
       ∀a1:{a1:ℝ(a1 ∈ [a, b]) ∧ (f(a1) ≤ c)} . ∀b1:{b1:ℝ(b1 ∈ [a, b]) ∧ (c ≤ f(b1)) ∧ (a1 < b1)} .
         ∃a2:{a2:ℝ(a2 ∈ [a, b]) ∧ (f(a2) ≤ c)} (∃b2:{b2:ℝ(b2 ∈ [a, b]) ∧ (c ≤ f(b2))}  [((a1 ≤ a2) ∧ (a2 < b2) ∧ (\000Cb2 ≤ b1) ∧ ((b2 a2) ≤ ((b1 a1) d)))]))
     (∃x:ℝ [((x ∈ [a, b]) ∧ (f(x) c))])) 
  supposing ∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  (f[x] f[y]))

Definition: locally-non-constant
locally-non-constant(f;a;b;c) ==  ∀u,v:ℝ.  ((a ≤ u)  (u < v)  (v ≤ b)  (∃z:ℝ((u ≤ z) ∧ (z ≤ v) ∧ f(z) ≠ c)))

Lemma: locally-non-constant_wf
[a,b,c:ℝ]. ∀[f:[a, b] ⟶ℝ].  (locally-non-constant(f;a;b;c) ∈ ℙ)

Definition: locally-non-constant-rational
locally-non-constant-rational(f;a;b;c) ==
  ∀u,v:ℝ.  ((a ≤ u)  (u < v)  (v ≤ b)  (∃n:ℕ+. ∃z:ℤ((u ≤ (r(z))/n) ∧ ((r(z))/n ≤ v) ∧ f((r(z))/n) ≠ c)))

Lemma: locally-non-constant-rational_wf
[a,b,c:ℝ]. ∀[f:[a, b] ⟶ℝ].  (locally-non-constant-rational(f;a;b;c) ∈ ℙ)

Lemma: locally-non-constant-via-rational
a,b,c:ℝ. ∀f:[a, b] ⟶ℝ.
  (f[x] continuous for x ∈ [a, b]  locally-non-constant(f;a;b;c)  locally-non-constant-rational(f;a;b;c))

Lemma: IVT-locally-non-constant
a:ℝ. ∀b:{b:ℝa < b} . ∀f:[a, b] ⟶ℝ.
  ∀c:ℝ((f(a) ≤ c)  (c ≤ f(b))  locally-non-constant(f;a;b;c)  (∃x:ℝ [((x ∈ [a, b]) ∧ (f(x) c))])) 
  supposing ∀x,y:{x:ℝx ∈ [a, b]} .  ((x y)  (f[x] f[y]))

Lemma: strict-monotonic-is-locally-non-constant
I:Interval. ∀f:I ⟶ℝ.
  (((∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f x) < (f y)))) ∨ (∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f y) < (f x)))))
   (∀a,b:{x:ℝx ∈ I} . ∀x:ℝ.  locally-non-constant(f;a;b;x)))

Lemma: IVT-strict-increasing
I:Interval. ∀f:I ⟶ℝ.
  ((∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f x) < (f y))))
   (∀x,y:{t:ℝt ∈ I} .  ((x y)  ((f x) (f y))))
   (∀a,b:{x:ℝx ∈ I} .
        ((a < b)  (∀x:ℝ((((f a) ≤ x) ∧ (x ≤ (f b)))  (∃c:ℝ [(((a ≤ c) ∧ (c ≤ b)) ∧ ((f c) x))]))))))

Lemma: IVT-strict-decreasing
I:Interval. ∀f:I ⟶ℝ.
  ((∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f y) < (f x))))
   (∀x,y:{t:ℝt ∈ I} .  ((x y)  ((f x) (f y))))
   (∀a,b:{x:ℝx ∈ I} .
        ((a < b)  (∀x:ℝ((((f b) ≤ x) ∧ (x ≤ (f a)))  (∃c:ℝ(((a ≤ c) ∧ (c ≤ b)) ∧ ((f c) x))))))))

Lemma: inverse-of-strict-increasing-function
I:Interval. ∀f:I ⟶ℝ. ∀J:Interval. ∀g:x:{x:ℝx ∈ J}  ⟶ {x:ℝx ∈ I} .
  ((∀t:{t:ℝt ∈ I} (f t ∈ J))
   (∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f x) < (f y))))
   (∀x,y:{t:ℝt ∈ I} .  ((x y)  ((f x) (f y))))
   (∀x:{x:ℝx ∈ J} ((f (g x)) x))
   ((∀x:{x:ℝx ∈ I} ((g (f x)) x))
     ∧ (∀x,y:{x:ℝx ∈ J} .  ((x < y)  ((g x) < (g y))))
     ∧ (∀x,y:{t:ℝt ∈ J} .  ((x y)  ((g x) (g y))))))

Lemma: inverse-of-strict-increasing-function-exists
I:Interval. ∀f:I ⟶ℝ.
  ((∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f x) < (f y))))
   (∀x,y:{t:ℝt ∈ I} .  ((x y)  ((f x) (f y))))
   (∀J:Interval
        ((∀t:{t:ℝt ∈ I} (f t ∈ J))
         (∀x:{x:ℝx ∈ J} . ∃a,b:{t:ℝt ∈ I} ((a < b) ∧ ((f a) ≤ x) ∧ (x ≤ (f b))))
         (∃g:{x:ℝx ∈ J}  ⟶ {x:ℝx ∈ I} 
             ((∀x:{x:ℝx ∈ J} ((f (g x)) x))
             ∧ (∀x:{x:ℝx ∈ I} ((g (f x)) x))
             ∧ (∀x,y:{x:ℝx ∈ J} .  ((x < y)  ((g x) < (g y))))
             ∧ (∀x,y:{t:ℝt ∈ J} .  ((x y)  ((g x) (g y)))))))))

Lemma: inverse-of-strict-decreasing-function
I:Interval. ∀f:I ⟶ℝ. ∀J:Interval. ∀g:x:{x:ℝx ∈ J}  ⟶ {x:ℝx ∈ I} .
  ((∀t:{t:ℝt ∈ I} (f t ∈ J))
   (∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f y) < (f x))))
   (∀x,y:{t:ℝt ∈ I} .  ((x y)  ((f x) (f y))))
   (∀x:{x:ℝx ∈ J} ((f (g x)) x))
   ((∀x:{x:ℝx ∈ I} ((g (f x)) x))
     ∧ (∀x,y:{x:ℝx ∈ J} .  ((x < y)  ((g y) < (g x))))
     ∧ (∀x,y:{t:ℝt ∈ J} .  ((x y)  ((g x) (g y))))))

Lemma: inverse-of-strict-decreasing-function-exists
I:Interval. ∀f:I ⟶ℝ.
  ((∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f y) < (f x))))
   (∀x,y:{t:ℝt ∈ I} .  ((x y)  ((f x) (f y))))
   (∀J:Interval
        ((∀t:{t:ℝt ∈ I} (f t ∈ J))
         (∀x:{x:ℝx ∈ J} . ∃a,b:{t:ℝt ∈ I} ((a < b) ∧ ((f b) ≤ x) ∧ (x ≤ (f a))))
         (∃g:{x:ℝx ∈ J}  ⟶ {x:ℝx ∈ I} 
             ((∀x:{x:ℝx ∈ J} ((f (g x)) x))
             ∧ (∀x:{x:ℝx ∈ I} ((g (f x)) x))
             ∧ (∀x,y:{x:ℝx ∈ J} .  ((x < y)  ((g y) < (g x))))
             ∧ (∀x,y:{t:ℝt ∈ J} .  ((x y)  ((g x) (g y)))))))))

Lemma: IVT-strict-monotonic
I:Interval. ∀f:I ⟶ℝ.
  (((∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f x) < (f y)))) ∨ (∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f y) < (f x)))))
   (∀x,y:{t:ℝt ∈ I} .  ((x y)  ((f x) (f y))))
   (∀a,b:{x:ℝx ∈ I} .
        (a ≠ b
         (∀x:ℝ
              (((((f a) ≤ x) ∧ (x ≤ (f b))) ∨ (((f b) ≤ x) ∧ (x ≤ (f a))))
               (∃c:ℝ((((a ≤ c) ∧ (c ≤ b)) ∨ ((b ≤ c) ∧ (c ≤ a))) ∧ ((f c) x))))))))

Lemma: inverse-of-strict-monotonic-function
I:Interval. ∀f:I ⟶ℝ.
  (((∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f x) < (f y)))) ∨ (∀x,y:{x:ℝx ∈ I} .  ((x < y)  ((f y) < (f x)))))
   (∀x,y:{t:ℝt ∈ I} .  ((x y)  ((f x) (f y))))
   (∀J:Interval
        ((∀t:{t:ℝt ∈ I} (f t ∈ J))
         (∀x:{x:ℝx ∈ J} . ∃a,b:{t:ℝt ∈ I} (a ≠ b ∧ ((((f a) ≤ x) ∧ (x ≤ (f b))) ∨ (((f b) ≤ x) ∧ (x ≤ (f a))))))
         (∃g:{x:ℝx ∈ J}  ⟶ {x:ℝx ∈ I} 
             ((∀x:{x:ℝx ∈ J} ((f (g x)) x))
             ∧ (∀x:{x:ℝx ∈ I} ((g (f x)) x))
             ∧ ((∀x,y:{x:ℝx ∈ J} .  ((x < y)  ((g x) < (g y)))) ∨ (∀x,y:{x:ℝx ∈ J} .  ((x < y)  ((g y) < (g x)))\000C))
             ∧ (∀x,y:{t:ℝt ∈ J} .  ((x y)  ((g x) (g y)))))))))

Lemma: near-fixpoint-on-0-1
f:[r0, r1] ⟶ℝ
  ((∀x:ℝ((x ∈ [r0, r1])  (f[x] ∈ [r0, r1])))
   f[x] continuous for x ∈ [r0, r1]
   (∀e:ℝ((r0 < e)  (∃x:ℝ((x ∈ [r0, r1]) ∧ (|f[x] x| < e))))))

Definition: separated-partitions
separated-partitions(P;Q) ==  frs-increasing(P) ∧ frs-increasing(Q) ∧ frs-separated(P;Q)

Lemma: separated-partitions_wf
I:Interval. ∀P,Q:partition(I).  (separated-partitions(P;Q) ∈ ℙsupposing icompact(I)

Lemma: separated-partitions-have-common-refinement
I:Interval
  ∀P,Q:partition(I).
    (separated-partitions(P;Q)  (∃R:partition(I). (frs-increasing(R) ∧ frs-refines(R;P) ∧ frs-refines(R;Q)))) 
  supposing icompact(I)

Lemma: nearby-separated-partitions
I:Interval
  ((icompact(I) ∧ iproper(I))
   (∀p,q:partition(I). ∀e:{e:ℝr0 < e} .
        ∃p',q':partition(I). (separated-partitions(p';q') ∧ nearby-partitions(e;p;p') ∧ nearby-partitions(e;q;q'))))

Lemma: nearby-partition-choice
I:Interval
  (icompact(I)
   (∀p,q:partition(I). ∀e:{e:ℝr0 < e} .
        (nearby-partitions(e;p;q)
         (∀x:partition-choice(full-partition(I;p))
              ∃y:partition-choice(full-partition(I;q)). ∀i:ℕ||p|| 1. (|x[i] y[i]| ≤ e)))))

Definition: decimal-digits
(d digits of x) ==  eval 10^d in eval N2 N ÷ in eval N2 in eval rem in eval D ÷ in   <Q, R>

Lemma: decimal-digits_wf
d:ℕ+. ∀x:ℝ.  ((d digits of x) ∈ {p:ℤ × ℤlet n,m in |x r(n) (r(m)/r(10^d))| ≤ (r(2)/r(10^d))} )

Definition: dd
For (x N), a/2N is within 1/N of x. So, if (10^n)/2,  will be the
integer that has decimal digits of x.
We actually compute <"display-as", "decimal-rational", x, n, a>
because the "display hook" will display this result as =mmm.kkkkkkk
by using the and the to place the decimal point appropriately.
(One can explode the term so displayed to see the computed x, n, and a)⋅

decimal digits of x  ==
  let answer ⟵ <"display-as", "decimal-rational", x, n, eval 10^n in eval N2 N ÷ in   N2>
  in answer

Lemma: dd_wf
d:ℕ+. ∀x:ℝ.
  (d decimal digits of x  ∈ {a:Atom| "display-as" ∈ Atom} 
   × {a:Atom| "decimal-rational" ∈ Atom} 
   × {z:ℝx} 
   × {n:ℕ+d ∈ ℤ
   × {n:ℤ|x (r(n)/r(10^d))| ≤ (r(2)/r(10^d))} )

Lemma: rnexp-convex
a,b:ℝ.  ((r0 ≤ b)  (b ≤ a)  (∀n:ℕ+(a b^n ≤ (a^n b^n))))

Lemma: rnexp-convex2
a,b:ℝ.  ((r0 ≤ a)  (r0 ≤ b)  (∀n:ℕ+(|a b|^n ≤ |a^n b^n|)))

Lemma: rnexp-convex3
a,b:ℝ.  ((((r0 ≤ a) ∧ (r0 ≤ b)) ∨ ((a ≤ r0) ∧ (b ≤ r0)))  (∀n:ℕ+(|a b|^n ≤ |a^n b^n|)))

Comment: constructing rroot
We didn't know whether the simple program for the i-th root of real x
was correct. Attempts to prove it correct had failed, so we made the
construction in this subdirectory and extracted correct-by-construction
program  extracted-rroot.

This program is fairly efficient. Using it, we compute
50 decimal digits of the 12-root of 2,  ⌜50 decimal digits of extracted-rroot(12;r(2)) ⌝
in 531,134  computation steps, taking about 30 seconds.

However, the program rroot is much simpler, and computes
50 decimal digits of rroot(12;r(2)) ⌝
in only 7,690 steps, almost instantly.

We used the computations find-real-neq and find-real-neq2
to experimentally check that the known correct ⌜extracted-rroot([i];[x])⌝
and the simpler ⌜rroot([i];[x])⌝ always 
(on many reals of various sizes and signs) computed the same integers
(or integers that differed by at most one).

This gave us confidence that the simpler program was also correct
and we were then able to prove that is was correct!.



Lemma: near-root-rational
k:{2...}. ∀p:{p:ℤ(0 ≤ p) ∨ (↑isOdd(k))} . ∀q,n:ℕ+.
  (∃r:ℤ × ℕ+ [let a,b 
              in (0 ≤ ⇐⇒ 0 ≤ a) ∧ (|(r(a))/b^k (r(p)/r(q))| < (r1/r(n)))])

Definition: near-root
near-root(k;p;q;n) ==
  if q=1
  then if n=1
       then eval 2^k in
            eval in
            eval (2 c) in
            eval iroot(k;|a| d) in
            eval ((k M^k 1) ÷ d) in
            eval iroot(k;(|a| d) y^k) ÷ in
              <if (p) < (0)  then -x  else x, y>
       else eval in
            eval b^k in
            eval in
            eval in
            eval iroot(k;|a| d) in
            eval ((k M^k 1) ÷ d) in
            eval iroot(k;(|a| d) y^k) ÷ in
              <if (p) < (0)  then -x  else x, y>
  else eval in
       eval b^k in
       eval in
       eval in
       eval iroot(k;|a| d) in
       eval ((k M^k 1) ÷ d) in
       eval iroot(k;(|a| d) y^k) ÷ in
         <if (p) < (0)  then -x  else x, y>

Lemma: near-root-rational-ext
k:{2...}. ∀p:{p:ℤ(0 ≤ p) ∨ (↑isOdd(k))} . ∀q,n:ℕ+.
  (∃r:ℤ × ℕ+ [let a,b 
              in (0 ≤ ⇐⇒ 0 ≤ a) ∧ (|(r(a))/b^k (r(p)/r(q))| < (r1/r(n)))])

Lemma: near-root_wf
k:{2...}. ∀p:{p:ℤ(0 ≤ p) ∨ (↑isOdd(k))} . ∀q,n:ℕ+.
  (near-root(k;p;q;n) ∈ {r:ℤ × ℕ+let a,b in (0 ≤ ⇐⇒ 0 ≤ a) ∧ (|(r(a))/b^k (r(p)/r(q))| < (r1/r(n)))} )

Lemma: rroot-exists-part1
i:{2...}. ∀x:{x:ℝ(↑isEven(i))  (r0 ≤ x)} .
  ∃q:ℕ ⟶ ℝ
   (lim n→∞.q n^i x
   ∧ (∀n,m:ℕ.  (((r0 ≤ (q n)) ∧ (r0 ≤ (q m))) ∨ (((q n) ≤ r0) ∧ ((q m) ≤ r0))))
   ∧ ((↑isEven(i))  (∀m:ℕ(r0 ≤ (q m)))))

Lemma: rroot-exists1
i:{2...}. ∀x:{x:ℝ(↑isEven(i))  (r0 ≤ x)} .
  ∃q:{q:ℕ ⟶ ℝ
      (∀n,m:ℕ.  (((r0 ≤ (q n)) ∧ (r0 ≤ (q m))) ∨ (((q n) ≤ r0) ∧ ((q m) ≤ r0))))
      ∧ ((↑isEven(i))  (∀m:ℕ(r0 ≤ (q m))))} 
   lim n→∞.q n^i x

Lemma: rroot-exists1-ext
i:{2...}. ∀x:{x:ℝ(↑isEven(i))  (r0 ≤ x)} .
  ∃q:{q:ℕ ⟶ ℝ
      (∀n,m:ℕ.  (((r0 ≤ (q n)) ∧ (r0 ≤ (q m))) ∨ (((q n) ≤ r0) ∧ ((q m) ≤ r0))))
      ∧ ((↑isEven(i))  (∀m:ℕ(r0 ≤ (q m))))} 
   lim n→∞.q n^i x

Lemma: rroot-exists-part2
i:{2...}. ∀x:{x:ℝ(↑isEven(i))  (r0 ≤ x)} . ∀q:{q:ℕ ⟶ ℝ
                                                   (∀n,m:ℕ.
                                                      (((r0 ≤ (q n)) ∧ (r0 ≤ (q m))) ∨ (((q n) ≤ r0) ∧ ((q m) ≤ r0))))
                                                   ∧ ((↑isEven(i))  (∀m:ℕ(r0 ≤ (q m))))} .
  (lim n→∞.q n^i  cauchy(n.q n))

Lemma: rroot-exists
i:{2...}. ∀x:{x:ℝ(↑isEven(i))  (r0 ≤ x)} .  (∃y:ℝ [(((↑isEven(i))  (r0 ≤ y)) ∧ (y^i x))])

Lemma: rroot-exists-ext
i:{2...}. ∀x:{x:ℝ(↑isEven(i))  (r0 ≤ x)} .  (∃y:ℝ [(((↑isEven(i))  (r0 ≤ y)) ∧ (y^i x))])

Definition: extracted-rroot
extracted-rroot(i;x) ==  TERMOF{rroot-exists-ext:o, 1:l} x

Lemma: extracted-rroot_wf
[i:{2...}]. ∀[x:{x:ℝ(↑isEven(i))  (r0 ≤ x)} ].
  (extracted-rroot(i;x) ∈ {y:ℝ((↑isEven(i))  (r0 ≤ y)) ∧ (y^i x)} )

Definition: find-real-neq
find-real-neq(a;b) ==
  eval in
  eval in
    mu-ge-print(λn.eval xn in
                   eval yn in
                   eval xn yn in
                     1 <|d|;0)

Definition: find-real-neq2
find-real-neq2(a;b) ==
  eval in
  eval in
    find-xover-print(λn.eval xn in
                        eval yn in
                        eval xn yn in
                          1 <|d|;0;0;1)

Lemma: exp-2-3-fact
n:ℕ((2 ≤ n)  2^n < 3^n)

Lemma: rroot-regularity-lemma
[k:{2...}]. ∀[n,m:ℕ+]. ∀[a,b,c,d:ℤ].
  (((m ≤ a) ∨ ((a 0 ∈ ℤ) ∧ (c 0 ∈ ℤ)))
   ((n ≤ b) ∨ ((b 0 ∈ ℤ) ∧ (d 0 ∈ ℤ)))
   (a^k ≤ c)
   c < (a m)^k
   (b^k ≤ d)
   d < (b n)^k
   (|c d| ≤ (2^k (n^k m^k)))
   (|a b| ≤ (2 (n m))))

Definition: rroot-abs
rroot-abs(i;x) ==  eval 2^i in λn.eval n^i in eval |x| in   iroot(i;b z)

Lemma: rroot-abs_wf
i:{2...}. ∀x:ℝ.  (rroot-abs(i;x) ∈ ℝ)

Lemma: rroot-abs-property
i:{2...}. ∀x:ℝ.  (rroot-abs(i;x)^i |x|)

Lemma: rroot-abs-non-neg
i:{2...}. ∀x:ℝ. ∀n:ℕ+.  (0 ≤ (rroot-abs(i;x) n))

Definition: rroot-odd
rroot-odd(i;x) ==
  eval 2^i in
  λn.eval n^i in
     eval in
       if z <then -iroot(i;b (-z)) else iroot(i;b z) fi 

Lemma: rroot-odd_wf
i:{2...}. ∀x:ℝ.  (rroot-odd(i;x) ∈ ℕ+ ⟶ ℤ)

Lemma: rroot-odd-2-regular
i:{2...}. ∀x:ℝ.  2-regular-seq(rroot-odd(i;x))

Definition: rroot
rroot(i;x) ==  if isEven(i) then rroot-abs(i;x) else accelerate(2;rroot-odd(i;x)) fi 

Lemma: rroot_wf
[i:{2...}]. ∀[x:{x:ℝ(↑isEven(i))  (r0 ≤ x)} ].  (rroot(i;x) ∈ {y:ℝ((↑isEven(i))  (r0 ≤ y)) ∧ (y^i x)} )

Lemma: rroot_functionality
[i:{2...}]. ∀[x:{x:ℝ(↑isEven(i))  (r0 ≤ x)} ]. ∀[y:ℝ].  rroot(i;x) rroot(i;y) supposing y

Definition: rsqrt
rsqrt(x) ==  rroot(2;x)

Lemma: rsqrt_wf
[x:{x:ℝr0 ≤ x} ]. (rsqrt(x) ∈ {r:ℝ(r0 ≤ r) ∧ ((r r) x)} )

Lemma: rsqrt_functionality
[x:{x:ℝr0 ≤ x} ]. ∀[y:ℝ].  rsqrt(x) rsqrt(y) supposing y

Lemma: rsqrt_nonneg
[x:{x:ℝr0 ≤ x} ]. (r0 ≤ rsqrt(x))

Lemma: rsqrt_squared
[x:{x:ℝr0 ≤ x} ]. ((rsqrt(x) rsqrt(x)) x)

Lemma: rsqrt-rleq-iff
[x:{x:ℝr0 ≤ x} ]. ∀[c:ℝ].  uiff(rsqrt(x) ≤ c;(r0 ≤ c) ∧ (x ≤ c^2))

Lemma: rsqrt-is-zero
[x:{x:ℝr0 ≤ x} ]. (rsqrt(x) r0 ⇐⇒ r0)

Lemma: rsqrt-rnexp-2
[x:{x:ℝr0 ≤ x} ]. (rsqrt(x)^2 x)

Lemma: rsqrt-unique
[x,s:{x:ℝr0 ≤ x} ].  rsqrt(x) supposing (s s) x

Lemma: rsqrt-unique2
[x:{x:ℝr0 ≤ x} ]. ∀[s:ℝ].  uiff((s s) x;¬¬((s rsqrt(x)) ∨ (s -(rsqrt(x)))))

Lemma: rsqrt1
rsqrt(r1) r1

Lemma: rsqrt-is-one
[x:ℝ]. uiff((r0 ≤ x) ∧ (rsqrt(x) r1);x r1)

Lemma: implies-rsqrt-is-one
[x:ℝ]. rsqrt(x) r1 supposing r1

Lemma: rsqrt0
rsqrt(r0) r0

Lemma: rsqrt-of-square
[x:{x:ℝr0 ≤ x} ]. (rsqrt(x x) x)

Lemma: rsqrt_square
[x:ℝ]. (rsqrt(x x) |x|)

Lemma: rsqrt-positive
x:{x:ℝr0 < x} (r0 < rsqrt(x))

Lemma: rsqrt-positive-iff
x:{x:ℝr0 ≤ x} (r0 < rsqrt(x) ⇐⇒ r0 < x)

Lemma: rsqrt-rmul
x:{x:ℝr0 ≤ x} . ∀[y:{x:ℝr0 ≤ x} ]. ((rsqrt(x) rsqrt(y)) rsqrt(x y))

Lemma: rsqrt-rdiv
x:{x:ℝr0 ≤ x} . ∀y:{x:ℝr0 < x} .  ((rsqrt(x)/rsqrt(y)) rsqrt((x/y)))

Lemma: rsqrt_functionality_wrt_rleq
[x:{x:ℝr0 ≤ x} ]. ∀[y:ℝ].  rsqrt(x) ≤ rsqrt(y) supposing x ≤ y

Lemma: rsqrt_functionality_wrt_rless
x:{x:ℝr0 ≤ x} . ∀y:ℝ.  ((x < y)  (rsqrt(x) < rsqrt(y)))

Lemma: rsqrt-rless-iff
x,y:{x:ℝr0 ≤ x} .  (x < ⇐⇒ rsqrt(x) < rsqrt(y))

Lemma: rsqrt-rneq
x,y:{x:ℝr0 ≤ x} .  (rsqrt(x) ≠ rsqrt(y)  x ≠ y)

Definition: quadratic1
quadratic1(a;b;c) ==  (-(b) rsqrt((b b) r(4) c)/r(2) a)

Lemma: quadratic1_wf
[a,b,c:ℝ].  (quadratic1(a;b;c) ∈ ℝsupposing ((r0 ≤ ((b b) r(4) c)) and a ≠ r0)

Lemma: quadratic1_functionality
[a,b,c,a',b',c':ℝ].
  (quadratic1(a;b;c) quadratic1(a';b';c')) supposing 
     ((c c') and 
     (b b') and 
     (a a') and 
     (r0 ≤ ((b b) r(4) c)) and 
     a ≠ r0)

Lemma: quadratic1-zero
[a,b,c:ℝ].  (quadratic1(a;b;c) r0) supposing ((r0 ≤ b) and (c r0) and a ≠ r0)

Lemma: quadratic1-one
[a,b,c:ℝ].  (quadratic1(a;b;c) r1) supposing ((c ≤ a) and (b -(a c)) and (r0 < a))

Definition: quadratic2
quadratic2(a;b;c) ==  (-(b) rsqrt((b b) r(4) c)/r(2) a)

Lemma: quadratic2_wf
[a,b,c:ℝ].  (quadratic2(a;b;c) ∈ ℝsupposing ((r0 ≤ ((b b) r(4) c)) and a ≠ r0)

Lemma: quadratic2_functionality
[a,b,c,a',b',c':ℝ].
  (quadratic2(a;b;c) quadratic2(a';b';c')) supposing 
     ((c c') and 
     (b b') and 
     (a a') and 
     (r0 ≤ ((b b) r(4) c)) and 
     a ≠ r0)

Lemma: quadratic2-zero
[a,b,c:ℝ].  (quadratic2(a;b;c) r0) supposing ((b ≤ r0) and (c r0) and a ≠ r0)

Lemma: quadratic-1-2-equal
[a,b,c:ℝ].
  ({(quadratic1(a;b;c) quadratic2(a;b;c)) ∧ (quadratic1(a;b;c) (-(b)/r(2) a))}) supposing 
     (((b b) (r(4) c)) and 
     a ≠ r0)

Lemma: quadratic-1-2-equal-implies
[a,b,c:ℝ].
  ((b b) (r(4) c)) supposing 
     ((quadratic1(a;b;c) quadratic2(a;b;c)) and 
     (r0 ≤ ((b b) r(4) c)) and 
     a ≠ r0)

Lemma: quadratic-1-2-equal-implies2
[a,b,c:ℝ].
  ({((b b) (r(4) c)) ∧ (quadratic1(a;b;c) (-(b)/r(2) a))}) supposing 
     ((quadratic1(a;b;c) quadratic2(a;b;c)) and 
     (r0 ≤ ((b b) r(4) c)) and 
     a ≠ r0)

Lemma: quadratic-formula1
a,b,c:ℝ.
  (a ≠ r0
   (r0 ≤ ((b b) r(4) c))
   (∀x:ℝ(((x quadratic1(a;b;c)) ∨ (x quadratic2(a;b;c)))  (((a x^2) (b x) c) r0))))

Lemma: quadratic-formula-simple
a,b,c:ℝ.
  ((r0 < a)
   (c < r0)
   (∀x:ℝ(((x quadratic1(a;b;c)) ∨ (x quadratic2(a;b;c)))  (((a x^2) (b x) c) r0))))

Lemma: quadratic-formula2
a,b,c:ℝ.
  (a ≠ r0
   (r0 ≤ ((b b) r(4) c))
   (∀x:ℝ
        ((((a x^2) (b x) c) r0)
         ((¬¬((x quadratic1(a;b;c)) ∨ (x quadratic2(a;b;c))))
           ∧ ((((r(2) a) x) b ≠ r0 ∨ (r0 < ((b b) r(4) c)))
              ((x quadratic1(a;b;c)) ∨ (x quadratic2(a;b;c))))))))

Lemma: arith-geom-mean-inequality-simple
x:{t:ℝr0 ≤ t} . ∀y:ℝ.  ((r(2) y) ≤ (x^2 y^2))

Lemma: arith-geom-mean-inequality
x,y:{t:ℝr0 ≤ t} .  ((rsqrt(x y) ≤ (x y/r(2))) ∧ (x ≠  (rsqrt(x y) < (x y/r(2)))))

Lemma: Cauchy-Schwarz3
[n:ℕ]. ∀[x,y:ℕn ⟶ ℝ].
  (|Σ{x[i] y[i] 0≤i≤1}| ≤ (rsqrt(Σ{x[i] x[i] 0≤i≤1}) rsqrt(Σ{y[i] y[i] 0≤i≤1})))

Lemma: Cauchy-Schwarz3-strict
n:ℕ. ∀x,y:ℕn ⟶ ℝ.
  (∃i,j:ℕn. x[j] y[i] ≠ x[i] y[j]
  ⇐⇒ {x[i] y[i] 0≤i≤1}| < (rsqrt(Σ{x[i] x[i] 0≤i≤1}) rsqrt(Σ{y[i] y[i] 0≤i≤1})))

Definition: real-vec
^n ==  ℕn ⟶ ℝ

Lemma: real-vec_wf
[n:ℕ]. (ℝ^n ∈ Type)

Lemma: real-vec-subtype
[n,m:ℕ].  ℝ^m ⊆r ℝ^n supposing n ≤ m

Definition: req-vec
req-vec(n;x;y) ==  ∀i:ℕn. ((x i) (y i))

Lemma: req-vec_wf
[n:ℕ]. ∀[x,y:ℝ^n].  (req-vec(n;x;y) ∈ ℙ)

Lemma: stable__req-vec
n:ℕ. ∀x,y:ℝ^n.  Stable{req-vec(n;x;y)}

Lemma: req-vec_functionality
[n:ℕ]. ∀[x1,x2,y1,y2:ℝ^n].  (uiff(req-vec(n;x1;y1);req-vec(n;x2;y2))) supposing (req-vec(n;y1;y2) and req-vec(n;x1;x2))

Lemma: req-vec_weakening
[n:ℕ]. ∀[x,y:ℝ^n].  req-vec(n;x;y) supposing y ∈ ℝ^n

Lemma: req-vec_inversion
[n:ℕ]. ∀[x,y:ℝ^n].  req-vec(n;y;x) supposing req-vec(n;x;y)

Lemma: req-vec_transitivity
[n:ℕ]. ∀[x,y,z:ℝ^n].  (req-vec(n;x;z)) supposing (req-vec(n;y;z) and req-vec(n;x;y))

Definition: real-vec-add
==  λs.((X s) (Y s))

Lemma: real-vec-add_wf
[n:ℕ]. ∀[X,Y:ℝ^n].  (X Y ∈ ℝ^n)

Lemma: real-vec-add-com
[n:ℕ]. ∀[X,Y:ℝ^n].  req-vec(n;X Y;Y X)

Lemma: real-vec-add-assoc
[n:ℕ]. ∀[X,Y,Z:ℝ^n].  req-vec(n;X Z;X Z)

Lemma: real-vec-add_functionality
[n:ℕ]. ∀[X1,Y1,X2,Y2:ℝ^n].  (req-vec(n;X1 Y1;X2 Y2)) supposing (req-vec(n;X1;X2) and req-vec(n;Y1;Y2))

Definition: real-vec-sub
==  λs.((X s) s)

Lemma: real-vec-sub_wf
[n:ℕ]. ∀[X,Y:ℝ^n].  (X Y ∈ ℝ^n)

Lemma: real-vec-sub_functionality
[n:ℕ]. ∀[X1,Y1,X2,Y2:ℝ^n].  (req-vec(n;X1 Y1;X2 Y2)) supposing (req-vec(n;X1;X2) and req-vec(n;Y1;Y2))

Lemma: real-vec-add-cancel
[n:ℕ]. ∀[p,a,b:ℝ^n].  req-vec(n;a;b) supposing req-vec(n;p a;p b)

Definition: real-vec-mul
a*X ==  λi.(a (X i))

Lemma: real-vec-mul_wf
[n:ℕ]. ∀[X:ℝ^n]. ∀[a:ℝ].  (a*X ∈ ℝ^n)

Lemma: real-vec-mul-linear
[n:ℕ]. ∀[X,Y:ℝ^n]. ∀[a:ℝ].  req-vec(n;a*X Y;a*X a*Y)

Lemma: real-vec-mul-linear-sub
[n:ℕ]. ∀[X,Y:ℝ^n]. ∀[a:ℝ].  req-vec(n;a*X Y;a*X a*Y)

Lemma: real-vec-mul-mul
[n:ℕ]. ∀[X:ℝ^n]. ∀[a,b:ℝ].  req-vec(n;a*b*X;a b*X)

Lemma: real-vec-mul_functionality
[n:ℕ]. ∀[X,Y:ℝ^n]. ∀[a,b:ℝ].  (req-vec(n;a*X;b*Y)) supposing ((a b) and req-vec(n;X;Y))

Definition: real-vec-sum
Σ{x[k] n≤k≤m} ==  λi.Σ{x[k] n≤k≤m}

Lemma: real-vec-sum_wf
[n,m:ℤ]. ∀[k:ℕ]. ∀[x:{n..m 1-} ⟶ ℝ^k].  {x[k] n≤k≤m} ∈ ℝ^k)

Lemma: real-vec-sum_functionality
[n,m:ℤ]. ∀[k:ℕ]. ∀[x,y:{n..m 1-} ⟶ ℝ^k].
  req-vec(k;Σ{x[k] n≤k≤m};Σ{y[k] n≤k≤m}) supposing ∀i:ℤ((n ≤ i)  (i ≤ m)  req-vec(k;x[i];y[i]))

Lemma: real-vec-sum-single
[n,m:ℤ]. ∀[k:ℕ]. ∀[x:{n..m 1-} ⟶ ℝ^k].  req-vec(k;Σ{x[k] n≤k≤m};x[n]) supposing n ∈ ℤ

Lemma: real-vec-sum-empty
[n,m:ℤ]. ∀[x:Top].  Σ{x[k] n≤k≤m} ~ λi.r0 supposing m < n

Lemma: real-vec-sum-split
[j,n,m:ℤ]. ∀[k:ℕ]. ∀[x:{n..m 1-} ⟶ ℝ^k].
  (req-vec(k;Σ{x[i] n≤i≤m};Σ{x[i] n≤i≤j} + Σ{x[i] 1≤i≤m})) supposing ((j ≤ m) and (n ≤ j))

Lemma: real-vec-sum-shift
[k,n,m:ℤ]. ∀[x:Top].  {x[i] n≤i≤m} ~ Σ{x[i k] k≤i≤k})

Lemma: real-vec-sum-split-first
[n,m:ℤ]. ∀[k:ℕ]. ∀[x:{n..m 1-} ⟶ ℝ^k].  req-vec(k;Σ{x[i] n≤i≤m};x[n] + Σ{x[i 1] n≤i≤1}) supposing n ≤ m

Definition: real-vec-indep
real-vec-indep(k;L) ==  ∀a:ℕ||L|| ⟶ ℝ(req-vec(k;Σ{a i*L[i] 0≤i≤||L|| 1};λi.r0)  (∀i:ℕ||L||. ((a i) r0)))

Lemma: real-vec-indep_wf
[k:ℕ]. ∀[L:ℝ^k List].  (real-vec-indep(k;L) ∈ ℙ)

Lemma: stable__real-vec-indep
[k:ℕ]. ∀[L:ℝ^k List].  Stable{real-vec-indep(k;L)}

Lemma: sq_stable__real-vec-indep
[k:ℕ]. ∀[L:ℝ^k List].  SqStable(real-vec-indep(k;L))

Definition: real-vec-between
a-b-c ==  ∃t:ℝ((t ∈ (r0, r1)) ∧ req-vec(n;b;t*a r1 t*c))

Lemma: real-vec-between_wf
[n:ℕ]. ∀[a,b,c:ℝ^n].  (a-b-c ∈ ℙ)

Lemma: real-vec-between_functionality
n:ℕ. ∀a1,a2,b1,b2,c1,c2:ℝ^n.  (req-vec(n;a1;a2)  req-vec(n;b1;b2)  req-vec(n;c1;c2)  (a1-b1-c1 ⇐⇒ a2-b2-c2))

Lemma: real-vec-between-symmetry
n:ℕ. ∀a,b,c:ℝ^n.  (a-b-c  c-b-a)

Lemma: real-vec-between-inner-trans
n:ℕ. ∀a,b,c,d:ℝ^n.  (a-b-d  b-c-d  a-b-c)

Definition: real-vec-be
real-vec-be(n;a;b;c) ==  ∃t:ℝ((t ∈ [r0, r1]) ∧ req-vec(n;b;t*a r1 t*c))

Lemma: real-vec-be_wf
[n:ℕ]. ∀[a,b,c:ℝ^n].  (real-vec-be(n;a;b;c) ∈ ℙ)

Definition: dot-product
x⋅==  Σ{(x i) (y i) 0≤i≤1}

Lemma: dot-product_wf
[n:ℕ]. ∀[x,y:ℝ^n].  (x⋅y ∈ ℝ)

Lemma: dot-product_functionality
[n:ℕ]. ∀[x1,y1,x2,y2:ℝ^n].  (x1⋅y1 x2⋅y2) supposing (req-vec(n;x1;x2) and req-vec(n;y1;y2))

Lemma: dot-product-comm
[n:ℕ]. ∀[x,y:ℝ^n].  (x⋅y⋅x)

Lemma: dot-product-linearity1
[n:ℕ]. ∀[x,y,z:ℝ^n].  ((x y⋅(x⋅y⋅z)) ∧ (z⋅(z⋅z⋅y)))

Lemma: dot-product-linearity1-sub
[n:ℕ]. ∀[x,y,z:ℝ^n].  ((x y⋅(x⋅y⋅z)) ∧ (z⋅(z⋅z⋅y)))

Lemma: dot-product-linearity2
[n:ℕ]. ∀[x,y:ℝ^n]. ∀[a:ℝ].  ((x⋅a*y (a x⋅y)) ∧ (a*x⋅(a x⋅y)))

Lemma: dot-product-nonneg
[n:ℕ]. ∀[x:ℝ^n].  (r0 ≤ x⋅x)

Lemma: dot-product-zero
[n:ℕ]. ∀[x:ℝ^n].  (x⋅λi.r0 r0)

Lemma: r2-dot-product
[a,b:ℝ^2].  (a⋅(((a 0) (b 0)) ((a 1) (b 1))))

Lemma: r3-dot-product
[a,b:ℝ^3].  (a⋅(((a 0) (b 0)) ((a 1) (b 1)) ((a 2) (b 2))))

Lemma: dot-product-split
[n:ℕ]. ∀[k:ℕn]. ∀[x,y:ℝ^n].  (x⋅(x⋅+ λi.(x (k i))⋅λi.(y (k i))))

Lemma: dot-product-split-last
[n:ℕ+]. ∀[x,y:ℝ^n].  (x⋅(x⋅((x (n 1)) (y (n 1)))))

Lemma: dot-product-split-first
[n:ℕ+]. ∀[x,y:ℝ^n].  (x⋅(((x 0) (y 0)) + λi.(x (i 1))⋅λi.(y (i 1))))

Definition: real-vec-norm
||x|| ==  rsqrt(x⋅x)

Lemma: real-vec-norm_wf
[n:ℕ]. ∀[x:ℝ^n].  (||x|| ∈ ℝ)

Lemma: real-vec-norm_functionality
[n:ℕ]. ∀[x,y:ℝ^n].  ||x|| ||y|| supposing req-vec(n;x;y)

Lemma: real-vec-norm-nonneg
[n:ℕ]. ∀[x:ℝ^n].  (r0 ≤ ||x||)

Lemma: real-vec-norm-positive-iff
n:ℕ. ∀x:ℝ^n.  (r0 < ||x|| ⇐⇒ ∃i:ℕn. r0 ≠ i)

Lemma: real-vec-norm-positive-iff2
n:ℕ. ∀x:ℝ^n.  (∃i:ℕn. i ≠ r0 ⇐⇒ r0 < ||x||)

Lemma: real-vec-norm-squared
[n:ℕ]. ∀[x:ℝ^n].  (||x||^2 x⋅x)

Lemma: real-vec-norm-diff-squared
[n:ℕ]. ∀[x,y:ℝ^n].  (||x y||^2 ((||x||^2 ||y||^2) (r(-2) x⋅y)))

Lemma: real-vec-norm-eq-iff
[n:ℕ]. ∀[x:ℝ^n]. ∀[r:ℝ].  uiff(||x|| r;(x⋅r^2) ∧ (r0 ≤ r))

Lemma: real-vec-norm-equal-iff
[n:ℕ]. ∀[x,y:ℝ^n].  uiff(||x|| ||y||;x⋅y⋅y)

Lemma: real-vec-norm-mul
[n:ℕ]. ∀[x:ℝ^n]. ∀[a:ℝ].  (||a*x|| (|a| ||x||))

Lemma: component-rleq-real-vec-norm
n:ℕ. ∀i:ℕn. ∀x:ℝ^n.  (|x i| ≤ ||x||)

Lemma: real-vec-norm-is-0
[n:ℕ]. ∀[x:ℝ^n].  uiff(||x|| r0;req-vec(n;x;λi.r0))

Lemma: real-vec-norm-0
[n:ℕ]. (||λi.r0|| r0)

Lemma: real-vec-norm-1-exists
[n:ℕ+]. ∃b:ℝ^n. (||b|| r1)

Lemma: real-vec-norm-diff
[n:ℕ]. ∀[x,y:ℝ^n].  (||x y|| ||y x||)

Lemma: real-vec-norm-dim1
[x:ℝ^1]. (||x|| |x 0|)

Lemma: implies-real-vec-norm-rleq
[n:ℕ]. ∀[x:ℝ^n]. ∀[c:ℝ].  ((∀i:ℕn. (|x i| ≤ c))  (||x|| ≤ (rsqrt(r(n)) c)))

Comment: real_matrix_theory
We have some matrix theory developed for matrices over commutative ring.
The reals do not form that kind of commutative ring 
So we form ⌜real-ring()⌝ that uses the quotient type ⌜x,y:ℝ//(x y)⌝.
Now we can instantiate theorems in the matrix theory
and use them to prove things about the unquotiented reals,
by introducing the equivalence relation ⌜X ≡ Y⌝ on the type ⌜ℝ(n × m)⌝
of by real matrices.

For example, from the general Cramer's rule proved for commutative rings
we get version for the reals:
real-Cramers-rule

Definition: real-ring
real-ring() ==  <x,y:ℝ//(x y), λx,y. ff, λx,y. ff, λx,y. (x y), r0, λx.-(x), λx,y. (x y), r1, λx,y. (inr ⋅ )>

Lemma: real-ring_wf
real-ring() ∈ CRng

Lemma: real-rng_lsum-sq
[L,A:Top].  {real-ring()} x ∈ L. A[x] reduce(λx,y. (x y);r0;map(λx.A[x];L)))

Definition: rmatrix
(a × b) ==  ℕa ⟶ ℕb ⟶ ℝ

Lemma: rmatrix_wf
[a,b:ℕ].  (ℝ(a × b) ∈ Type)

Definition: rcolumn
col(b) ==  λi,j. (b i)

Lemma: rcolumn_wf
[n:ℕ]. ∀[b:ℝ^n].  (col(b) ∈ ℝ(n × 1))

Definition: reqmatrix
X ≡ ==  ∀i:ℕa. ∀j:ℕb.  ((X j) (Y j))

Lemma: reqmatrix_wf
[a,b:ℕ]. ∀[X,Y:ℝ(a × b)].  (X ≡ Y ∈ ℙ)

Lemma: stable__reqmatrix
[a,b:ℕ]. ∀[X,Y:ℝ(a × b)].  Stable{X ≡ Y}

Lemma: reqmatrix_weakening
[a,b:ℕ]. ∀[X,Y:ℝ(a × b)].  X ≡ supposing Y ∈ ℝ(a × b)

Lemma: reqmatrix_inversion
[a,b:ℕ]. ∀[X,Y:ℝ(a × b)].  Y ≡ supposing X ≡ Y

Lemma: reqmatrix_transitivity
[a,b:ℕ]. ∀[X,Y,Z:ℝ(a × b)].  (X ≡ Z) supposing (X ≡ and Y ≡ Z)

Lemma: reqmatrix_functionality
[a,b:ℕ]. ∀[X1,X2,Y1,Y2:ℝ(a × b)].  (uiff(X1 ≡ Y1;X2 ≡ Y2)) supposing (Y1 ≡ Y2 and X1 ≡ X2)

Definition: real-matrix-times
(A*B) ==  λx,y. Σ{(A i) (B y) 0≤i≤1}

Lemma: real-matrix-times_wf
[n,a,b:ℕ]. ∀[A:ℝ(a × n)]. ∀[B:ℝ(n × b)].  ((A*B) ∈ ℝ(a × b))

Lemma: real-matrix-times_functionality
[n,a,b:ℕ]. ∀[A1,A2:ℝ(a × n)]. ∀[B1,B2:ℝ(n × b)].  ((A1*B1) ≡ (A2*B2)) supposing (A1 ≡ A2 and B1 ≡ B2)

Lemma: matrix-times-req-real-matrix-times
[n,a,b:ℕ]. ∀[A:ℝ(a × n)]. ∀[B:ℝ(n × b)].  (A*B) ≡ (A*B)

Definition: real-matrix-add
==  λx,y. ((A y) (B y))

Lemma: real-matrix-add_wf
[a,b:ℕ]. ∀[A,B:ℝ(a × b)].  (A B ∈ ℝ(a × b))

Lemma: real-matrix-add_functionality
[a,b:ℕ]. ∀[A1,A2,B1,B2:ℝ(a × b)].  (A1 B1 ≡ A2 B2) supposing (A1 ≡ A2 and B1 ≡ B2)

Lemma: matrix-plus-sq-real-matrix-add
[A,B:Top].  (A B)

Definition: real-matrix-scalar-mul
c*A ==  λx,y. (c (A y))

Lemma: real-matrix-scalar-mul_wf
[a,b:ℕ]. ∀[c:ℝ]. ∀[A:ℝ(a × b)].  (c*A ∈ ℝ(a × b))

Lemma: real-matrix-scalar-mul_functionality
[a,b:ℕ]. ∀[c1,c2:ℝ]. ∀[A,B:ℝ(a × b)].  (c1*A ≡ c2*B) supposing ((c1 c2) and A ≡ B)

Lemma: matrix-scalar-mul-sq-real-matrix-scalar-mul
[c,A:Top].  (c*A c*A)

Definition: real-det
|M| ==
  r-list-sum(map(λf.let rprod(0;n 1;i.M (f i)) in
                        if permutation-sign(n;f)=1 then else -(k);permutations-list(n)))

Lemma: real-det_wf
[n:ℕ]. ∀[M:ℕn ⟶ ℕn ⟶ ℝ].  (|M| ∈ ℝ)

Lemma: real-det-sq-matrix-det
n:ℕ. ∀M:Top.  (|M| |M|)

Lemma: real-det-req-matrix-det
n:ℕ. ∀M:ℕn ⟶ ℕn ⟶ ℝ.  (|M| |M|)

Lemma: real-Cramers-rule1
[n:ℕ]. ∀[A:ℕn ⟶ ℕn ⟶ ℝ]. ∀[c:{c:ℝ(c |A|) r1} ]. ∀[b:ℝ^n].
  (A*c*col(λj.|λx,y. if y=j then else (A y)|)) ≡ col(b)

Lemma: real-Cramers-rule
[n:ℕ]. ∀[A:ℝ(n × n)].  ∀[b:ℝ^n]. (A*col(λj.(|λx,y. if y=j then else (A y)|/|A|))) ≡ col(b) supposing |A| ≠ r0

Definition: r2-unit-circle
r2-unit-circle(p) ==  (p 0^2 1^2) r1

Lemma: r2-unit-circle_wf
[p:ℝ^2]. (r2-unit-circle(p) ∈ ℙ)

Definition: circle-param
circle-param(t) ==  λi.if (i =z 0) then (r1 t/r1 (t t)) else (r(2) t/r1 (t t)) fi 

Lemma: circle-param_wf
[t:ℝ]. (circle-param(t) ∈ {p:ℝ^2| r2-unit-circle(p)} )

Lemma: circle-param-one-one
[t1,t2:ℝ].  t1 t2 supposing req-vec(2;circle-param(t1);circle-param(t2))

Lemma: circle-param-onto
p:{p:ℝ^2| r2-unit-circle(p)} ((r(-1) < (p 0))  (∃t:ℝreq-vec(2;circle-param(t);p)))

Definition: real-vec-extend
a++z ==  λi.if i <then else fi 

Lemma: real-vec-extend_wf
[k:ℕ+]. ∀[a:ℝ^k 1]. ∀[z:ℝ].  (a++z ∈ ℝ^k)

Lemma: req-vec-extend
[k:ℕ+]. ∀[a,b:ℝ^k 1]. ∀[z,u:ℝ].  uiff(req-vec(k;a++z;b++u);(z u) ∧ req-vec(k 1;a;b))

Definition: interval-vec
I^n ==  {v:ℝ^n| ∀i:ℕn. (v i ∈ I)} 

Lemma: interval-vec_wf
[I:Interval]. ∀[n:ℕ].  (I^n ∈ Type)

Lemma: interval-vec-subtype
[I:Interval]. ∀[n,m:ℕ].  I^m ⊆I^n supposing n ≤ m

Lemma: real-vec-extend_wf_interval
[I:Interval]. ∀[k:ℕ]. ∀[a:I^k]. ∀[z:{z:ℝz ∈ I} ].  (a++z ∈ I^k 1)

Definition: infn
infn(n;I) ==  primrec(n;λf.(f ⋅);λi,r,f. inf{r a.(f a++z)) z ∈ I})

Lemma: infn_wf
[I:{I:Interval| icompact(I)} ]
  ∀n:ℕ
    (infn(n;I) ∈ {F:{f:I^n ⟶ ℝ| ∀a,b:I^n.  (req-vec(n;a;b)  ((f a) (f b)))}  ⟶ ℝ
                  ∀f,g:{f:I^n ⟶ ℝ| ∀a,b:I^n.  (req-vec(n;a;b)  ((f a) (f b)))} .
                    ((∀x:I^n. ((f x) (g x)))  ((F f) (F g)))} )

Lemma: infn_functionality
n:ℕ. ∀I:{I:Interval| icompact(I)} . ∀f,g:I^n ⟶ ℝ.
  ((∀x,y:I^n.  (req-vec(n;x;y)  ((f x) (f y))))  (∀x:I^n. ((f x) (g x)))  ((infn(n;I) f) (infn(n;I) g)))

Lemma: rleq-infn
[I:{I:Interval| icompact(I)} ]
  ∀n:ℕ. ∀f:{f:I^n ⟶ ℝ| ∀a,b:I^n.  (req-vec(n;a;b)  ((f a) (f b)))} . ∀c:ℝ.
    c ≤ (infn(n;I) f) supposing ∀a:I^n. (c ≤ (f a))

Lemma: infn-rleq
[I:{I:Interval| icompact(I)} ]
  ∀n:ℕ. ∀f:{f:I^n ⟶ ℝ| ∀a,b:I^n.  (req-vec(n;a;b)  ((f a) (f b)))} . ∀x:I^n.  ((infn(n;I) f) ≤ (f x))

Lemma: infn-nonneg
[I:{I:Interval| icompact(I)} ]
  ∀n:ℕ. ∀f:{f:I^n ⟶ ℝ| ∀a,b:I^n.  (req-vec(n;a;b)  ((f a) (f b)))} .
    r0 ≤ (infn(n;I) f) supposing ∀a:I^n. (r0 ≤ (f a))

Lemma: infn-property
I:{I:Interval| icompact(I)} . ∀n:ℕ. ∀f:{f:I^n ⟶ ℝ| ∀a,b:I^n.  (req-vec(n;a;b)  ((f a) (f b)))} . ∀e:{e:ℝr0 < e} \000C.
  ∃x:I^n. ((f x) ≤ ((infn(n;I) f) e))

Lemma: approx-zero
I:{I:Interval| icompact(I)} . ∀n:ℕ. ∀f:{f:I^n ⟶ ℝ| ∀a,b:I^n.  (req-vec(n;a;b)  ((f a) (f b)))} .
  ((¬(∀x:I^n. x ≠ r0))  (∀e:{e:ℝr0 < e} . ∃x:I^n. (|f x| < e)))

Lemma: classical-exists-n-implies-approx
I:{I:Interval| icompact(I)} . ∀n:ℕ. ∀f:{f:I^n ⟶ ℝ| ∀a,b:I^n.  (req-vec(n;a;b)  ((f a) (f b)))} .
  ((¬¬(∃x:I^n. ((f x) r0)))  (∀e:{e:ℝr0 < e} . ∃x:I^n. (|f x| < e)))

Lemma: classical-exists-implies-approx
I:{I:Interval| icompact(I)} . ∀f:{f:I ⟶ℝifun(f;I)} .
  ((¬¬(∃x:{x:ℝx ∈ I} (f[x] r0)))  (∀e:{e:ℝr0 < e} . ∃x:{x:ℝx ∈ I} (|f[x]| < e)))

Lemma: Cauchy-Schwarz
[n:ℕ]. ∀[x,y:ℝ^n].  (|x⋅y| ≤ (||x|| ||y||))

Lemma: Cauchy-Schwarz-strict
n:ℕ. ∀x,y:ℝ^n.  (∃i,j:ℕn. (x j) (y i) ≠ (x i) (y j) ⇐⇒ |x⋅y| < (||x|| ||y||))

Lemma: Cauchy-Schwarz-not-strict
[n:ℕ]. ∀[x,y:ℝ^n].  (|x⋅y| < (||x|| ||y||)) ⇐⇒ ∀i,j:ℕn.  (((x j) (y i)) ((x i) (y j))))

Lemma: Cauchy-Schwarz-equality
[n:ℕ]. ∀[x,y:ℝ^n].  ((r0 < ||y||)  (|x⋅y| (||x|| ||y||))  req-vec(n;x;(x⋅y/||y||^2)*y))

Lemma: Cauchy-Schwarz-proof2
[n:ℕ]. ∀[x,y:ℝ^n].  (|x⋅y| ≤ (||x|| ||y||))

Lemma: Cauchy-Schwarz-equality2
n:ℕ. ∀x,y:ℝ^n.  (|x⋅y| < (||x|| ||y||)) ⇐⇒ (r0 < ||y||)  (∃t:ℝreq-vec(n;x;t*y)))

Definition: rv-pos-angle
rv-pos-angle(n;a;b;c) ==  |a b⋅b| < (||a b|| ||c b||)

Lemma: rv-pos-angle_wf
[n:ℕ]. ∀[a,b,c:ℝ^n].  (rv-pos-angle(n;a;b;c) ∈ ℙ)

Lemma: rv-pos-angle_functionality
n:ℕ. ∀a1,b1,c1,a2,b2,c2:ℝ^n.
  (req-vec(n;a1;a2)  req-vec(n;b1;b2)  req-vec(n;c1;c2)  {rv-pos-angle(n;a1;b1;c1) ⇐⇒ rv-pos-angle(n;a2;b2;c2)})

Lemma: Minkowski-inequality1
[n:ℕ]. ∀[x,y:ℝ^n].  (||x y|| ≤ (||x|| ||y||))

Lemma: Minkowski-inequality2
[n:ℕ]. ∀[x,y:ℝ^n].  (||x y|| ≤ (||x|| ||y||))

Lemma: Minkowski-equality
n:ℕ. ∀x,y:ℝ^n.
  ((r0 < ||y||)  (||x y|| (||x|| ||y||))  (∃t:ℝ((r0 ≤ t) ∧ req-vec(n;x;t*y) ∧ ((r0 < ||x||)  (r0 < t)))))

Definition: real-vec-dist
d(x;y) ==  ||x y||

Lemma: real-vec-dist_wf
[n:ℕ]. ∀[x,y:ℝ^n].  (d(x;y) ∈ {d:ℝr0 ≤ d} )

Lemma: real-vec-dist-dim0
[x,y:Top].  (d(x;y) r0)

Lemma: real-vec-dist-dim1
[x,y:ℝ^1].  (d(x;y) |(x 0) 0|)

Lemma: real-vec-dist-nonneg
[n:ℕ]. ∀[x,y:ℝ^n].  (r0 ≤ d(x;y))

Lemma: real-vec-dist_functionality
[n:ℕ]. ∀[x1,x2,y1,y2:ℝ^n].  (d(x1;y1) d(x2;y2)) supposing (req-vec(n;x1;x2) and req-vec(n;y1;y2))

Lemma: real-vec-dist-symmetry
[n:ℕ]. ∀[x,y:ℝ^n].  (d(x;y) d(y;x))

Lemma: real-vec-dist-identity
[n:ℕ]. ∀[x,y:ℝ^n].  uiff(d(x;y) r0;req-vec(n;x;y))

Lemma: real-vec-dist-same-zero
[n:ℕ]. ∀[x:ℝ^n].  (d(x;x) r0)

Lemma: real-vec-dist-equal-iff
[n:ℕ]. ∀[x,y,a,b:ℝ^n].  uiff(d(x;y) d(a;b);x y⋅b⋅b)

Lemma: real-vec-dist-translation
[n:ℕ]. ∀[x,y,z:ℝ^n].  (d(x z;y z) d(x;y))

Lemma: real-vec-dist-translation2
[n:ℕ]. ∀[x,y,z:ℝ^n].  (d(z x;z y) d(x;y))

Lemma: real-vec-dist-sub-zero
[n:ℕ]. ∀[p,q:ℝ^n].  (d(p q;λi.r0) d(p;q))

Lemma: real-vec-dist-from-zero
[n:ℕ]. ∀[p:ℝ^n].  (d(p;λi.r0) ||p||)

Lemma: real-vec-dist-from-zero2
[n:ℕ]. ∀[p:ℝ^n].  (d(λi.r0;p) ||p||)

Lemma: real-vec-dist-dilation
[n:ℕ]. ∀[x,y:ℝ^n]. ∀[a:ℝ].  (d(a*x;a*y) (|a| d(x;y)))

Lemma: real-vec-dist-minus
[n:ℕ]. ∀[x,y:ℝ^n].  (d(r(-1)*x;r(-1)*y) d(x;y))

Lemma: real-vec-dist-vec-mul
[n:ℕ]. ∀[x:ℝ^n]. ∀[a,b:ℝ].  (d(a*x;b*x) (|a b| ||x||))

Lemma: implies-real-vec-dist-rleq
[n:ℕ]. ∀[x,y:ℝ^n]. ∀[c:ℝ].  ((∀i:ℕn. (|(x i) i| ≤ c))  (d(x;y) ≤ (rsqrt(r(n)) c)))

Lemma: rleq-real-vec-dist
[n:ℕ]. ∀[x,y:ℝ^n]. ∀[i:ℕn].  (|(x i) i| ≤ d(x;y))

Lemma: real-vec-dist-monotone-in-dim
[m:ℕ+]. ∀[p,q:ℝ^m].  (d(p;q) ≤ d(p;q))

Lemma: real-vec-triangle-inequality
[n:ℕ]. ∀[x,y,z:ℝ^n].  (d(x;z) ≤ (d(x;y) d(y;z)))

Definition: rn-metric
rn-metric(n) ==  λx,y. d(x;y)

Lemma: rn-metric_wf
[n:ℕ]. (rn-metric(n) ∈ metric(ℝ^n))

Lemma: rn-metric-meq
[n:ℕ]. ∀[x,y:ℝ^n].  uiff(x ≡ y;req-vec(n;x;y))

Lemma: real-vec-dist-lower-bound
[n:ℕ]. ∀[x,y:ℝ^n].  (|||y|| ||x||| ≤ d(x;y))

Definition: rn-prod-metric
rn-prod-metric(n) ==  prod-metric(n;λi.rmetric())

Lemma: rn-prod-metric_wf
[n:ℕ]. (rn-prod-metric(n) ∈ metric(ℝ^n))

Lemma: mcomplete-rn-prod-metric
n:ℕmcomplete(ℝ^n with rn-prod-metric(n))

Lemma: mdist-rn-prod-metric
[k,x,y:Top].  (mdist(rn-prod-metric(k);x;y) ~ Σ{|(x i) i| 0≤i≤1})

Lemma: meq-rn-prod-metric
[k:ℕ]. ∀[x,y:ℝ^k].  uiff(x ≡ y;req-vec(k;x;y))

Definition: max-metric
max-metric(n) ==  λp,q. primrec(n;r0;λi,r. rmax(r;|(p i) i|))

Lemma: max-metric_wf
[n:ℕ]. (max-metric(n) ∈ metric(ℝ^n))

Lemma: mdist-max-metric-ub
[n:ℕ]. ∀[x,y:ℝ^n].  ∀i:ℕn. (|(x i) i| ≤ mdist(max-metric(n);x;y))

Lemma: mdist-max-metric-strict-lb
c:{c:ℝr0 < c} . ∀n:ℕ. ∀x,y:ℝ^n.  ((∀i:ℕn. (|(x i) i| < c))  (mdist(max-metric(n);x;y) < c))

Lemma: max-metric-mdist-from-zero-rless
c:{c:ℝr0 < c} . ∀n:ℕ. ∀x:ℝ^n.  ((∀i:ℕn. (|x i| < c))  (mdist(max-metric(n);x;λi.r0) < c))

Lemma: max-metric-mdist-from-zero
[c:{c:ℝr0 ≤ c} ]. ∀[n:ℕ]. ∀[x:ℝ^n].  uiff(mdist(max-metric(n);x;λi.r0) ≤ c;∀i:ℕn. (x i ∈ [-(c), c]))

Lemma: max-metric-mdist-from-zero-strict
c:{c:ℝr0 < c} . ∀n:ℕ. ∀x:ℝ^n.  (mdist(max-metric(n);x;λi.r0) < ⇐⇒ ∀i:ℕn. (x i ∈ (-(c), c)))

Lemma: max-metric-mdist-from-zero-2
[c:{c:ℝr0 ≤ c} ]. ∀[n:ℕ]. ∀[x:ℝ^n].  uiff(mdist(max-metric(n);λi.r0;x) ≤ c;∀i:ℕn. (x i ∈ [-(c), c]))

Lemma: max-metric-leq-rn-metric
[n:ℕ]. max-metric(n) ≤ rn-metric(n)

Lemma: rn-metric-leq-rn-prod-metric
[n:ℕ]. rn-metric(n) ≤ rn-prod-metric(n)

Lemma: rn-prod-metric-le-max-metric
[n:ℕ]. rn-prod-metric(n) ≤ r(n)*max-metric(n)

Lemma: rn-metric-leq-max-metric
[n:ℕ]. rn-metric(n) ≤ r(n)*max-metric(n)

Lemma: max-metric-complete
n:ℕmcomplete(ℝ^n with max-metric(n))

Lemma: rn-metric-complete
n:ℕmcomplete(ℝ^n with rn-metric(n))

Lemma: mdist-max-metric-mul2
[n:ℕ]. ∀[p,q:ℝ^n]. ∀[c:ℝ].  (mdist(max-metric(n);c*p;c*q) (|c| mdist(max-metric(n);p;q)))

Lemma: mdist-max-metric-mul-rleq
[n:ℕ]. ∀[x:ℝ^n]. ∀[c,c':ℝ].  (mdist(max-metric(n);c*x;c'*x) ≤ (|c c'| mdist(max-metric(n);x;λi.r0)))

Lemma: mdist-rn-metric-mul
[n:ℕ]. ∀[p:ℝ^n]. ∀[c:ℝ].  (mdist(rn-metric(n);c*p;λi.r0) (|c| mdist(rn-metric(n);p;λi.r0)))

Lemma: meq-max-metric-iff-meq-rn-metric
[n:ℕ]. ∀[x,y:ℝ^n].  uiff(x ≡ y;x ≡ y)

Lemma: meq-max-metric
[n:ℕ]. ∀[x,y:ℝ^n].  uiff(x ≡ y;req-vec(n;x;y))

Lemma: mdist-max-metric-mul
[n:ℕ]. ∀[p:ℝ^n]. ∀[c:ℝ].  (mdist(max-metric(n);c*p;λi.r0) (|c| mdist(max-metric(n);p;λi.r0)))

Lemma: req-vec-meq-max-metric
[n:ℕ]. ∀[v,w:ℝ^n].  v ≡ supposing req-vec(n;v;w)

Lemma: real-vec-norm-diff-bound
[n:ℕ]. ∀[x,y:ℝ^n].  (|||x|| ||y||| ≤ d(x;y))

Lemma: real-vec-angle-lemma
n:ℕ. ∀x,z:ℝ^n.  (d(z;x) < d(z;r(-1)*x) ⇐⇒ r0 < x⋅z)

Lemma: real-vec-angle-lemma2
n:ℕ. ∀x,z:ℝ^n.  (d(z;r(-1)*x) < d(z;x) ⇐⇒ x⋅z < r0)

Lemma: real-vec-right-angle-lemma
[n:ℕ]. ∀[x,z:ℝ^n].  uiff(x⋅r0;d(z;x) d(z;r(-1)*x))

Lemma: rv-pos-angle-symmetry
n:ℕ. ∀a,b,c:ℝ^n.  (rv-pos-angle(n;a;b;c)  rv-pos-angle(n;c;b;a))

Lemma: rv-pos-angle-permute-lemma
n:ℕ. ∀x,y:ℝ^n.  ((|x⋅y| < (||x|| ||y||))  (|x⋅x| < (||x|| ||y x||)))

Lemma: rv-pos-angle-permute
n:ℕ. ∀a,b,c:ℝ^n.  (rv-pos-angle(n;a;b;c)  rv-pos-angle(n;c;a;b))

Lemma: rv-pos-angle-linearity
n:ℕ. ∀a,b,c:ℝ^n.  (rv-pos-angle(n;a;b;c)  (∀t:ℝ((r0 < |t|)  rv-pos-angle(n;b t*a b;b;c))))

Lemma: not-rv-pos-angle
n:ℕ. ∀a,b,c:ℝ^n.
  ((r0 < d(a;b))  (r0 < d(c;b))  rv-pos-angle(n;a;b;c))  (∃t:ℝ((r0 < |t|) ∧ req-vec(n;c;b t*a b))))

Definition: rv-congruent
ab=cd ==  d(a;b) d(c;d)

Lemma: rv-congruent_wf
[n:ℕ]. ∀[a,b,c,d:ℝ^n].  (ab=cd ∈ ℙ)

Lemma: rv-congruent-sym
[n:ℕ]. ∀[a,b:ℝ^n].  ab=ba

Lemma: rv-congruent-trans
[n:ℕ]. ∀[a,b,p,q,r,s:ℝ^n].  (pq=rs) supposing (ab=rs and ab=pq)

Lemma: stable_rv-congruent
[n:ℕ]. ∀[a,b,c,d:ℝ^n].  Stable{ab=cd}

Lemma: rv-congruent_functionality
n:ℕ. ∀a1,a2,b1,b2,c1,c2,d1,d2:ℝ^n.
  (req-vec(n;a1;a2)  req-vec(n;b1;b2)  req-vec(n;c1;c2)  req-vec(n;d1;d2)  (a1b1=c1d1 ⇐⇒ a2b2=c2d2))

Definition: real-vec-sep
a ≠ ==  r0 < d(a;b)

Lemma: real-vec-sep_wf
[n:ℕ]. ∀[a,b:ℝ^n].  (a ≠ b ∈ ℙ)

Lemma: sq_stable__real-vec-sep
n:ℕ. ∀a,b:ℝ^n.  SqStable(a ≠ b)

Lemma: sq_stable-dist-rless
n:ℕ. ∀a,b,c,d:ℝ^n.  SqStable(d(c;d) < d(a;b))

Lemma: real-vec-sep_functionality
n:ℕ. ∀a1,a2,b1,b2:ℝ^n.  (req-vec(n;a1;a2)  req-vec(n;b1;b2)  (a1 ≠ b1 ⇐⇒ a2 ≠ b2))

Lemma: real-vec-sep-symmetry
n:ℕ. ∀a,b:ℝ^n.  (a ≠ ⇐⇒ b ≠ a)

Lemma: real-vec-sep_inversion
n:ℕ. ∀x,y:ℝ^n.  (x ≠  y ≠ x)

Lemma: real-vec-sep-cases
n:ℕ. ∀a,b:ℝ^n.  (a ≠  (∀c:ℝ^n. (a ≠ c ∨ b ≠ c)))

Lemma: real-vec-sep-cases-alt
n:ℕ. ∀x,y,z:ℝ^n.  (x ≠  (x ≠ z ∨ y ≠ z))

Lemma: r2-sep-or
a:ℝ^2. ∀b:{b:ℝ^2| d(a;a) < d(a;b)} . ∀c:ℝ^2.  ((d(a;a) < d(a;c)) ∨ (d(b;b) < d(b;c)))

Lemma: not-real-vec-sep-iff-eq
[n:ℕ]. ∀[a,b:ℝ^n].  uiff(¬a ≠ b;req-vec(n;a;b))

Lemma: not-real-vec-sep-refl
[n:ℕ]. ∀[a:ℝ^n].  a ≠ a)

Lemma: real-vec-sep-implies
n:ℕ. ∀a,c:ℝ^n.  (a ≠  (∃i:ℕn. (r0 < |(a i) i|)))

Lemma: real-vec-sep-iff
n:ℕ. ∀a,c:ℝ^n.  (a ≠ ⇐⇒ ∃i:ℕn. (r0 < |(a i) i|))

Lemma: real-vec-sep-msep-prod-metric
n:ℕ. ∀a,c:ℝ^n.  (a ≠ ⇐⇒ c)

Lemma: real-vec-sep-iff-rneq
n:ℕ. ∀a,c:ℝ^n.  (a ≠ ⇐⇒ ∃i:ℕn. i ≠ i)

Lemma: rn-metric-sep
n:ℕ. ∀x,y:ℝ^n.  (r0 < mdist(rn-metric(n);x;y) ⇐⇒ ∃i:ℕn. i ≠ i)

Lemma: rn-metric-msep
n:ℕ. ∀x,y:ℝ^n.  (x ⇐⇒ x ≠ y)

Lemma: max-metric-sep
n:ℕ. ∀x,y:ℝ^n.  (r0 < mdist(max-metric(n);x;y) ⇐⇒ ∃i:ℕn. i ≠ i)

Lemma: real-vec-sep-add
n:ℕ. ∀x,x',y,y':ℝ^n.  (x y ≠ x' y'  (x ≠ x' ∨ y ≠ y'))

Lemma: real-vec-sep-mul
n:ℕ. ∀a,b:ℝ. ∀y,y':ℝ^n.  (a*y ≠ b*y'  (a ≠ b ∨ y ≠ y'))

Lemma: real-vec-sep-0-iff
n:ℕ. ∀x:ℝ^n.  (x ≠ λi.r0 ⇐⇒ r0 < x⋅x)

Lemma: real-vec-sep-iff-dot-product
n:ℕ. ∀x,y:ℝ^n.  (x ≠ ⇐⇒ r0 < x⋅x)

Definition: real-vec-free
real-vec-free(k;L) ==  ∀a:ℕ||L|| ⟶ ℝ(a ≠ λi.r0  Σ{a i*L[i] 0≤i≤||L|| 1} ≠ λi.r0)

Lemma: real-vec-free_wf
[k:ℕ]. ∀[L:ℝ^k List].  (real-vec-free(k;L) ∈ ℙ)

Lemma: sq_stable__real-vec-free
k:ℕ. ∀L:ℝ^k List.  SqStable(real-vec-free(k;L))

Definition: is-simplex
is-simplex(k;L) ==  real-vec-free(k;map(λv.v hd(L);tl(L)))

Lemma: is-simplex_wf
[k:ℕ]. ∀[L:ℝ^k List].  is-simplex(k;L) ∈ ℙ supposing 0 < ||L||

Lemma: sq_stable__is-simplex
k:ℕ. ∀L:ℝ^k List.  SqStable(is-simplex(k;L)) supposing 0 < ||L||

Definition: geometric-simplex
geometric-simplex(k;n) ==  {L:ℝ^k List| (||L|| (n 1) ∈ ℤ) ∧ is-simplex(k;L)} 

Lemma: geometric-simplex_wf
[k,n:ℕ].  (geometric-simplex(k;n) ∈ Type)

Lemma: real-vec-mul-cancel
[n:ℕ]. ∀[v:ℝ^n]. ∀[x,y:ℝ].  (v ≠ λi.r0  supposing req-vec(n;x*v;y*v))

Lemma: length-rneq-real-vec-sep
n:ℕ. ∀x,v:ℝ^n.  (||v|| ≠ ||x||  x ≠ v)

Definition: real-vec-infinity-norm
||v||∞ ==  mdist(max-metric(n);v;λi.r0)

Lemma: real-vec-infinity-norm_wf
[n:ℕ+]. ∀[v:ℝ^n].  (||v||∞ ∈ ℝ)

Lemma: real-vec-infinity-norm-req
[n:ℕ+]. ∀[v:ℝ^n].  (||v||∞ if (n =z 1) then |v 0| else rmax(||v||∞;|v (n 1)|) fi )

Lemma: infinity-norm-rneq-real-vec-sep
n:ℕ+. ∀x,v:ℝ^n.  (||v||∞ ≠ ||x||∞  x ≠ v)

Lemma: Cauchy-Schwarz-non-equality1
n:ℕ. ∀x,y:ℝ^n.  ((r0 < ||y||)  (∀a:ℝx ≠ a*y)  (|x⋅y| < (||x|| ||y||)))

Lemma: Cauchy-Schwarz-non-equality
n:ℕ. ∀x,y:ℝ^n.  ((r0 < ||y||)  x ≠ (||x||/||y||)*y  x ≠ (-(||x||)/||y||)*y  (|x⋅y| < (||x|| ||y||)))

Lemma: real-vec-perp-exists
n:{2...}. ∀x:ℝ^n.  (x ≠ λi.r0  (∃y:ℝ^n. (y ≠ λi.r0 ∧ (x⋅r0))))

Lemma: rv-congruent_functionality2
n:ℕ. ∀a1,a2,b1,b2,c1,c2,d1,d2:ℝ^n.  ((¬a1 ≠ a2)  b1 ≠ b2)  c1 ≠ c2)  d1 ≠ d2)  (a1b1=c1d1 ⇐⇒ a2b2=c2d2))

Lemma: rv-congruent-preserves-sep
n:ℕ. ∀a,b,c,d:ℝ^n.  (ab=cd  a ≠  c ≠ d)

Lemma: rv-congruent-implies-eq
n:ℕ. ∀a,b,c:ℝ^n.  (aa=bc  req-vec(n;b;c))

Definition: rv-between
a-b-c ==  a-b-c ∧ a ≠ c

Lemma: rv-between_wf
[n:ℕ]. ∀[a,b,c:ℝ^n].  (a-b-c ∈ ℙ)

Lemma: rv-pos-angle-implies-separated
n:ℕ. ∀a,b,c:ℝ^n.  (rv-pos-angle(n;a;b;c)  a ≠ c)

Lemma: rv-pos-angle-implies-separated2
n:ℕ. ∀a,b,c:ℝ^n.  (rv-pos-angle(n;a;b;c)  a ≠ b)

Lemma: rv-pos-angle-shift
n:ℕ. ∀a,b,c:ℝ^n.  (rv-pos-angle(n;a;b;c)  (∀z:ℝ^n. (z ≠  rv-pos-angle(n;a;b;z))  rv-pos-angle(n;z;b;c))))

Lemma: rv-pos-angle-not-between
[n:ℕ]. ∀[a,b,c:ℝ^n].  ¬a-b-c supposing rv-pos-angle(n;a;b;c)

Lemma: rv-pos-angle-lemma
n:ℕ. ∀x,y:ℝ^n.  ((||x|| ||y||)  (r0 < d(x;y))  (r0 < d(r(-1)*x;y))  (|x⋅y| < (||x|| ||y||)))

Lemma: implies-rv-pos-angle
n:ℕ. ∀a,b,c,a':ℝ^n.  (a-b-a'  ab=a'b  ab=cb  c ≠  c ≠ a'  rv-pos-angle(n;a;b;c))

Lemma: rv-pos-angle-implies
n:ℕ. ∀a,b,c:ℝ^n.
  (rv-pos-angle(n;a;b;c)
   (rv-pos-angle(n;c;b;a)
     ∧ rv-pos-angle(n;c;a;b)
     ∧ a ≠ c
     ∧ a-b-c)
     ∧ (∀z:ℝ^n. (z ≠  rv-pos-angle(n;a;b;z))  rv-pos-angle(n;z;b;c)))))

Lemma: not-rv-pos-angle-implies
[n:ℕ]. ∀a,b,c:ℝ^n.  ((¬rv-pos-angle(n;a;b;c))  (a ≠ b ∧ b ≠ c ∧ c ≠ a ∧ a-b-c) ∧ b-c-a) ∧ c-a-b))))

Lemma: not-rv-pos-angle-iff
[n:ℕ]. ∀[a,b,c:ℝ^n].  uiff(¬rv-pos-angle(n;a;b;c);¬(a ≠ b ∧ b ≠ c ∧ c ≠ a ∧ a-b-c) ∧ b-c-a) ∧ c-a-b)))

Lemma: rv-between_functionality
n:ℕ. ∀a1,a2,b1,b2,c1,c2:ℝ^n.  (req-vec(n;a1;a2)  req-vec(n;b1;b2)  req-vec(n;c1;c2)  (a1-b1-c1 ⇐⇒ a2-b2-c2))

Lemma: rv-between_functionality2
n:ℕ. ∀a1,a2,b1,b2,c1,c2:ℝ^n.  ((¬a1 ≠ a2)  b1 ≠ b2)  c1 ≠ c2)  (a1-b1-c1 ⇐⇒ a2-b2-c2))

Lemma: rv-between_functionality3
n:ℕ. ∀a1,a2,b1,b2,c1,c2:ℝ^n.  ((¬a1 ≠ a2)  b1 ≠ b2)  c1 ≠ c2)  a1-b1-c1  a2-b2-c2)

Lemma: rv-between-symmetry
n:ℕ. ∀a,b,c:ℝ^n.  (a-b-c  c-b-a)

Lemma: rv-between-translation
n:ℕ. ∀a,b,c,z:ℝ^n.  (z a-z b-z ⇐⇒ a-b-c)

Lemma: rv-between-vec-mul
n:ℕ. ∀a,b,c:ℝ. ∀z:ℝ^n.  ((r0 < ||z||)  (a*z-b*z-c*z ⇐⇒ ((a < b) ∧ (b < c)) ∨ ((c < b) ∧ (b < a))))

Lemma: real-vec-dist-between-1
n:ℕ. ∀a,c:ℝ^n. ∀t:ℝ.  (d(a;t*a r1 t*c) (|r1 t| d(a;c)))

Lemma: real-vec-dist-between-2
n:ℕ. ∀a,c:ℝ^n. ∀t:ℝ.  (d(t*a r1 t*c;c) (|t| d(a;c)))

Lemma: real-vec-dist-between
n:ℕ. ∀a,b,c:ℝ^n.  (a-b-c  (d(a;c) (d(a;b) d(b;c))))

Lemma: real-vec-dist-be
n:ℕ. ∀a,b,c:ℝ^n.  (real-vec-be(n;a;b;c)  (d(a;c) (d(a;b) d(b;c))))

Lemma: real-vec-triangle-equality
n:ℕ. ∀x,y,z:ℝ^n.  ((r0 < d(y;z))  (d(x;z) (d(x;y) d(y;z)))  (real-vec-be(n;x;y;z) ∧ ((r0 < d(x;y))  x-y-z)))

Lemma: real-vec-right-angle
[n:ℕ]. ∀[x,y,z:ℝ^n].
  uiff(x y⋅r0;∀x':ℝ^n. ((d(x';y) d(x;y))  real-vec-be(n;x;y;x')  (d(z;x) d(z;x'))))

Lemma: real-vec-acute-angle
n:ℕ. ∀x,y,z:ℝ^n.  (r0 < y⋅⇐⇒ ∀x':ℝ^n. ((d(x';y) d(x;y))  real-vec-be(n;x;y;x')  (d(z;x) < d(z;x'))))

Lemma: real-vec-obtuse-angle
n:ℕ. ∀x,y,z:ℝ^n.  (x y⋅y < r0 ⇐⇒ ∀x':ℝ^n. ((d(x';y) d(x;y))  real-vec-be(n;x;y;x')  (d(z;x') < d(z;x))))

Lemma: rv-between-simple
n:ℕ. ∀c,d:ℝ^n.  ((r0 < ||d||)  d-c-c d)

Lemma: rv-between-small-expand
n:ℕ+. ∀a,c:ℝ^n. ∀k:ℕ+.  (a ≠  (r(k 1)/r(k))*a (r(-1)/r(k))*c-a-(r(k 1)/r(k))*c (r(-1)/r(k))*a)

Lemma: rv-between-sep
n:ℕ. ∀a,b,c:ℝ^n.  (a-b-c  a ≠ b)

Lemma: rv-sep-between
n:ℕ. ∀a,b:ℝ^n.  (a ≠  (∃m:ℝ^n. a-m-b))

Lemma: rv-Tsep
n:ℕ. ∀a,b:ℝ^n. ∀c:{c:ℝ^n| ¬(b ≠ c ∧ a-b-c))} .  (a ≠  a ≠ c)

Lemma: rv-between-inner-trans
n:ℕ. ∀a,b,c,d:ℝ^n.  (a-b-d  b-c-d  a-b-c)

Lemma: rv-nontrivial
n:{2...}. ∃a,b,c:ℝ^n. (a ≠ b ∧ b ≠ c ∧ c ≠ a ∧ a-b-c) ∧ b-c-a) ∧ c-a-b))

Lemma: rv-sep-exists
n:{1...}. ∃a,b:ℝ^n. a ≠ b

Lemma: r2-nontrivial
a:ℝ^2. (∃b:ℝ^2 [(d(a;a) < d(a;b))])

Lemma: rv-extend-1
n:ℕ. ∀a,b:ℝ^n.  (a ≠  (∀s:{s:ℝr0 ≤ s} . ∃x:ℝ^n. ((d(b;x) s) ∧ ((r0 < s)  a-b-x))))

Lemma: rv-extend
n:ℕ. ∀a,b,c,d:ℝ^n.  (a ≠  (∃x:{x:ℝ^n| bx=cd} (c ≠  a-b-x)))

Lemma: rv-extend-2
n:ℕ. ∀a,b,c,d:ℝ^n.  (a ≠  (∃x:{x:ℝ^n| bx=cd} (c ≠  a-b-x)))

Lemma: rv-inner-Pasch
n:ℕ. ∀a,b,c,p,q:ℝ^n.  (a-p-c  b-q-c  (∃x:ℝ^n. ((a ≠  a-x-q) ∧ (b ≠  b-x-p))))

Lemma: rv-tarski-parallel
n:ℕ. ∀a,b,c:ℝ^n.  (a ≠  a ≠  (∀d,p:ℝ^n.  (b-d-c  a-d-p  (∃x,y:ℝ^n. (a-b-x ∧ x-p-y ∧ a-c-y)))))

Lemma: rv-five-segment-lemma
n:ℕ. ∀A,B,C,D:ℝ^n. ∀t:ℝ.
  ((t ∈ (r0, r1))
   req-vec(n;B;t*A r1 t*C)
   let d(D;B) in
      let d(A;D) in
      let d(B;C) in
      let d(A;B) in
      let d(D;C) in
      let (t/t r1) in
      x^2 (c^2 a^2 (s (b^2 d^2 a^2))))

Lemma: rv-five-segment
[n:ℕ]. ∀[a,b,c,d,A,B,C,D:ℝ^n].  (cd=CD) supposing (bd=BD and ad=AD and bc=BC and ab=AB and A-B-C and a-b-c)

Lemma: stable_real-vec-be
n:ℕ. ∀a,c:ℝ^n.  (a ≠  (∀b:ℝ^n. Stable{real-vec-be(n;a;b;c)}))

Lemma: rv-non-strict-between-iff
n:ℕ. ∀a,b,c:ℝ^n.  (a ≠  (a ≠ b ∧ b ≠ c ∧ a-b-c)) ⇐⇒ real-vec-be(n;a;b;c)))

Definition: rv-T
rv-T(n;a;b;c) ==  (a ≠  real-vec-be(n;a;b;c)) ∧ ((¬a ≠ c)  b ≠ c))

Lemma: rv-T_wf
[n:ℕ]. ∀[a,b,c:ℝ^n].  (rv-T(n;a;b;c) ∈ ℙ)

Lemma: rv-T-iff
n:ℕ. ∀a,b,c:ℝ^n.  (rv-T(n;a;b;c) ⇐⇒ ¬(a ≠ b ∧ b ≠ c ∧ a-b-c)))

Lemma: rv-T-dist
n:ℕ. ∀a,b,c:ℝ^n.  (rv-T(n;a;b;c)  (d(a;c) (d(a;b) d(b;c))))

Lemma: rv-between-iff
n:ℕ. ∀a,b,c:ℝ^n.  (a-b-c ⇐⇒ a ≠ b ∧ b ≠ c ∧ a ≠ c ∧ rv-T(n;a;b;c))

Lemma: sq_stable__rv-between
n:ℕ. ∀a,b,c:ℝ^n.  SqStable(a-b-c)

Lemma: rv-Tsep-alt
n:ℕ. ∀a,b,c,d:ℝ^n.  ((¬¬(∃u,v,w:ℝ^n. (rv-T(n;u;v;w) ∧ ab=uv ∧ cd=uw)))  a ≠  c ≠ d)

Lemma: rv-Tsep-alt'
n:ℕ. ∀a,b,c,d:ℝ^n.  ((¬¬(∃u,v,w:ℝ^n. ((¬(u ≠ v ∧ v ≠ w ∧ u-v-w))) ∧ ab=uv ∧ cd=uw)))  a ≠  c ≠ d)

Lemma: not-rv-pos-angle-implies2
n:ℕ. ∀a,b,c:ℝ^n.  ((¬rv-pos-angle(n;a;b;c))  ((¬rv-T(n;a;b;c)) ∧ rv-T(n;b;c;a)) ∧ rv-T(n;c;a;b)))))

Lemma: rv-Gsep
n:ℕ. ∀a,b,c:ℝ^n. ∀d:{d:ℝ^n| ¬¬(∃u,v,w:ℝ^n. ((¬(u ≠ v ∧ v ≠ w ∧ u-v-w))) ∧ ab=uv ∧ cd=uw))} .  (a ≠  c ≠ d)

Lemma: rv-inner-Pasch3
n:ℕ. ∀a,b,c,p,q:ℝ^n.
  (a ≠ p
   b ≠ c
   rv-T(n;a;p;c)
   rv-T(n;b;q;c)
   (∃x:ℝ^n
       (rv-T(n;a;x;q)
       ∧ rv-T(n;b;x;p)
       ∧ (a ≠  x ≠ a)
       ∧ ((a ≠ q ∧ p ≠ c ∧ b ≠ q)  x ≠ q)
       ∧ ((b ≠ p ∧ b ≠ q)  x ≠ b)
       ∧ ((b ≠ p ∧ q ≠ c)  x ≠ p))))

Lemma: rv-inner-Pasch2
n:ℕ. ∀a,b,c,p,q:ℝ^n.
  (a ≠ p
   b ≠ c
   rv-T(n;a;p;c)
   rv-T(n;b;q;c)
   (∃x:ℝ^n
       ((((a ≠ q ∧ p ≠ c ∧ b ≠ q)  a-x-q) ∧ ((b ≠ p ∧ q ≠ c ∧ b ≠ q)  b-x-p)) ∧ rv-T(n;a;x;q) ∧ rv-T(n;b;x;p))))

Lemma: rv-inner-Pasch'
n:ℕ. ∀a,b,c,p,q:ℝ^n.
  (a-p-c  b-q-c  (∃x:ℝ^n. (((a ≠  a-x-q) ∧ (b ≠  b-x-p)) ∧ rv-T(n;a;x;q) ∧ rv-T(n;b;x;p))))

Lemma: rv-inner-Pasch''
n:ℕ. ∀a,b,c,p,q:ℝ^n.
  (a-p-c
   b-q-c
   (∃x:ℝ^n. ((¬(a ≠ x ∧ x ≠ q ∧ a-x-q))) ∧ (b ≠ x ∧ x ≠ p ∧ b-x-p))) ∧ (a ≠  a-x-q) ∧ (b ≠  b-x-p))))

Definition: rv-T'
rv-T'(n;a;b;c) ==  ∀x,y:ℝ^n.  (x-a-y  x-c-y  x-b-y)

Lemma: rv-T'_wf
[n:ℕ]. ∀[a,b,c:ℝ^n].  (rv-T'(n;a;b;c) ∈ ℙ)

Lemma: rv-T'-implies-rv-T
n:ℕ+. ∀a,b,c:ℝ^n.  (rv-T'(n;a;b;c)  rv-T(n;a;b;c))

Lemma: rv-T-partially-implies-rv-T'
n:ℕ+. ∀a,b,c:ℝ^n.  ((a ≠ c ∨ a ≠ c))  rv-T(n;a;b;c)  rv-T'(n;a;b;c))

Lemma: rv-line-circle-lemma0
n:ℕ. ∀r:ℝ. ∀p,q:ℝ^n.  ((||p|| ≤ r)  (r0 ≤ (p⋅p^2 ||q p||^2 (||p||^2 r^2))))

Lemma: rv-line-circle-lemma
n:ℕ. ∀r:ℝ. ∀p,q:ℝ^n.
  (p ≠ q
   (||p|| ≤ r)
   let in
         (r0 ≤ (((r(2) p⋅v) r(2) p⋅v) r(4) ||v||^2 (||p||^2 r^2)))
         ∧ (||p quadratic1(||v||^2;r(2) p⋅v;||p||^2 r^2)*v|| r)
         ∧ (||p quadratic2(||v||^2;r(2) p⋅v;||p||^2 r^2)*v|| r))

Lemma: punctured-ball-boundary-retraction
n:ℕ. ∀p:{p:ℝ^n| ||p|| < r1} .  Retract({x:ℝ^n| x ≠ p}  ⟶ {x:ℝ^n| ||x|| r1} )

Lemma: rv-line-circle-0
n:ℕ. ∀a,b,p,q:ℝ^n.
  (p ≠ q
   (d(a;p) ≤ d(a;b))
   (d(a;b) ≤ d(a;q))
   (∃u:{u:ℝ^n| ab=au ∧ (q ≠ u ∧ u ≠ p ∧ q-u-p)))} 
       (∃v:ℝ^n [(ab=av
               ∧ (q ≠ p ∧ p ≠ v ∧ q-p-v)))
               ∧ ((d(a;p) < d(a;b))  (q-p-v ∧ ((d(a;b) < d(a;q))  q-u-p)))
               ∧ ((d(a;p) d(a;b))
                  ((u ≠  ((req-vec(n;u;p) ∧ (r0 < a⋅p)) ∨ (req-vec(n;v;p) ∧ (p a⋅p < r0))))
                    ∧ (req-vec(n;u;v)  ((p a⋅r0) ∧ req-vec(n;u;p))))))])))

Definition: rvlinecircle0
rvlinecircle0(n;a;b;p;q) ==
  <quadratic1(d(q a;p a)^2;r(2) a⋅a;||p a||^2 d(a;b)^2)*q a
  quadratic2(d(q a;p a)^2;r(2) a⋅a;||p a||^2 d(a;b)^2)*q a
  >

Lemma: rvlinecircle0_wf
[n:ℕ]. ∀[a,b,p,q:ℝ^n].
  (rvlinecircle0(n;a;b;p;q) ∈ u:{u:ℝ^n| ab=au ∧ (q ≠ u ∧ u ≠ p ∧ q-u-p)))}  × {v:ℝ^n| 
                                                     ab=av
                                                     ∧ (q ≠ p ∧ p ≠ v ∧ q-p-v)))
                                                     ∧ ((d(a;p) < d(a;b))  (q-p-v ∧ ((d(a;b) < d(a;q))  q-u-p)))
                                                     ∧ ((d(a;p) d(a;b))
                                                        ((u ≠ v
                                                           ((req-vec(n;u;p) ∧ (r0 < a⋅p))
                                                             ∨ (req-vec(n;v;p) ∧ (p a⋅p < r0))))
                                                          ∧ (req-vec(n;u;v)
                                                             ((p a⋅r0) ∧ req-vec(n;u;p)))))} supposing 
     ((d(a;b) ≤ d(a;q)) and 
     (d(a;p) ≤ d(a;b)) and 
     p ≠ q)

Lemma: rv-line-circle-0-ext
n:ℕ. ∀a,b,p,q:ℝ^n.
  (p ≠ q
   (d(a;p) ≤ d(a;b))
   (d(a;b) ≤ d(a;q))
   (∃u:∃u:ℝ^n [(ab=au ∧ (q ≠ u ∧ u ≠ p ∧ q-u-p))))]
       (∃v:ℝ^n [(ab=av
               ∧ (q ≠ p ∧ p ≠ v ∧ q-p-v)))
               ∧ ((d(a;p) < d(a;b))  (q-p-v ∧ ((d(a;b) < d(a;q))  q-u-p)))
               ∧ ((d(a;p) d(a;b))
                  ((u ≠  ((req-vec(n;u;p) ∧ (r0 < a⋅p)) ∨ (req-vec(n;v;p) ∧ (p a⋅p < r0))))
                    ∧ (req-vec(n;u;v)  ((p a⋅r0) ∧ req-vec(n;u;p))))))])))

Lemma: rv-line-circle-1
n:ℕ. ∀a,b,p,q:ℝ^n.
  (a ≠ b
   p ≠ q
   (d(a;p) ≤ d(a;b))
   (d(a;b) ≤ d(a;q))
   (∃u:{u:ℝ^n| ab=au ∧ (q ≠ u ∧ u ≠ p ∧ q-u-p)))} 
       ∃v:{v:ℝ^n| ab=av ∧ (q ≠ p ∧ p ≠ v ∧ q-p-v)))} 
        ((d(a;p) < d(a;b))  (q-p-v ∧ ((d(a;b) < d(a;q))  q-u-p)))))

Lemma: rv-line-circle
n:ℕ. ∀a,b,p,q:ℝ^n.
  (p ≠ q
   (∀x:{x:ℝ^n| ap=ax ∧ (a ≠ x ∧ x ≠ b ∧ a-x-b)))} . ∀y:{y:ℝ^n| aq=ay ∧ (a ≠ b ∧ b ≠ y ∧ a-b-y)))} .
        ∃u:{u:ℝ^n| ab=au ∧ (q ≠ u ∧ u ≠ p ∧ q-u-p)))} 
         ∃v:{v:ℝ^n| ab=av ∧ (q ≠ p ∧ p ≠ v ∧ q-p-v)))} (a-x-b  (q-p-v ∧ (a-b-y  q-u-p)))))

Definition: rv-ge
cd ≥ ab ==  ¬¬(∃u,v,w:ℝ^n. ((¬(u ≠ v ∧ v ≠ w ∧ u-v-w))) ∧ ab=uv ∧ cd=uw))

Lemma: rv-ge_wf
[n:ℕ]. ∀[c,d,a,b:ℝ^n].  (cd ≥ ab ∈ ℙ)

Definition: rv-be
a_b_c ==  ¬(a ≠ b ∧ b ≠ c ∧ a-b-c))

Lemma: rv-be_wf
[n:ℕ]. ∀[a,b,c:ℝ^n].  (a_b_c ∈ ℙ)

Lemma: stable__rv-be
[n:ℕ]. ∀[a,b,c:ℝ^n].  Stable{a_b_c}

Lemma: sq_stable__rv-be
[n:ℕ]. ∀[a,b,c:ℝ^n].  SqStable(a_b_c)

Lemma: rv-be-inner-trans
a,b,c,d:ℝ^2.  (a_b_d  b_c_d  a_b_c)

Lemma: rv-be-symmetry
n:ℕ. ∀a,b,c:ℝ^n.  (a_b_c  c_b_a)

Lemma: rv-be-five-segment
a,b,c,d,A,B,C,D:ℝ^2.  (a ≠  a_b_c  A_B_C  ab=AB  bc=BC  ad=AD  bd=BD  cd=CD)

Lemma: rv-be-dist
[n:ℕ]. ∀[a,b,c:ℝ^n].  (a_b_c  (d(a;c) (d(a;b) d(b;c))))

Lemma: rv-pos-angle-not-be
[n:ℕ]. ∀[a,b,c:ℝ^n].  ¬a_b_c supposing rv-pos-angle(n;a;b;c)

Lemma: rv-ge-dist
n:ℕ. ∀a,b,c,p:ℝ^n.  (d(a;b) ≤ d(c;p) ⇐⇒ cp ≥ ab)

Lemma: rv-gt-dist
n:ℕ. ∀a,b,q:ℝ^n.  ((d(a;q) < d(a;b))  (∃w:ℝ^n. (a_w_b ∧ aw=aq ∧ w ≠ b)))

Lemma: rv-line-circle-2
n:ℕ. ∀a,b:ℝ^n. ∀p:{p:ℝ^n| ab ≥ ap} . ∀q:{q:ℝ^n| aq ≥ ab} .
  (a ≠  p ≠  (∃u:{u:ℝ^n| ab=au ∧ q_u_p} (∃v:ℝ^n [(ab=av ∧ q_p_v)])))

Lemma: rv-line-circle-3
n:ℕ. ∀c,b,d,q:ℝ^n.
  (c_b_d
   (d(c;d) < d(c;q))
   (∃u:{u:ℝ^n| cd=cu ∧ q_u_b} 
       (∃v:ℝ^n [(cd=cv
               ∧ q_b_v
               ∧ (b ≠  (u ≠ v ∧ u ≠ b ∧ v ≠ b))
               ∧ (req-vec(n;b;d)
                  ((u ≠  ((req-vec(n;u;b) ∧ (r0 < c⋅b)) ∨ (req-vec(n;v;b) ∧ (b c⋅b < r0))))
                    ∧ (req-vec(n;u;v)  ((b c⋅r0) ∧ req-vec(n;u;b))))))])))

Lemma: rv-line-circle-3-ext
n:ℕ. ∀c,b,d,q:ℝ^n.
  (c_b_d
   (d(c;d) < d(c;q))
   (∃u:{u:ℝ^n| cd=cu ∧ q_u_b} 
       (∃v:ℝ^n [(cd=cv
               ∧ q_b_v
               ∧ (b ≠  (u ≠ v ∧ u ≠ b ∧ v ≠ b))
               ∧ (req-vec(n;b;d)
                  ((u ≠  ((req-vec(n;u;b) ∧ (r0 < c⋅b)) ∨ (req-vec(n;v;b) ∧ (b c⋅b < r0))))
                    ∧ (req-vec(n;u;v)  ((b c⋅r0) ∧ req-vec(n;u;b))))))])))

Lemma: better-r2-straightedge-compass
c,d,a:ℝ^2. ∀b:{b:ℝ^2| b ≠ a ∧ c_b_d} .
  ∃u:{u:ℝ^2| cu=cd ∧ a_b_u} 
   (∃v:ℝ^2 [(cv=cd
           ∧ v_b_u
           ∧ ((¬a_b_v) ∧ b_v_a) ∧ v_a_b)))
           ∧ (b ≠  v ≠ u)
           ∧ ((¬d ≠ b)
              ((v ≠  ((¬u ≠ b) ∨ v ≠ b)))
                ∧ ((¬v ≠ u)  ((∀a':ℝ^2. ((d(a';b) d(a;b))  a'_b_a  (d(c;a) d(c;a')))) ∧ u ≠ b))))))])

Lemma: r2-straightedge-compass
c,d,a:ℝ^2. ∀b:{b:ℝ^2| b ≠ a ∧ c_b_d} .
  ∃u:{u:ℝ^2| cu=cd ∧ a_b_u} 
   (∃v:ℝ^2 [(cv=cd ∧ v_b_u ∧ ((¬a_b_v) ∧ b_v_a) ∧ v_a_b))) ∧ (b ≠  (v ≠ u ∧ u ≠ b ∧ v ≠ b)))])

Lemma: r2-straightedge-compass-ext
c,d,a:ℝ^2. ∀b:{b:ℝ^2| b ≠ a ∧ c_b_d} .
  ∃u:{u:ℝ^2| cu=cd ∧ a_b_u} 
   (∃v:ℝ^2 [(cv=cd ∧ v_b_u ∧ ((¬a_b_v) ∧ b_v_a) ∧ v_a_b))) ∧ (b ≠  (v ≠ u ∧ u ≠ b ∧ v ≠ b)))])

Lemma: r2-straightedge-compass-simple1
c,d,a:ℝ^2. ∀b:{b:ℝ^2| b ≠ a ∧ c_b_d} .  (∃u:ℝ^2 [(cu=cd ∧ a_b_u ∧ (b ≠  b ≠ u))])

Lemma: rv-weak-triangle-inequality
n:ℕ. ∀a,b,x,p:ℝ^n.  (ax=ab  a-x-p  p ≠ b)

Lemma: rv-weak-triangle-inequality2
n:ℕ. ∀a,b,x,p:ℝ^n.  (ax=ab  a_x_p  bp ≥ xp)

Lemma: rv-circle-circle-lemma
n:ℕ. ∀r1,r2:{r:ℝr0 ≤ r} . ∀b:ℝ^n.
  ((r0 < ||b||)
   (∀b':ℝ^n
        ((b⋅b' r0)
         (||b'|| ||b||)
         ((r1^2 r2^2) ||b||^2^2 ≤ (r(4) ||b||^2 r1^2))
         let ((r1^2 r2^2) ||b||^2/r(2)) in
            let (||b||^2 r1^2) c^2 in
            ∀x:ℝ^n
              ((req-vec(n;x;(r1/||b||^2)*c*b rsqrt(d)*b') ∨ req-vec(n;x;(r1/||b||^2)*c*b rsqrt(d)*b'))
               ((||x|| r1) ∧ (||x b|| r2))))))

Lemma: rv-circle-circle-lemma2
n:{2...}. ∀r1,r2:{r:ℝr0 ≤ r} . ∀b:ℝ^n.
  ((r0 < ||b||)
   ((r1^2 r2^2) ||b||^2^2 ≤ (r(4) ||b||^2 r1^2))
   (∃u,v:ℝ^n
       (((||u|| r1) ∧ (||u b|| r2))
       ∧ ((||v|| r1) ∧ (||v b|| r2))
       ∧ (((r1^2 r2^2) ||b||^2^2 < (r(4) ||b||^2 r1^2))  u ≠ v))))

Lemma: rv-circle-circle-lemma3
n:{2...}. ∀a,b,c,d:ℝ^n. ∀p:{p:ℝ^n| ab=ap} . ∀q:{q:ℝ^n| cd=cq} . ∀x:{x:ℝ^n| cp=cx ∧ (c ≠ x ∧ x ≠ d ∧ c-x-d)))} .
y:{y:ℝ^n| aq=ay ∧ (a ≠ y ∧ y ≠ b ∧ a-y-b)))} .
  (a ≠  (∃u,v:{p:ℝ^n| ab=ap ∧ cd=cp} (((d(a;y) < d(a;b)) ∧ (d(c;x) < d(c;d)))  u ≠ v)))

Lemma: rv-circle-circle
n:{2...}. ∀a,b,c,d:ℝ^n. ∀p:{p:ℝ^n| ab=ap} . ∀q:{q:ℝ^n| cd=cq} . ∀x:{x:ℝ^n| cp=cx ∧ (c ≠ x ∧ x ≠ d ∧ c-x-d)))} .
y:{y:ℝ^n| aq=ay ∧ (a ≠ y ∧ y ≠ b ∧ a-y-b)))} .
  (a ≠  (∃u,v:{p:ℝ^n| ab=ap ∧ cd=cp} ((x ≠ d ∧ y ≠ b)  u ≠ v)))

Definition: r2-det
|pqr| ==  (((p 0) (q 1)) ((q 0) (r 1)) ((r 0) (p 1))) ((p 1) (q 0)) ((q 1) (r 0)) ((r 1) (p 0))

Lemma: r2-det_wf
[p,q,r:ℝ^2].  (|pqr| ∈ ℝ)

Lemma: r2-det-is-dot-product
[a,b,c:ℝ^2].  (|abc| b⋅λi.if (i =z 0) then -(c 1) else fi )

Lemma: r2-det_functionality
[p,q,r,p',q',r':ℝ^2].  (|pqr| |p'q'r'|) supposing (req-vec(2;r;r') and req-vec(2;q;q') and req-vec(2;p;p'))

Lemma: r2-det-symmetry
[p,q,r:ℝ^2].  (|pqr| |qrp|)

Lemma: r2-det-antisymmetry
[p,q,r:ℝ^2].  (|pqr| -(|prq|))

Lemma: r2-det-identity
[p,q,r,t:ℝ^2].  (|pqr| (|tqr| |ptr| |pqt|))

Lemma: another-r2-det-identity
[p,q,r,t,s:ℝ^2].  (((|tqr| |tsp|) (|ptr| |tsq|) (|pqt| |tsr|)) r0)

Lemma: r2-det-syzygy
[p,q,r,u,v,w,x,y,z:ℝ^2].
  (((|pqw| |rpv| |qry| |prx| |quz|)
  (|rpv| |qru| |rpz| |qpy| |qwx|)
  (|qru| |pqw| |rpz| |prx| |qvy|)
  (|qru| |pqw| |rpz| |qpy| |rvx|)
  (|pqw| |rpv| |pqx| |rqz| |ruy|)
  (|rpv| |qru| |pqx| |qpy| |rwz|)
  (|rpv| |qru| |pqx| |rqz| |pwy|)
  (|qru| |pqw| |qry| |prx| |pvz|)
  (|pqw| |rpv| |qry| |rqz| |pux|))
  r0)

Lemma: r2-det-syzygy-proof2
[p,q,r,u,v,w,x,y,z:ℝ^2].
  (((|pqw| |rpv| |qry| |prx| |quz|)
  (|rpv| |qru| |rpz| |qpy| |qwx|)
  (|qru| |pqw| |rpz| |prx| |qvy|)
  (|qru| |pqw| |rpz| |qpy| |rvx|)
  (|pqw| |rpv| |pqx| |rqz| |ruy|)
  (|rpv| |qru| |pqx| |qpy| |rwz|)
  (|rpv| |qru| |pqx| |rqz| |pwy|)
  (|qru| |pqw| |qry| |prx| |pvz|)
  (|pqw| |rpv| |qry| |rqz| |pux|))
  r0)

Lemma: r2-det-add
[p,q,r,t:ℝ^2].  (|p tqr| (|pqr| |tqr| (((q 1) (r 0)) (q 0) (r 1))))

Lemma: r2-det-mul
[p,q,r:ℝ^2]. ∀[a:ℝ].  (|a*pqr| ((a |pqr|) ((r1 a) (((q 0) (r 1)) (q 1) (r 0)))))

Lemma: r2-det-convex1
[p,q,r,t:ℝ^2]. ∀[a:ℝ].  (|a*p r1 a*tqr| ((a |pqr|) ((r1 a) |tqr|)))

Lemma: r2-det-convex2
[p,q,r,t,s:ℝ^2]. ∀[a,b,c:ℝ].
  |a*p b*q c*rts| ((a |pts|) (b |qts|) (c |rts|)) supposing ((a c) r1) ∧ r1 a ≠ r0

Lemma: r2-det-convex3
[p,q,r,t,s:ℝ^2]. ∀[a,b,c:ℝ].
  |a*p b*q c*rts| ((a |pts|) (b |qts|) (c |rts|)) supposing (a c) r1

Lemma: r2-det-nonzero
p,q,r:ℝ^2.  (rv-pos-angle(2;p;q;r)  |pqr| ≠ r0)

Definition: r2-left
r2-left(p;q;r) ==  r0 < |pqr|

Lemma: r2-left_wf
[p,q,r:ℝ^2].  (r2-left(p;q;r) ∈ ℙ)

Lemma: sq_stable__r2-left
a,b,c:ℝ^2.  SqStable(r2-left(a;b;c))

Lemma: sq_stable__r2-left-or
a,b,c:ℝ^2.  SqStable(r2-left(a;b;c) ∨ r2-left(a;c;b))

Lemma: r2-left-cases
a,b:ℝ^2. ∀c:{c:ℝ^2| |a r(-1)*b⋅r(-1)*b| < (||a r(-1)*b|| ||c r(-1)*b||)} .
  (r2-left(a;b;c) ∨ r2-left(a;c;b))

Lemma: r2-left-pos-angle
a,b,c:ℝ^2.  (r2-left(a;b;c)  rv-pos-angle(2;a;b;c))

Lemma: r2-left-or-right--pos-angle
a,b,c:ℝ^2.  (r2-left(a;b;c) ∨ r2-left(a;c;b) ⇐⇒ rv-pos-angle(2;a;b;c))

Lemma: rv-circle-circle-lemma2'
r1,r2:{r:ℝr0 ≤ r} . ∀b:ℝ^2.
  ((r0 < ||b||)
   ((r1^2 r2^2) ||b||^2^2 ≤ (r(4) ||b||^2 r1^2))
   (∃u,v:ℝ^2
       (((||u|| r1) ∧ (||u b|| r2))
       ∧ ((||v|| r1) ∧ (||v b|| r2))
       ∧ (((r1^2 r2^2) ||b||^2^2 < (r(4) ||b||^2 r1^2))  (r2-left(u;b;λi.r0) ∧ r2-left(v;λi.r0;b))))))

Lemma: rv-circle-circle-lemma3'
a,b,c,d:ℝ^2. ∀p:{p:ℝ^2| ab=ap} . ∀q:{q:ℝ^2| cd=cq} . ∀x:{x:ℝ^2| cp=cx ∧ (c ≠ x ∧ x ≠ d ∧ c-x-d)))} .
y:{y:ℝ^2| aq=ay ∧ (a ≠ y ∧ y ≠ b ∧ a-y-b)))} .
  (a ≠ c
   (∃u,v:{p:ℝ^2| ab=ap ∧ cd=cp} (((d(a;y) < d(a;b)) ∧ (d(c;x) < d(c;d)))  (r2-left(u;c;a) ∧ r2-left(v;a;c)))))

Lemma: r2-circle-circle
a,b,c,d:ℝ^2. ∀p:{p:ℝ^2| ab=ap} . ∀q:{q:ℝ^2| cd=cq} . ∀x:{x:ℝ^2| cp=cx ∧ (c ≠ x ∧ x ≠ d ∧ c-x-d)))} .
y:{y:ℝ^2| aq=ay ∧ (a ≠ y ∧ y ≠ b ∧ a-y-b)))} .
  (a ≠  (∃u,v:{p:ℝ^2| ab=ap ∧ cd=cp} ((x ≠ d ∧ y ≠ b)  (r2-left(u;c;a) ∧ r2-left(v;a;c)))))

Lemma: rv-compass-compass-lemma
a,b,c,d:ℝ^2.
  (a ≠ c
   (↓∃p,q:ℝ^2. (((d(a;b) d(a;p)) ∧ (d(c;d) d(c;q))) ∧ (d(c;p) ≤ d(c;d)) ∧ (d(a;q) ≤ d(a;b))))
   (∃u,v:{p:ℝ^2| ab=ap ∧ cd=cp} 
       ((↓∃p,q:ℝ^2. (((d(a;b) d(a;p)) ∧ (d(c;d) d(c;q))) ∧ (d(c;p) < d(c;d)) ∧ (d(a;q) < d(a;b))))
        (r2-left(u;c;a) ∧ r2-left(v;a;c)))))

Lemma: r2-compass-compass
a,b:ℝ^2. ∀c:{c:ℝ^2| a ≠ c} . ∀d:{d:ℝ^2| 
                                 ↓∃p,q:ℝ^2
                                   ((ab=ap ∧ (¬¬(∃w:ℝ^2. (c_w_d ∧ cw=cp)))) ∧ cd=cq ∧ (¬¬(∃w:ℝ^2. (a_w_b ∧ aw=aq))))} .
  ∃u:{u:ℝ^2| ab=au ∧ cd=cu} 
   (∃v:ℝ^2 [((ab=av ∧ cd=cv)
           ∧ ((↓∃p,q:ℝ^2. ((ab=ap ∧ (↓∃w:ℝ^2. (c_w_d ∧ cw=cp ∧ w ≠ d))) ∧ cd=cq ∧ (↓∃w:ℝ^2. (a_w_b ∧ aw=aq ∧ w ≠ b))))
              (r2-left(u;a;c) ∧ r2-left(v;c;a))))])

Lemma: r2-compass-compass-simple1
a,b:ℝ^2. ∀c:{c:ℝ^2| a ≠ c} . ∀d:{d:ℝ^2| 
                                 ↓∃p,q:ℝ^2
                                   ((ab=ap ∧ (↓∃w:ℝ^2. (c_w_d ∧ cw=cp ∧ w ≠ d)))
                                   ∧ cd=cq
                                   ∧ (↓∃w:ℝ^2. (a_w_b ∧ aw=aq ∧ w ≠ b)))} .
  (∃u:ℝ^2 [(ab=au ∧ cd=cu ∧ r2-left(u;a;c))])

Lemma: r2-left-sep
a,b,c:ℝ^2.  (r2-left(a;b;c)  b ≠ c)

Lemma: r2-left-separated
a,b,c,d:ℝ^2.  (r2-left(a;c;d)  r2-left(b;d;c)  a ≠ b)

Lemma: r2-left-convex
a,b,x,y,z:ℝ^2.  (r2-left(x;a;b)  r2-left(z;a;b)  rv-T(2;x;y;z)  r2-left(y;a;b))

Lemma: r2-left-extend
a,b,x,y:ℝ^2.  (r2-left(x;a;b)  rv-T(2;b;x;y)  r2-left(y;a;b))

Lemma: r2-left-between
a,b,x,y:ℝ^2.  (r2-left(x;a;b)  rv-T(2;b;y;x)  b ≠  r2-left(y;a;b))

Lemma: r2-left-right-lemma
a,b,x,y:ℝ^2.  (r2-left(x;a;b)  r2-left(y;b;a)  (∃t:ℝ((t ∈ [r0, r1]) ∧ (|t*x r1 t*yab| r0))))

Lemma: r2-left-right
a,b,x,y:ℝ^2.  (r2-left(x;a;b)  r2-left(y;b;a)  (∃z:ℝ^2. (rv-T(2;x;z;y) ∧ rv-pos-angle(2;z;a;b)))))

Lemma: r2-plane-separation1
a,b:ℝ^2. ∀u:{u:ℝ^2| r2-left(u;a;b)} . ∀v:{v:ℝ^2| r2-left(v;b;a)} .
  (∃x:ℝ^2 [((¬((¬a_b_x) ∧ b_x_a) ∧ x_a_b))) ∧ u_x_v)])

Lemma: r2-not-left-right
a,b,c:ℝ^2.  ((¬r2-left(a;b;c))  r2-left(a;c;b))  ((¬rv-T(2;a;b;c)) ∧ rv-T(2;b;c;a)) ∧ rv-T(2;c;a;b)))))

Lemma: r2-not-left-right-iff
a,b,c:ℝ^2.  (r2-left(a;b;c) ∨ r2-left(a;c;b)) ⇐⇒ ¬((¬a_b_c) ∧ b_c_a) ∧ c_a_b)))

Lemma: r2-be-iff
u,v,x:ℝ^2.  (u_x_v ⇐⇒ ((¬(d(u;v) < d(u;x))) ∧ (d(u;v) < d(x;v)))) ∧ (r2-left(u;x;v) ∨ r2-left(u;v;x))))

Lemma: real-vec-sep-iff2
n:ℕ. ∀a,b:ℝ^n.  (b ≠ ⇐⇒ d(b;b) < d(b;a))

Lemma: rv-congruent-iff2
n:ℕ. ∀a,b,c,d:ℝ^n.  (ab=cd ⇐⇒ ¬((d(a;b) < d(c;d)) ∨ (d(c;d) < d(a;b))))

Lemma: r2-plane-separation
a,b:ℝ^2. ∀u:{u:ℝ^2| r2-left(u;a;b)} . ∀v:{v:ℝ^2| r2-left(v;b;a)} .
  (∃x:ℝ^2 [((¬(r2-left(a;b;x) ∨ r2-left(a;x;b)))
          ∧ ((¬(d(u;v) < d(u;x))) ∧ (d(u;v) < d(x;v))))
          ∧ (r2-left(u;x;v) ∨ r2-left(u;v;x))))])

Lemma: r2-straightedge-compass-simple
c,d,a:ℝ^2. ∀b:{b:ℝ^2| 
                (d(b;b) < d(b;a))
                ∧ ((¬(d(c;d) < d(c;b))) ∧ (d(c;d) < d(b;d))))
                ∧ (r2-left(c;b;d) ∨ r2-left(c;d;b)))} .
  (∃u:ℝ^2 [((¬((d(c;u) < d(c;d)) ∨ (d(c;d) < d(c;u))))
          ∧ (((¬(d(a;u) < d(a;b))) ∧ (d(a;u) < d(b;u)))) ∧ (r2-left(a;b;u) ∨ r2-left(a;u;b))))
          ∧ ((d(b;b) < d(b;d))  (d(b;b) < d(b;u))))])

Lemma: r2-compass-compass-simple
a,b:ℝ^2. ∀c:{c:ℝ^2| d(a;a) < d(a;c)} . ∀d:{d:ℝ^2| 
                                           ↓∃p,q:ℝ^2
                                             (((¬((d(a;b) < d(a;p)) ∨ (d(a;p) < d(a;b)))) ∧ (d(c;p) < d(c;d)))
                                             ∧ ((d(c;d) < d(c;q)) ∨ (d(c;q) < d(c;d))))
                                             ∧ (d(a;q) < d(a;b)))} .
  (∃u:ℝ^2 [((¬((d(a;b) < d(a;u)) ∨ (d(a;u) < d(a;b)))) ∧ ((d(c;d) < d(c;u)) ∨ (d(c;u) < d(c;d)))) ∧ r2-left(u;a;c))])

Definition: rcp
(a b) ==  λi.[((a 1) (b 2)) (a 2) (b 1); ((a 2) (b 0)) (a 0) (b 2); ((a 0) (b 1)) (a 1) (b 0)][i]

Lemma: rcp_wf
[a,b:ℝ^3].  ((a b) ∈ ℝ^3)

Lemma: rcp_functionality
[a1,b1,a2,b2:ℝ^3].  (req-vec(3;(a1 b1);(a2 b2))) supposing (req-vec(3;b1;b2) and req-vec(3;a1;a2))

Lemma: rcp-same
[a:ℝ^3]. req-vec(3;(a a);λi.r0)

Lemma: rcp-0
[a:ℝ^3]. req-vec(3;(a x λi.r0);λi.r0)

Lemma: rcp-anti-comm
[a,b:ℝ^3].  req-vec(3;(b a);r(-1)*(a b))

Lemma: rcp-Jacobi
[a,b,c:ℝ^3].  req-vec(3;(a (b c)) (b (c a)) (c (a b));λi.r0)

Lemma: rcp-distrib1
[a,b,c:ℝ^3].  req-vec(3;(a c);(a b) (a c))

Lemma: rcp-distrib2
[a,b,c:ℝ^3].  req-vec(3;(b a);(b a) (c a))

Lemma: rcp-mul1
[a,b:ℝ^3]. ∀[c:ℝ].  req-vec(3;(c*a b);c*(a b))

Lemma: rcp-mul2
[a,b:ℝ^3]. ∀[c:ℝ].  req-vec(3;(a c*b);c*(a b))

Lemma: rcp-Binet-Cauchy
[a,b,c,d:ℝ^3].  ((a b)⋅(c d) ((a⋅b⋅d) a⋅b⋅c))

Lemma: rcp-Binet-Cauchy-corollary
[a,b:ℝ^3].  ((a b)⋅(a b) ((a⋅b⋅b) a⋅b⋅a))

Lemma: rcp-perp1
[a,b:ℝ^3].  (a⋅(a b) r0)

Lemma: rcp-perp2
[a,b:ℝ^3].  (b⋅(a b) r0)

Definition: r3-dp-prim
r3-dp-prim() ==  (vec={v:ℝ^3| v ≠ λi.r0} sep=λa,b. (a b) ≠ λi.r0, perp=λa,b. (a⋅r0))

Lemma: r3-dp-prim_wf
r3-dp-prim() ∈ DualPlanePrimitives

Definition: r-rational
r-rational(x) ==  ∃a:ℤ. ∃b:ℕ+(x (r(a)/r(b)))

Lemma: r-rational_wf
[x:ℝ]. (r-rational(x) ∈ ℙ)

Definition: irrational
irrational(x) ==  ∀a:ℤ. ∀b:ℕ+.  x ≠ (r(a)/r(b))

Lemma: irrational_wf
[x:ℝ]. (irrational(x) ∈ ℙ)

Lemma: irrational-sqrt-number-lemma
a:ℤ. ∀b:ℕ+. ∀n:ℕ.  (((a a) (n b) ∈ ℤ (∃m:ℕ1. ((m m) n ∈ ℤ)))

Lemma: rsqrt-irrational
n:ℕ(irrational(rsqrt(r(n))) ∨ (∃m:ℕ1. ((m m) n ∈ ℤ)))

Lemma: rsqrt_2-irrational
irrational(rsqrt(r(2)))

Lemma: rsqrt2-repels-rationals
n:ℕ+. ∀m:ℤ.  ((r1/r(3 n)) ≤ |rsqrt(r(2)) (r(m)/r(n))|)

Definition: fun-converges-to
Like our definition of continuouswe quantify only over ∈ of the form (r1/r(k))
  and compact subintervals of the form i-approx(I;m).⋅

lim n→∞.f[n; x] = λy.g[y] for x ∈ ==
  ∀m:{m:ℕ+icompact(i-approx(I;m))} . ∀k:ℕ+.
    ∃N:ℕ+. ∀x:{x:ℝx ∈ i-approx(I;m)} . ∀n:{N...}.  (|f[n; x] g[x]| ≤ (r1/r(k)))

Lemma: fun-converges-to_wf
[I:Interval]. ∀[f:ℕ ⟶ I ⟶ℝ]. ∀[g:I ⟶ℝ].  (lim n→∞.f[n;x] = λy.g[y] for x ∈ I ∈ ℙ)

Lemma: fun-converges-to_functionality
I:Interval. ∀f1,f2:ℕ ⟶ I ⟶ℝ.
  ∀[g:I ⟶ℝ]
    ((∀n:ℕ. ∀x:{x:ℝx ∈ I} .  (f1[n;x] f2[n;x]))
     {lim n→∞.f1[n;x] = λy.g[y] for x ∈  lim n→∞.f2[n;x] = λy.g[y] for x ∈ I})

Lemma: fun-converges-to_functionality2
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ. ∀g1:I ⟶ℝ.
  ∀[g2:I ⟶ℝ]
    ((∀x:{x:ℝx ∈ I} (g1[x] g2[x]))  {lim n→∞.f[n;x] = λy.g1[y] for x ∈  lim n→∞.f[n;x] = λy.g2[y] for x ∈ I}\000C)

Lemma: fun-converges-to-rsub
I:Interval. ∀f1,f2:ℕ ⟶ I ⟶ℝ. ∀g1,g2:I ⟶ℝ.
  (lim n→∞.f1[n;x] = λy.g1[y] for x ∈ I
   lim n→∞.f2[n;x] = λy.g2[y] for x ∈ I
   lim n→∞.f1[n;x] f2[n;x] = λy.g1[y] g2[y] for x ∈ I)

Lemma: fun-converges-to-rminus
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ. ∀g:I ⟶ℝ.  (lim n→∞.f[n;x] = λy.g[y] for x ∈  lim n→∞.-(f[n;x]) = λy.-(g[y]) for x ∈ I)

Lemma: fun-converges-to-pointwise
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ. ∀g:I ⟶ℝ.
  (lim n→∞.f[n;x] = λy.g[y] for x ∈  {∀x:ℝ((x ∈ I)  lim n→∞.f[n;x] g[x])})

Lemma: fun-converges-to-continuous
[I:Interval]. ∀[f:ℕ ⟶ I ⟶ℝ]. ∀[g:I ⟶ℝ].
  (lim n→∞.f[n;x] = λy.g[y] for x ∈  (∀n:ℕf[n;x] continuous for x ∈ I)  g[y] continuous for y ∈ I)

Definition: fun-cauchy
λn.f[n; x] is cauchy for x ∈ ==
  ∀a:{a:ℕ+icompact(i-approx(I;a))} . ∀k:ℕ+.
    ∃N:ℕ+. ∀x:{x:ℝx ∈ i-approx(I;a)} . ∀n,m:{N...}.  (|f[n; x] f[m; x]| ≤ (r1/r(k)))

Lemma: fun-cauchy_wf
[I:Interval]. ∀[f:ℕ ⟶ I ⟶ℝ].  n.f[n;x] is cauchy for x ∈ I ∈ ℙ)

Definition: fun-converges
λn.f[n; x]↓ for x ∈ I) ==  ∃g:I ⟶ℝlim n→∞.f[n; x] = λy.g for x ∈ I

Lemma: fun-converges_wf
[I:Interval]. ∀[f:ℕ ⟶ I ⟶ℝ].  n.f[n;x]↓ for x ∈ I) ∈ ℙ)

Lemma: fun-converges-converges-to
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ. ∀g:I ⟶ℝ.
  ((∀x:{x:ℝx ∈ I} lim n→∞.f[n;x] g[x])  λn.f[n;x]↓ for x ∈ I)  lim n→∞.f[n;x] = λx.g[x] for x ∈ I)

Lemma: fun-converges_functionality
[I:Interval]. ∀f,g:ℕ ⟶ I ⟶ℝ.  ((∀n:ℕrfun-eq(I;f n;g n))  λn.f[n;x]↓ for x ∈ I)  λn.g[n;x]↓ for x ∈ I))

Lemma: fun-converges-iff-cauchy
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ.  n.f[n;x]↓ for x ∈ I) ⇐⇒ λn.f[n;x] is cauchy for x ∈ I)

Lemma: fun-converges-on-compact
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ.
  ((∀m:{m:ℕ+icompact(i-approx(I;m))} . λn.f[n;x]↓ for x ∈ i-approx(I;m)))  λn.f[n;x]↓ for x ∈ I))

Lemma: const-fun-converges
I:Interval. ∀f:ℕ ⟶ ℝ.  (f[n]↓ as n→∞  λn.f[n]↓ for x ∈ I))

Lemma: fun-converges-rmul
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ.
  n.f[n;x]↓ for x ∈ I)  (∀g:I ⟶ℝ(g[x] continuous for x ∈  λn.f[n;x] g[x]↓ for x ∈ I))))

Definition: fun-series-converges
Σn.f[n; x]↓ for x ∈ ==  λn.Σ{f[i; x] 0≤i≤n}↓ for x ∈ I)

Lemma: fun-series-converges_wf
[I:Interval]. ∀[f:ℕ ⟶ I ⟶ℝ].  n.f[n;x]↓ for x ∈ I ∈ ℙ)

Lemma: fun-series-converges-on-compact
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ.
  ((∀m:{m:ℕ+icompact(i-approx(I;m))} . Σn.f[n;x]↓ for x ∈ i-approx(I;m))  Σn.f[n;x]↓ for x ∈ I)

Definition: fun-series-sum
Σn.f[n](z) ==  (fst(cnv)) z

Lemma: fun-series-sum_wf
[I:Interval]. ∀[f:ℕ ⟶ I ⟶ℝ].  ∀cnv:Σn.f[n;x]↓ for x ∈ I. ∀z:{z:ℝz ∈ I} .  n.f[n](z) ∈ ℝ)

Lemma: continuous-series-sum
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ. ∀cnv:Σn.f[n;x]↓ for x ∈ I.
  ((∀n:ℕf[n;x] continuous for x ∈ I)  Σn.f[n](y) continuous for y ∈ I)

Definition: fun-series-converges-absolutely
Σn.f[n; x]↓ absolutely for x ∈ ==  Σn.|f[n; x]|↓ for x ∈ I

Lemma: fun-series-converges-absolutely_wf
[I:Interval]. ∀[f:ℕ ⟶ I ⟶ℝ].  n.f[n;x]↓ absolutely for x ∈ I ∈ ℙ)

Lemma: fun-comparison-test
I:Interval. ∀f,g:ℕ ⟶ I ⟶ℝ.
  n.g[n;x]↓ for x ∈  (∀n:ℕ. ∀x:{x:ℝx ∈ I} .  (|f[n;x]| ≤ g[n;x]))  Σn.f[n;x]↓ for x ∈ I)

Lemma: fun-series-converges-absolutely-converges
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ.  n.f[n;x]↓ absolutely for x ∈  Σn.f[n;x]↓ for x ∈ I)

Lemma: fun-series-converges-tail
M:ℕ. ∀I:Interval. ∀f:ℕ ⟶ I ⟶ℝ.  n.f[n M;x]↓ for x ∈  Σn.f[n;x]↓ for x ∈ I)

Lemma: fun-ratio-test
I:Interval. ∀f:ℕ ⟶ I ⟶ℝ.
  ((∀n:ℕf[n;x] continuous for x ∈ I)
   (∀m:{m:ℕ+icompact(i-approx(I;m))} 
        ∃c:ℝ((r0 ≤ c) ∧ (c < r1) ∧ (∃N:ℕ. ∀n:{N...}. ∀x:{x:ℝx ∈ i-approx(I;m)} .  (|f[n 1;x]| ≤ (c |f[n;x]|)))))
   Σn.f[n;x]↓ absolutely for x ∈ I)

Lemma: fun-ratio-test-everywhere
f:ℕ ⟶ ℝ ⟶ ℝ
  ((∀n:ℕ. ∀x,y:ℝ.  ((x y)  (f[n;x] f[n;y])))
   (∀m:ℕ+. ∃c:ℝ((r0 ≤ c) ∧ (c < r1) ∧ (∃N:ℕ. ∀n:{N...}. ∀x:{x:ℝ|x| ≤ r(m)} .  (|f[n 1;x]| ≤ (c |f[n;x]|)))))
   Σn.f[n;x]↓ absolutely for x ∈ (-∞, ∞))

Lemma: fun-series-converges-to-everywhere
f:ℕ ⟶ ℝ ⟶ ℝ
  ((∀n:ℕ. ∀x,y:ℝ.  ((x y)  (f[n;x] f[n;y])))
   (∀m:ℕ+. ∃c:ℝ((r0 ≤ c) ∧ (c < r1) ∧ (∃N:ℕ. ∀n:{N...}. ∀x:{x:ℝ|x| ≤ r(m)} .  (|f[n 1;x]| ≤ (c |f[n;x]|)))))
   (∀g:ℝ ⟶ ℝ((∀x:ℝlim n→∞{f[i;x] 0≤i≤n} g[x])  lim n→∞{f[i;x] 0≤i≤n} = λx.g[x] for x ∈ (-∞, ∞))))

Lemma: power-series-converges
a:ℕ ⟶ ℝ. ∀b:ℝ. ∀r:{r:ℝr0 < r} . ∀N:ℕ.
  ((∀n:{N...}. (|a[n 1]| ≤ (|a[n]|/r)))  Σn.a[n] b^n↓ absolutely for x ∈ (b r, r))

Lemma: power-series-converges-to
a:ℕ ⟶ ℝ. ∀b:ℝ. ∀r:{r:ℝr0 < r} . ∀N:ℕ.
  ((∀n:{N...}. (|a[n 1]| ≤ (|a[n]|/r)))
   (∀g:(b r, r) ⟶ℝ
        ((∀x:{x:ℝx ∈ (b r, r)} . Σi.a[i] b^i g[x])
         lim n→∞{a[i] b^i 0≤i≤n} = λx.g[x] for x ∈ (b r, r))))

Lemma: power-series-converges-everywhere-to
a:ℕ ⟶ ℝ. ∀b:ℝ. ∀g:ℝ ⟶ ℝ.
  ((∀x:ℝ. Σi.a[i] b^i g[x])
   (∀r:{r:ℝr0 < r} . ∃N:ℕ. ∀n:{N...}. (|a[n 1]| ≤ (|a[n]|/r)))
   lim n→∞{a[i] b^i 0≤i≤n} = λx.g[x] for x ∈ (-∞, ∞))

Lemma: power-series-converges-everywhere
a:ℕ ⟶ ℝ. ∀g:ℝ ⟶ ℝ.
  ((∀x:ℝ. Σi.a[i] x^i g[x])
   (∀r:{r:ℝr0 < r} . ∃N:ℕ. ∀n:{N...}. (|a[n 1]| ≤ (|a[n]|/r)))
   lim n→∞{a[i] x^i 0≤i≤n} = λx.g[x] for x ∈ (-∞, ∞))

Lemma: fun-converges-to-rexp
lim n→∞{(x^i)/(i)! 0≤i≤n} = λx.e^x for x ∈ (-∞, ∞)

Lemma: fun-converges-to-cosine
lim n→∞{-1^i (x^2 i)/(2 i)! 0≤i≤n} = λx.cosine(x) for x ∈ (-∞, ∞)

Lemma: fun-converges-to-sine
lim n→∞{-1^i (x^(2 i) 1)/((2 i) 1)! 0≤i≤n} = λx.sine(x) for x ∈ (-∞, ∞)



Home Index