Definition: experimental-impliesFunctionality def
experimental{impliesFunctionality}(possibleextract) ==  possibleextract

Definition: experimental-allFunctionality def
experimental{allFunctionality}(possibleextract) ==  possibleextract

Definition: experimental-uallFunctionality def
experimental{uallFunctionality}(possibleextract) ==  possibleextract

Definition: experimental-orFunctionality def
experimental{orFunctionality}(possibleextract) ==  possibleextract

Definition: square-board
square-board(n;T) ==  {s:T List List| (||s|| n ∈ ℤ) ∧ (∀i:ℕn. (||s[i]|| n ∈ ℤ))} 

Lemma: square-board_wf
[n:ℕ]. ∀[T:Type].  (square-board(n;T) ∈ Type)

Lemma: square-board-subtype
[n:ℕ]. ∀[T,S:Type].  square-board(n;T) ⊆square-board(n;S) supposing T ⊆S

Definition: map-square-board
map-square-board(i,j,v.f[i;
                         j;
                         v];b) ==
  map-index(λi,r. map-index(λj,v. f[i;
                                    j;
                                    v];r);b)

Lemma: map-square-board_wf
[n:ℕ]. ∀[T1,T2:Type]. ∀[f:ℕn ⟶ ℕn ⟶ T1 ⟶ T2]. ∀[b:square-board(n;T1)].
  (map-square-board(i,j,v.f[i;j;v];b) ∈ square-board(n;T2))

Definition: board-cell
board-cell(b;i;j) ==  b[i][j]

Lemma: board-cell_wf
[T:Type]. ∀[n:ℕ]. ∀[b:square-board(n;T)]. ∀[i,j:ℕn].  (board-cell(b;i;j) ∈ T)

Lemma: map-square-board-cell
[n:ℕ]. ∀[T1,T2:Type]. ∀[f:ℕn ⟶ ℕn ⟶ T1 ⟶ T2]. ∀[b:square-board(n;T1)]. ∀[i,j:ℕn].
  (board-cell(map-square-board(i,j,v.f[i;j;v];b);i;j) f[i;j;board-cell(b;i;j)] ∈ T2)

Definition: sudoku-cell
sudoku-cell(s;i;j) ==  board-cell(s;i;j)

Lemma: sudoku-cell_wf
[s:square-board(9;ℕ+10)]. ∀[i,j:ℕ9].  (sudoku-cell(s;i;j) ∈ ℕ+10)

Definition: Sudoku
Sudoku() ==
  {s:square-board(9;ℕ+10)| 
   (∀i,j,k:ℕ9.  ((sudoku-cell(s;i;j) sudoku-cell(s;i;k) ∈ ℤ (j k ∈ ℤ)))
   ∧ (∀i,j,k:ℕ9.  ((sudoku-cell(s;i;k) sudoku-cell(s;j;k) ∈ ℤ (j k ∈ ℤ)))
   ∧ (∀i,j,k,l:ℕ9.
        ((sudoku-cell(s;i;k) sudoku-cell(s;j;l) ∈ ℤ)
         ((i ÷ 3) (j ÷ 3) ∈ ℤ)
         ((k ÷ 3) (l ÷ 3) ∈ ℤ)
         ((i j ∈ ℤ) ∧ (k l ∈ ℤ))))} 

Lemma: Sudoku_wf
Sudoku() ∈ Type

Definition: SudokuPuzzle
SudokuPuzzle() ==  square-board(9;ℕ+10 List)

Lemma: SudokuPuzzle_wf
SudokuPuzzle() ∈ Type

Definition: sudoku-allowed
sudoku-allowed(P;i;j) ==  board-cell(P;i;j)

Lemma: sudoku-allowed_wf
[P:SudokuPuzzle()]. ∀[i,j:ℕ9].  (sudoku-allowed(P;i;j) ∈ ℕ+10 List)

Definition: Sudoku-size
Sudoku-size(P) ==  Σ(||sudoku-allowed(P;i;j)|| j < 9) i < 9)

Lemma: Sudoku-size_wf
[P:SudokuPuzzle()]. (Sudoku-size(P) ∈ ℕ)

Definition: SudokuSolution
SudokuSolution(s;P) ==  ∀i,j:ℕ9.  (sudoku-cell(s;i;j) ∈ sudoku-allowed(P;i;j))

Lemma: SudokuSolution_wf
[s:Sudoku()]. ∀[P:SudokuPuzzle()].  (SudokuSolution(s;P) ∈ ℙ)

Lemma: Sudoku-induction
P:SudokuPuzzle() ⟶ ℙ
  ((∀p:SudokuPuzzle(). ((∀q:SudokuPuzzle(). P[q] supposing Sudoku-size(q) < Sudoku-size(p))  P[p]))
   (∀p:SudokuPuzzle(). P[p]))

Lemma: sq_stable_and_decidable
[P:ℙ]. (Dec(P)  (∀[Q:⋂x:P. ℙ]. (SqStable(P)  (P  SqStable(Q))  SqStable(P ∧ Q))))

Lemma: sq_stable_and_left_false
[P:ℙ]. ∀Q:Top. ((¬P)  SqStable(P ∧ Q))

Lemma: extract-decider-of-decidable-prop
[T:Type]. ∀[P:T ⟶ ℙ].  ((∀t:T. ((P t) ∨ (P t))))  (∃f:T ⟶ 𝔹. ∀t:T. (↑(f t) ⇐⇒ t)))

Lemma: assert_of_band2
p:𝔹. ∀q:⋂:↑p. 𝔹.  uiff(↑(p ∧b q);(↑p) ∧ (↑q))

Lemma: sqequal-ff-to-assert
[t:𝔹]. uiff(t ff;¬↑t)

Lemma: sqequal-tt-to-assert
[t:𝔹]. uiff(t tt;↑t)

Comment: vr-list-of-added-rules
sqequalnIntensionalEquality
sqlenIntensionalEquality
sqlenSubtypeRel
sqequalnSqlen
sqlenSqequaln
sqequalnSymm


Lemma: sq_stable__sqle
[a,b:Base].  SqStable(a ≤ b)

Lemma: has-value_functionality
[a,b:Base].  {(a)↓  (b)↓supposing a ≤ b

Definition: proof-tree
proof-tree(Sequent;Rule;effect) ==
  W(Sequent × Rule;sr.case effect sr of inl(subgoals) => ℕ||subgoals|| inr(x) => Void)

Lemma: proof-tree_wf
[Sequent,Rule:Type]. ∀[effect:(Sequent × Rule) ⟶ (Sequent List?)].  (proof-tree(Sequent;Rule;effect) ∈ Type)

Lemma: proof-tree-ext
[Sequent,Rule:Type]. ∀[effect:(Sequent × Rule) ⟶ (Sequent List?)].
  proof-tree(Sequent;Rule;effect) ≡ sr:Sequent × Rule × (case effect sr
                                                         of inl(subgoals) =>
                                                         ℕ||subgoals||
                                                         inr(x) =>
                                                         Void ⟶ proof-tree(Sequent;Rule;effect))

Definition: make-proof-tree
make-proof-tree(s;r;L) ==  <<s, r>, λi.L[i]>

Lemma: make-proof-tree_wf
[Sequent,Rule:Type]. ∀[effect:(Sequent × Rule) ⟶ (Sequent List?)]. ∀[s:Sequent]. ∀[r:Rule].
[L:proof-tree(Sequent;Rule;effect) List].
  (make-proof-tree(s;r;L) ∈ proof-tree(Sequent;Rule;effect)) supposing 
     ((||L|| ||outl(effect <s, r>)|| ∈ ℤand 
     (↑isl(effect <s, r>)))

Definition: proof-abort
proof-abort(s;r) ==  <<s, r>, λx.x>

Lemma: proof-abort_wf
[Sequent,Rule:Type]. ∀[effect:(Sequent × Rule) ⟶ (Sequent List?)]. ∀[s:Sequent]. ∀[r:Rule].
  proof-abort(s;r) ∈ proof-tree(Sequent;Rule;effect) supposing ↑isr(effect <s, r>)

Lemma: proof-tree-induction
[Sequent,Rule:Type].
  ∀effect:(Sequent × Rule) ⟶ (Sequent List?)
    ∀[Q:proof-tree(Sequent;Rule;effect) ⟶ ℙ]
      ((∀s:Sequent. ∀r:Rule.  Q[proof-abort(s;r)] supposing ↑isr(effect <s, r>))
       (∀s:Sequent. ∀r:Rule.
            ∀L:proof-tree(Sequent;Rule;effect) List
              (∀pf∈L.Q[pf])  Q[make-proof-tree(s;r;L)] supposing ||L|| ||outl(effect <s, r>)|| ∈ ℤ 
            supposing ↑isl(effect <s, r>))
       (∀pf:proof-tree(Sequent;Rule;effect). Q[pf]))

Definition: proof_tree_ind
proof_tree_ind(effect;abort;progress;pf) ==
  W_ind(λa,f,q. let s,r in case effect <s, r> of inl(x) => progress mklist(||x||;f) i.(q i)) inr(x) => abort\000C r;pf)

Lemma: proof-tree-induction-ext
[Sequent,Rule:Type].
  ∀effect:(Sequent × Rule) ⟶ (Sequent List?)
    ∀[Q:proof-tree(Sequent;Rule;effect) ⟶ ℙ]
      ((∀s:Sequent. ∀r:Rule.  Q[proof-abort(s;r)] supposing ↑isr(effect <s, r>))
       (∀s:Sequent. ∀r:Rule.
            ∀L:proof-tree(Sequent;Rule;effect) List
              (∀pf∈L.Q[pf])  Q[make-proof-tree(s;r;L)] supposing ||L|| ||outl(effect <s, r>)|| ∈ ℤ 
            supposing ↑isl(effect <s, r>))
       (∀pf:proof-tree(Sequent;Rule;effect). Q[pf]))

Lemma: proof_tree_ind_wf
[Sequent,Rule:Type]. ∀[effect:(Sequent × Rule) ⟶ (Sequent List?)]. ∀[Q:proof-tree(Sequent;Rule;effect) ⟶ ℙ].
[abort:∀s:Sequent. ∀r:Rule.  Q[proof-abort(s;r)] supposing ↑isr(effect <s, r>)].
[progress:∀s:Sequent. ∀r:Rule.
             ∀L:proof-tree(Sequent;Rule;effect) List
               (∀pf∈L.Q[pf])  Q[make-proof-tree(s;r;L)] supposing ||L|| ||outl(effect <s, r>)|| ∈ ℤ 
             supposing ↑isl(effect <s, r>)]. ∀[pf:proof-tree(Sequent;Rule;effect)].
  (proof_tree_ind(effect;abort;progress;pf) ∈ Q[pf])

Definition: correct_proof
correct_proof(Sequent;effect;s;pf) ==
  fix((λcorrect_proof,s,pf. let sr,f pf 
                            in ((fst(sr)) s ∈ Sequent)
                               ∧ case effect sr
                                  of inl(subgoals) =>
                                  ∀i:ℕ||subgoals||. (correct_proof subgoals[i] (f i))
                                  inr(x) =>
                                  False)) 
  
  pf

Lemma: correct_proof-wf
True

Lemma: correct_proof_wf
[Sequent,Rule:Type]. ∀[effect:(Sequent × Rule) ⟶ (Sequent List?)]. ∀[s:Sequent].
[pf:proof-tree(Sequent;Rule;effect)].
  (correct_proof(Sequent;effect;s;pf) ∈ ℙ)

Lemma: sq_stable__correct_proof
[Sequent,Rule:Type].
  ∀effect:(Sequent × Rule) ⟶ (Sequent List?)
    ∀[s:Sequent]. ∀pf:proof-tree(Sequent;Rule;effect). SqStable(correct_proof(Sequent;effect;s;pf))

Definition: proveable
generic definition of provability.
It says that sequent is provable if there is correct proof tree.
proof tree proof-tree is well-founded tree 
(for those we use Martin-Lof's, type, W)
with sequent and rule labeling each node.
The proof is correct for sequent if the sequent labeling the top node is
and for every node, the effect of the given rule on the given sequent
does not fail and the successors of the node
are correct proofs of the generated subgoal sequents.⋅

proveable(Sequent;Rule;effect;s) ==  ∃pf:proof-tree(Sequent;Rule;effect) [correct_proof(Sequent;effect;s;pf)]

Lemma: proveable_wf
[Sequent,Rule:Type]. ∀[effect:(Sequent × Rule) ⟶ (Sequent List?)]. ∀[s:Sequent].
  (proveable(Sequent;Rule;effect;s) ∈ ℙ)

Definition: hint
hint(t) ==  True

Lemma: hint_wf
[t:Type]. ∀[t:t].  (hint(t) ∈ ℙ)

Lemma: test-evd1
[A,B,C,D,E:ℙ].  ((((A ∧ B)  (A ∧ B))  (C ∨ D))  (C  E)  (D  E)  E)

Lemma: test-evd1'
[A,B,C,D,E:ℙ].  ((((A ∧ B)  (A ∧ B))  (C ∨ D))  (C  E)  (D  E)  E)

Lemma: test-evd2
[D:ℙ]. ∀[P,Q:D ⟶ ℙ].  ((∀x:D. ((P x)  (Q x)))  (∃x:D. (P x))  (∃x:D. (Q x)))

Lemma: test-evd3
[D:ℙ]. ∀[P,Q:D ⟶ ℙ].  ((∀x:D. ((P x)  (Q x)))  (∃x:D. (P x))  (∃x:D. (Q x)))

Lemma: test-exists-example1
[D:ℙ]. ∀[P,Q,R:D ⟶ ℙ].
  ((∀x:D. ((R x)  (Q x)))  (∀x:D. ((P x)  ((Q x) ∨ (R x))))  (∃x:D. (P x))  (∃x:D. (Q x)))

Lemma: test-exists-example2
[D:ℙ]. ∀[P,Q,R:D ⟶ ℙ].
  ((∀x:D. ((R x)  (Q x)))  (∀x:D. ((P x)  ((Q x) ∨ (R x))))  (∃x:D. (P x))  (∀x:D. ((R x) ∨ (∃x:D. (Q x)))))

Lemma: test-evd-middle
[P,A:ℙ].  (((P ∨ (P  A))  A)  A)

Lemma: not-not-A-B-example
[A,B,F:ℙ].  ((((A ∧ B) ∨ (A  F) ∨ (B  F))  F)  F)

Lemma: Peirce's-law-iff-xmiddle
[P,B:ℙ].  (((P  B)  P)  P) ⇐⇒ ∀[P,B:ℙ].  (P ∨ (P  B))

Lemma: Peirce-subtype-dneg-elim
(∀[P,B:ℙ].  (((P  B)  P)  P)) ⊆(∀[P:ℙ]. ((¬¬P)  P))

Lemma: double-negation-iff-xmiddle
[A:Type]. (∀[P:ℙ]. (((P  A)  A)  (P ∨ A)) ⇐⇒ ∀[P:ℙ]. (P ∨ (P  A)))

Lemma: no-uniform-xmiddle
¬(∀[P:ℙ]. (P ∨ P)))

Lemma: no-uniform-double-negation-elim
¬(∀[P:ℙ]. ((¬¬P)  P))

Lemma: no-uniform-Peirce's-law
¬(∀[P,B:ℙ].  (((P  B)  P)  P))

Lemma: test-red-black-example
[A,D:Type]. ∀[Red,Black:D ⟶ ℙ]. ∀[R:D ⟶ D ⟶ ℙ].
  ((∀x:D. (Red[x] ∨ Black[x]))
   (∀x,y,z:D.  (R[x;y]  R[y;z]  R[x;z]))
   (∀x:D. (R[x;x]  A))
   (∀x:D. (Red[x]  (∃y:D. (Black[y] ∧ R[x;y]))))
   (∀x:D. (Black[x]  (∃y:D. (Red[y] ∧ R[x;y]))))
   (∃m:D. ((∀x:D. (Red[x]  R[x;m])) ∨ (∀x:D. (Black[x]  R[x;m]))))
   A)

Lemma: test-infinite-domain-example
[A,D:Type]. ∀[R,Eq:D ⟶ D ⟶ ℙ].
  ((∀x,y,z:D.  (R[x;y]  (R[y;z] ∨ Eq[y;z])  R[x;z]))
   (∀x:D. (R[x;x]  A))
   (∀x:D. ∃y:D. R[x;y])
   (∃m:D. ∀x:D. ((Eq[x;m]  A)  R[x;m]))
   A)

Lemma: test-hidden-ap
[A,B,C:Type].  (((A  A)  (B ∨ C))  (C  B)  B)

Lemma: test-example4
[Dom:Type]. ∀[A,B,C,D,P,Q,G:Dom ⟶ ℙ]. ∀[R1,R2:Dom ⟶ Dom ⟶ ℙ].
  ((∀x:Dom. (A[x] ∨ B[x]))
   (∀x:Dom. (C[x] ∨ D[x]))
   (∀x:Dom. ((A[x] ∧ C[x])  (∃y:Dom. (R1[x;y] ∧ P[y]))))
   (∀x:Dom. ((A[x] ∧ D[x])  (∃y:Dom. (R1[x;y] ∧ Q[y]))))
   (∀x:Dom. ((B[x] ∧ C[x])  (∃y:Dom. (R1[x;y] ∧ P[y]))))
   (∀x:Dom. ((B[x] ∧ D[x])  (∃y:Dom. (R1[x;y] ∧ Q[y]))))
   (∀x:Dom. ((A[x] ∧ C[x])  (∃y:Dom. (R2[x;y] ∧ Q[y]))))
   (∀x:Dom. ((A[x] ∧ D[x])  (∃y:Dom. (R2[x;y] ∧ P[y]))))
   (∀x:Dom. ((B[x] ∧ C[x])  (∃y:Dom. (R2[x;y] ∧ Q[y]))))
   (∀x:Dom. ((B[x] ∧ D[x])  (∃y:Dom. (R2[x;y] ∧ P[y]))))
   (∀x,y,z:Dom.  (((R1[x;y] ∧ R2[x;z]) ∧ ((P[y] ∧ Q[z]) ∨ (Q[y] ∧ P[z])))  G[x]))
   (∀x:Dom. G[x]))

Lemma: test-model
[Dom:Type]. ∀[A,B:Dom ⟶ ℙ]. ∀[R:Dom ⟶ Dom ⟶ ℙ].
  ((∀x:Dom. ∃y:Dom. (R[x;y] ∧ (A[y] ∨ B[y])))
   (∀x,y,z:Dom.  ((R[x;y] ∧ R[x;z] ∧ A[y] ∧ A[z])  A[x]))
   (∀x,y,z:Dom.  ((R[x;y] ∧ R[x;z] ∧ B[y] ∧ B[z])  A[x]))
   (∀x:Dom. A[x]))

Lemma: test-model2
[Dom:Type]. ∀[A,B:Dom ⟶ ℙ].
  ((∀x:Dom. ∃y:Dom. (A[y] ∨ B[y]))
   (∀x,y,z:Dom.  ((A[y] ∧ A[z])  A[x]))
   (∀x,y,z:Dom.  ((B[y] ∧ B[z])  A[x]))
   (∀x:Dom. A[x]))

Lemma: test-model3
[Dom:Type]. ∀[A,B:Dom ⟶ ℙ].
  ((∀x:Dom. ∃y:Dom. (A[y] ∨ B[y]))
   (∀x,y,z:Dom.  ((A[y] ∧ A[z])  A[x]))
   (∀x,y,z:Dom.  ((B[y] ∧ B[z])  A[x]))
   (∀x:Dom. A[x]))

Lemma: finite-double-negation-shift
[A:ℙ]. ∀[B:ℕ ⟶ ℙ].  ∀n:ℕ((∀i:ℕn. (((B i)  A)  A))  ((∀i:ℕn. (B i))  A)  A)

Lemma: finite-double-negation-shift-extract
[A:ℙ]. ∀[B:ℕ ⟶ ℙ].  ∀n:ℕ((∀i:ℕn. (((B i)  A)  A))  ((∀i:ℕn. (B i))  A)  A)

Lemma: finite-double-negation-shift2
[A:ℙ]. ∀n:ℕ. ∀[B:ℕn ⟶ ℙ]. ((∀i:ℕn. (((B i)  A)  A))  ((∀i:ℕn. (B i))  A)  A)

Definition: double-negation-shift
DNSi.A[i] ==  (∀i:ℕ(¬¬A[i]))  (¬¬(∀i:ℕA[i]))

Lemma: double-negation-shift_wf
A:ℕ ⟶ ℙ(DNSi.A[i] ∈ ℙ)

Definition: generalized-markov-principle
GMPi.A[i]) ==  (∀i:ℕA[i]))  (∃n:ℕA[n]))

Lemma: generalized-markov-principle_wf
[A:ℕ ⟶ ℙ]. (GMPi.A[i]) ∈ ℙ)

Lemma: DNS-iff-not-not-GMP
[A:ℕ ⟶ ℙ]. (DNSi.A[i] ⇐⇒ ¬¬GMPi.A[i]))

Lemma: finite-partition-property
k:ℕ. ∀f:ℕ ⟶ ℕk.  (¬¬(∃i:ℕk. ∀n:ℕ(¬¬(∃m:ℕ(n < m ∧ ((f m) i ∈ ℤ))))))

Definition: known_false
known_false(t) ==  t

Definition: known_not_false
known_not_false(t) ==  t

Lemma: exists-simp
[T:Type]. ∀a:T. ∀P:T ⟶ ℙ.  ((∀x:T. (P[x]  (x a ∈ T)))  (∃x:T. P[x] ⇐⇒ P[a]))

Lemma: exists-simp-test
T:Type. ∀P:T ⟶ ℙ'. ∀a:T.  (∃x:T. (P[x] ∧ (x a ∈ T)) ⇐⇒ P[a])

Lemma: assert-equal-test
[A,B:Type]. ∀[f:A ⟶ B]. ∀[a1,a2:A].  (f a1) (f a2) ∈ supposing a1 a2 ∈ A

Lemma: isect-subtype-1
[F:Type ⟶ Type]. ∀[A:Type].  ((⋂A:Type. F[A]) ⊆F[A])

Lemma: isect-subtype-2
[G:Type ⟶ Type]. ∀[A,B:Type]. ∀[F:G[A] ⟶ G[B] ⟶ Type]. ∀[X:G[A]]. ∀[Y:G[B]].
  ((⋂X:G[A]. ⋂Y:G[B].  F[X;Y]) ⊆F[X;Y])

Lemma: ppcc-test
[a,b,c:ℤ].  ((a c) (c c) ∈ ℤsupposing (((b c) (c c) ∈ ℤand (a b ∈ ℤ))

Lemma: ppcc-test2
[T:Type]
  ∀f:T ⟶ T
    ∀[Q:T ⟶ ℙ]. ∀[P:T ⟶ T ⟶ ℙ].  ((∀z:T. (Q[z]  P[z;f[z]]))  (∀x,y:T.  Q[x]  P[x;y] supposing f[x] ∈ T))

Lemma: ppcc-test3
[T:Type]
  ∀f:T ⟶ T
    ∀[Q,P:T ⟶ T ⟶ ℙ].
      ((∀a,b:T.  (Q[f[a];b] ⇐⇒ P[a;f[b]]))
       Trans(T;a,b.P[a;b])
       (∀a,b,c,d,e,x:T.  (P[a;c]  Q[d;b]  P[a;e]) supposing ((f[b] e ∈ T) and (f[x] d ∈ T) and (c x ∈ T))))

Lemma: append_is_nil_test
[a,b:Top List].
  ((([] (a b) ∈ (Top List)) ∨ ((a b) [] ∈ (Top List)))  ((b [] ∈ (Top List)) ∧ (a [] ∈ (Top List))))

Lemma: ppcc-test4
f:ℤ ⟶ ℤ
  ∀[Q,P:ℤ ⟶ ℤ ⟶ ℙ].
    ((∀a,b:ℤ.  (Q[f[a];b] ⇐⇒ P[a;f[b]]))
     Trans(ℤ;a,b.P[a;b])
     (∀z,w:𝔹.
          ∀a,b,c,d,e:ℤ. ∀x:{z:ℤf[(z 1) 1] d ∈ ℤ.
            ((f[(b 2) 2] if (w ∨bz) ∧b (c =z x) then else f[b] fi  ∈ ℤc∧ P[a;a 1])
             P[(a a) a;c]
             (∀x:ℤ × ℤQ[fst(x);b]  P[snd(x);e 1] supposing = <d, a> ∈ (ℤ × ℤ)) 
            supposing <<c, a>e> = <<x, a>e> ∈ (ℤ × ℤ × ℤ
          supposing (↑z) ∧ (¬↑w)))

Lemma: ppcc-problem
ff tt supposing (inl tt) (inl ff) ∈ (𝔹?)

Lemma: spread_wf
[A,B,C:Type]. ∀[p:A × B]. ∀[F:A ⟶ B ⟶ C].  (let x,y in F[x;y] ∈ C)

Lemma: list_ind-as-fix
[L:Top List]. ∀[x,F:Top].
  (rec-case(L) of
   [] => x
   b::bs =>
    r.F[r;b;bs] fix((λR,L. if null(L) then else (R tl(L)) hd(L) tl(L) fi )) L)

Lemma: spread-wf
[A:Type]. ∀[B:A ⟶ Type]. ∀[C:a:A ⟶ b:B[a] ⟶ Type]. ∀[p:a:A × B[a]]. ∀[F:a:A ⟶ b:B[a] ⟶ C[a;b]].
  (let x,y 
   in F[x;y] ∈ C[fst(p);snd(p)])

Lemma: fun_with_inv_is_bij2
[A,B:Type].  ∀f:A ⟶ B. ((∃g:B ⟶ A. InvFuns(A;B;f;g))  Bij(A;B;f))

Lemma: decide_wf
[T1,T2,T3:Type]. ∀[x:T1 T2]. ∀[l:T1 ⟶ T3]. ∀[r:T2 ⟶ T3].  (case of inl(a) => l[a] inr(b) => r[b] ∈ T3)

Lemma: squash-assert
[b:𝔹]. uiff(↓↑b;↑b)

Definition: church-Nat
cNat ==  ⋂X:Type. (X ⟶ (X ⟶ X) ⟶ X)

Lemma: church-Nat_wf
cNat ∈ 𝕌'

Definition: church-zero
cZ ==  λz,s. z

Lemma: church-zero_wf
cZ ∈ cNat

Definition: church-succ
cS ==  λx,z,s. (s (x s))

Lemma: church-succ_wf
cS ∈ cNat ⟶ cNat

Definition: church-inductive
church-inductive{i:l}(x) ==  ∀[P:cNat ⟶ ℙ]. ((P cZ)  (∀[y:cNat]. ((P y)  (P (cS y))))  (P x))

Lemma: church-inductive_wf
[x:cNat]. (church-inductive{i:l}(x) ∈ ℙ')

Lemma: church-iZ
cZ ∈ church-inductive{i:l}(cZ)

Lemma: church-iS
[x:cNat]. (cS ∈ church-inductive{i:l}(x) ⟶ church-inductive{i:l}(cS x))

Definition: indNat
INat ==  x:cNat ⋂ church-inductive{i:l}(x)

Lemma: indNat_wf
INat ∈ 𝕌'

Lemma: church-zero_wf2
cZ ∈ INat

Lemma: church-succ_wf2
cS ∈ INat ⟶ INat

Definition: church-true
church-true() ==  λx,y. x

Lemma: church-true_wf
[T:Type]. (church-true() ∈ T ⟶ Top ⟶ T)

Definition: church-false
church-false() ==  λx,y. y

Lemma: church-false_wf
[T:Type]. (church-false() ∈ Top ⟶ T ⟶ T)

Definition: church-ite
church-ite() ==  λp,a,b. (p b)

Lemma: church-ite_wf
[A,B,C:Type].  (church-ite() ∈ (A ⟶ B ⟶ C) ⟶ A ⟶ B ⟶ C)

Definition: church-pair
church-pair() ==  λx,y,f. (f y)

Lemma: church-pair_wf
[A,B,C:Type].  (church-pair() ∈ A ⟶ B ⟶ (A ⟶ B ⟶ C) ⟶ C)

Definition: church-fst
church-fst() ==  λp.(p church-true())

Lemma: church-fst_wf
[A,T:Type].  (church-fst() ∈ ((T ⟶ Top ⟶ T) ⟶ A) ⟶ A)

Definition: church-snd
church-snd() ==  λp.(p church-false())

Lemma: church-snd_wf
[A,T:Type].  (church-snd() ∈ ((Top ⟶ T ⟶ T) ⟶ A) ⟶ A)

Definition: church-nil
church-nil() ==  λx.church-true()

Lemma: church-nil_wf
[T:Type]. (church-nil() ∈ Top ⟶ T ⟶ Top ⟶ T)

Definition: church-null
church-null() ==  λp.(p x,y. church-false()))

Lemma: church-null_wf
[A,T:Type].  (church-null() ∈ ((Top ⟶ Top ⟶ Top ⟶ T ⟶ T) ⟶ A) ⟶ A)

Lemma: church_ite_true_lemma
y,x:Top.  (church-ite() church-true() x)

Lemma: church_ite_false_lemma
y,x:Top.  (church-ite() church-false() y)

Lemma: church_fst_lemma
y,x:Top.  (church-fst() (church-pair() y) x)

Lemma: church_snd_lemma
y,x:Top.  (church-snd() (church-pair() y) y)

Lemma: church_null_nil_lemma
church-null() church-nil() church-true()

Lemma: church_null_pair_lemma
y,x:Top.  (church-null() (church-pair() y) church-false())

Lemma: exists-product1
[A,B:Type].  ∀P:(A × B) ⟶ ℙ'. {∃x:A × B. P[x] ⇐⇒ ∃a:A. ∃b:B. P[<a, b>]}

Lemma: equal-product1
[A:Type]. ∀[B:A ⟶ Type]. ∀[a1,a2:A]. ∀[b1:B[a1]]. ∀[b2:B[a2]].
  {<a1, b1> = <a2, b2> ∈ (a:A × B[a]) ⇐⇒ (a1 a2 ∈ A) ∧ (b1 b2 ∈ B[a1])}

Lemma: exists-product2
[A,B,C:Type].  ∀P:(A × B × C) ⟶ ℙ'. {∃x:A × B × C. P[x] ⇐⇒ ∃a:A. ∃b:B. ∃c:C. P[<a, b, c>]}

Lemma: exists-product3
[A,B,C,D:Type].  ∀P:(A × B × C × D) ⟶ ℙ'. {∃x:A × B × C × D. P[x] ⇐⇒ ∃a:A. ∃b:B. ∃c:C. ∃d:D. P[<a, b, c, d>]}

Lemma: exists-product4
[A,B,C,D,E:Type].
  ∀P:(A × B × C × D × E) ⟶ ℙ'. {∃x:A × B × C × D × E. P[x] ⇐⇒ ∃a:A. ∃b:B. ∃c:C. ∃d:D. ∃e:E. P[<a, b, c, d, e>]}

Lemma: exists-product5
[A,B,C,D,E,F:Type].
  ∀P:(A × B × C × D × E × F) ⟶ ℙ'
    {∃x:A × B × C × D × E × F. P[x] ⇐⇒ ∃a:A. ∃b:B. ∃c:C. ∃d:D. ∃e:E. ∃f:F. P[<a, b, c, d, e, f>]}

Lemma: exists-elim
[T:Type]. ∀[P:T ⟶ ℙ'].  ∀a:T. ((∀x:T. (P[x]  (x a ∈ T)))  {∃x:T. P[x] ⇐⇒ P[a]})

Definition: exists_uni
!x:T. P[x] ==  ∃x:T. (P[x] ∧ (∀y:T. (P[y]  (x y ∈ T))))

Lemma: exists_uni_wf
[T:Type]. ∀[P:T ⟶ ℙ].  (∃!x:T. P[x] ∈ ℙ)

Lemma: test-eq-product
[A,B1,B2,C,D:Type]. ∀[a1,a2:A]. ∀[b1:B1]. ∀[b1':B2]. ∀[b2:B1]. ∀[b2':B2]. ∀[c1,c2:C]. ∀[d1,d2:D].
  (<a1, <b1, b1'>c1, d1> = <a2, <b2, b2'>c2, d2> ∈ (a:A × b:B1 × B2 × c:C × D)
  ⇐⇒ (a1 a2 ∈ A) ∧ ((b1 b2 ∈ B1) ∧ (b1' b2' ∈ B2)) ∧ (c1 c2 ∈ C) ∧ (d1 d2 ∈ D))

Lemma: trivial-eq
[T:Type]. ∀[x:T].  {x x ∈ ⇐⇒ True}

Lemma: trivial-int-eq1
[x,y:ℤ].  (((x y) x) ∧ (y (x y) x) ∧ ((x y) x) ∧ (y x))

Lemma: trivial-int-eq-test
P:ℤ ⟶ ℙ
  ∀[x,y:ℤ].
    (((P ((x y) y))  (P x))
    ∧ ((P (y (x y)))  (P x))
    ∧ ((P ((x y) y))  (P x))
    ∧ ((P (y x))  (P x)))

Lemma: test-exists-elim
[A:Type]. ∀[Q:A ⟶ A ⟶ ℙ].  ∀x,b:A.  (∃z:A × A. (Q[fst(z);snd(z)] ∧ (z = <x, b> ∈ (A × A))) ⇐⇒ Q[x;b] ∧ (x x ∈ A))

Lemma: sq_stable__assert
[b:𝔹]. SqStable(↑b)

Definition: gensearch
gensearch(f;g;x) ==
  fix((λgensearch,x. case of inl(a) => inl inr(b) => case of inl(y) => gensearch inr(z) => inr )) x

Lemma: gensearch_wf
[A,B:Type]. ∀[r:A ⟶ ℕ]. ∀[f:A ⟶ (B Top)]. ∀[g:A ⟶ (A Top)].
  ∀[a:A]. (gensearch(f;g;a) ∈ Top) supposing ∀a:A. ((↑isl(g a))  outl(g a) < a)

Definition: typed
x:T ==  x

Lemma: typed_wf
[T:Type]. ∀[x:T].  (x:T ∈ T)

Definition: type-mismatch
type-mismatch(f; A; x; B) ==  ⋅

Lemma: type-mismatch_wf
[T,T1,T2,T3:Type]. ∀[f:T]. ∀[A:T1]. ∀[x:T2]. ∀[B:T3].  (type-mismatch(f; A; x; B) ∈ Unit)

Lemma: all-union
[A,B:Type].  ∀P:(A B) ⟶ ℙ(∀x:A B. P[x] ⇐⇒ (∀a:A. P[inl a]) ∧ (∀b:B. P[inr ]))

Lemma: exists-union
[A,B:Type].  ∀P:(A B) ⟶ ℙ(∃x:A B. P[x] ⇐⇒ (∃a:A. P[inl a]) ∨ (∃b:B. P[inr ]))

Lemma: uall-union
[A,B:Type].  ∀P:(A B) ⟶ ℙ((∀[x:A B]. P[x])  ((∀[a:A]. P[inl a]) ∧ (∀[b:B]. P[inr ])))

Lemma: uall-unit
P:Unit ⟶ ℙ(∀[x:Unit]. P[x] ⇐⇒ P[⋅])

Lemma: all-unit
P:Unit ⟶ ℙ(∀x:Unit. P[x] ⇐⇒ P[⋅])

Lemma: exists-unit
P:Unit ⟶ ℙ(∃x:Unit. P[x] ⇐⇒ P[⋅])

Lemma: first_index_bounds
[T:Type]. ∀[P:T ⟶ 𝔹]. ∀[L:T List].  (index-of-first in L.P[x] ∈ ℕ||L|| 1)

Lemma: first_index-positive
[T:Type]. ∀P:T ⟶ 𝔹. ∀L:T List.  (0 < index-of-first in L.P[x] ⇐⇒ (∃x∈L. ↑P[x]))

Lemma: decide-trivial
[x:Top Top]. ∀[y:Top].  (case of inl(z) => inr(z) => y)

Lemma: spread-trivial
[x:Top × Top]. ∀[y:Top].  (let a,b in y)

Lemma: rec-case-map-sq
[f:Top]. ∀[l:Top List].  (rec-case(l) of [] => [] a::as => r.[f[a] r] map(λa.f[a];l))

Definition: prior
prior(n;f) ==
  letrec rec(n)=if (n =z 0) then inr ⋅  else eval in case of inl(x) => inl <m, x> inr(z) => rec fi  in
   rec 
  n

Lemma: prior_wf
[T:Type]. ∀[f:ℕ ⟶ (T Top)]. ∀[n:ℕ].  (prior(n;f) ∈ ℕn × T?)

Lemma: prior-cases
[T:Type]
  ∀f:ℕ ⟶ (T Top). ∀n:ℕ.
    case prior(n;f)
     of inl(p) =>
     let m,x 
     in ((f m) (inl x) ∈ (T Top)) ∧ (∀k:{m 1..n-}. (¬↑isl(f k)))
     inr(q) =>
     ∀k:ℕn. (¬↑isl(f k))

Lemma: isl-prior
[T:Type]
  ∀f:ℕ ⟶ (T Top). ∀n:ℕ.
    let m,x outl(prior(n;f)) 
    in ((f m) (inl x) ∈ (T Top)) ∧ (∀k:{m 1..n-}. (¬↑isl(f k))) 
    supposing ↑isl(prior(n;f))

Lemma: isl-prior-iff
[T:Type]. ∀f:ℕ ⟶ (T Top). ∀n:ℕ.  (↑isl(prior(n;f)) ⇐⇒ ∃k:ℕn. (↑isl(f k)))

Definition: longest-prefix
longest-prefix(P;L) ==
  fix((λlongest-prefix,P,L. if null(L)
                           then []
                           else let longest-prefix L'.(P [hd(L) L'])) tl(L) in
                                    if null(p)
                                    then if null(tl(L)) then []
                                         if [hd(L)] then [hd(L)]
                                         else []
                                         fi 
                                    else [hd(L) p]
                                    fi 
                           fi )) 
  
  L

Lemma: longest-prefix_wf
[T:Type]. ∀[L:T List]. ∀[P:T List+ ⟶ 𝔹].  (longest-prefix(P;L) ∈ List)

Lemma: eq_bool_tt
[b:𝔹]. =b tt b

Lemma: eq_bool_ff
[b:𝔹]. =b ff = ¬bb

Definition: ite
ite(b;x;y) ==  if then else fi 

Lemma: ite_wf
[b:𝔹]. ∀[T:Type]. ∀[x:T supposing ↑b]. ∀[y:T supposing ¬↑b].  (ite(b;x;y) ∈ T)

Definition: nfoldunion
nfoldunion(n) ==  primrec(n;Top;λn,T. (Top T))

Lemma: nfoldunion_wf
[n:ℕ]. (nfoldunion(n) ∈ Type)

Lemma: band-to-and
[a,b:𝔹].  {(a tt) ∧ (b tt)} supposing a ∧b tt

Lemma: bor-to-and
[a,b:𝔹].  {(a ff) ∧ (b ff)} supposing a ∨bff

Lemma: bnot-tt
[a:𝔹]. ff supposing ¬btt

Lemma: bnot-ff
[a:𝔹]. tt supposing ¬bff

Definition: bool-decider
bool-decider(b) ==  TERMOF{decidable__assert:o, 1:l} b

Lemma: bool-decider_wf
[b:𝔹]. (bool-decider(b) ∈ Dec(↑b))

Lemma: isl-ite
[x:𝔹]. ∀[a,b:Top Top].  (isl(if then else fi (isl(a) ∧b x) ∨b(isl(b) ∧b bx)))

Definition: projn
projn(n;x) ==  fix((λprojn,n,x. if (n =z 1) then fst(x) else projn (n 1) (snd(x)) fi )) x

Definition: invert-union
invert-union(x) ==  case of inl(a) => inr a  inr(a) => inl a

Lemma: invert-union_wf
[A,B:Type]. ∀[x:A B].  (invert-union(x) ∈ A)

Definition: invertunion
invertunion(x) ==  case of inl(a) => inr a  inr(a) => inl a

Lemma: invertunion_wf
[A,B:Type]. ∀[x:A B].  (invertunion(x) ∈ A)

Lemma: not-false
uiff(¬False;True)

Lemma: not-true
uiff(¬True;False)

Lemma: not-assert
[b:𝔹]. uiff(¬↑b;b ff)

Lemma: not-not-assert
[b:𝔹]. uiff(¬¬↑b;↑b)

Lemma: equal-bnot
[x,y:𝔹].  uiff(x = ¬by;¬y)

Lemma: all-nsub2
[P:ℕ2 ⟶ ℙ]. (∀x:ℕ2. P[x] ⇐⇒ P[0] ∧ P[1])

Lemma: inconsistent-bool-eq
uiff(tt ff;False)

Lemma: inconsistent-bool-eq2
uiff(ff tt;False)

Lemma: inconsistent-bool-eq3
[b:𝔹]. uiff(b = ¬bb;False)

Lemma: inconsistent-bool-eq4
[b:𝔹]. uiff(¬bb;False)

Lemma: test-rewrite-dcdr
[P:ℙ]. ∀d:Dec(P). ((↑[d]b) ∨ (¬↑[d]b))

Lemma: bool-to-dcdr-aux
[A:Type]. ∀f:A ⟶ 𝔹. ∀x:A.  Dec(f tt)

Lemma: bool-to-neg-dcdr-aux
[A:Type]. ∀f:A ⟶ 𝔹. ∀x:A.  Dec(f ff)

Definition: bool-to-neg-dcdr
{f}q ==  TERMOF{bool-to-neg-dcdr-aux:o, 1:l, i:l} f

Lemma: bool-to-neg-dcdr_wf
[A:Type]. ∀[f:A ⟶ 𝔹]. ∀[x:A].  ({f}q x ∈ Dec(f ff))

Definition: bool-to-dcdr
{f}b ==  TERMOF{bool-to-dcdr-aux:o, 1:l, i:l} f

Lemma: bool-to-dcdr_wf
[A:Type]. ∀[f:A ⟶ 𝔹]. ∀[x:A].  ({f}b x ∈ Dec(f tt))

Definition: branch
if p:P then A[p] else fi  ==  case of inl(p) => A[p] inr(x) => B

Lemma: branch_wf
[P:ℙ]. ∀[d:Dec(P)]. ∀[T:Type]. ∀[A:P ⟶ T]. ∀[B:T].  (if p:P then A[p] else fi  ∈ T)

Lemma: branch_wf2
[P:ℙ]. ∀[d:Dec(P)]. ∀[T:Type]. ∀[A:P ⟶ T]. ∀[B:⋂q:¬P. T].  (if p:P then A[p] else fi  ∈ T)

Lemma: branch-ifthenelse
[b:𝔹]. ∀[P,Q:Top].  (if x:↑then else fi  if then else fi )

Lemma: decidable-filter
[T:Type]
  ∀L:T List
    ∀[P:{x:T| (x ∈ L)}  ⟶ ℙ]. ((∀x∈L.Dec(P[x]))  (∃L':T List. (L' ⊆ L ∧ (∀x:T. ((x ∈ L') ⇐⇒ (x ∈ L) ∧ P[x])))))

Definition: order-preserving
order-preserving(A;B;a1,a2.R1[a1; a2];b1,b2.R2[b1; b2];f) ==  ∀a1,a2:A.  (R1[a1; a2]  R2[f a1; a2])

Lemma: order-preserving_wf
[A,B:Type]. ∀[R1:A ⟶ A ⟶ ℙ]. ∀[R2:B ⟶ B ⟶ ℙ]. ∀[f:A ⟶ B].
  (order-preserving(A;B;a1,a2.R1[a1;a2];b1,b2.R2[b1;b2];f) ∈ ℙ)

Definition: can-apply
can-apply(f;x) ==  isl(f x)

Lemma: can-apply_wf
[A:Type]. ∀[f:A ⟶ (Top Top)]. ∀[x:A].  (can-apply(f;x) ∈ 𝔹)

Definition: do-apply
do-apply(f;x) ==  outl(f x)

Lemma: do-apply_wf
[A,B:Type]. ∀[f:A ⟶ (B Top)]. ∀[x:A].  do-apply(f;x) ∈ supposing ↑can-apply(f;x)

Definition: apply?
f(x)?d ==  if isl(f x) then outl(f x) else fi 

Lemma: apply?_wf
[A,T:Type]. ∀[f:A ⟶ (T Top)]. ∀[x:A]. ∀[d:T].  (f(x)?d ∈ T)

Lemma: inl-do-apply
[A,B:Type]. ∀[f:A ⟶ (B Top)]. ∀[x:A].  inl do-apply(f;x) supposing ↑can-apply(f;x)

Definition: p-compose
==  λx.if can-apply(g;x) then do-apply(g;x) else fi 

Lemma: p-compose_wf
[A,B,C:Type]. ∀[g:A ⟶ (B Top)]. ∀[f:B ⟶ (C Top)].  (f g ∈ A ⟶ (C Top))

Definition: p-compose'
o' ==  λx.if can-apply(g;x) then inl (f do-apply(g;x)) else fi 

Lemma: p-compose'_wf
[A,B,C:Type]. ∀[g:A ⟶ (B Top)]. ∀[f:A ⟶ B ⟶ C].  (f o' g ∈ A ⟶ (C Top))

Definition: p-lift
p-lift(d;f) ==  λx.case of inl(a) => inl (f x) inr(a) => inr 

Lemma: p-lift_wf
[A,B:Type]. ∀[P:A ⟶ ℙ]. ∀[d:x:A ⟶ Dec(P[x])]. ∀[f:{x:A| P[x]}  ⟶ B].  (p-lift(d;f) ∈ A ⟶ (B Top))

Lemma: can-apply-p-lift
[A,B:Type]. ∀[P:A ⟶ ℙ].  ∀d:x:A ⟶ Dec(P[x]). ∀f:{x:A| P[x]}  ⟶ B. ∀x:A.  (↑can-apply(p-lift(d;f);x) ⇐⇒ P[x])

Lemma: do-apply-p-lift
[A,B:Type]. ∀[P:A ⟶ ℙ]. ∀[d:x:A ⟶ Dec(P[x])]. ∀[f:{x:A| P[x]}  ⟶ B]. ∀[x:A].
  do-apply(p-lift(d;f);x) (f x) ∈ supposing ↑can-apply(p-lift(d;f);x)

Lemma: can-apply-compose-sq
[A,B,C:Type]. ∀[g:A ⟶ (B Top)]. ∀[f:B ⟶ (C Top)]. ∀[x:A].
  (can-apply(f g;x) can-apply(g;x) ∧b can-apply(f;do-apply(g;x)))

Lemma: can-apply-compose
[A,B,C:Type]. ∀[g:A ⟶ (B Top)]. ∀[f:B ⟶ (C Top)]. ∀[x:A].
  {(↑can-apply(g;x)) ∧ (↑can-apply(f;do-apply(g;x)))} supposing ↑can-apply(f g;x)

Lemma: can-apply-compose-iff
[A,B,C:Type]. ∀[g:A ⟶ (B Top)]. ∀[f:B ⟶ (C Top)]. ∀[x:A].
  uiff(↑can-apply(f g;x);(↑can-apply(g;x)) ∧ (↑can-apply(f;do-apply(g;x))))

Lemma: do-apply-compose
[A,B,C:Type]. ∀[g:A ⟶ (B Top)]. ∀[f:B ⟶ (C Top)]. ∀[x:A].
  do-apply(f g;x) do-apply(f;do-apply(g;x)) supposing ↑can-apply(f g;x)

Lemma: can-apply-compose'
[A,B,C:Type]. ∀[g:A ⟶ (B Top)]. ∀[f:A ⟶ B ⟶ C]. ∀[x:A].  (can-apply(f o' g;x) can-apply(g;x))

Lemma: do-apply-compose'
[A,B,C:Type]. ∀[g:A ⟶ (B Top)]. ∀[f:A ⟶ B ⟶ C]. ∀[x:A].
  do-apply(f o' g;x) do-apply(g;x) supposing ↑can-apply(f o' g;x)

Definition: p-id
p-id() ==  λx.(inl x)

Lemma: p-id_wf
[T:Type]. (p-id() ∈ T ⟶ (T Top))

Lemma: p-compose-id
[A,B:Type]. ∀[f:A ⟶ (B Top)].  (f p-id() f ∈ (A ⟶ (B Top)))

Lemma: p-id-compose
[A,B:Type]. ∀[f:A ⟶ (B Top)].  (p-id() f ∈ (A ⟶ (B Top)))

Lemma: p-compose-associative
[A,B,C,D:Type]. ∀[h:A ⟶ (B Top)]. ∀[g:B ⟶ (C Top)]. ∀[f:C ⟶ (D Top)].
  (f h ∈ (A ⟶ (D Top)))

Definition: p-first
p-first(L) ==
  λx.accumulate (with value and list item f):
      case of inl(z) => inr(z) => x
     over list:
       L
     with starting value:
      inr ⋅ )

Lemma: p-first_wf
[A,B:Type]. ∀[L:(A ⟶ (B Top)) List].  (p-first(L) ∈ A ⟶ (B Top))

Lemma: p-first-singleton
[A,B:Type]. ∀[f:A ⟶ (B Top)].  (p-first([f]) f ∈ (A ⟶ (B Top)))

Lemma: p_first_nil_lemma
x:Top. (can-apply(p-first([]);x) ff)

Definition: p-conditional
[f?g] ==  λx.if can-apply(f;x) then else fi 

Lemma: p-conditional_wf
[A,B:Type]. ∀[f,g:A ⟶ (B Top)].  ([f?g] ∈ A ⟶ (B Top))

Lemma: p-conditional-domain
[A,B:Type].  ∀f,g:A ⟶ (B Top). ∀x:A.  (↑can-apply([f?g];x) ⇐⇒ (↑can-apply(f;x)) ∨ (↑can-apply(g;x)))

Lemma: p-conditional-to-p-first
[A,B:Type]. ∀[f,g:A ⟶ (B Top)].  ([f?g] p-first([f; g]) ∈ (A ⟶ (B Top)))

Definition: p-filter
p-filter(f) ==  λx.case of inl(p) => inl inr(p) => inr 

Lemma: p-filter_wf
[T:Type]. ∀[P:T ⟶ ℙ]. ∀[f:∀x:T. Dec(P[x])].  (p-filter(f) ∈ T ⟶ (T Top))

Definition: p-co-filter
p-co-filter(f) ==  λx.case of inl(p) => inr p  inr(p) => inl x

Lemma: p-co-filter_wf
[T:Type]. ∀[P:T ⟶ ℙ]. ∀[f:∀x:T. Dec(P[x])].  (p-co-filter(f) ∈ T ⟶ (T Top))

Lemma: can-apply-p-filter
[T:Type]. ∀[P:T ⟶ ℙ].  ∀f:∀x:T. Dec(P[x]). ∀x:T.  (↑can-apply(p-filter(f);x) ⇐⇒ P[x])

Lemma: can-apply-p-co-filter
[T:Type]. ∀[P:T ⟶ ℙ]. ∀[f:∀x:T. Dec(P[x])]. ∀[x:T].  uiff(↑can-apply(p-co-filter(f);x);¬P[x])

Lemma: do-apply-p-filter
[T:Type]. ∀[P:T ⟶ ℙ]. ∀[f:∀x:T. Dec(P[x])]. ∀[x:T].
  do-apply(p-filter(f);x) x ∈ supposing ↑can-apply(p-filter(f);x)

Lemma: do-apply-p-co-filter
[T:Type]. ∀[P:T ⟶ ℙ]. ∀[f:∀x:T. Dec(P[x])]. ∀[x:T].
  do-apply(p-co-filter(f);x) x ∈ supposing ↑can-apply(p-co-filter(f);x)

Definition: p-restrict
p-restrict(f;p) ==  p-filter(p)

Lemma: p-restrict_wf
[A,B:Type]. ∀[f:A ⟶ (B Top)]. ∀[P:A ⟶ ℙ]. ∀[p:∀x:A. Dec(P[x])].  (p-restrict(f;p) ∈ A ⟶ (B Top))

Definition: p-co-restrict
p-co-restrict(f;p) ==  p-co-filter(p)

Lemma: p-co-restrict_wf
[A,B:Type]. ∀[f:A ⟶ (B Top)]. ∀[P:A ⟶ ℙ]. ∀[p:∀x:A. Dec(P[x])].  (p-co-restrict(f;p) ∈ A ⟶ (B Top))

Lemma: can-apply-p-restrict
[A,B:Type].
  ∀f:A ⟶ (B Top)
    ∀[P:A ⟶ ℙ]. ∀p:∀x:A. Dec(P[x]). ∀x:A.  (↑can-apply(p-restrict(f;p);x) ⇐⇒ (↑can-apply(f;x)) ∧ P[x])

Lemma: can-apply-p-co-restrict
[A,B:Type]. ∀[f:A ⟶ (B Top)]. ∀[P:A ⟶ ℙ]. ∀[p:∀x:A. Dec(P[x])]. ∀[x:A].
  uiff(↑can-apply(p-co-restrict(f;p);x);(↑can-apply(f;x)) ∧ P[x]))

Lemma: do-apply-p-restrict
[A,B:Type]. ∀[f:A ⟶ (B Top)]. ∀[P:A ⟶ ℙ]. ∀[p:∀x:A. Dec(P[x])]. ∀[x:A].
  do-apply(p-restrict(f;p);x) do-apply(f;x) ∈ supposing ↑can-apply(p-restrict(f;p);x)

Lemma: do-apply-p-co-restrict
[A,B:Type]. ∀[f:A ⟶ (B Top)]. ∀[P:A ⟶ ℙ]. ∀[p:∀x:A. Dec(P[x])]. ∀[x:A].
  do-apply(p-co-restrict(f;p);x) do-apply(f;x) ∈ supposing ↑can-apply(p-co-restrict(f;p);x)

Definition: p-fun-exp
f^n ==  primrec(n;p-id();λi,g. g)

Lemma: p-fun-exp_wf
[A:Type]. ∀[f:A ⟶ (A Top)]. ∀[n:ℕ].  (f^n ∈ A ⟶ (A Top))

Lemma: p-fun-exp-one
[A:Type]. ∀[f:A ⟶ (A Top)].  (f^1 f ∈ (A ⟶ (A Top)))

Lemma: p-fun-exp-compose
[T:Type]. ∀[n:ℕ]. ∀[h,f:T ⟶ (T Top)].  (f^n primrec(n;h;λi,g. g) ∈ (T ⟶ (T Top)))

Lemma: p-fun-exp-add
[T:Type]. ∀[n,m:ℕ]. ∀[f:T ⟶ (T Top)].  (f^n f^n f^m ∈ (T ⟶ (T Top)))

Lemma: can-apply-fun-exp-add
[A:Type]. ∀[n,m:ℕ]. ∀[f:A ⟶ (A Top)]. ∀[x:A].
  {(↑can-apply(f^m;x)) ∧ (↑can-apply(f^n;do-apply(f^m;x))) ∧ (do-apply(f^n m;x) do-apply(f^n;do-apply(f^m;x)) ∈ A)} 
  supposing ↑can-apply(f^n m;x)

Lemma: can-apply-fun-exp-add-iff
[A:Type]. ∀[n,m:ℕ]. ∀[f:A ⟶ (A Top)]. ∀[x:A].
  uiff(↑can-apply(f^n m;x);(↑can-apply(f^m;x)) ∧ (↑can-apply(f^n;do-apply(f^m;x))))

Lemma: p-fun-exp-add1-sq
[A:Type]. ∀[f:A ⟶ (A Top)]. ∀[x:A]. ∀[n:ℕ].  f^n f^n do-apply(f;x) supposing ↑can-apply(f;x)

Lemma: can-apply-fun-exp
[A:Type]. ∀[f:A ⟶ (A Top)]. ∀[n:ℕ]. ∀[y:A].  ∀[m:ℕ]. ↑can-apply(f^m;y) supposing m ≤ supposing ↑can-apply(f^n;y)

Lemma: p-fun-exp-add-sq
[A:Type]. ∀[f:A ⟶ (A Top)]. ∀[x:A]. ∀[m,n:ℕ].  f^n f^n do-apply(f^m;x) supposing ↑can-apply(f^m;x)

Definition: p-mu
p-mu(P;x) ==  case of inl(n) => (↑(P n)) ∧ (∀i:ℕn. (¬↑(P i))) inr(z) => ∀i:ℕ(¬↑(P i))

Lemma: p-mu_wf
[P:ℕ ⟶ 𝔹]. ∀[x:ℕ Top].  (p-mu(P;x) ∈ ℙ)

Lemma: p-mu-exists
P:ℕ ⟶ 𝔹(Dec(∃n:ℕ(↑(P n)))  (∃x:ℕ Top. p-mu(P;x)))

Lemma: p-mu-decider
[A:Type]. ∀P:A ⟶ ℕ ⟶ 𝔹((∀x:A. Dec(∃n:ℕ(↑(P n))))  (∀x:A. ∃y:ℕ Top. p-mu(P x;y)))

Definition: mu'
mu'(P) ==  λx.(fst((TERMOF{p-mu-decider:o, 1:l, 1:l} x)))

Lemma: mu'_wf
[A:Type]. ∀[P:A ⟶ ℕ ⟶ 𝔹]. ∀[d:∀x:A. Dec(∃n:ℕ(↑(P n)))].  (mu'(P) ∈ A ⟶ (ℕ Top))

Lemma: can-apply-mu'
[A:Type]. ∀P:A ⟶ ℕ ⟶ 𝔹. ∀d:∀x:A. Dec(∃n:ℕ(↑(P n))). ∀x:A.  (↑can-apply(mu'(P);x) ⇐⇒ ∃n:ℕ(↑(P n)))

Lemma: do-apply-mu'
[A:Type]. ∀[P:A ⟶ ℕ ⟶ 𝔹]. ∀[d:∀x:A. Dec(∃n:ℕ(↑(P n)))]. ∀[x:A].
  {(↑(P do-apply(mu'(P);x))) ∧ (∀[i:ℕdo-apply(mu'(P);x)]. (¬↑(P i)))} supposing ↑can-apply(mu'(P);x)

Definition: quick-find
quick-find(p;n)==r if then else eval n' in quick-find(p;n') fi 

Lemma: quick-find_wf
[n:ℕ+]. ∀[p:{n...} ⟶ 𝔹].  quick-find(p;n) ∈ {m:{n...}| ↑(p m)}  supposing ∃N:{n...}. ∀m:{N...}. (↑(p m))

Lemma: member-assert
[b:𝔹]. Ax ∈ ↑supposing ↑b

Lemma: select-nthtl0
[n:ℕ]. ∀[L:Top List].  (L[n] nth_tl(n;L)[0])

Lemma: pair_wf_l_member
[T:Type]. ∀[L:T List]. ∀[x:T]. ∀[i:ℕ||L||].  <i, Ax, Ax> ∈ (x ∈ L) supposing L[i] x ∈ T

Lemma: member_nth_tl
[T:Type]. ∀n:ℕ. ∀x:T. ∀L:T List.  ((x ∈ nth_tl(n;L))  (x ∈ L))

Lemma: sq_stable__iseg
[T:Type]. ∀l1,l2:T List.  SqStable(l1 ≤ l2)

Definition: sublist-rec
sublist-rec(T;l1;l2) ==
  fix((λsublist-rec,l1,l2. case l1 of 
                             [] => True 
                             a::as =>
                              case l2 of 
                                [] => False 
                                b::bs =>
                                 ((a b ∈ T) ∧ (sublist-rec as bs)) ∨ (sublist-rec l1 bs) 
                             esac 
                          esac)) 
  l1 
  l2

Lemma: sublist-rec_wf
[T:Type]. ∀[l1,l2:T List].  (sublist-rec(T;l1;l2) ∈ ℙ)

Lemma: sublist-rec-nil
[T:Type]. ∀[l:T List].  sublist-rec(T;[];l)

Lemma: sublist-rec-nil-iff
[T:Type]. ∀[l:T List].  uiff(sublist-rec(T;[];l);True)

Lemma: sublist-rec-iff-sublist
[T:Type]. ∀l1,l2:T List.  (l1 ⊆ l2 ⇐⇒ sublist-rec(T;l1;l2))

Lemma: decidable__sublist-rec
[T:Type]. ∀l1,l2:T List.  ((∀x,y:T.  Dec(x y ∈ T))  Dec(sublist-rec(T;l1;l2)))

Lemma: sq_stable__sublist-rec
[T:Type]. ∀l1,l2:T List.  ((∀x,y:T.  Dec(x y ∈ T))  SqStable(sublist-rec(T;l1;l2)))

Lemma: sq_stable__sublist
[T:Type]. ∀l1,l2:T List.  ((∀x,y:T.  Dec(x y ∈ T))  SqStable(l1 ⊆ l2))

Lemma: fseg_append
[T:Type]. ∀l1,l2,l3:T List.  (fseg(T;l1;l2)  fseg(T;l1;l3 l2))

Lemma: fseg_cons
[T:Type]. ∀x:T. ∀[L:T List]. fseg(T;L;[x L])

Lemma: fseg_cons2
[T:Type]. ∀x:T. ∀[L1,L2:T List].  (fseg(T;L1;L2)  fseg(T;L1;[x L2]))

Lemma: fseg_cons_left
[T:Type]. ∀x:T. ∀[L1,L2:T List].  (fseg(T;[x L1];L2)  fseg(T;L1;L2))

Lemma: fseg_extend
[T:Type]
  ∀l1:T List. ∀v:T. ∀l2:T List.
    (fseg(T;l1;l2)  fseg(T;[v l1];l2) supposing ||l1|| < ||l2|| c∧ (l2[||l2|| ||l1|| 1] v ∈ T))

Lemma: fseg_transitivity
[T:Type]. ∀l1,l2,l3:T List.  (fseg(T;l1;l2)  fseg(T;l2;l3)  fseg(T;l1;l3))

Lemma: fseg_weakening
[T:Type]. ∀l1,l2:T List.  fseg(T;l1;l2) supposing l1 l2 ∈ (T List)

Lemma: nil_fseg
[T:Type]. ∀l:T List. fseg(T;[];l)

Lemma: fseg_nil
[T:Type]. ∀L:T List. (fseg(T;L;[]) ⇐⇒ ↑null(L))

Lemma: fseg_length
[T:Type]. ∀[l1,l2:T List].  ||l1|| ≤ ||l2|| supposing fseg(T;l1;l2)

Lemma: filter_fseg
[T:Type]. ∀P:T ⟶ 𝔹. ∀L2,L1:T List.  (fseg(T;L1;L2)  fseg(T;filter(P;L1);filter(P;L2)))

Lemma: fseg_member
[T:Type]. ∀l1,l2:T List. ∀x:T.  (fseg(T;l1;l2)  (x ∈ l1)  (x ∈ l2))

Lemma: fseg_select
[T:Type]
  ∀l1,l2:T List.
    (fseg(T;l1;l2) ⇐⇒ (||l1|| ≤ ||l2||) c∧ (∀i:ℕl1[i] l2[(||l2|| ||l1||) i] ∈ supposing i < ||l1||))

Lemma: fseg-test
T:Type. ∀as,bs,cs:T List.
  ((fseg(T;as;as) ∧ (fseg(T;as;bs)  fseg(T;bs;cs)  fseg(T;as;cs))) ∧ as ≤ as ∧ (as ≤ bs  bs ≤ cs  as ≤ cs))

Definition: lastn
lastn(n;L) ==  nth_tl(||L|| n;L)

Lemma: lastn_wf
[A:Type]. ∀[L:A List]. ∀[n:ℤ].  (lastn(n;L) ∈ List)

Lemma: lastn-cases
[L:Top List]. ∀[n:ℤ].  (lastn(n;L) if ||L|| ≤then if n ≤then [] else lastn(n;tl(L)) fi )

Lemma: lastn-nil
[n:ℤ]. (lastn(n;[]) [])

Lemma: tl-lastn
[L:Top List]. ∀[n:ℤ].  (tl(lastn(n;L)) if n <||L|| then lastn(n 1;L) else lastn(n;tl(L)) fi )

Lemma: length-lastn
[A:Type]. ∀[L:A List]. ∀[n:ℕ].  ||lastn(n;L)|| n ∈ ℤ supposing n ≤ ||L||

Lemma: lastn-0
[L:Top List]. ∀[n:ℤ].  lastn(n;L) [] supposing n ≤ 0

Lemma: map-is-top-list
[f:Top]. ∀[L:Top List].  (map(f;L) ∈ Top List)

Definition: listid
listid(L) ==  rec-case(L) of [] => [] a::as => r.[a r]

Lemma: listid_wf
[A:Type]. ∀[L:A List].  (listid(L) ∈ List)

Lemma: listid_nil_lemma
listid([]) []

Lemma: listid_cons_lemma
y,x:Top.  (listid([x y]) [x listid(y)])

Lemma: listid-sq
[L:Top List]. (listid(L) L)

Definition: adjacent
adjacent(T;L;x;y) ==  ∃i:ℕ||L|| 1. ((x L[i] ∈ T) ∧ (y L[i 1] ∈ T))

Lemma: adjacent_wf
[T:Type]. ∀[L:T List]. ∀[x,y:T].  (adjacent(T;L;x;y) ∈ ℙ)

Lemma: adjacent-nil
[T:Type]. ∀[x,y:T].  False supposing adjacent(T;[];x;y)

Lemma: adjacent-singleton
[T:Type]. ∀[x,y,u:T].  False supposing adjacent(T;[u];x;y)

Lemma: adjacent-cons
[T:Type]
  ∀x,y,u:T. ∀L:T List.  (adjacent(T;[u L];x;y) ⇐⇒ 0 < ||L|| ∧ (((x u ∈ T) ∧ (y hd(L) ∈ T)) ∨ adjacent(T;L;x;y)))

Lemma: simplify-equal-imp
[T:Type]. ∀[x,y,z:T].  uiff(x z ∈ supposing y ∈ T;¬(x y ∈ T)) supposing ¬(y z ∈ T)

Definition: p-inject
p-inject(A;B;f) ==
  ∀x,y:A.  ((↑can-apply(f;x))  (↑can-apply(f;y))  (do-apply(f;x) do-apply(f;y) ∈ B)  (x y ∈ A))

Lemma: p-inject_wf
[A,B:Type]. ∀[f:A ⟶ (B Top)].  (p-inject(A;B;f) ∈ ℙ)

Lemma: p-compose-inject
[A,B,C:Type]. ∀[g:A ⟶ (B Top)]. ∀[f:B ⟶ (C Top)].
  (p-inject(A;C;f g)) supposing (p-inject(B;C;f) and p-inject(A;B;g))

Lemma: p-fun-exp-injection
[A:Type]. ∀[f:A ⟶ (A Top)].  ∀[n:ℕ]. p-inject(A;A;f^n) supposing p-inject(A;A;f)

Lemma: decidable__implies_better
[P:ℙ]. ∀Q:⋂x:P. ℙ(Dec(P)  (P  Dec(Q))  Dec(P  Q))

Lemma: decidable__uimplies
[P:ℙ]. ∀Q:⋂x:P. ℙ(Dec(P)  Dec(Q) supposing  Dec(Q supposing P))

Lemma: biject_functionality
[A1,B1,A2,B2:Type].  ∀f:A1 ⟶ B1. (Bij(A1;B1;f) ⇐⇒ Bij(A2;B2;f)) supposing (B1 ≡ B2 and A1 ≡ A2)

Lemma: nil_member-variant
[T,A:Type].  ∀x:T. (x ∈ []) ⇐⇒ False supposing A ⊆T

Lemma: sub-equality
[T:Type]. ∀[P:T ⟶ ℙ]. ∀[i,u:T].  (i u ∈ {j:T| {j:T| j} supposing ((P u) and (i u ∈ T))

Lemma: l_all_wf2
[T:Type]. ∀[L:T List]. ∀[P:{x:T| (x ∈ L)}  ⟶ ℙ].  ((∀x∈L.P[x]) ∈ ℙ)

Definition: first-member
first-member(T;x;L;P) ==  ∃i:ℕ||L||. ((x L[i] ∈ T) ∧ (↑(P x)) ∧ (∀j:ℕi. (¬↑(P L[j]))))

Lemma: first-member_wf
[T:Type]. ∀[P:T ⟶ 𝔹]. ∀[L:T List]. ∀[x:T].  (first-member(T;x;L;P) ∈ ℙ)

Lemma: first-member-cons
[T:Type]
  ∀P:T ⟶ 𝔹. ∀x,u:T. ∀L:T List.  (first-member(T;x;[u L];P) ⇐⇒ if then u ∈ else first-member(T;x;L;P) fi )

Lemma: first-member-iff
[T:Type]
  ∀L:T List. ∀P:T ⟶ 𝔹. ∀x:T.
    (first-member(T;x;L;P) ⇐⇒ ∃K,J:T List. ((L (K [x J]) ∈ (T List)) ∧ (↑(P x)) ∧ (∀y∈K.¬↑(P y))))

Lemma: can-find-first
[T:Type]. ∀P:T ⟶ 𝔹. ∀L:T List.  ((∃x:T [first-member(T;x;L;P)]) ∨ (∀x∈L.¬↑(P x)))

Lemma: can-find-first1-ext
[T:Type]. ∀P:T ⟶ 𝔹. ∀L:T List.  ((∃x:T [first-member(T;x;L;P)]) ∨ (∀x∈L.¬↑(P x)))

Lemma: can-find-first2
[T:Type]. ∀L:T List. ∀P:{x:T| (x ∈ L)}  ⟶ 𝔹.  ((∃x:T [first-member(T;x;L;P)]) ∨ (∀x∈L.¬↑(P x)))

Lemma: can-find-first-ext
[T:Type]. ∀L:T List. ∀P:{x:T| (x ∈ L)}  ⟶ 𝔹.  ((∃x:T [first-member(T;x;L;P)]) ∨ (∀x∈L.¬↑(P x)))

Definition: find-first
find-first(P;L) ==  TERMOF{can-find-first-ext:o, 1:l, 1:l} P

Lemma: find-first_wf
[T:Type]. ∀[L:T List]. ∀[P:{x:T| (x ∈ L)}  ⟶ 𝔹].  (find-first(P;L) ∈ (∃x:T [first-member(T;x;L;P)]) ∨ (∀x∈L.¬↑(P x)))

Definition: l_find
l_find(L;P) ==  reduce(λh,r. if then inl else fi ;inr ⋅ ;L)

Lemma: l_find_wf
[T:Type]. ∀[L:T List]. ∀[P:{x:T| (x ∈ L)}  ⟶ 𝔹].
  (l_find(L;P) ∈ (∃x:T [(∃i:ℕ||L||. ((x L[i] ∈ T) ∧ (↑(P x)) ∧ (∀j:ℕi. (¬↑(P L[j])))))]) ∨ (↓∀i:ℕ||L||. (¬↑(P L[i]))))

Lemma: null-ite
[b:𝔹]. ∀[x,y:Top].  (null(if then else fi if then null(x) else null(y) fi )

Lemma: typed-null-ite
[x,y:Top List]. ∀[b:𝔹].  null(if then else fi if then null(x) else null(y) fi 

Lemma: reduce-id
[a:Top]. ∀[L:Top List].  (reduce(λx,y. y;a;L) a)

Lemma: filter-commutes
[T:Type]. ∀[P1,P2:T ⟶ 𝔹]. ∀[L:T List].  (filter(P1;filter(P2;L)) filter(P2;filter(P1;L)))

Lemma: count-as-filter
[T:Type]. ∀[P:T ⟶ 𝔹]. ∀[L:T List].  (count(P;L) ||filter(P;L)||)

Lemma: count-all
[T:Type]. ∀[P:T ⟶ 𝔹]. ∀[L:T List].  count(P;L) ||L|| supposing (∀x∈L.↑(P x))

Lemma: null_wf3
[as:Top List]. (null(as) ∈ 𝔹)

Lemma: adjacent-append
[T:Type]
  ∀x,y:T. ∀L1,L2:T List.
    (adjacent(T;L1 L2;x;y)
    ⇐⇒ adjacent(T;L1;x;y) ∨ (0 < ||L1|| ∧ 0 < ||L2|| ∧ (x last(L1) ∈ T) ∧ (y hd(L2) ∈ T)) ∨ adjacent(T;L2;x;y))

Lemma: adjacent-reverse
[T:Type]. ∀L:T List. ∀x,y:T.  (adjacent(T;rev(L);x;y) ⇐⇒ adjacent(T;L;y;x))

Lemma: adjacent-before
[T:Type]. ∀L:T List. ∀x,y:T.  (adjacent(T;L;x;y)  before y ∈ L)

Lemma: adjacent-member
[T:Type]. ∀L:T List. ∀x,y:T.  (adjacent(T;L;x;y)  {(x ∈ L) ∧ (y ∈ L)})

Lemma: adjacent-sublist
[T:Type]. ∀L1,L2:T List.  (L1 ⊆ L2  (∀x,y:T.  (adjacent(T;L1;x;y)  before y ∈ L2)))

Lemma: before-adjacent
[T:Type]
  ∀L:T List. ∀x,y:T.
    adjacent(T;L;x;y)  (∀z:T. (z before y ∈  (z before x ∈ L ∨ (z x ∈ T)))) supposing no_repeats(T;L)

Lemma: before-adjacent2
[T:Type]
  ∀L:T List. ∀x,y:T.
    adjacent(T;L;x;y)  (∀z:T. (x before z ∈  (y before z ∈ L ∨ (z y ∈ T)))) supposing no_repeats(T;L)

Lemma: adjacent-to-same
[T:Type]. ∀[L:T List]. ∀[a,b,c:T].
  (b c ∈ T) supposing (adjacent(T;L;a;c) and adjacent(T;L;a;b) and no_repeats(T;L))

Lemma: adjacent-to-same2
[T:Type]. ∀[L:T List]. ∀[a,b,c:T].
  (b c ∈ T) supposing (adjacent(T;L;c;a) and adjacent(T;L;b;a) and no_repeats(T;L))

Lemma: adjacent-to-same-sublist
[T:Type]
  ∀L1,L2:T List. ∀a,b,c:T.
    L1 ⊆ L2  adjacent(T;L1;b;a)  adjacent(T;L2;c;a)  (b before c ∈ L2 ∨ (b c ∈ T)) supposing no_repeats(T;L2)

Lemma: adjacent-to-same-sublist2
[T:Type]
  ∀L1,L2:T List. ∀a,b,c:T.
    L1 ⊆ L2  adjacent(T;L1;a;b)  adjacent(T;L2;a;c)  (c before b ∈ L2 ∨ (b c ∈ T)) supposing no_repeats(T;L2)

Lemma: adjacent-to-last
[T:Type]. ∀L:T List. (∀a:T. (adjacent(T;L;last(L);a) ⇐⇒ False)) supposing (no_repeats(T;L) and 0 < ||L||)

Lemma: sublist-same-last
[T:Type]. ∀[L,L':T List].
  (last(L') last(L) ∈ T) supposing ((last(L) ∈ L') and L' ⊆ and (¬↑null(L')) and no_repeats(T;L) and (¬↑null(L)))

Lemma: decidable__no_repeats
[T:Type]. ((∀x,y:T.  Dec(x y ∈ T))  (∀L:T List. Dec(no_repeats(T;L))))

Lemma: longest-prefix-decomp
[T:Type]. ∀[L:T List]. ∀[P:(T List) ⟶ 𝔹].  (L longest-prefix(P;L) nth_tl(||longest-prefix(P;L)||;L))

Lemma: longest-prefix_property
[T:Type]
  ∀L:T List. ∀P:(T List) ⟶ 𝔹.
    (longest-prefix(P;L) ≤ L
    ∧ longest-prefix(P;L) < supposing 0 < ||L||
    ∧ (((longest-prefix(P;L) [] ∈ (T List)) ∧ (∀L':T List. (L' <  (¬↑(P L')))))
      ∨ ((↑(P longest-prefix(P;L))) ∧ (∀L':T List. (longest-prefix(P;L) < L'  L' <  (¬↑(P L')))))))

Lemma: longest-prefix_property'
[T:Type]
  ∀L:T List. ∀P:T List+ ⟶ 𝔹.
    (longest-prefix(P;L) ≤ L
    ∧ longest-prefix(P;L) < supposing 0 < ||L||
    ∧ (((longest-prefix(P;L) [] ∈ (T List)) ∧ (∀L':T List. ([] < L'  L' <  (¬↑(P L')))))
      ∨ (0 < ||longest-prefix(P;L)||
        ∧ (↑(P longest-prefix(P;L)))
        ∧ (∀L':T List. (longest-prefix(P;L) < L'  L' <  (¬↑(P L')))))))

Lemma: longest-prefix_property2
[T:Type]
  ∀L:T List. ∀P:(T List) ⟶ 𝔹. ∀L2:T List.
    0 < ||L2|| supposing 0 < ||L||
    ∧ (∀L':T List. ([] < L'  L' < L2  (¬↑(P (longest-prefix(P;L) L')))))
    ∧ ((↑(P longest-prefix(P;L))) ∨ (↑null(longest-prefix(P;L)))) 
    supposing (longest-prefix(P;L) L2) ∈ (T List)

Lemma: longest-prefix-is-nil
[T:Type]. ∀[L:T List]. ∀[P:(T List) ⟶ 𝔹].
  ∀[L':T List]. (¬↑(P L')) supposing (L' < and [] < L') supposing longest-prefix(P;L) [] ∈ (T List)

Lemma: longest-prefix-singleton
[x,P:Top].  (longest-prefix(P;[x]) [])

Lemma: longest-prefix-list_accum
[T,B:Type]. ∀[as:T List]. ∀[P:B ⟶ 𝔹]. ∀[b0:B]. ∀[acc:B ⟶ T ⟶ B]. ∀[a:T].
  (longest-prefix(λL.P[accumulate (with value and list item a):
                        acc[t;a]
                       over list:
                         L
                       with starting value:
                        b0)];[a as]) if null(longest-prefix(λL.P[accumulate (with value and list item a):
                                                                      acc[t;a]
                                                                     over list:
                                                                       L
                                                                     with starting value:
                                                                      acc[b0;a])];as))
  then if bnull(as)) ∧b P[acc[b0;a]] then [a] else [] fi 
  else [a 
        longest-prefix(λL.P[accumulate (with value and list item a):
                             acc[t;a]
                            over list:
                              L
                            with starting value:
                             acc[b0;a])];as)]
  fi )

Lemma: member_list_accum_l_subset
[T:Type]
  ∀f:(T List) ⟶ T ⟶ (T List). ∀L,a:T List. ∀x:T.
    ((∀a:T List. ∀x:T.  l_subset(T;a;f[a;x]))
     ((x ∈ a) ∨ (∃z:T. ((z ∈ L) ∧ (∀l:T List. (x ∈ f[l;z])))))
     (x ∈ accumulate (with value and list item x):
             f[a;x]
            over list:
              L
            with starting value:
             a)))

Lemma: member_list_accum_l_subset2
[T:Type]
  ∀f:(T List) ⟶ T ⟶ (T List). ∀L,a:T List. ∀x:T.
    ((∀a:T List. ∀x:T.  l_subset(T;f[a;x];[x a]))
     (x ∈ accumulate (with value and list item x):
             f[a;x]
            over list:
              L
            with starting value:
             a))
     ((x ∈ a) ∨ (x ∈ L)))

Lemma: list_accum_iseg_inv
[T,A:Type].
  ∀f:A ⟶ T ⟶ A
    ∀[R:A ⟶ A ⟶ ℙ]
      (Refl(A;a,b.R[a;b])
       Trans(A;a,b.R[a;b])
       (∀a:A. ∀x:T.  R[a;f[a;x]])
       (∀a:A. ∀L1,L2:T List.
            (L1 ≤ L2
             R[accumulate (with value and list item x):
                  f[a;x]
                 over list:
                   L1
                 with starting value:
                  a);accumulate (with value and list item x):
                      f[a;x]
                     over list:
                       L2
                     with starting value:
                      a)])))

Lemma: list_accum_proper_iseg_inv
[T,A:Type].
  ∀f:A ⟶ T ⟶ A
    ∀[R:A ⟶ A ⟶ ℙ]
      (Trans(A;a,b.R[a;b])
       (∀a:A. ∀x:T.  R[a;f[a;x]])
       (∀a:A. ∀L1,L2:T List.
            (L1 < L2
             R[accumulate (with value and list item x):
                  f[a;x]
                 over list:
                   L1
                 with starting value:
                  a);accumulate (with value and list item x):
                      f[a;x]
                     over list:
                       L2
                     with starting value:
                      a)])))

Lemma: lastn-as-accum
[A:Type]. ∀[n:ℤ]. ∀[L:A List].
  (lastn(n;L) accumulate (with value and list item x):
                 if ||b|| <then [x] else tl(b [x]) fi 
                over list:
                  L
                with starting value:
                 []))

Definition: remove_leading
remove_leading(a.P[a];L) ==  rec-case(L) of [] => [] a::as => r.if P[a] then else [a as] fi 

Lemma: remove_leading_wf
[T:Type]. ∀[L:T List]. ∀[P:T ⟶ 𝔹].  (remove_leading(x.P[x];L) ∈ {L:T List| (¬↑null(L))  (¬↑P[hd(L)])} )

Lemma: remove_leading_property
[T:Type]. ∀L:T List. ∀P:T ⟶ 𝔹.  ∃xs:{x:T| ↑P[x]}  List. (L (xs remove_leading(x.P[x];L)) ∈ (T List))

Lemma: remove_leading_property2
[T:Type]. ∀[L:T List]. ∀[P:T ⟶ 𝔹].  ¬↑P[hd(remove_leading(x.P[x];L))] supposing ¬↑null(remove_leading(x.P[x];L))

Lemma: null_remove_leading
[T:Type]. ∀[L:T List]. ∀[P:T ⟶ 𝔹].  uiff(↑null(remove_leading(x.P[x];L));(∀x∈L.↑P[x]))

Lemma: remove_leading_eq_nil
[T:Type]. ∀[L:T List]. ∀[P:T ⟶ 𝔹].  uiff(remove_leading(x.P[x];L) [] ∈ (T List);(∀x∈L.↑P[x]))

Lemma: p-first-append
[A,B:Type]. ∀[L1,L2:(A ⟶ (B Top)) List].
  (p-first(L1 L2) p-first([p-first(L1); p-first(L2)]) ∈ (A ⟶ (B Top)))

Lemma: p-first-cons
[A,B:Type]. ∀[L:(A ⟶ (B Top)) List]. ∀[f:A ⟶ (B Top)].  (p-first([f L]) [f?p-first(L)] ∈ (A ⟶ (B Top)))

Lemma: can-apply-p-first
[A,B:Type].  ∀L:(A ⟶ (B Top)) List. ∀x:A.  (↑can-apply(p-first(L);x) ⇐⇒ (∃f∈L. ↑can-apply(f;x)))

Lemma: do-apply-p-first
[A,B:Type]. ∀[L:(A ⟶ (B Top)) List]. ∀[x:A].
  do-apply(p-first(L);x) do-apply(hd(filter(λf.can-apply(f;x);L));x) ∈ supposing ↑can-apply(p-first(L);x)

Definition: p-disjoint
p-disjoint(A;f;g) ==  ∀x:A. ((↑can-apply(f;x)) ∧ (↑can-apply(g;x))))

Lemma: p-disjoint_wf
[A,B:Type]. ∀[f,g:A ⟶ (B Top)].  (p-disjoint(A;f;g) ∈ ℙ)

Lemma: compat-iff-common-iseg
[T:Type]. ∀l1,l2:T List.  (l1 || l2 ⇐⇒ ∃l:T List. (l1 ≤ l ∧ l2 ≤ l))

Lemma: compat-common-member
[T:Type]. ∀[L1,L2:T List].  ∀[k:ℕ]. (L1[k] L2[k] ∈ T) supposing (k < ||L2|| and k < ||L1||) supposing L1 || L2

Lemma: compat-no_repeats_common-member
[T:Type]. ∀[L1,L2:T List].
  (∀[i:ℕ||L1||]. ∀[j:ℕ||L2||].  j ∈ ℤ supposing L1[i] L2[j] ∈ T) supposing 
     ((no_repeats(T;L1) ∧ no_repeats(T;L2)) and 
     L1 || L2)

Definition: l-all
x∈L.P[x] ==  reduce(λx,p. (P[x] ∧ p);True;L)

Lemma: l-all_wf
[T:Type]. ∀[L:T List]. ∀[P:{x:T| (x ∈ L)}  ⟶ ℙ].  (∀x∈L.P[x] ∈ ℙ)

Lemma: l-all-property
[T:Type]. ∀L:T List. ∀[P:{x:T| (x ∈ L)}  ⟶ ℙ']. (∀x∈L.P[x] ⇐⇒ (∀x∈L.P[x]))

Definition: l-all-and
B ∧ ∀x∈L.P[x] ==  reduce(λx,p. (P[x] ∧ p);B;L)

Lemma: l-all-and_wf
[T:Type]. ∀[L:T List]. ∀[P:{x:T| (x ∈ L)}  ⟶ ℙ]. ∀[B:ℙ].  (B ∧ ∀x∈L.P[x] ∈ ℙ)

Lemma: l-all-and-property
[T:Type]. ∀L:T List. ∀[P:{x:T| (x ∈ L)}  ⟶ ℙ']. ∀[B:ℙ].  (B ∧ ∀x∈L.P[x] ⇐⇒ B ∧ (∀x∈L.P[x]))

Lemma: l-all-iff
[T:Type]. ∀L:T List. ∀[P:{x:T| (x ∈ L)}  ⟶ ℙ]. (∀x∈L.P[x] ⇐⇒ (∀x∈L.P[x]))

Lemma: length-as-accum
[L:Top List]. (||L|| accumulate (with value and list item x): 1over list:  Lwith starting value: 0))

Lemma: exists-pair
[A,B:Type]. ∀[P:(A × B) ⟶ ℙ].  (∃p:A × B. P[p] ⇐⇒ ∃a:A. ∃b:B. P[<a, b>])

Definition: list_accum_pair
list_accum_pair(a,x.f[a; x];b,x.g[b; x];a0;b0;L) ==
  accumulate (with value and list item x):
   let a,b 
   in <f[a; x], g[b; x]>
  over list:
    L
  with starting value:
   <a0, b0>)

Lemma: list_accum_pair_wf
[T,A,B:Type]. ∀[a0:A]. ∀[b0:B]. ∀[f:A ⟶ T ⟶ A]. ∀[g:B ⟶ T ⟶ B]. ∀[L:T List].
  (list_accum_pair(a,x.f[a;x];b,x.g[b;x];a0;b0;L) ∈ A × B)

Lemma: list_accum_pair-sq
[T,A,B:Type]. ∀[f:A ⟶ T ⟶ A]. ∀[g:B ⟶ T ⟶ B]. ∀[L:T List]. ∀[a0:A]. ∀[b0:B].
  (list_accum_pair(a,x.f[a;x];b,x.g[b;x];a0;b0;L) ~ <accumulate (with value and list item x):
                                                      f[a;x]
                                                     over list:
                                                       L
                                                     with starting value:
                                                      a0)
                                                    accumulate (with value and list item x):
                                                       g[b;x]
                                                      over list:
                                                        L
                                                      with starting value:
                                                       b0)
                                                    >)

Definition: cps-accum
cps-accum(k,a.f[k; a];b;L) ==  rec-case(L) of [] => λk.(k b) a::as => r.λk.(r f[k; a])

Lemma: cps-accum_wf
[A,B,C:Type]. ∀[f:(B ⟶ C) ⟶ A ⟶ B ⟶ C]. ∀[b:B]. ∀[L:A List].  (cps-accum(k,a.f[k;a];b;L) ∈ (B ⟶ C) ⟶ C)

Definition: convolution
convolution(a,b,c.f[a;
                    b;
                    c];c0;L1;L2) ==
  cps-accum(k,a.λp.let remaining,c 
                   in rec-case(remaining) of
                      [] => p
                      b::bs =>
                       r.k 
                         <bs
                         f[a;
                             b;
                             c]
                         >;<L2, c0>;L1) 
  p.(snd(p)))

Lemma: convolution_wf
[A,B,C:Type]. ∀[f:A ⟶ B ⟶ C ⟶ C]. ∀[c0:C]. ∀[L1:A List]. ∀[L2:B List].  (convolution(a,b,c.f[a;b;c];c0;L1;L2) ∈ C)

Definition: rev-zip
rev-zip(L1;L2) ==  convolution(a,b,c.[<a, b> c];[];L1;L2)

Lemma: rev-zip_wf
[A,B:Type]. ∀[L1:A List]. ∀[L2:B List].  (rev-zip(L1;L2) ∈ (A × B) List)

Definition: bl-rev-exists
(∃x∈rev(L).P[x])_b ==  rec-case(L) of [] => ff a::as => r.case of inl(x) => tt inr(x) => P[a]

Lemma: bl-rev-exists-sq
[T:Type]. ∀[L:T List]. ∀[P:{x:T| (x ∈ L)}  ⟶ 𝔹].  ((∃x∈rev(L).P[x])_b (∃x∈L.P[x])_b)

Lemma: bl-rev-exists_wf
[T:Type]. ∀[L:T List]. ∀[P:{x:T| (x ∈ L)}  ⟶ 𝔹].  ((∃x∈rev(L).P[x])_b ∈ 𝔹)

Definition: update
f[x:=v] ==  λy.if eq then else fi 

Lemma: update_wf
[A,B:Type]. ∀[eq:A ⟶ A ⟶ 𝔹]. ∀[f:A ⟶ B]. ∀[x:A]. ∀[v:B].  (f[x:=v] ∈ A ⟶ B)

Definition: list_update
l[i:=x] ==  mklist(||l||;λi.l[i][i:=x])

Lemma: list_update_wf
[T:Type]. ∀[l:T List]. ∀[i:ℤ]. ∀[x:T].  (l[i:=x] ∈ List)

Lemma: list_update_select
[T:Type]. ∀[l:T List]. ∀[i:ℤ]. ∀[j:ℕ||l||]. ∀[x:T].  (l[i:=x][j] if (j =z i) then else l[j] fi  ∈ T)

Lemma: list_update_length
[l:Top List]. ∀[i:ℤ]. ∀[x:Top].  (||l[i:=x]|| ||l||)

Lemma: iseg_antisymmetry
[T:Type]. ∀[as,bs:T List].  (as bs ∈ (T List)) supposing (bs ≤ as and as ≤ bs)

Lemma: compat-cons
[T:Type]. ∀as,bs:T List. ∀a,b:T.  ([a as] || [b bs] ⇐⇒ (a b ∈ T) ∧ as || bs)

Lemma: compat-append
[T:Type]. ∀as,cs,bs,ds:T List.  (as bs || cs ds  as || cs)

Lemma: compat-append2
[T:Type]. ∀as,cs,bs,ds:T List.  (as bs || cs ds  bs || ds supposing as cs ∈ (T List))

Lemma: compat_symmetry
[T:Type]. ∀as,bs:T List.  (as || bs ⇐⇒ bs || as)

Lemma: compat-iseg
[T:Type]. ∀L1,L2,L3:T List.  (L1 ≤ L2  L2 || L3  L1 || L3)

Lemma: compat-iseg2
[T:Type]. ∀L1,L2,L3:T List.  (L1 ≤ L2  L3 || L2  L3 || L1)

Lemma: merge-strict-exists
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (Trans(T;a,b.R b)
   (∀sa,sb:T List.
        ((∀a,b:T.  ((a ∈ sa)  (b ∈ sb)  ((R b) ∨ (R a))))
         sorted-by(R;sa)
         sorted-by(R;sb)
         (∃sc:T List. (sorted-by(R;sc) ∧ sa ⊆ sc ∧ sb ⊆ sc ∧ sc ⊆ sa sb)))))

Lemma: iseg-as-filter
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ]. ∀[sa,sb:T List].
  (∀[dR:T ⟶ T ⟶ 𝔹]
     (sa filter(λx.(dR last(sa));sb) ∈ (T List)) supposing 
        ((¬↑null(sa)) and 
        (∀x,y:T.  (↑(dR y) ⇐⇒ (R y) ∨ (x y ∈ T))))) supposing 
     (Trans(T;a,b.R b) and 
     sorted-by(R;sb) and 
     sa ≤ sb and 
     AntiSym(T;x,y.R y) and 
     Irrefl(T;x,y.R y))

Lemma: iseg-remainder-as-filter
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ]. ∀[sa,s,sb:T List].
  (∀[dR:T ⟶ T ⟶ 𝔹]
     (sb filter(dR last(sa);s) ∈ (T List)) supposing ((¬↑null(sa)) and (∀x,y:T.  (↑(dR y) ⇐⇒ y)))) supposing 
     (Trans(T;a,b.R b) and 
     sorted-by(R;s) and 
     (s (sa sb) ∈ (T List)) and 
     AntiSym(T;x,y.R y) and 
     Irrefl(T;x,y.R y))

Definition: s-filter
s-filter(p;as) ==  reduce(λa,l. if then s-insert(a;l) else fi ;[];as)

Lemma: s-filter_wf
[T:Type]. ∀[as:T List]. ∀[P:{a:T| (a ∈ as)}  ⟶ 𝔹].  s-filter(P;as) ∈ List supposing T ⊆r ℤ

Definition: merge
merge(as;bs) ==  reduce(λb,l. s-insert(b;l);as;bs)

Lemma: merge_wf
[T:Type]. ∀[as,bs:T List].  (merge(as;bs) ∈ List) supposing T ⊆r ℤ

Lemma: member-merge
[T:Type]. ∀bs,as:T List. ∀x:T.  ((x ∈ merge(as;bs)) ⇐⇒ (x ∈ as) ∨ (x ∈ bs)) supposing T ⊆r ℤ

Lemma: sorted-merge
[T:Type]. ∀[bs,as:T List].  sorted(merge(as;bs)) supposing sorted(as) supposing T ⊆r ℤ

Lemma: no_repeats-merge
[T:Type]. ∀[bs,as:T List].  (no_repeats(T;merge(as;bs))) supposing (sorted(as) and no_repeats(T;as)) supposing T ⊆r ℤ

Lemma: strict-sorted
[T:Type]. ∀[as:T List]. uiff(sorted(as) ∧ no_repeats(T;as);∀[i:ℕ||as||]. ∀[j:ℕi].  as[j] < as[i]) supposing T ⊆r ℤ

Definition: priority-select
priority-select(f;g;as) ==
  accumulate (with value and list item m):
   if isl(x) then x
   if then inl tt
   if then inl ff
   else x
   fi 
  over list:
    as
  with starting value:
   inr ⋅ )

Lemma: priority-select_wf
[T:Type]. ∀[as:T List]. ∀[f,g:T ⟶ 𝔹].  (priority-select(f;g;as) ∈ 𝔹?)

Lemma: priority-select-property
[T:Type]
  ∀as:T List. ∀f,g:T ⟶ 𝔹.
    ((priority-select(f;g;as) (inl tt) ∈ (𝔹?) ⇐⇒ ∃i:ℕ||as||. ((↑(f as[i])) ∧ (∀j:ℕi. (¬↑(g as[j])))))
    ∧ (priority-select(f;g;as) (inl ff) ∈ (𝔹?) ⇐⇒ ∃i:ℕ||as||. ((↑(g as[i])) ∧ (∀j:ℕ1. (¬↑(f as[j])))))
    ∧ (priority-select(f;g;as) (inr ⋅ ) ∈ (𝔹?) ⇐⇒ ∀i:ℕ||as||. ((¬↑(f as[i])) ∧ (¬↑(g as[i])))))

Lemma: priority-select-inr
[T:Type]. ∀[as:T List]. ∀[f,g:T ⟶ 𝔹].  uiff(priority-select(f;g;as) (inr ⋅ ) ∈ (𝔹?);(∀a∈as.(¬↑(f a)) ∧ (¬↑(g a))))

Lemma: not-isl-priority-select
[T:Type]. ∀[as:T List]. ∀[f,g:T ⟶ 𝔹].  uiff(¬↑isl(priority-select(f;g;as));(∀a∈as.(¬↑(f a)) ∧ (¬↑(g a))))

Lemma: priority-select-tt
[T:Type]
  ∀as:T List. ∀f,g:T ⟶ 𝔹.
    (priority-select(f;g;as) (inl tt) ∈ (𝔹?)
       ⇐⇒ (∃a∈as. (↑(f a)) ∧ (∀b:T. ((b ∈ as)  ¬↑(g b) supposing b < a)))) supposing 
       (no_repeats(T;as) and 
       sorted(as) and 
       (T ⊆r ℤ))

Lemma: priority-select-ff
[T:Type]
  ∀as:T List. ∀f,g:T ⟶ 𝔹.
    (priority-select(f;g;as) (inl ff) ∈ (𝔹?)
       ⇐⇒ (∃a∈as. (↑(g a)) ∧ (∀b:T. ((b ∈ as)  ¬↑(f b) supposing b ≤ a)))) supposing 
       (sorted(as) and 
       (T ⊆r ℤ))

Lemma: fun_exp_add_sq_again
[n,m:ℕ]. ∀[f,i:Top].  (f^n f^n (f^m i))

Definition: list-maximals
list-maximals(eq;ls;f;L) ==
  accumulate (with value pr and list item x):
   eval in
   let m,xs pr 
   in if eq then <n, [x xs]>
      if ls then <n, [x]>
      else pr
      fi 
  over list:
    tl(L)
  with starting value:
   <hd(L), [hd(L)]>)

Lemma: member-le-max
[L:ℤ List]. ∀[x:ℤ].  x ≤ imax-list(L) supposing (x ∈ L)

Lemma: do-apply-p-first-disjoint
[A,B:Type]. ∀[L:(A ⟶ (B Top)) List]. ∀[x:A].
  ∀[f:A ⟶ (B Top)]. (do-apply(p-first(L);x) do-apply(f;x) ∈ B) supposing ((↑can-apply(f;x)) and (f ∈ L)) 
  supposing (∀f,g∈L.  p-disjoint(A;f;g))

Definition: inv-rel
inv-rel(A;B;f;finv) ==
  (∀a:A. ∀b:B.  (((finv b) (inl a) ∈ (A?))  (b (f a) ∈ B))) ∧ (∀a:A. ((finv (f a)) (inl a) ∈ (A?)))

Lemma: inv-rel_wf
[A,B:Type]. ∀[f:A ⟶ B]. ∀[finv:B ⟶ (A?)].  (inv-rel(A;B;f;finv) ∈ ℙ)

Lemma: inv-rel-inject
[A,B:Type]. ∀[f:A ⟶ B]. ∀[finv:B ⟶ (A?)].  Inj(A;B;f) supposing inv-rel(A;B;f;finv)

Lemma: hd-filter
[T:Type]
  ∀P:T ⟶ 𝔹. ∀as:T List.
    ((∃a:T. ((a ∈ as) ∧ (↑P[a])))
     ((hd(filter(λa.P[a];as)) ∈ T) c∧ ((hd(filter(λa.P[a];as)) ∈ as) ∧ (↑P[hd(filter(λa.P[a];as))]))))

Lemma: find-hd-filter
[T:Type]. ∀[P:T ⟶ 𝔹]. ∀[as:T List]. ∀[d:Top].
  (first a ∈ as s.t. P[a] else d) hd(filter(λa.P[a];as)) ∈ supposing ∃a:T. ((a ∈ as) ∧ (↑P[a]))

Lemma: sq_stable__sorted-by
[T:Type]
  ∀L:T List
    ∀[R:{x:T| (x ∈ L)}  ⟶ {x:T| (x ∈ L)}  ⟶ ℙ]. ((∀x,y:{x:T| (x ∈ L)} .  SqStable(R y))  SqStable(sorted-by(R;L)))

Lemma: pair-list-set-type
[T:Type]. ∀[B:T ⟶ Type]. ∀[L:(t:T × B[t]) List].  (L ∈ (t:{t:T| (t ∈ map(λp.(fst(p));L))}  × B[t]) List)

Lemma: general-pigeon-hole
[n,m,k:ℕ]. ∀[f:ℕn ⟶ ℕm].
  n ≤ (k m) supposing ∀L:ℕList. (no_repeats(ℕn;L)  (∃i:ℕm. (∀x∈L.f[x] i ∈ ℕm))  (||L|| ≤ k))

Lemma: hd-as-reduce
[L:Top List]. (hd(L) reduce(λu,z. u;hd([]);L))

Lemma: select-as-reduce
[n:ℕ]. ∀[L:Top List].
  L[||L|| 1] 
  outr(reduce(λu,x. case of inl(i) => if (i =z n) then inr u  else inl (i 1) fi  inr(u) => x;inl 0;L)) 
  supposing n < ||L||

Lemma: combine-reduce
[F,f,g,x,y:Top]. ∀[L:Top List].
  (F reduce(f;x;L) reduce(g;y;L) z.let x,y in y) reduce(λu,z. let x,y in <x, y>;<x, y>;L))

Lemma: wellfounded-anti-reflexive
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  ∀[a:T]. R[a;a]) supposing WellFnd{i}(T;x,y.R[x;y])

Lemma: mapfilter-as-accum-aux
[A:Type]. ∀[p:A ⟶ 𝔹]. ∀[L:A List]. ∀[X:Top List]. ∀[f:Top].
  (X mapfilter(f;p;L) accumulate (with value and list item x):
                           if p[x] then [f[x]] else fi 
                          over list:
                            L
                          with starting value:
                           X))

Lemma: mapfilter-as-accum
[A:Type]. ∀[L:A List]. ∀[p:A ⟶ 𝔹]. ∀[f:Top].
  (mapfilter(f;p;L) accumulate (with value and list item x):
                       if p[x] then [f[x]] else fi 
                      over list:
                        L
                      with starting value:
                       []))

Lemma: mapfilter-as-reduce
[L,p,f:Top].  (mapfilter(λx.f[x];λx.p[x];L) reduce(λx,a. if p[x] then [f[x] a] else fi ;[];L))

Lemma: mapfilter-as-reduce2
[L,p,f:Top].  (mapfilter(f;p;L) reduce(λx,a. if then [f a] else fi ;[];L))

Definition: fast-mapfilter
fast-mapfilter(p;f;L) ==  reduce(λx,a. if p[x] then [f[x] a] else fi ;[];L)

Lemma: fast-mapfilter_wf
[A,B:Type]. ∀[L:A List]. ∀[p:{x:A| (x ∈ L)}  ⟶ 𝔹]. ∀[f:{x:A| (x ∈ L) ∧ (↑p[x])}  ⟶ B].
  (fast-mapfilter(p;f;L) ∈ List)

Lemma: mapfilter-as-fast-mapfilter
[L,p,f:Top].  (mapfilter(f;p;L) fast-mapfilter(p;f;L))

Lemma: reduce-fast-mapfilter1
[L,z,k,d,f,p:Top].
  (reduce(λx,a. case d[x;a] of inl() => inr() => a;k;fast-mapfilter(p;f;L)) reduce(λx,a. if x
                                                                                          then case d[f x;a]
                                                                                                of inl() =>
                                                                                                z
                                                                                                inr() =>
                                                                                                a
                                                                                          else a
                                                                                          fi ;k;L))

Lemma: reduce-fast-mapfilter2
[L,z,k,c,d,f,p:Top].
  (reduce(λx,a. if c[x;a]=d[x;a] then else a;k;fast-mapfilter(p;f;L)) reduce(λx,a. if x
                                                                                  then if c[f x;a]=d[f x;a]
                                                                                       then z
                                                                                       else a
                                                                                  else a
                                                                                  fi ;k;L))

Lemma: filter-as-accum-aux
[A:Type]. ∀[p:A ⟶ 𝔹]. ∀[L:A List]. ∀[X:Top List].
  (filter(p;rev(L)) accumulate (with value and list item x):
                           if p[x] then [x a] else fi 
                          over list:
                            L
                          with starting value:
                           X))

Lemma: filter-as-accum-aux2
[A:Type]. ∀[p:A ⟶ 𝔹]. ∀[L:A List]. ∀[X:Top List].
  (X filter(p;L) accumulate (with value and list item x):
                      if p[x] then [x] else fi 
                     over list:
                       L
                     with starting value:
                      X))

Lemma: nil-iff-no-member
[T:Type]. ∀[L:T List].  uiff(L [] ∈ (T List);∀[x:T]. (x ∈ L)))

Lemma: tl_sublist
[T:Type]. ∀a:T. ∀L1,L2:T List.  ([a L1] ⊆ L2  L1 ⊆ L2)

Lemma: sublist_accum
[T:Type]
  ∀L,l1,l2:T List. ∀f:(T List) ⟶ T ⟶ (T List).
    (l1 ⊆ l2
     (∀x:T. ∀l:T List.  l ⊆ f[l;x])
     l1 ⊆ accumulate (with value and list item x):
             f[l;x]
            over list:
              L
            with starting value:
             l2))

Lemma: unzip-as-accum
[as:(Top × Top) List]
  (unzip(as) accumulate (with value and list item a):
                let p1,p2 
                in let a1,a2 
                   in <p1 [a1], p2 [a2]>
               over list:
                 as
               with starting value:
                <[], []>))

Definition: un-zip
un-zip(as) ==  reduce(λa,p. let p1,p2 in let a1,a2 in <[a1 p1], [a2 p2]>;<[], []>;as)

Lemma: un-zip_wf
[A,B:Type]. ∀[as:(A × B) List].  (un-zip(as) ∈ List × (B List))

Lemma: unzip-un-zip
[as:(Top × Top) List]. (unzip(as) un-zip(as))

Definition: dectt
dectt(d) ==  isl(d)

Lemma: dectt_wf
[p:ℙ]. ∀[d:Dec(p)].  (dectt(d) ∈ 𝔹)

Lemma: assert-dectt
[p:ℙ]. ∀d:Dec(p). (↑dectt(d) ⇐⇒ p)

Lemma: inr_equal
[A,B:Type]. ∀[x,y:B].  uiff((inr (inr ) ∈ (A B);x y ∈ B)

Lemma: inl_equal
[A,B:Type]. ∀[x,y:A].  uiff((inl x) (inl y) ∈ (A B);x y ∈ A)

Lemma: inl_eq_inr
[A,B:Type]. ∀[x:A]. ∀[y:B].  uiff((inl x) (inr ) ∈ (A B);False)

Lemma: inr_eq_inl
[A,B:Type]. ∀[x:A]. ∀[y:B].  uiff((inr (inl x) ∈ (A B);False)

Lemma: length-zip-map
[L:Top List]. ∀[f,g:Top].  (||zip(map(f;L);map(g;L))|| ||L||)

Lemma: select-zip
[as,bs:Top List]. ∀[i:ℕ].  zip(as;bs)[i] ~ <as[i], bs[i]> supposing i < ||as|| ∧ i < ||bs||

Lemma: member-listify
[T:Type]. ∀m:ℤ. ∀n:{n:ℤn ≥ . ∀f:{m..n-} ⟶ T.  ∀[x:T]. ((x ∈ listify(f;m;n)) ⇐⇒ ∃i:{m..n-}. (x (f i) ∈ T))

Lemma: map-permute_list
[g:Top]. ∀[L:Top List]. ∀[f:ℕ||L|| ⟶ ℕ||L||].  (map(g;(L f)) (map(g;L) f))

Lemma: pairwise-map1
[T,T':Type].  ∀f:T ⟶ T'. ∀L:T List.  ∀[P:T' ⟶ T' ⟶ ℙ']. ((∀x,y∈map(f;L).  P[x;y]) ⇐⇒ (∀x,y∈L.  P[f x;f y]))

Lemma: pairwise-map2
[T,T':Type].
  ∀L:T List. ∀f:{t:T| (t ∈ L)}  ⟶ T'.  ∀[P:T' ⟶ T' ⟶ ℙ']. ((∀x,y∈map(f;L).  P[x;y]) ⇐⇒ (∀x,y∈L.  P[f x;f y]))

Lemma: general_length_nth_tl
[r:ℕ]. ∀[L:Top List].  (||nth_tl(r;L)|| if r <||L|| then ||L|| else fi  ∈ ℤ)

Definition: next
(next i > s.t. ↑p[i]) ==  fix((λnext,k. eval in if p[j] then else next fi )) k

Lemma: next-unfold
[k,p:Top].  ((next i > s.t. ↑p[i]) eval in if p[j] then else (next i > s.t. ↑p[i]) fi )

Lemma: next_wf_bound
[b:ℕ]. ∀[k:ℤ]. ∀[p:{i:ℤk < i}  ⟶ 𝔹].
  (next i > s.t. ↑p[i]) ∈ {i:ℤ(k < i ∧ (i ≤ (k b))) ∧ (↑p[i]) ∧ (∀j:{k 1..i-}. (¬↑p[j]))}  
  supposing ∃n:{i:ℤk < i ∧ (i ≤ (k b))} (↑p[n])

Lemma: next_wf
[k:ℤ]. ∀[p:{i:ℤk < i}  ⟶ 𝔹].  (next i > s.t. ↑p[i]) ∈ {i:ℤk < i ∧ (↑p[i]) ∧ (∀j:{k 1..i-}. (¬↑p[j]))}  supposi\000Cng ∃n:{i:ℤk < i} (↑p[n])

Lemma: primality-test
b:ℕ(prime(b)) supposing ((∀p:ℕ(prime(p)  ((p p) ≤ b)  (p b)))) and (2 ≤ b))

Definition: factorit
factorit(x;b;tried;facs) ==
  fix((λfactorit,x,b,tried,facs. if x <then if x <then facs else [x facs] fi 
                                if (∃p∈tried.(b rem =z 0))_b then eval b' in factorit b' tried facs
                                if (x rem =z 0) then eval x' x ÷ in factorit x' tried [b facs]
                                else eval b' in
                                     factorit b' [b tried] facs
                                fi )) 
  
  
  tried 
  facs

Lemma: factorit_wf
[x:ℕ+]. ∀[b:ℕ].
  ∀[tried:{L:{p:ℕprime(p) ∧ p < b}  List| ∀p:{p:ℕprime(p)} (p <  ((p ∈ L) ∧ (p x))))} ].
  ∀[facs:{p:ℕprime(p)}  List].
    (factorit(x;b;tried;facs) ∈ {L:{p:ℕprime(p)}  List| reduce(λp,q. (p q);1;L) (x reduce(λp,q. (p q);1;facs))\000C ∈ ℤ
  supposing 2 ≤ b

Definition: mul-list
Π(ns)  ==  reduce(λx,y. (x y);1;ns)

Lemma: mul-list_wf
[ns:ℤ List]. (ns)  ∈ ℤ)

Lemma: mul_list_nil_lemma
Π([])  1

Lemma: mul-list-append
[ns1,ns2:ℤ List].  (ns1 ns2)  ~ Π(ns1)  * Π(ns2) )

Lemma: mul-list-insert-int
[ns:ℤ List]. ∀[x:ℤ].  (insert-int(x;ns))  (x * Π(ns) ) ∈ ℤ)

Lemma: mul-list-merge
[ns2,ns1:ℤ List].  (merge-int(ns1;ns2))  (ns1)  * Π(ns2) ) ∈ ℤ)

Lemma: mul-list-bag-product
L:ℤ List. (L)  = Π(L) ∈ ℤ)

Lemma: mul-list-positive
L:ℕ+ List. 0 < Π(L) 

Lemma: prime-factors
n:{2...}. (∃factors:{m:{2...}| prime(m)}  List [(n = Π(factors)  ∈ ℤ)])

Lemma: reducible-nat
n:ℤreducible(n)  (∃n1:ℕ(n1 < n ∧ (2 ≤ n1) ∧ (n1 n))) supposing 2 ≤ n

Definition: factorization
factorization(n) ==  TERMOF{prime-factors:o, 1:l} n

Lemma: factorization_wf
n:{2...}. (factorization(n) ∈ {factors:{m:{2...}| prime(m)}  List| = Π(factors)  ∈ ℤ)

Definition: Prime
Prime ==  {p:{2...}| prime(p)} 

Lemma: Prime_wf
Prime ∈ Type

Lemma: Prime-isOdd
p:Prime. ((¬(p 2 ∈ ℤ))  (↑isOdd(p)))

Definition: factors
factors(n) ==  if (n =z 1) then {} else factorization(n) fi 

Lemma: factors_wf
[n:ℕ+]. (factors(n) ∈ bag(Prime))

Lemma: product-factors
[n:ℕ+]. (factors(n)) n ∈ ℤ)

Definition: norm-factors
norm-factors(L) ==
  let L,x,n accumulate (with value and list item y):
               let L,x,n in 
               if (x =z y) then eval n' in <L, x, n'>
               if (x =z 1) then <L, y, 1>
               if (n =z 1) then <[x L], y, 1>
               else <[x^n L], y, 1>
               fi 
              over list:
                L
              with starting value:
               <[], 1, 0>in 
  if (n =z 1) then [x L] else [x^n L] fi 

Definition: least-factor
least-factor(n) ==  mu(λi.(n rem =z 0)) 2

Lemma: least-factor_wf
[n:ℤ]. least-factor(n) ∈ {p:ℕ1 < p ∧ prime(p) ∧ (p n) ∧ (∀q:ℕ(prime(q)  (q n)  (p ≤ q)))}  supposing 1 < |n\000C|

Definition: is_prime
is_prime(n) ==  isl(TERMOF{decidable__prime:o, 1:l} n)

Lemma: is_prime_wf
[n:ℕ]. (is_prime(n) ∈ 𝔹)

Lemma: assert-is_prime
n:ℕ(↑is_prime(n) ⇐⇒ prime(n))

Comment: tail-recursive-extracts
Here are some results from paper called
"New development in extracting tail recursive programs from proofs"
by Luca Chiarabini and Philippe Audebaud.


They use the notion of "non-computational" forall quantifier.
That is the same as our "uniform"-all ⌜∀[x:T]. P⌝.⋅

Lemma: primrec-induction
[P:ℕ ⟶ ℙ]. (P[0]  (∀n:ℕ(P[n]  P[n 1]))  (∀n:ℕP[n]))

Lemma: primrec-induction-factorial
n:ℕ(∃x:ℤ [((n)! x ∈ ℤ)])

Lemma: cont-induction-lemma
[P:ℕ ⟶ ℙ]. (P[0]  (∀n:ℕ(P[n]  P[n 1]))  (∀n:ℕ. ∀[m:ℕ]. ((P[n]  P[n m])  P[n m])))

Lemma: cont-induction
[P:ℕ ⟶ ℙ]. (P[0]  (∀n:ℕ(P[n]  P[n 1]))  (∀n:ℕP[n]))

Lemma: cont-induction-ext
[P:ℕ ⟶ ℙ]. (P[0]  (∀n:ℕ(P[n]  P[n 1]))  (∀n:ℕP[n]))

Lemma: cont-induction-factorial
n:ℕ(∃x:ℤ [((n)! x ∈ ℤ)])

Lemma: accum-induction-lemma
[P:ℕ ⟶ ℙ]. (P[0]  (∀n:ℕ(P[n]  P[n 1]))  (∀n,m:ℕ.  (P[m]  P[n m])))

Lemma: accum-induction
[P:ℕ ⟶ ℙ]. (P[0]  (∀n:ℕ(P[n]  P[n 1]))  (∀n:ℕP[n]))

Lemma: accum-induction-ext
[P:ℕ ⟶ ℙ]. (P[0]  (∀n:ℕ(P[n]  P[n 1]))  (∀n:ℕP[n]))

Lemma: accum-induction-factorial
n:ℕ(∃x:ℤ [((n)! x ∈ ℤ)])

Lemma: def-cont-induction-lemma
[P:ℕ ⟶ ℙ]
  ((∀n:ℕ(P[n]  P[n 1]))  (∀x:ℤ List. ∀[n,m:ℕ].  P[n]  P[m] supposing (x [n, m) ∈ (ℤ List)) ∧ (n ≤ m)))

Lemma: def-cont-induction-lemma-ext
[P:ℕ ⟶ ℙ]
  ((∀n:ℕ(P[n]  P[n 1]))  (∀x:ℤ List. ∀[n,m:ℕ].  P[n]  P[m] supposing (x [n, m) ∈ (ℤ List)) ∧ (n ≤ m)))

Lemma: list-eq-set-type
[T:Type]. ∀[P:T ⟶ ℙ]. ∀[A,B:T List].
  (A B ∈ ({x:T| P[x]}  List)) supposing ((∀i:ℕ||A||. P[A[i]]) and (A B ∈ (T List)))

Lemma: connection-bound
[T:Type]
  ((∀x,y:T.  Dec(x y ∈ T))
   (∀k:ℕ(|T| ≤  (∀f:T ⟶ T. ∀a,b:T.  (∃n:ℕ(b (f^n a) ∈ T) ⇐⇒ ∃n:ℕk. (b (f^n a) ∈ T))))))

Lemma: decidable__connection
[T:Type]. ((∀x,y:T.  Dec(x y ∈ T))  finite-type(T)  (∀f:T ⟶ T. ∀a,b:T.  Dec(∃n:ℕ(b (f^n a) ∈ T))))

Lemma: first_index_property
[T:Type]. ∀[P:T ⟶ 𝔹]. ∀[L:T List].
  (↑P[L[index-of-first in L.P[x] 1]]) ∧ (∃x∈firstn(index-of-first in L.P[x] 1;L). ↑P[x])) 
  supposing 0 < index-of-first in L.P[x]

Definition: last_index
last_index(L;x.P[x]) ==
  snd(accumulate (with value and list item x):
       let n,lst 
       in <1, if P[x] then else lst fi >
      over list:
        L
      with starting value:
       <0, 0>))

Lemma: last_index_wf
[T:Type]. ∀[L:T List]. ∀[P:T ⟶ 𝔹].  (last_index(L;x.P[x]) ∈ ℕ||L|| 1)

Lemma: last_index_append
[T:Type]. ∀[as,bs:T List]. ∀[P:T ⟶ 𝔹].
  (last_index(as bs;x.P[x])
  if 0 <last_index(bs;x.P[x]) then ||as|| last_index(bs;x.P[x]) else last_index(as;x.P[x]) fi 
  ∈ ℤ)

Lemma: last_index_cons
[T:Type]. ∀[P:T ⟶ 𝔹]. ∀[L:T List]. ∀[a:T].
  (last_index([a L];x.P[x])
  if 0 <last_index(L;x.P[x]) then last_index(L;x.P[x])
    if P[a] then 1
    else 0
    fi 
  ∈ ℤ)

Lemma: last_index_property
[T:Type]. ∀[P:T ⟶ 𝔹]. ∀[L:T List].
  ((↑P[L[last_index(L;x.P[x]) 1]]) ∧ (∃x∈nth_tl(last_index(L;x.P[x]);L). ↑P[x])) supposing 0 < last_index(L;x.P[x])
  ∧ ¬(∃x∈L. ↑P[x]) supposing last_index(L;x.P[x]) 0 ∈ ℤ)

Lemma: fix_wf1
[F:{F:ℕ ⟶ Type| Top ⊆(F 0)} ]. ∀[G:⋂n:ℕ((F n) ⟶ (F (n 1)))].  (fix(G) ∈ ⋂n:ℕ(F n))

Definition: DCC
DCC (descending chain condition) says that there is no infinte
descending chain for the given relation.⋅

DCC(T;<==  ∀f:ℕ ⟶ T. (∀i:ℕ((f (i 1)) < (f i))))

Lemma: DCC_wf
[T:Type]. ∀[<:T ⟶ T ⟶ ℙ].  (DCC(T;<) ∈ ℙ)

Definition: order-type-less
WFO has strictly smaller order type than WFO when
there is an order-preserving map from to whose range
is bounded by some member of y.⋅

order-type-less() ==
  λx,y. let A,R,wf1 in 
       let B,S,wf2 in 
       ∃f:A ⟶ B. ∃b:B. (order-preserving(A;B;a1,a2.a1 a2;b1,b2.b1 b2;f) ∧ (∀a:A. ((f a) b)))

Definition: WFO
WFO (well founded order) is type together with relation on the type
that satisfies the DCC (descending chain condition).⋅

WFO{i:l}() ==  A:Type × <:A ⟶ A ⟶ ℙ × DCC(A;<)

Lemma: WFO_wf
WFO{i:l}() ∈ 𝕌'

Lemma: order-type-less_wf
order-type-less() ∈ WFO{i:l}() ⟶ WFO{i:l}() ⟶ ℙ'

Definition: TRO
TRO{i:l}() ==  {tr:A:Type × (A ⟶ A ⟶ ℙ)| Trans(fst(tr);a1,a2.a1 (snd(tr)) a2)} 

Lemma: TRO_wf
TRO{i:l}() ∈ 𝕌'

Lemma: order-type-less-transitive
Trans(WFO{i:l}();x,y.order-type-less() y)

Definition: ot-less-trans
ot-less-trans() ==
  λa,b,c,%,%1. let A,<,c2 in 
              let A1,<1,b2 in 
              let A2,<2,a2 in 
              let f,b,%4,%5 %1 
              in let f1,b3,%8,%9 
                 in <f1, b, λa1,a3,%. (%4 (f1 a1) (f1 a3) (%8 a1 a3 %)), λa.(%5 (f1 a))>    

Lemma: order-type-less-transitive-ext
Trans(WFO{i:l}();x,y.order-type-less() y)

Lemma: ot-less-trans_wf
ot-less-trans() ∈ Trans(WFO{i:l}();x,y.order-type-less() y)

Lemma: DCC-order-type-less
DCC(WFO{i:l}();order-type-less())

Definition: DCC-order-type
DCC-order-type() ==
  λf,%. let x,x,y in 
       
       n.(primrec(n;λx.x;λm,H,x. (H let x1,x1,v3 (m 1) in let x1,x1,v5 in let x1,v2 in x1 x)) 
            let A,<,v3 (n 1) in 
            let A1,<1,v5 in 
            let f@0,f@0,v1 in 
            f@0)) 
       i.if (i 1) < (1)
              then Ax
              else (letrec f@0(n)=λa1,a2,%13. if (n) < (1)
                                                 then %13
                                                 else (f@0 (n 1) 
                                                       let A,<,v3 ((n 1) 1) in 
                                                       let A1,<1,v5 (n 1) in 
                                                       let A,v2 (n 1) 
                                                       in a1 
                                                       let A,<,v3 ((n 1) 1) in 
                                                       let A1,<1,v5 (n 1) in 
                                                       let A,v2 (n 1) 
                                                       in a2 
                                                       let A,<,v3 ((n 1) 1) in 
                                                       let A1,<1,v5 (n 1) in 
                                                       let f@0,b1,%22,%23 (n 1) 
                                                       in %22 a1 a2 %13  in
                     f@0(i) 
                    let A,<,v3 (i 1) in 
                    let A1,<1,v5 in 
                    let A1,v4 
                    in A1 let A1,<1,v4 ((i 1) 1) in let f@0,f@0,v1 (i 1) in f@0 
                    let A,<,v3 (i 1) in 
                    let A1,<1,v5 in 
                    let f@0,f@0,v1 in 
                    f@0 
                    let x,<,y1 (i 1) in 
                    let A1,<1,v5 in 
                    let x1,x1,%11,%12 
                    in %12 let A,<1,v3 ((i 1) 1) in let f@0,f@0,v1 (i 1) in f@0  ))

Lemma: DCC-order-type-less-ext
DCC(WFO{i:l}();order-type-less())

Lemma: DCC-order-type_wf
DCC-order-type() ∈ DCC(WFO{i:l}();order-type-less())

Definition: WFTRO
WFTRO{i:l}() ==  A:Type × <:A ⟶ A ⟶ ℙ × DCC(A;<) × Trans(A;x,y.x < y)

Lemma: WFTRO_wf
WFTRO{i:l}() ∈ 𝕌'

Definition: max-WO
This is member of WFO{i':l} in which all members of WFO{i:l}() will embed.
(If ⌜Type ∈ Type⌝ this would lead to contradiction.)⋅

max-WO{i:l}() ==  <WFO{i:l}(), order-type-less(), DCC-order-type()>

Lemma: max-WO_wf
max-WO{i:l}() ∈ WFO{i':l}

Definition: max-WFTO
This is the same maximal WFO as max-WO{i:l}() together with the
witness that its ordering is transitive.⋅

max-WFTO{i:l}() ==  <WFO{i:l}(), order-type-less(), DCC-order-type(), ot-less-trans()>

Lemma: max-WFTO_wf
max-WFTO{i:l}() ∈ WFTRO{i':l}

Lemma: order-type-less-maximal
x:WFTRO{i:l}(). (x order-type-less() max-WO{i:l}())

Definition: Maximal_order_type
Maximal_order_type() ==
  λx.let A,<,x3,x4 
     in <λa.<b:A × (< a), λp1,p2. (< (fst(p1)) (fst(p2))), λf,%. Ax>, <A, <x3>, λa1,a2,%. <λp1.let a,p2 p1 in <a, \000Cx4 a1 a2 p2 %>, <a1, %>, λa1@0,a2,%. %, λa.(snd(a))>, λa.<λx.(fst(x)), a, λa1,a2,%. %, λa.(snd(a))>>  

Lemma: order-type-less-maximal-ext
x:WFTRO{i:l}(). (x order-type-less() max-WO{i:l}())

Lemma: Maximal_order_type_wf
Maximal_order_type() ∈ ∀x:WFTRO{i:l}(). (x order-type-less() max-WO{i:l}())

Lemma: Girard-theorem
¬(Type ∈ Type)

Lemma: Girard-theorem-extract
¬(Type ∈ Type)

Definition: descending
descending(a,b.<[a; b];L) ==  ∀i:ℕ||L|| 1. <[L[i 1]; L[i]]

Lemma: descending_wf
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ]. ∀[L:A List].  (descending(a,b.<[a;b];L) ∈ ℙ)

Lemma: descending-append
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].
  ∀L1,L2:A List.
    (descending(a,b.<[a;b];L1 L2)
    ⇐⇒ descending(a,b.<[a;b];L1)
        ∧ descending(a,b.<[a;b];L2)
        ∧ (<[hd(L2);last(L1)]) supposing (0 < ||L2|| and 0 < ||L1||))

Definition: Des
Des(A;a,b.<[a; b]) ==  {L:A List| descending(a,b.<[a; b];L)} 

Lemma: Des_wf
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].  (Des(A;a,b.<[a;b]) ∈ ℙ)

Definition: llex
llex(A;a,b.<[a; b]) ==
  λL1,L2. ((||L1|| < ||L2|| ∧ (∀i:ℕ||L1||. (L1[i] L2[i] ∈ A)))
         ∨ (∃i:ℕ(i < ||L1|| ∧ i < ||L2|| ∧ (∀j:ℕi. (L1[j] L2[j] ∈ A)) ∧ <[L1[i]; L2[i]])))

Lemma: llex_wf
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].  (llex(A;a,b.<[a;b]) ∈ (A List) ⟶ (A List) ⟶ ℙ)

Lemma: llex_transitivity
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].
  (Trans(A;a,b.<[a;b])
   (∀as,bs,cs:A List.  ((as llex(A;a,b.<[a;b]) bs)  (bs llex(A;a,b.<[a;b]) cs)  (as llex(A;a,b.<[a;b]) cs))))

Lemma: nil-llex
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].  ∀L:A List. ((L [] ∈ (A List)) ∨ ([] llex(A;a,b.<[a;b]) L))

Lemma: decidable__llex
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].
  ((∀a,b:A.  (Dec(<[a;b]) ∧ Dec(a b ∈ A)))  (∀L1,L2:A List.  Dec(L1 llex(A;a,b.<[a;b]) L2)))

Lemma: llex-linear
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].
  ((∀a,b:A.  (<[a;b] ∨ (a b ∈ A) ∨ <[b;a]))
   (∀L1,L2:A List.  ((L1 llex(A;a,b.<[a;b]) L2) ∨ (L1 L2 ∈ (A List)) ∨ (L2 llex(A;a,b.<[a;b]) L1))))

Lemma: llex-irreflexive
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].  ((∀a:A. (¬<[a;a]))  (∀[L:A List]. (L llex(A;a,b.<[a;b]) L))))

Lemma: llex-append1
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].
  ∀L1,L2:A List. ∀a:A.
    ((L1 llex(A;a,b.<[a;b]) (L2 [a]))
     ((L1 llex(A;a,b.<[a;b]) L2) ∨ (∃L3:A List. ((L1 (L2 L3) ∈ (A List)) ∧ <[hd(L3);a] supposing 0 < ||L3||))))

Definition: llex-le
llex-le(A;a,b.<[a; b]) ==  λL1,L2. ((L1 llex(A;a,b.<[a; b]) L2) ∨ (L1 L2 ∈ (A List)))

Lemma: llex-le_wf
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].  (llex-le(A;a,b.<[a;b]) ∈ (A List) ⟶ (A List) ⟶ ℙ)

Lemma: llex-le-order
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].
  ((∀a:A. (¬<[a;a]))  Trans(A;a,b.<[a;b])  Order(A List;as,bs.as llex-le(A;a,b.<[a;b]) bs))

Lemma: llex-le-lin-order
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].
  ((∀a:A. (¬<[a;a]))
   Trans(A;a,b.<[a;b])
   (∀a,b:A.  (<[a;b] ∨ (a b ∈ A) ∨ <[b;a]))
   Linorder(A List;as,bs.as llex-le(A;a,b.<[a;b]) bs))

Lemma: decidable__llex-le
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].
  ((∀a,b:A.  (Dec(<[a;b]) ∧ Dec(a b ∈ A)))  (∀L1,L2:A List.  Dec(L1 llex-le(A;a,b.<[a;b]) L2)))

Lemma: wellfounded-llex
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].
  ((∀a,b:A.  SqStable(<[a;b]))
   WellFnd{i}(A;a,b.<[a;b])
   WellFnd{i}(Des(A;a,b.<[a;b]);L1,L2.L1 llex(A;a,b.<[a;b]) L2))

Lemma: wellfounded-llex-ext
[A:Type]. ∀[<:A ⟶ A ⟶ ℙ].
  ((∀a,b:A.  SqStable(<[a;b]))
   WellFnd{i}(A;a,b.<[a;b])
   WellFnd{i}(Des(A;a,b.<[a;b]);L1,L2.L1 llex(A;a,b.<[a;b]) L2))

Lemma: wellfounded-lex
[A:Type]. ∀[<A:A ⟶ A ⟶ ℙ].
  (WellFnd{i}(A;a,b.<A[a;b])
   (∀[B:Type]. ∀[<B:B ⟶ B ⟶ ℙ].
        (WellFnd{i}(B;a,b.<B[a;b])
         WellFnd{i}(A × B;p,q.<A[fst(p);fst(q)] ∨ (((fst(p)) (fst(q)) ∈ A) ∧ <B[snd(p);snd(q)])))))

Lemma: map_is_append
[A,B:Type]. ∀[f:A ⟶ B]. ∀[L:A List]. ∀[L1,L2:B List].
  {(map(f;firstn(||L1||;L)) L1 ∈ (B List)) ∧ (map(f;nth_tl(||L1||;L)) L2 ∈ (B List))} 
  supposing map(f;L) (L1 L2) ∈ (B List)

Lemma: map_is_cons
[A,B:Type]. ∀[f:A ⟶ B]. ∀[L:A List]. ∀[L2:B List]. ∀[b:B].
  {((f hd(L)) b ∈ B) ∧ (map(f;tl(L)) L2 ∈ (B List))} supposing map(f;L) [b L2] ∈ (B List)

Lemma: decidable-last-rel
[T:Type]. ∀[P:(T List) ⟶ T ⟶ ℙ].
  ((∀L:T List. ∀x:T.  Dec(P[L;x]))  (∀L:T List. Dec(∃L':T List. ∃x:T. ((L (L' [x]) ∈ (T List)) ∧ P[L';x]))))

Lemma: decidable-exists-iseg
[T:Type]. ∀[P:(T List) ⟶ ℙ].  ((∀L:T List. Dec(P[L]))  (∀L:T List. Dec(∃L':T List. (L' ≤ L ∧ P[L']))))

Lemma: decidable__l_exists_better-extract
[A:Type]. ∀[F:A ⟶ ℙ].  ∀L:A List. ((∀k:A. Dec(F[k]))  Dec((∃k∈L. F[k])))

Lemma: decidable__l_all-better-extract
[A:Type]. ∀L:A List. ∀[F:{a:A| (a ∈ L)}  ⟶ ℙ]. ((∀k:{a:A| (a ∈ L)} Dec(F[k]))  Dec((∀k∈L.F[k])))

Lemma: first-iseg
[T:Type]. ∀[P:(T List) ⟶ ℙ].
  ((∀L:T List. Dec(P[L]))
   (∀L:T List
        (P[L]  (∃L':T List. (L' ≤ L ∧ P[L'] ∧ (∀L'':T List. (L'' ≤ L'  P[L'']  (L'' L' ∈ (T List)))))))))

Lemma: iseg-transition-lemma
[T:Type]. ∀[P:(T List) ⟶ ℙ].
  ∀L:T List. ∀x:T.
    ((∃L1:T List. (L1 ≤ [x] ∧ P[L1])) ∧ (∃L1:T List. (L1 ≤ L ∧ P[L1])))
    ⇐⇒ P[L [x]] ∧ (∃L1:T List. (L1 ≤ L ∧ P[L1]))))

Lemma: list-list-concat-type
[T:Type]. ∀[L:T List List].  (L ∈ {x:T| (x ∈ concat(L))}  List List)

Lemma: prime-factors-unique
ps:{m:ℕprime(m)}  List. ∀qs:{qs:{m:ℕprime(m)}  List| Π(ps)  = Π(qs)  ∈ ℤ.  permutation(ℤ;ps;qs)

Lemma: bag-product-primes
b:bag(Prime). 0 < Π(b)

Lemma: factors-prime-product
[b:bag(Prime)]. (factors(Π(b)) b ∈ bag(Prime))

Lemma: prime-product-injection
Inj(bag(Prime);ℕ+b.Π(b))

Lemma: fund-theorem-arithmetic
Bij(ℕ+;bag(Prime);λn.factors(n))

Lemma: append-factors
n,m:ℕ+.  (factors(n m) (factors(n) factors(m)) ∈ bag(Prime))

Lemma: nat-plus-ind-primes
[P:ℕ+ ⟶ ℙ]. (P[1]  (∀p:Prime. P[p])  (∀n,m:ℕ+.  (P[n]  P[m]  P[n m]))  (∀n:ℕ+P[n]))

Lemma: Euler-four-square-identity
[a1,b1,c1,d1,a,b,c,d:ℤ].
  ((((a1 a1) (b1 b1) (c1 c1) (d1 d1)) ((a a) (b b) (c c) (d d)))
  ((((a1 a) ((-b1) b) ((-c1) c) ((-d1) d)) ((a1 a) ((-b1) b) ((-c1) c) ((-d1) d)))
    (((a1 b) (b1 a) (c1 d) ((-d1) c)) ((a1 b) (b1 a) (c1 d) ((-d1) c)))
    (((a1 c) (c1 a) (d1 b) ((-b1) d)) ((a1 c) (c1 a) (d1 b) ((-b1) d)))
    (((a1 d) (d1 a) (b1 c) ((-c1) b)) ((a1 d) (d1 a) (b1 c) ((-c1) b))))
  ∈ ℤ)

Lemma: prime-sum-of-two-squares-lemma
p:Prime. ∀c:ℕ. ∀a,b:ℤ.
  (((((a a) (b b)) (p c) ∈ ℤ) ∧ 0 < c ∧ c < p)  (∃a,b:ℤ(p ((a a) (b b)) ∈ ℤ)))

Lemma: prime-sum-of-two-squares
p:Prime
  ((∃a,b:ℤ((¬((a ≡ mod p) ∧ (b ≡ mod p))) ∧ (((a a) (b b)) ≡ mod p)))
   (∃a,b:ℤ(p ((a a) (b b)) ∈ ℤ)))

Lemma: prime-sum-of-four-squares
p:Prime. ∀r:ℕ. ∀a,b,c,d:ℤ.
  (((((a a) (b b) (c c) (d d)) (p r) ∈ ℤ) ∧ 0 < r ∧ r < p)
   (∃a,b,c,d:ℤ(p ((a a) (b b) (c c) (d d)) ∈ ℤ)))

Lemma: four-squares
n:ℕ+. ∃a,b,c,d:ℤ(n ((a a) (b b) (c c) (d d)) ∈ ℤ)

Lemma: implies-sum-of-two-squares
n:ℕ((∃x:ℤ-o. ∃w,y:ℤ((n x) ((w w) (y y)) ∈ ℤ))  (∃a,b:ℤ(n ((a a) (b b)) ∈ ℤ)))

Lemma: divides-iff-factors
n,m:ℕ+.  (n ⇐⇒ sub-bag(Prime;factors(n);factors(m)))

Definition: divisors-sum
Σ i|n. f[i]  ==  Σ(if (n rem =z 0) then f[i 1] else fi  i < n)

Lemma: divisors-sum_wf
[n:ℕ+]. ∀[f:ℕ+1 ⟶ ℤ].  (Σ i|n. f[i]  ∈ ℤ)

Definition: gen-divisors-sum
Σ i|n. f[i] ==  Σ(i∈[1, 1)). if (n rem =z 0) then f[i] else fi 

Lemma: gen-divisors-sum_wf
[r:CRng]. ∀[n:ℕ+]. ∀[f:ℕ+1 ⟶ |r|].  (Σ i|n. f[i] ∈ |r|)

Lemma: gen-divisors-sum-int-ring
[n:ℕ+]. ∀[f:ℕ+1 ⟶ ℤ].  (Σ i|n. f[i] = Σ i|n. f[i]  ∈ ℤ)

Definition: int-list-member
int-list-member(i;xs) ==  reduce(λj,b. ((i =z j) ∨bb);ff;xs)

Lemma: int-list-member_wf
i:ℤ. ∀xs:ℤ List.  (int-list-member(i;xs) ∈ 𝔹)

Lemma: assert-int-list-member
i:ℤ. ∀xs:ℤ List.  (↑int-list-member(i;xs) ⇐⇒ (i ∈ xs))

Lemma: int-list-member-append
[i:ℤ]. ∀[xs,ys:ℤ List].  (int-list-member(i;xs ys) int-list-member(i;xs) ∨bint-list-member(i;ys))

Definition: permute-to-front
permute-to-front(L;idxs) ==
  (L o λi.filter(λi.int-list-member(i;idxs);upto(||L||)) filter(λi.(¬bint-list-member(i;idxs));upto(||L||))[i])

Lemma: permute-to-front_wf
[T:Type]. ∀[L:T List]. ∀[idxs:ℕ List].  (permute-to-front(L;idxs) ∈ List)

Lemma: permute-to-front-permutation
[T:Type]. ∀L:T List. ∀idxs:ℕ List.  permutation(T;L;permute-to-front(L;idxs))

Definition: int-list-index
int-list-index(L;x) ==  rec-case(L) of [] => a::as => r.if (a =z x) then else fi 

Lemma: int-list-index_wf
[x:ℤ]. ∀[L:ℤ List].  (int-list-index(L;x) ∈ ℕ||L|| 1)

Lemma: int-list-index-property
x:ℤ. ∀L:ℤ List.  (((x ∈ L) ⇐⇒ int-list-index(L;x) < ||L||) ∧ ((x ∈ L)  (L[int-list-index(L;x)] x ∈ ℤ)))

Lemma: int-list-index-append
[x:ℤ]. ∀[L1,L2:ℤ List].
  (int-list-index(L1 L2;x) if int-list-member(x;L1)
  then int-list-index(L1;x)
  else ||L1|| int-list-index(L2;x)
  fi )

Definition: index-split
index-split(L;idxs) ==
  let ||filter(λi.int-list-member(i;idxs);upto(||L||))|| in
   let L' permute-to-front(L;idxs) in
   <firstn(n;L'), nth_tl(n;L')>

Lemma: index-split_wf
[T:Type]. ∀[L:T List]. ∀[ids:ℕ List].  (index-split(L;ids) ∈ List × (T List))

Lemma: index-split-permutation
[T:Type]. ∀L:T List. ∀ids:ℕ List.  let L1,L2 index-split(L;ids) in permutation(T;L;L1 L2)

Lemma: index-split_property
[T:Type]
  ∀L:T List. ∀idxs:ℕ List.
    let L1 fst(index-split(L;idxs)) in
        ∃f:ℕ||L1|| ⟶ {i:ℕ||L||| (i ∈ idxs)} (Bij(ℕ||L1||;{i:ℕ||L||| (i ∈ idxs)} ;f) ∧ (∀j:ℕ||L1||. (L1[j] L[f j] ∈ \000CT)))

Lemma: split-by-indices
[T:Type]
  ∀L:T List. ∀ids:ℕ List.
    ∃L1,L2:T List. (permutation(T;L;L1 L2) ∧ (∃f:ℕ||L1|| ⟶ {i:ℕ||L||| (i ∈ ids)} (Bij(ℕ||L1||;{i:ℕ||L||| (i ∈ ids)}\000C ;f) ∧ (∀j:ℕ||L1||. (L1[j] L[f j] ∈ T)))))

Definition: add-indices
add-indices(L) ==  map(λi.<i, L[i]>;upto(||L||))

Lemma: add-indices_wf
[T:Type]. ∀[L:T List].  (add-indices(L) ∈ (ℕ||L|| × T) List)

Lemma: select-add-indices
[T:Type]. ∀[L:T List]. ∀[i:ℕ||L||].  (add-indices(L)[i] = <i, L[i]> ∈ (ℕ||L|| × T))

Definition: hd?
hd?(L) ==  if null(L) then inr ⋅  else inl hd(L) fi 

Lemma: hd?_wf
[T:Type]. ∀[L:T List].  (hd?(L) ∈ T?)

Lemma: finite-type-product
[A:Type]. ∀[B:A ⟶ Type].  (finite-type(A)  (∀a:A. finite-type(B[a]))  finite-type(a:A × B[a]))

Lemma: finite-type-union
[A,B:Type].  (finite-type(A)  finite-type(B)  finite-type(A B))

Lemma: finite-type-unit
finite-type(Unit)

Definition: star-append
star-append(T;P;Q) ==  λL.∃LL:T List List. ∃L':T List. ((∀X∈LL.P X) ∧ (Q L') ∧ (L (concat(LL) L') ∈ (T List)))

Lemma: star-append_wf
[T:Type]. ∀[P,Q:(T List) ⟶ ℙ].  (star-append(T;P;Q) ∈ (T List) ⟶ ℙ)

Lemma: star-append-iff
[T:Type]. ∀[P,Q:(T List) ⟶ ℙ].
  ∀L:T List
    (star-append(T;P;Q) ⇐⇒ (Q L) ∨ (∃L1,L2:T List. ((L (L1 L2) ∈ (T List)) ∧ (P L1) ∧ (star-append(T;P;Q) L2))))

Lemma: finite-set-type-cases
[T:Type]
  ∀L:(T ⟶ ℙList
    ∀[P:T ⟶ ℙ]
      ((∀x:T. Dec(P[x]))
       (∀Q∈L.∀x:T. Dec(Q[x]))
       (∀Q∈L.finite-type({x:T| Q[x]} ))
       (∀x:T. (P[x]  (∃Q∈L. Q[x])))
       finite-type({x:T| P[x]} ))

Definition: quicksort
quicksort(cmp;L) ==
  fix((λquicksort,L. if null(L)
                    then L
                    else let x ⟵ hd(L)
                         in let L1 ⟵ filter(λz.0 <cmp x;L)
                            in let L2 ⟵ filter(λz.0 <cmp z;L)
                               in let L3 ⟵ filter(λz.(0 =z cmp z);L)
                                  in (quicksort L1) L3 (quicksort L2)
                    fi )) 
  L

Lemma: quicksort_wf
[T:Type]. ∀[cmp:comparison(T)].
  ∀L:T List. (quicksort(cmp;L) ∈ {srtd:T List| sorted-by(λx,y. (0 ≤ (cmp y));srtd) ∧ permutation(T;srtd;L)} supposin\000Cg valueall-type(T)

Definition: cpsquicksort
cpsquicksort(cmp;L;k) ==
  fix((λcpsquicksort,L,k. if null(L)
                         then L
                         else let hd(L) in
                               let L1 filter(λz.0 <cmp x;L) in
                               let L2 filter(λz.(0 =z cmp z);L) in
                               let L3 filter(λz.0 <cmp z;L) in
                               cpsquicksort L1 S1.(cpsquicksort L3 S3.(k (S1 L2 S3)))))
                         fi )) 
  
  k

Lemma: cpsquicksort_wf
[T,A:Type]. ∀[cmp:comparison(T)].  ∀L:T List. ∀[k:(T List) ⟶ A]. (cpsquicksort(cmp;L;k) ∈ A)

Lemma: cpsquicksort-quicksort
[T:Type]
  ∀[A:Type]. ∀[cmp:comparison(T)].  ∀L:T List. ∀[k:(T List) ⟶ A]. (cpsquicksort(cmp;L;k) quicksort(cmp;L)) 
  supposing valueall-type(T)

Definition: quicksort-int
quicksort-int(L) ==  quicksort(λi,j. (j i);L)

Lemma: quicksort-int_wf
L:ℤ List. (quicksort-int(L) ∈ {srtd:ℤ List| sorted-by(λx,y. (x ≤ y);srtd) ∧ permutation(ℤ;srtd;L)} )

Lemma: quicksort-int-length
[L:ℤ List]. (||L|| ||quicksort-int(L)|| ∈ ℤ)

Lemma: quicksort-int-member
L:ℤ List. ∀x:ℤ.  ((x ∈ L) ⇐⇒ (x ∈ quicksort-int(L)))

Lemma: quicksort-int-prop1
[L:ℤ List]. ∀[i:ℕ||L||]. ∀[j:ℕi].  (quicksort-int(L)[j] ≤ quicksort-int(L)[i])

Lemma: quicksort-int-iseg
[L,L':ℤ List]. ∀[n:ℤ].  ∀[x:ℤ]. x ≤ supposing (x ∈ L') supposing L' [n] ≤ quicksort-int(L)

Lemma: quicksort-int-nil
quicksort-int([]) []

Lemma: quicksort-int-single
[n:ℤ]. (quicksort-int([n]) [n])

Lemma: quicksort-int-last
[L:ℤ List]. ∀n:ℤn ≤ last(quicksort-int(L)) supposing (n ∈ L)

Lemma: quicksort-int-null
[L:ℤ List]. uiff(↑null(quicksort-int(L));↑null(L))

Lemma: map-rev-append-top
[f,a,b:Top].  (map(f;rev(a) b) rev(map(f;a)) map(f;b))

Lemma: map-reverse-top
[f,L:Top].  (map(f;rev(L)) rev(map(f;L)))

Lemma: select-front-as-reduce
[n:ℕ]. ∀[L:Top List].
  L[n] hd(reduce(λu,x. if ||x|| <then [u] else tl(x) [u] fi ;[];L)) supposing n < ||L||

Definition: sub-mset
sub-mset(T; L1; L2) ==  ∃L:T List. permutation(T;L L1;L2)

Lemma: sub-mset_wf
[T:Type]. ∀[L1,L2:T List].  (sub-mset(T; L1; L2) ∈ ℙ)

Lemma: sub-mset-contains
[T:Type]. ∀L1,L2:T List.  (sub-mset(T; L1; L2)  L1 ⊆ L2)

Lemma: member-sub-mset
[T:Type]. ∀L1,L2:T List.  (sub-mset(T; L1; L2)  (∀x:T. ((x ∈ L1)  (x ∈ L2))))

Lemma: sub-mset_weakening
[T:Type]. ∀L1,L2:T List.  (permutation(T;L1;L2)  sub-mset(T; L1; L2))

Lemma: sub-mset_transitivity
[T:Type]. ∀A,B,C:T List.  (sub-mset(T; A; B)  sub-mset(T; B; C)  sub-mset(T; A; C))

Lemma: sub-mset-map
[A,B:Type].  ∀f:A ⟶ B. ∀L1,L2:A List.  (sub-mset(A; L1; L2)  sub-mset(B; map(f;L1); map(f;L2)))

Definition: CV
CV(F) ==  fix((λCV,t. (F CV)))

Lemma: CV_wf
[A:Type]. ∀[F:t:ℕ ⟶ (ℕt ⟶ A) ⟶ A].  (CV(F) ∈ ℕ ⟶ A)

Lemma: CV_property
[F,t:Top].  (CV(F) CV(F))

Definition: accum_filter_rel
accum(z,x.f[z; x],a,{x∈X|P[x]}) ==
  ∀L:T List
    ((L ⊆ X ∧ (∀x:T. ((x ∈ L) ⇐⇒ (x ∈ X) ∧ P[x])))
     (b accumulate (with value and list item x): f[z; x]over list:  Lwith starting value: a) ∈ A))

Lemma: accum_filter_rel_wf
[T,A:Type]. ∀[a,b:A]. ∀[X:T List]. ∀[P:{x:T| (x ∈ X)}  ⟶ ℙ]. ∀[f:A ⟶ {x:T| (x ∈ X)}  ⟶ A].
  (b accum(z,x.f[z;x],a,{x∈X|P[x]}) ∈ ℙ)

Lemma: accum_filter_rel_nil
[T,A:Type]. ∀[a,b:A]. ∀[P,f:Top].  uiff(b accum(z,x.f[z;x],a,{x∈[]|P[x]});b a ∈ A)

Lemma: concat-map-decide
[T:Type]. ∀[B:T ⟶ (Top Top)]. ∀[L:T List].
  (concat(map(λx.case B[x] of inl(a) => [a] inr(a) => [];L)) mapfilter(λx.outl(B[x]);λx.isl(B[x]);L))

Lemma: map-decide
[b,f,L1,L2:Top].
  (map(f;case of inl(a) => L1[a] inr(a) => L2[a]) case of inl(a) => map(f;L1[a]) inr(a) => map(f;L2[a]))

Lemma: map-spread
[x,f,L:Top].  (map(f;let a,b in L[a;b]) let a,b in map(f;L[a;b]))

Lemma: concat-map-map-decide
[T:Type]. ∀[g:Top]. ∀[L:T List]. ∀[f:{x:T| (x ∈ L)}  ⟶ (Top Top)].
  (concat(map(λx.map(λy.g[x;y];case f[x] of inl(m) => [m] inr(x) => []);L)) mapfilter(λx.g[x;outl(f[x])];
                                                                                          λx.isl(f[x]);
                                                                                          L))

Lemma: void-list-equality
[x,y:Void List].  (x y)

Lemma: void-list-equality2
[x,y:Void List]. ∀[T:Type].  (x y ∈ (T List))

Lemma: void-list-equality3
[x,y:Void List].  {(↑null(x)) ∧ (↑null(y))} supposing y ∈ (Void List)

Lemma: equal-nil-lists
[T:Type]. ∀[x,y:Top List].  (x y ∈ (T List)) supposing ((↑null(y)) and (↑null(x)))

Definition: p-graph
p-graph(A;f) ==  λx,y. ((↑can-apply(f;x)) c∧ (y do-apply(f;x) ∈ A))

Lemma: p-graph_wf
[A,B:Type]. ∀[f:A ⟶ (B Top)].  (p-graph(B;f) ∈ A ⟶ B ⟶ ℙ)

Lemma: p-graph_wf2
[A:Type]. ∀[f:A ⟶ (A Top)].  (p-graph(A;f) ∈ A ⟶ A ⟶ ℙ)

Definition: final-iterate
final-iterate(f;x) ==  fix((λfinal-iterate,x. if can-apply(f;x) then final-iterate do-apply(f;x) else fi )) x

Lemma: final-iterate_wf
[A:Type]. ∀[f:A ⟶ (A Top)].  ∀[x:A]. (final-iterate(f;x) ∈ A) supposing SWellFounded(p-graph(A;f) x)

Lemma: final-iterate-property
[A:Type]
  ∀f:A ⟶ (A Top)
    (SWellFounded(p-graph(A;f) x)
     (∀x:A
          ∃n:ℕ
           ((↑can-apply(f^n;x)) c∧ ((final-iterate(f;x) do-apply(f^n;x) ∈ A) ∧ (¬↑can-apply(f;final-iterate(f;x)))))))

Lemma: same-final-iterate-one-one
[A:Type]
  ∀f:A ⟶ (A Top)
    (SWellFounded(p-graph(A;f) x)
     ∀x,y:A.
         ∃n:ℕ((p-graph(A;f^n) y) ∨ (p-graph(A;f^n) x)) supposing final-iterate(f;x) final-iterate(f;y) ∈ 
       supposing p-inject(A;A;f))

Lemma: map-upto-length
[T:Type]. ∀[L:T List]. ∀[f:ℕ||L|| ⟶ T].  map(f;upto(||L||)) ∈ (T List) supposing ∀i:ℕ||L||. ((f i) L[i] ∈ T)

Lemma: map-upto
[n:ℕ+]. ∀[f:Top].  (map(f;upto(n)) map(f;upto(n 1)) [f (n 1)])

Lemma: map-as-map-upto
[F:Top]. ∀[L:Top List].  (map(λx.F[x];L) map(λi.F[L[i]];upto(||L||)))

Lemma: mklist-general-fun
[T:Type]. ∀[h:(T List) ⟶ T]. ∀[n:ℕ].  (mklist-general(n;h) map(λn.mklist-general(n 1;h)[n];upto(n)))

Definition: l-ordered
l-ordered(T;x,y.R[x; y];L) ==  ∀x,y:T.  (x before y ∈  R[x; y])

Lemma: l-ordered_wf
[T:Type]. ∀[L:T List]. ∀[R:T ⟶ T ⟶ ℙ].  (l-ordered(T;x,y.R[x;y];L) ∈ ℙ)

Lemma: sq_stable__l-ordered
[T:Type]. ∀[L:T List]. ∀[R:T ⟶ T ⟶ ℙ].  ((∀x,y:T.  SqStable(R[x;y]))  SqStable(l-ordered(T;x,y.R[x;y];L)))

Lemma: l-ordered-nil
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  l-ordered(T;x,y.R[x;y];[])

Lemma: l-ordered-nil-true
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  (l-ordered(T;x,y.R[x;y];[]) ⇐⇒ True)

Lemma: l-ordered-append
[T:Type]
  ∀L1,L2:T List.
    ∀[R:T ⟶ T ⟶ ℙ]
      (l-ordered(T;x,y.R[x;y];L1 L2)
      ⇐⇒ l-ordered(T;x,y.R[x;y];L1) ∧ l-ordered(T;x,y.R[x;y];L2) ∧ (∀x,y:T.  ((x ∈ L1)  (y ∈ L2)  R[x;y])))

Lemma: l-ordered-single
[T:Type]. ∀x:T. ∀[R:T ⟶ T ⟶ ℙ]. l-ordered(T;x,y.R[x;y];[x])

Lemma: l-ordered-cons
[T:Type]
  ∀x:T. ∀L:T List.
    ∀[R:T ⟶ T ⟶ ℙ]. (l-ordered(T;x,y.R[x;y];[x L]) ⇐⇒ l-ordered(T;x,y.R[x;y];L) ∧ (∀y:T. ((y ∈ L)  R[x;y])))

Lemma: l-ordered-inst
[T:Type]. ∀L:T List. ∀R:T ⟶ T ⟶ ℙ. ∀i:ℕ||L||. ∀j:ℕi.  (l-ordered(T;x,y.R[x;y];L)  R[L[j];L[i]])

Lemma: l-ordered-from-upto-lt
[n,m:ℤ].  l-ordered(ℤ;x,y.x < y;[n, m))

Lemma: l-ordered-from-upto-lt-true
[n,m:ℤ].  (l-ordered(ℤ;x,y.x < y;[n, m)) ⇐⇒ True)

Lemma: l-ordered-from-upto-lt-nat
[n,m:ℕ].  l-ordered(ℕ;x,y.x < y;[n, m))

Lemma: l-ordered-from-upto-lt-nat-true
[n,m:ℕ].  (l-ordered(ℕ;x,y.x < y;[n, m)) ⇐⇒ True)

Lemma: l-ordered-remove-combine
T:Type. ∀R:T ⟶ T ⟶ ℙ. ∀cmp:T ⟶ ℤ. ∀L:T List.
  (l-ordered(T;x,y.R[x;y];L)  l-ordered(T;x,y.R[x;y];remove-combine(cmp;L)))

Lemma: l-ordered-insert-combine
T:Type. ∀R:T ⟶ T ⟶ ℙ. ∀cmp:comparison(T). ∀f:T ⟶ T ⟶ T. ∀x:T.
  ((∀u,x,y:T.  (R[u;x]  R[x;y]  R[u;y]))
   (∀u,x,y:T.  (((cmp u) 0 ∈ ℤ R[x;y]  R[u;y]))
   (∀u,x,y:T.  (((cmp u) 0 ∈ ℤ R[x;y]  R[x;u]))
   (∀u,x:T.  (((cmp u) 0 ∈ ℤ ((cmp (f u)) 0 ∈ ℤ)))
   (∀x,y:T.  (0 < cmp  R[x;y]))
   (∀L:T List. (l-ordered(T;x,y.R[x;y];L)  l-ordered(T;x,y.R[x;y];insert-combine(cmp;f;x;L)))))

Lemma: l-ordered-is-sorted-by
[T:Type]. ∀R:T ⟶ T ⟶ ℙ. ∀L:T List.  (l-ordered(T;x,y.R[x;y];L) ⇐⇒ sorted-by(λx,y. R[x;y];L))

Lemma: l-ordered-eq-rels
T:Type. ∀R1,R2:T ⟶ T ⟶ ℙ. ∀L:T List.
  ((∀x∈L.(∀y∈L.R1[x;y] ⇐⇒ R2[x;y]))  l-ordered(T;x,y.R1[x;y];L)  l-ordered(T;x,y.R2[x;y];L))

Lemma: sorted-by-eq-rels
T:Type. ∀R1,R2:T ⟶ T ⟶ ℙ. ∀L:T List.
  ((∀x∈L.(∀y∈L.R1[x;y] ⇐⇒ R2[x;y]))  sorted-by(λx,y. R1[x;y];L)  sorted-by(λx,y. R2[x;y];L))

Lemma: l-ordered-filter
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  ∀P:T ⟶ 𝔹. ∀L:T List.  (l-ordered(T;x,y.R[x;y];L)  l-ordered(T;x,y.R[x;y];filter(P;L)))

Lemma: l-ordered-filter2
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  ∀L:T List. ∀P:{x:T| (x ∈ L)}  ⟶ 𝔹.  (l-ordered(T;x,y.R[x;y];L)  l-ordered(T;x,y.R[x;y];filter(P;L)))

Lemma: l-ordered-decomp
[T:Type]. ∀[R:T ⟶ T ⟶ 𝔹]. ∀[x:T].
  ∀[L:T List]. (filter(λy.R[y;x];L) filter(λy.(¬bR[y;x]);L)) ∈ (T List) supposing l-ordered(T;x,y.↑R[x;y];L) 
  supposing Trans(T;x,y.↑R[x;y])

Lemma: l-ordered-decomp2
[T:Type]. ∀[R:T ⟶ T ⟶ 𝔹]. ∀[x:T].
  (∀[L:T List]
     (L (filter(λy.R[y;x];L) [x filter(λy.R[x;y];L)]) ∈ (T List)) supposing 
        (l-ordered(T;x,y.↑R[x;y];L) and 
        (x ∈ L))) supposing 
     (Trans(T;x,y.↑R[x;y]) and 
     Irrefl(T;x,y.↑R[x;y]) and 
     StAntiSym(T;x,y.↑R[x;y]))

Lemma: l-ordered-sublist
[A:Type]. ∀R:A ⟶ A ⟶ ℙ. ∀as,bs:A List.  (as ⊆ bs  l-ordered(A;x,y.R[x;y];bs)  l-ordered(A;x,y.R[x;y];as))

Lemma: l-ordered-remove-repeats-fun
[A,B:Type].
  ∀R:A ⟶ A ⟶ ℙ. ∀eq:EqDecider(B). ∀f:A ⟶ B. ∀L:A List.
    (l-ordered(A;x,y.R[x;y];L)  l-ordered(A;x,y.R[x;y];remove-repeats-fun(eq;f;L)))

Lemma: iseg-l-ordered
[T:Type]. ∀R:T ⟶ T ⟶ ℙ. ∀a,b:T List.  (a ≤  l-ordered(T;x,y.R[x;y];b)  l-ordered(T;x,y.R[x;y];a))

Lemma: sublist-if-no_repeats
[A:Type]
  ∀R:A ⟶ A ⟶ 𝔹. ∀as,bs:A List.
    (StAntiSym(A;x,y.↑R[x;y])
     Irrefl(A;x,y.↑R[x;y])
     Trans(A;x,y.↑R[x;y])
     no_repeats(A;as)
     l-ordered(A;x,y.↑R[x;y];as)
     l-ordered(A;x,y.↑R[x;y];bs)
     (∀a∈as.(a ∈ bs))
     as ⊆ bs)

Lemma: sub-bag-list-if-bag-no-repeats-sq
[A:Type]. ∀as:bag(A). ∀bs:A List.  (bag-no-repeats(A;as)  b_all(A;as;x.(x ∈ bs))  (↓sub-bag(A;as;bs)))

Lemma: remove-combine-l-ordered-implies-member
[T:Type]
  ∀cmp:T ⟶ ℤ. ∀x:T. ∀l:T List.
    (l-ordered(T;x,y.cmp x < cmp y;l)  (x ∈ remove-combine(cmp;l))  ((cmp x) 0 ∈ ℤ)))

Lemma: no_repeats-before-equality
[T:Type]
  ∀as,bs:T List.
    (as bs ∈ (T List)
       ⇐⇒ (∀x:T. ((x ∈ as) ⇐⇒ (x ∈ bs))) ∧ (∀x,y:T.  (x before y ∈ as ⇐⇒ before y ∈ bs))) supposing 
       (no_repeats(T;bs) and 
       no_repeats(T;as))

Lemma: l-ordered-no_repeats
[T:Type]. ∀[as:T List]. ∀[R:T ⟶ T ⟶ ℙ].
  (no_repeats(T;as)) supposing (l-ordered(T;x,y.R[x;y];as) and (∀x:T. R[x;x])))

Lemma: no_repeats-permute
[T:Type]. ∀[as,bs:T List].  uiff(no_repeats(T;as bs);no_repeats(T;bs as))

Lemma: l_member-permute
[T:Type]. ∀as,bs:T List. ∀x:T.  ((x ∈ as bs) ⇐⇒ (x ∈ bs as))

Lemma: l-ordered-reorder
[A:Type]
  ∀R:A ⟶ A ⟶ 𝔹. ∀L:A List.
    (Trans(A;x,y.↑R[x;y])
     (∀x∈L.(∀y∈L.(¬↑R[x;y])  (↑R[y;x])))
     (∃L':A List. (l-ordered(A;x,y.↑R[x;y];L') ∧ permutation(A;L;L'))))

Lemma: split-at-first
[T:Type]. ∀[P:T ⟶ ℙ].
  ((∀x:T. Dec(P[x]))
   (∀L:T List. ∃X,Y:T List. ((L (X Y) ∈ (T List)) ∧ (∀x∈X.¬P[x]) ∧ P[hd(Y)] supposing ||Y|| ≥ )))

Lemma: split-at-first-rel
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  ((∀x,y:T.  Dec(R[x;y]))
   (∀L:T List
        (∃XY:T List × (T List) [let X,Y XY 
                                in (L (X Y) ∈ (T List))
                                   ∧ (∀i:ℕ||X|| 1. R[X[i];X[i 1]])
                                   ∧ ((¬↑null(L))  ((¬↑null(X)) ∧ ¬R[last(X);hd(Y)] supposing ||Y|| ≥ ))])))

Lemma: split-at-first-gap
[T:Type]
  ∀f:T ⟶ ℤ. ∀L:T List.
    (∃XY:T List × (T List) [let X,Y XY 
                            in (L (X Y) ∈ (T List))
                               ∧ (∀i:ℕ||X|| 1. ((f X[i 1]) ((f X[i]) 1) ∈ ℤ))
                               ∧ ((¬↑null(L))
                                  ((¬↑null(X)) ∧ ¬((f hd(Y)) ((f last(X)) 1) ∈ ℤsupposing ||Y|| ≥ ))])

Lemma: split-at-first-gap-ext
[T:Type]
  ∀f:T ⟶ ℤ. ∀L:T List.
    (∃XY:T List × (T List) [let X,Y XY 
                            in (L (X Y) ∈ (T List))
                               ∧ (∀i:ℕ||X|| 1. ((f X[i 1]) ((f X[i]) 1) ∈ ℤ))
                               ∧ ((¬↑null(L))
                                  ((¬↑null(X)) ∧ ¬((f hd(Y)) ((f last(X)) 1) ∈ ℤsupposing ||Y|| ≥ ))])

Definition: split-gap
split-gap(f;L) ==
  rec-case(L) of
  [] => <Ax, Ax>
  a::L =>
   r.if Ax then <<a, Ax>Ax> otherwise let u1,v1 
                             in if u1=(f a) then let X1,X2 in <<a, X1>X2> else <<a, Ax>u1, v1>

Lemma: split-gap_wf
[T:Type]
  ∀f:T ⟶ ℤ. ∀L:T List.
    (split-gap(f;L) ∈ ∃XY:T List × (T List) [let X,Y XY 
                                             in (L (X Y) ∈ (T List))
                                                ∧ (∀i:ℕ||X|| 1. ((f X[i 1]) ((f X[i]) 1) ∈ ℤ))
                                                ∧ ((¬↑null(L))
                                                   ((¬↑null(X))
                                                     ∧ ¬((f hd(Y)) ((f last(X)) 1) ∈ ℤsupposing ||Y|| ≥ ))])

Lemma: l-ordered-equality
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (∀as,bs:T List.
     (l-ordered(T;x,y.R[x;y];as)
      l-ordered(T;x,y.R[x;y];bs)
      (as bs ∈ (T List) ⇐⇒ ∀x:T. ((x ∈ as) ⇐⇒ (x ∈ bs))))) supposing 
     ((∀x,y:T.  (R[x;y]  R[y;x]))) and 
     (∀x:T. R[x;x])))

Lemma: transitive-loop
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (Trans(T;x,y.R[x;y])
   (∀L:T List. ((∀i:ℕ||L|| 1. R[L[i];L[i 1]])  R[last(L);hd(L)] supposing ¬↑null(L)  (∀a∈L.(∀b∈L.R[a;b])))))

Lemma: transitive-loop2
[T:Type]
  ∀L:T List
    ∀[R:{x:T| (x ∈ L)}  ⟶ {x:T| (x ∈ L)}  ⟶ ℙ]
      (Trans({x:T| (x ∈ L)} ;x,y.R[x;y])
       (∀i:ℕ||L|| 1. R[L[i];L[i 1]])
       R[last(L);hd(L)] supposing ¬↑null(L)
       (∀a∈L.(∀b∈L.R[a;b])))

Definition: generic
Generic{f:ℕ⟶T|S[f]} ==
  ∃R:ℕ ⟶ (T List) ⟶ ℙ
   ((∀i:ℕ. ∀s:T List.  ∃s':T List. (s ≤ s' ∧ (R s')))
   ∧ (∀f:ℕ ⟶ T. ((∀i:ℕ. ∃s:T List. ((R s) ∧ (∀n:ℕ||s||. (s[n] (f n) ∈ T))))  S[f])))

Lemma: generic_wf
[T:Type]. ∀[S:(ℕ ⟶ T) ⟶ ℙ'].  (Generic{f:ℕ⟶T|S[f]} ∈ ℙ')

Lemma: generic-non-empty
[T:Type]. ∀[S:(ℕ ⟶ T) ⟶ ℙ'].  (T  Generic{f:ℕ⟶T|S[f]}  (∃f:ℕ ⟶ T. S[f]))

Definition: ratio-dist
|a/b p/q| < 1/m ==  |(a q) b| < q

Lemma: ratio-dist_wf
[a,p:ℕ]. ∀[b,q:ℕ+]. ∀[m:ℕ].  (|a/b p/q| < 1/m ∈ ℙ)

Definition: bool-size
𝔹size(k;f) ==  primrec(k;0;λn,m. (if then else fi  m))

Lemma: bool-size_wf
[k:ℕ]. ∀[f:ℕk ⟶ 𝔹].  (𝔹size(k;f) ∈ ℕ1)

Definition: seq-count
#{i<j|f eq x} ==  𝔹size(j;(eq x) f)

Lemma: seq-count_wf
[T:Type]. ∀[eq:T ⟶ T ⟶ 𝔹]. ∀[f:ℕ ⟶ T]. ∀[x:T]. ∀[j:ℕ].  (#{i<j|f eq x} ∈ ℕ1)

Definition: frequency
frequency(f;x) (p/q) ==  ∀m,k:ℕ.  ∃j:ℕ(k < c∧ |#{i<j|f eq x}/j p/q| < 1/m)

Lemma: frequency_wf
[T:Type]. ∀[eq:T ⟶ T ⟶ 𝔹]. ∀[f:ℕ ⟶ T]. ∀[x:T]. ∀[p:ℕ]. ∀[q:ℕ+].  (frequency(f;x) (p/q) ∈ ℙ)

Definition: derived-seq
derived-seq(f;s) ==  let k,g in λn.<k, λi.(f (n (g i)))>

Lemma: derived-seq_wf
[T:Type]. ∀[f:ℕ ⟶ T]. ∀[s:k:ℕ × (ℕk ⟶ ℕ)].  (derived-seq(f;s) ∈ ℕ ⟶ (k:ℕ × (ℕk ⟶ T)))

Definition: eq_seq
eq_seq(eq) ==  λs1,s2. let k1,g1 s1 in let k2,g2 s2 in (k1 =z k2) ∧b primrec(k1;tt;λi,b. ((eq (g1 i) (g2 i)) ∧b b))

Lemma: eq_seq_wf
[T:Type]. ∀[eq:T ⟶ T ⟶ 𝔹].  (eq_seq(eq) ∈ (k:ℕ × (ℕk ⟶ T)) ⟶ (k:ℕ × (ℕk ⟶ T)) ⟶ 𝔹)

Lemma: norm-factors_wf
[L:ℕ List]. (norm-factors(L) ∈ ℕ List)

Definition: spread8
let a,b,c,d,e,f,g,h in 
v[a;
  b;
  c;
  d;
  e;
  f;
  g;
  h] ==
  let a,zz1 
  in let b,zz2 zz1 
     in let c,zz3 zz2 
        in let d,zz4 zz3 
           in let e,zz5 zz4 
              in let f,zz6 zz5 
                 in let g,h zz6 
                    in v[a;
                         b;
                         c;
                         d;
                         e;
                         f;
                         g;
                         h]

Lemma: decidable__wellfound-bounded-exists
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ]. ∀[P:T ⟶ ℙ].
  ((∀x,y:T.  Dec(R y))
   (∀x:T. Dec(P[x]))
   (∀y:T. ∃L:T List. ∀x:T. ((R y)  (x ∈ L)))
   WellFnd{i}(T;x,y.R y)
   (∀y:T. Dec(∃x:T. ((R+ y) ∧ P[x]))))

Lemma: wellfounded-minimal
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ]. ∀[P:T ⟶ ℙ].
  ((∀x,y:T.  Dec(R y))
   (∀x:T. Dec(P x))
   (∀y:T. ∃L:T List. ∀x:T. ((R y)  (x ∈ L)))
   WellFnd{i}(T;x,y.R y)
   (∀z:T. ((P z)  (∃y:T. ((P y) ∧ ((R^*) z) ∧ (∀x:T. ((R+ y)  (P x)))))))))

Lemma: wellfounded-minimal-field
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (WellFnd{i}(T;x,y.R y)
   (∀x,y:T.  Dec(R y))
   (∀y:T. ∃L:T List. ∀x:T. ((R y)  (x ∈ L)))
   (∀y:T. ∃x:T. (((R^*) y) ∧ (∀z:T. (R x))))))

Lemma: closure-well-founded-total
well-founded, one-one, decidable relation
which is "retracable" (in that everything
has only finitely many predecessors)
with at most one minimal element has the
property that its transitive closure totally
orders its field.⋅

[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (WellFnd{i}(T;x,y.R y)
   (∀x,y:T.  Dec(R y))
   (∀y:T. ∃L:T List. ∀x:T. ((R y)  (x ∈ L)))
   (∀a,b:T.  (((R^*) b) ∨ ((R^*) a))) supposing 
        ((∀y,z:T.  ((∀x:T. ((¬(R y)) ∧ (R z))))  (y z ∈ T))) and 
        one-one(T;T;R)))

Definition: rel-immediate
R! ==  λx,y. ((R y) ∧ (∀z:T. ((R z) ∧ (R y)))))

Lemma: rel-immediate_wf
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  (R! ∈ T ⟶ T ⟶ ℙ)

Lemma: rel-immediate_functionality_wrt_iff
[T:Type]. ∀[R,Q:T ⟶ T ⟶ ℙ].  ((∀x,y:T.  (R ⇐⇒ y))  (∀x,y:T.  (R! ⇐⇒ Q! y)))

Lemma: rel-immediate_functionality_wrt_breqv
[T:Type]. ∀[R,Q:T ⟶ T ⟶ ℙ].  ((R <≡>{T} Q)  (R! <≡>{T} Q!))

Lemma: rel-plus-rel-immediate
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  R!+ => R+

Lemma: rel-immediate-rel-plus
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  (SWellFounded(R y)  (∀x,y:T.  Dec(∃z:T. ((R z) ∧ (R y))))  R+ => R!+)

Lemma: rel-immediate-exists
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (SWellFounded(R y)  (∀x,y:T.  Dec(∃z:T. ((R z) ∧ (R y))))  (∀y:T. ((∃x:T. (R y))  (∃x:T. (R! y)))))

Definition: sum_of_torder
sum_of_torder(T;R) ==  ∀a,b,x:T.  ((((R x) ∧ (R x)) ∨ ((R a) ∧ (R b)))  (((R b) ∨ (a b ∈ T)) ∨ (R a)))

Lemma: rel-is-immediate
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (∀x,y:T.  (R ⇐⇒ R+y)) supposing 
     ((∀a,b,c:T.  (((R b) ∧ (R c))  (b c ∈ T))) and 
     (∀x,y:T.  ((R+ y)  (R+ x)))))

Lemma: sum_of_torder_wf
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  (sum_of_torder(T;R) ∈ ℙ)

Lemma: rel-immediate-property
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (sum_of_torder(T;R)  (∀x,y,x',y':T.  ((R y)  (R! y' y)  ((R y') ∨ (x y' ∈ T)))))

Lemma: rel-immediate-preserves-order
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (Trans(T;x,y.R y)  sum_of_torder(T;R)  (∀x,y,x',y':T.  ((R y)  (R! x' x)  (R! y' y)  (R x' y'))))

Lemma: mutual-primitive-recursion
[A,B:Type].
  ∀f0:A. ∀g0:B. ∀F:ℕ ⟶ A ⟶ B ⟶ A. ∀G:ℕ ⟶ A ⟶ B ⟶ B.
    ∃f:ℕ ⟶ A
     ∃g:ℕ ⟶ B
      (((f 0) f0 ∈ A)
      ∧ ((g 0) g0 ∈ B)
      ∧ (∀s:ℕ+(((f s) F[s 1;f (s 1);g (s 1)] ∈ A) ∧ ((g s) G[s 1;f (s 1);g (s 1)] ∈ B))))

Definition: fun-path
y=f*(x) via ==
  0 < ||L||
  ∧ (y hd(L) ∈ T)
  ∧ (x last(L) ∈ T)
  ∧ (∀i:ℕ||L|| 1. ((L[i] (f L[i 1]) ∈ T) ∧ (L[i] L[i 1] ∈ T))))

Lemma: fun-path_wf
[T:Type]. ∀[f:T ⟶ T]. ∀[x,y:T]. ∀[L:T List].  (x=f*(y) via L ∈ ℙ)

Lemma: fun-path-member
[T:Type]. ∀f:T ⟶ T. ∀x,y:T. ∀L:T List.  {(x ∈ L) ∧ (y ∈ L)} supposing x=f*(y) via L

Lemma: fun-path-cons
[T:Type]. ∀[f:T ⟶ T]. ∀[L:T List]. ∀[x,y,z:T].
  uiff(z=f*(x) via [y L];{(z y ∈ T)
  ∧ ((y (f hd(L)) ∈ T) ∧ (y hd(L) ∈ T))) ∧ hd(L)=f*(x) via supposing 0 < ||L||
  ∧ y ∈ supposing ¬0 < ||L||})

Lemma: fun-path-fixedpoint
[T:Type]. ∀[f:T ⟶ T]. ∀[L:T List]. ∀[x,y,z:T].  (y z ∈ T) supposing (((f y) y ∈ T) and (y ∈ L) and z=f*(x) via L)

Lemma: fun-path-append
[T:Type]. ∀[f:T ⟶ T]. ∀[L1,L2:T List]. ∀[x,y,z:T].
  uiff(z=f*(x) via L1 [y L2];{y=f*(x) via [y L2] ∧ z=f*(y) via L1 [y]})

Lemma: fun-path-append1
[T:Type]. ∀[f:T ⟶ T]. ∀[L:T List]. ∀[x,y,z:T].
  (z=f*(x) via [x]) supposing ((¬(y x ∈ T)) and (y (f x) ∈ T) and z=f*(y) via L)

Definition: fun-connected
is f*(x) ==  ∃L:T List. y=f*(x) via L

Lemma: fun-connected_wf
[T:Type]. ∀[f:T ⟶ T]. ∀[x,y:T].  (y is f*(x) ∈ ℙ)

Lemma: fun-connected-induction
[T:Type]
  ∀f:T ⟶ T
    ∀[R:T ⟶ T ⟶ ℙ]
      ((∀x:T. R[x;x])
       (∀x,y,z:T.  (y is f*(z)  R[y;z]  R[x;z]) supposing ((¬(x y ∈ T)) and (x (f y) ∈ T)))
       {∀x,y:T.  (x is f*(y)  R[x;y])})

Lemma: fun-connected-induction2
[T:Type]
  ∀f:T ⟶ T
    ∀[R:T ⟶ T ⟶ ℙ]
      ((∀x:T. R[x;x])
       (∀x,y:T.  is f*(f y)  R[x;f y]  R[x;y] supposing ¬((f y) y ∈ T))
       {∀x,y:T.  (x is f*(y)  R[x;y])})

Lemma: fun-path-induction
[T:Type]
  ∀f:T ⟶ T
    ∀[R:T ⟶ T ⟶ (T List) ⟶ ℙ]
      ((∀x:T. R[x;x;[x]])
       (∀L:T List. ∀x,y,z:T.  (R[y;z;[y L]]  R[x;z;[x; [y L]]]) supposing ((¬(x y ∈ T)) and (x (f y) ∈ T)))
       {∀L:T List. ∀x,y:T.  R[x;y;L] supposing x=f*(y) via L})

Definition: strict-fun-connected
f+(x) ==  (x y ∈ T)) ∧ is f*(x)

Lemma: strict-fun-connected_wf
[T:Type]. ∀[f:T ⟶ T]. ∀[x,y:T].  (y f+(x) ∈ ℙ)

Lemma: strict-fun-connected_irreflexivity
[T:Type]. ∀[f:T ⟶ T]. ∀[x:T].  False supposing f+(x)

Lemma: fun-connected_weakening_eq
[T:Type]. ∀f:T ⟶ T. ∀x,y:T.  is f*(x) supposing y ∈ T

Lemma: fun-connected_weakening
[T:Type]. ∀f:T ⟶ T. ∀x,y:T.  (y f+(x)  is f*(x))

Lemma: fun-connected-step
[T:Type]. ∀f:T ⟶ T. ∀x:T.  (Dec((f x) x ∈ T)  is f*(x))

Lemma: fun-connected-step-back
[T:Type]. ∀f:T ⟶ T. ∀x,y:T.  (x is f*(y)  is f*(f y) supposing ¬(x y ∈ T))

Lemma: strict-fun-connected-step
[T:Type]. ∀f:T ⟶ T. ∀x:T.  f+(x) supposing ¬((f x) x ∈ T)

Lemma: strict-fun-connected-induction
[T:Type]
  ∀f:T ⟶ T
    ∀[R:T ⟶ T ⟶ ℙ]
      ((∀x,y,z:T.  (y is f*(z)  (R[y;z] ∨ (y z ∈ T))  R[x;z]) supposing ((¬(x y ∈ T)) and (x (f y) ∈ T)))
       {∀x,y:T.  (x f+(y)  R[x;y])})

Lemma: fun-connected_transitivity
[T:Type]. ∀f:T ⟶ T. ∀x,y,z:T.  (y is f*(x)  is f*(y)  is f*(x))

Lemma: fun-connected-test
[T:Type]. ∀f:T ⟶ T. ∀x,y,z,w:T.  (y is f*(x)  is f*(y)  is f*(z)  is f*(x))

Lemma: fun-connected-test2
[T:Type]. ∀f:T ⟶ T. ∀x:T.  is f*(x)

Lemma: fun-connected-iff-fun_exp
[T:Type]. ∀f:T ⟶ T. ((∀x:T. Dec((f x) x ∈ T))  (∀x,y:T.  (x is f*(y) ⇐⇒ ∃n:ℕ(x (f^n y) ∈ T))))

Lemma: fun-connected-tree
[T:Type]
  ∀f:T ⟶ T
    ∀x,y:T.  (x is f*(y)  (∀z:T. (x is f*(z)  (z is f*(y) ∨ is f*(z))))) 
    supposing ∀a,b:T.  (((f a) (f b) ∈ T)  ((f a) a ∈ T))  ((f b) b ∈ T))  (a b ∈ T))

Lemma: fun-connected-fixedpoint
[T:Type]. ∀[f:T ⟶ T]. ∀[x,y:T].  (x y ∈ T) supposing (((f y) y ∈ T) and is f*(y))

Lemma: fun-path-member-connected
[T:Type]. ∀f:T ⟶ T. ∀L:T List. ∀x,y:T.  ∀a:T. ((a ∈ L)  {x is f*(a) ∧ is f*(y)}) supposing x=f*(y) via L

Lemma: fun-path-before
[T:Type]. ∀f:T ⟶ T. ∀L:T List. ∀x,y,a,b:T.  before b ∈  is f*(b) supposing x=f*(y) via L

Definition: retraction
retraction(T;f) ==  ∃h:T ⟶ ℕ. ∀x:T. (((f x) x ∈ T) ∨ (f x) < x)

Lemma: retraction_wf
[T:Type]. ∀[f:T ⟶ T].  (retraction(T;f) ∈ ℙ)

Lemma: retraction-fun-path
[T:Type]
  ∀f:T ⟶ T. ∀h:T ⟶ ℕ.
    ((∀x:T. (((f x) x ∈ T) ∨ (f x) < x))  (∀L:T List. ∀x,y:T.  (x y ∈ T) ∨ y < supposing y=f*(x) via L))

Lemma: retraction-fun-path-squash
[T:Type]
  ∀f:T ⟶ T. ∀h:T ⟶ ℕ.
    ((∀x:T. (↓((f x) x ∈ T) ∨ (f x) < x))
     (∀L:T List. ∀x,y:T.  ↓(x y ∈ T) ∨ y < supposing y=f*(x) via L))

Definition: fix
f**(x) ==  fix,x. if eqof(eq) (f x) then else fix (f x) fi x

Lemma: fix_wf
[T:Type]. ∀[eq:EqDecider(T)]. ∀[f:T ⟶ T].  ∀[x:T]. (f**(x) ∈ T) supposing retraction(T;f)

Lemma: fix_property
[T:Type]. ∀eq:EqDecider(T). ∀f:T ⟶ T.  (retraction(T;f)  (∀x:T. (((f f**(x)) f**(x) ∈ T) ∧ f**(x) is f*(x))))

Lemma: fix-step
[T:Type]. ∀[eq:EqDecider(T)]. ∀[f:T ⟶ T]. ∀[x:T].  f**(f x) f**(x) ∈ supposing retraction(T;f)

Lemma: fix-connected
[T:Type]. ∀[eq:EqDecider(T)]. ∀[f:T ⟶ T].
  ∀[x,y:T].  f**(x) f**(y) ∈ supposing is f*(y) supposing retraction(T;f)

Lemma: fun-connected_antisymmetry
[T:Type]. ∀[f:T ⟶ T].  ∀[x,y:T].  (x y ∈ T) supposing (x is f*(y) and is f*(x)) supposing retraction(T;f)

Lemma: strict-fun-connected_transitivity1
[T:Type]. ∀f:T ⟶ T. (retraction(T;f)  (∀x,y,z:T.  (y f+(x)  is f*(y)  f+(x))))

Lemma: strict-fun-connected_transitivity2
[T:Type]. ∀f:T ⟶ T. (retraction(T;f)  (∀x,y,z:T.  (y is f*(x)  f+(y)  f+(x))))

Lemma: strict-fun-connected_transitivity3
[T:Type]. ∀f:T ⟶ T. (retraction(T;f)  (∀x,y,z:T.  (y f+(x)  is f*(y)  f+(x))))

Lemma: fun-path-no_repeats
[T:Type]. ∀[f:T ⟶ T].  ∀[L:T List]. ∀[x,y:T].  no_repeats(T;L) supposing x=f*(y) via supposing retraction(T;f)

Lemma: retraction-fun-path-before
[T:Type]. ∀f:T ⟶ T. (retraction(T;f)  (∀L:T List. ∀x,y,a,b:T.  before b ∈  f+(b) supposing x=f*(y) via L))

Lemma: retraction-fixedpoint
[T:Type]. ∀f:T ⟶ T. (retraction(T;f)  (∀x:T. ∃y:T. (((f y) y ∈ T) ∧ is f*(x))))

Lemma: strong-fun-connected-induction
[T:Type]
  ∀f:T ⟶ T
    ∀[R:T ⟶ T ⟶ ℙ]
      (retraction(T;f)
       (∀x:T. R[x;x])
       (∀x,y,z:T.
            (y is f*(z)  (∀u:T. (y is f*(u)  is f*(z)  R[u;z]))  R[x;z]) supposing 
               ((¬(x y ∈ T)) and 
               (x (f y) ∈ T)))
       {∀x,y:T.  (x is f*(y)  R[x;y])})

Lemma: decidable__fun-connected
[T:Type]. ∀f:T ⟶ T. (retraction(T;f)  (∀x,y:T.  Dec(x y ∈ T))  (∀x,y:T.  Dec(x is f*(y))))

Lemma: between-fun-connected
[T:Type]. ∀f:T ⟶ T. (retraction(T;f)  (∀y,x:T.  (y is f*(x)  is f*(y)  ((y x ∈ T) ∨ (y (f x) ∈ T)))))

Lemma: fun-connected-to-same
[T:Type]
  ∀f:T ⟶ T
    (retraction(T;f)
     (∀x,y:T.  Dec(x y ∈ T))
     (∀x,z:T.  (x is f*(z)  (∀y:T. (y is f*(z)  (x is f*(y) ∨ is f*(x)))))))

Definition: csm
CSM(V) ==
  C:Type
  × M:Type
  × A:V ⟶ Type
  × B:V ⟶ Type
  × G:V ⟶ V ⟶ 𝔹
  × i:V ⟶ (A i)
  × i:V ⟶ (A i) ⟶ (B i) ⟶ (C M) ⟶ (A i)
  × (i:V ⟶ {j:V| ↑(G j)}  ⟶ (A i) ⟶ (B i) ⟶ (C M) ⟶ (M Top))

Lemma: csm_wf
[V:Type]. (CSM(V) ∈ 𝕌')

Definition: csm-cmd
Cmd(sm) ==  fst(sm)

Lemma: csm-cmd_wf
[V:Type]. ∀[sm:CSM(V)].  (Cmd(sm) ∈ Type)

Definition: csm-msgtype
Msg(sm) ==  fst(snd(sm))

Lemma: csm-msgtype_wf
[V:Type]. ∀[sm:CSM(V)].  (Msg(sm) ∈ Type)

Definition: csm-type
Type(sm;i) ==  (fst(snd(snd(sm)))) i

Lemma: csm-type_wf
[V:Type]. ∀[sm:CSM(V)]. ∀[i:V].  (Type(sm;i) ∈ Type)

Definition: csm-aux
csm-aux(sm;i) ==  (fst(snd(snd(snd(sm))))) i

Lemma: csm-aux_wf
[V:Type]. ∀[sm:CSM(V)]. ∀[i:V].  (csm-aux(sm;i) ∈ Type)

Definition: csm-state
csm-state(sm;i;L) ==
  let C,M,A,B,G,Init,Nxt,Msg sm in 
  accumulate (with value and list item x):
   let b,c 
   in Nxt c
  over list:
    L
  with starting value:
   Init i)

Lemma: csm-state_wf
[V:Type]. ∀[sm:CSM(V)]. ∀[i:V]. ∀[L:(csm-aux(sm;i) × (Cmd(sm) Msg(sm))) List].  (csm-state(sm;i;L) ∈ Type(sm;i))

Definition: csm-sends
csm-sends(sm;i;j;a;b;c) ==  let C,M,A,B,G,Init,Nxt,Msg sm in (G j) ∧b can-apply(Msg b;c)

Lemma: csm-sends_wf
[V:Type]. ∀[sm:CSM(V)]. ∀[i,j:V]. ∀[a:Type(sm;i)]. ∀[b:csm-aux(sm;i)]. ∀[c:Cmd(sm) Msg(sm)].
  (csm-sends(sm;i;j;a;b;c) ∈ 𝔹)

Definition: csm-msg
csm-msg(sm;i;j;a;b;c) ==  let C,M,A,B,G,Init,Nxt,Msg sm in do-apply(Msg b;c)

Lemma: csm-msg_wf
[V:Type]. ∀[sm:CSM(V)]. ∀[i,j:V]. ∀[a:Type(sm;i)]. ∀[b:csm-aux(sm;i)]. ∀[c:Cmd(sm) Msg(sm)].
  csm-msg(sm;i;j;a;b;c) ∈ Msg(sm) supposing ↑csm-sends(sm;i;j;a;b;c)

Lemma: only-two-bools
f:⋂T:Type. (T ⟶ T ⟶ T). ((f x,y. x) ∈ (⋂T:Type. (T ⟶ T ⟶ T))) ∨ (f x,y. y) ∈ (⋂T:Type. (T ⟶ T ⟶ T))))

Lemma: pair-coding-exists
code:ℕ ⟶ (ℕ × ℕ). Surj(ℕ;ℕ × ℕ;code)

Lemma: finite-sequence-coding-exists
code:ℕ ⟶ (k:ℕ × (ℕk ⟶ ℕ)). Surj(ℕ;k:ℕ × (ℕk ⟶ ℕ);code)

Lemma: generic-countable-intersection
[T:Type]. ∀[S:ℕ ⟶ (ℕ ⟶ T) ⟶ ℙ'].  ((∀i:ℕGeneric{x:ℕ⟶T|S[i;x]})  Generic{x:ℕ⟶T|∀i:ℕS[i;x]})

Lemma: equipollent-nat-squared
ℕ ~ ℕ × ℕ

Lemma: equipollent-nat-sequences
ℕ k:ℕ × (ℕk ⟶ ℕ)

Lemma: equipollent-int_upper-nat
k:ℤ{k...} ~ ℕ

Lemma: equipollent-nat-prod-nsub
k:ℕ+. ℕ ~ ℕ × ℕk

Lemma: equipollent-nat-union-nsub
k:ℕ+. ℕ ~ ℕ + ℕk

Definition: cantor-to-fb
cantor-to-fb(b;g;n) ==  eval = Σ(b j < n) in eval (b n) in   mu(λi.(b <i ∨b(g (k i))))

Lemma: cantor-to-fb_wf
[b:ℕ ⟶ ℕ+]. ∀[g:ℕ ⟶ 𝔹]. ∀[n:ℕ].  (cantor-to-fb(b;g;n) ∈ ℕn)

Definition: fb-to-cantor
fb-to-cantor(b;f;n) ==
  eval mu(λm.n <z Σ(b j < m)) in
  eval = Σ(b j < m) in
  eval in
    (i =z m)

Lemma: fb-to-cantor_wf
[b:ℕ ⟶ ℕ+]. ∀[f:n:ℕ ⟶ ℕn]. ∀[k:ℕ].  (fb-to-cantor(b;f;k) ∈ 𝔹)

Lemma: surjection-cantor-finite-branching
b:ℕ ⟶ ℕ+. ∃F:(ℕ ⟶ 𝔹) ⟶ n:ℕ ⟶ ℕn. Surj(ℕ ⟶ 𝔹;n:ℕ ⟶ ℕn;F)

Definition: enumerate
enumerate(P;n) ==  primrec(n;mu(P);λi,r. eval r' in r' mu(λk.(P (r' k))))

Lemma: enumerate_wf
[P:ℕ ⟶ 𝔹]. ∀[n:ℕ].  enumerate(P;n) ∈ {k:ℕ| ↑(P k)}  supposing ∀n:ℕ. ∃k:ℕ((↑(P k)) ∧ (n ≤ k))

Lemma: enumerate-increases
[P:ℕ ⟶ 𝔹]. ∀[n,m:ℕ].  enumerate(P;n) < enumerate(P;m) supposing n < supposing ∀n:ℕ. ∃k:ℕ((↑(P k)) ∧ (n ≤ k))

Lemma: equipollent-nat-decidable-subset
P:ℕ ⟶ ℙ((∀n:ℕDec(P[n]))  (∀m:ℕ. ∃n:ℕ(P[n] ∧ (m ≤ n)))  ℕ {n:ℕP[n]} )

Lemma: divides-fact
m:ℕ. ∀c:{2..m 1-}.  (c (m)!)

Lemma: equipollent-primes
ℕ {p:ℕprime(p)} 

Lemma: equipollent-nat-subset
[T:Type]. ∀P:T ⟶ ℙ((∀x:T. Dec(P[x]))  (∀L:T List. ∃x:T. (P[x] ∧ (x ∈ L))))  ℕ  ℕ {x:T| P[x]} )

Lemma: equipollent-nat-subset-ext
[T:Type]. ∀P:T ⟶ ℙ((∀x:T. Dec(P[x]))  (∀L:T List. ∃x:T. (P[x] ∧ (x ∈ L))))  ℕ  ℕ {x:T| P[x]} )

Definition: all-large
large(n).P[n] ==  ∃N:ℕ. ∀n:ℕ((N ≤ n)  P[n])

Lemma: all-large_wf
[P:ℕ ⟶ ℙ]. (∀large(n).P[n] ∈ ℙ)

Lemma: all-large-and
[P,Q:ℕ ⟶ ℙ].  (∀large(n).P[n]  ∀large(n).Q[n]  ∀large(n).P[n] ∧ Q[n])

Definition: sq-all-large
large(n).{P[n]} ==  ∃N:ℕ [(∀n:ℕ((N ≤ n)  P[n]))]

Lemma: sq-all-large_wf
[P:ℕ ⟶ ℙ]. (∀large(n).{P[n]} ∈ ℙ)

Lemma: sq-all-large-and
[P,Q:ℕ ⟶ ℙ].  (∀large(n).{P[n]}  ∀large(n).{Q[n]}  ∀large(n).{P[n] ∧ Q[n]})

Definition: decide2
case x,y
of inl(a),inl(b) => ll[a; b]
 inl(a),inr(b) => lr[a; b]
 inr(a),inl(b) => rl[a; b]
 inr(a),inr(b) => rr[a; b] ==
  case x
   of inl(a) =>
   case of inl(b) => ll[a; b] inr(b) => lr[a; b]
   inr(a) =>
   case of inl(b) => rl[a; b] inr(b) => rr[a; b]

Lemma: decide2_wf
[T1,T2,T3,T4,T5:Type]. ∀[x:T1 T2]. ∀[y:T3 T4]. ∀[ll:T1 ⟶ T3 ⟶ T5]. ∀[lr:T1 ⟶ T4 ⟶ T5]. ∀[rl:T2 ⟶ T3 ⟶ T5].
[rr:T2 ⟶ T4 ⟶ T5].
  (case x,y
   of inl(a),inl(b) => ll[a;b]
    inl(a),inr(b) => lr[a;b]
    inr(a),inl(b) => rl[a;b]
    inr(a),inr(b) => rr[a;b] ∈ T5)

Definition: dontcare
(don't care) ==  ⋅

Lemma: dontcare_wf
(don't care) ∈ Top

Definition: one_or_both
one_or_both(A;B) ==  A × B

Lemma: one_or_both_wf
[A,B:Type].  (one_or_both(A;B) ∈ Type)

Lemma: oob-subtype
[A1,B1,A2,B2:Type].  (one_or_both(A1;B1) ⊆one_or_both(A2;B2)) supposing ((A1 ⊆A2) and (B1 ⊆B2))

Definition: oobboth
oobboth(bval) ==  inl bval

Lemma: oobboth_wf
[A,B:Type]. ∀[bval:A × B].  (oobboth(bval) ∈ one_or_both(A;B))

Definition: oobleft
oobleft(lval) ==  inr (inl lval) 

Lemma: oobleft_wf
[A,B:Type]. ∀[lval:A].  (oobleft(lval) ∈ one_or_both(A;B))

Definition: oobright
oobright(rval) ==  inr inr rval  

Lemma: oobright_wf
[A,B:Type]. ∀[rval:B].  (oobright(rval) ∈ one_or_both(A;B))

Definition: one_or_both_ind
one_or_both_ind(x;bval.both[bval];lval.left[lval];rval.right[rval]) ==
  case of inl(x) => both[x] inr(x) => case of inl(x) => left[x] inr(x) => right[x]

Lemma: one_or_both_ind_wf
[X:𝕌{j}]. ∀[A,B:Type]. ∀[x:one_or_both(A;B)]. ∀[both:bval:(A × B) ⟶ X]. ∀[left:lval:A ⟶ X]. ∀[right:rval:B ⟶ X].
  (one_or_both_ind(x;bval.both[bval];lval.left[lval];rval.right[rval]) ∈ X)

Lemma: one_or_both-induction
[A,B:Type]. ∀[P:one_or_both(A;B) ⟶ ℙ].
  ((∀bval:A × B. P[oobboth(bval)])
   (∀lval:A. P[oobleft(lval)])
   (∀rval:B. P[oobright(rval)])
   {∀x:one_or_both(A;B). P[x]})

Lemma: one_or_both_ind_oobboth_lemma
right,left,both,bval:Top.
  (one_or_both_ind(oobboth(bval);bval.both[bval];lval.left[lval];rval.right[rval]) both[bval])

Lemma: one_or_both_oobleft_lemma
right,left,both,lval:Top.
  (one_or_both_ind(oobleft(lval);bval.both[bval];lval.left[lval];rval.right[rval]) left[lval])

Lemma: one_or_both_ind_oobright_lemma
right,left,both,rval:Top.
  (one_or_both_ind(oobright(rval);bval.both[bval];lval.left[lval];rval.right[rval]) right[rval])

Definition: oobboth?
oobboth?(x) ==  one_or_both_ind(x;bval.tt;lval.ff;rval.ff)

Lemma: oobboth?_wf
[A,B:Type]. ∀[x:one_or_both(A;B)].  (oobboth?(x) ∈ 𝔹)

Definition: oobboth-bval
oobboth-bval(x) ==  one_or_both_ind(x;bval.bval;lval.⋅;rval.⋅)

Lemma: oobboth-bval_wf
[A,B:Type]. ∀[x:one_or_both(A;B)].  oobboth-bval(x) ∈ A × supposing ↑oobboth?(x)

Definition: oobleft?
oobleft?(x) ==  one_or_both_ind(x;bval.ff;lval.tt;rval.ff)

Lemma: oobleft?_wf
[A,B:Type]. ∀[x:one_or_both(A;B)].  (oobleft?(x) ∈ 𝔹)

Definition: oobleft-lval
oobleft-lval(x) ==  one_or_both_ind(x;bval.⋅;lval.lval;rval.⋅)

Lemma: oobleft-lval_wf
[A,B:Type]. ∀[x:one_or_both(A;B)].  oobleft-lval(x) ∈ supposing ↑oobleft?(x)

Definition: oobright?
oobright?(x) ==  one_or_both_ind(x;bval.ff;lval.ff;rval.tt)

Lemma: oobright?_wf
[A,B:Type]. ∀[x:one_or_both(A;B)].  (oobright?(x) ∈ 𝔹)

Definition: oobright-rval
oobright-rval(x) ==  one_or_both_ind(x;bval.⋅;lval.⋅;rval.rval)

Lemma: oobright-rval_wf
[A,B:Type]. ∀[x:one_or_both(A;B)].  oobright-rval(x) ∈ supposing ↑oobright?(x)

Definition: oob-hasleft
oob-hasleft(x) ==  oobleft?(x) ∨boobboth?(x)

Lemma: oob-hasleft_wf
[B,A:Type]. ∀[x:one_or_both(A;B)].  (oob-hasleft(x) ∈ 𝔹)

Definition: oob-getleft
oob-getleft(x) ==  if oobleft?(x) then oobleft-lval(x) else fst(oobboth-bval(x)) fi 

Lemma: oob-getleft_wf
[B,A:Type]. ∀[x:{x:one_or_both(A;B)| ↑oob-hasleft(x)} ].  (oob-getleft(x) ∈ A)

Definition: oob-hasright
oob-hasright(x) ==  oobright?(x) ∨boobboth?(x)

Lemma: oob-hasright_wf
[B,A:Type]. ∀[x:one_or_both(A;B)].  (oob-hasright(x) ∈ 𝔹)

Definition: oob-getright
oob-getright(x) ==  if oobright?(x) then oobright-rval(x) else snd(oobboth-bval(x)) fi 

Lemma: oob-getright_wf
[B,A:Type]. ∀[x:{x:one_or_both(A;B)| ↑oob-hasright(x)} ].  (oob-getright(x) ∈ B)

Lemma: oob-left-or-right
[B,A:Type].  ∀x:one_or_both(A;B). ((↑oob-hasleft(x)) ∨ (↑oob-hasright(x)))

Lemma: decide-ite
[b:𝔹]. ∀[x,y,f,g:Top].
  (case if then else fi  of inl(a) => f[a] inr(a) => g[a] if b
  then case of inl(a) => f[a] inr(a) => g[a]
  else case of inl(a) => f[a] inr(a) => g[a]
  fi )

Lemma: decide_bfalse_lemma
g,f:Top.  (case ff of inl(a) => f[a] inr(b) => g[b] g[⋅])

Lemma: bl-exists-singleton
[T:Type]. ∀[f:T ⟶ 𝔹]. ∀[a:T].  ((∃x∈[a].f[x])_b f[a])

Definition: transition-system
transition-system{i:l} ==  T:Type × init:T × R:T ⟶ T ⟶ ℙ × ({s:T| init (R^*) s}  ⟶ ℙ)

Lemma: transition-system_wf
transition-system{i:l} ∈ 𝕌{[2 i']}

Definition: ts-type
ts-type(ts) ==  fst(ts)

Lemma: ts-type_wf
[ts:transition-system{i:l}]. (ts-type(ts) ∈ Type)

Definition: ts-init
ts-init(ts) ==  fst(snd(ts))

Lemma: ts-init_wf
[ts:transition-system{i:l}]. (ts-init(ts) ∈ ts-type(ts))

Definition: ts-rel
ts-rel(ts) ==  fst(snd(snd(ts)))

Lemma: ts-rel_wf
[ts:transition-system{i:l}]. (ts-rel(ts) ∈ ts-type(ts) ⟶ ts-type(ts) ⟶ ℙ)

Definition: ts-reachable
ts-reachable(ts) ==  {s:ts-type(ts)| ts-init(ts) (ts-rel(ts)^*) s} 

Lemma: ts-reachable_wf
[ts:transition-system{i:l}]. (ts-reachable(ts) ∈ {T:Type| T ⊆ts-type(ts)} )

Definition: ts-final
ts-final(ts) ==  snd(snd(snd(ts)))

Lemma: ts-final_wf
[ts:transition-system{i:l}]. (ts-final(ts) ∈ ts-reachable(ts) ⟶ ℙ)

Definition: ts-refinement
ts-refinement(ts1;ts2;f) ==
  (ts-init(ts1) (ts-rel(ts1)^*) (f ts-init(ts2)))
  ∧ (∀x,y:ts-reachable(ts2).  ((x ts-rel(ts2) y)  ((f x) (ts-rel(ts1)^*) (f y))))
  ∧ (∀x:ts-reachable(ts1)
       ((ts-final(ts1) x)  (∃y:ts-reachable(ts2). ((ts-final(ts2) y) ∧ ((f y) x ∈ ts-type(ts1))))))

Lemma: ts-refinement_wf
[ts1,ts2:transition-system{i:l}]. ∀[f:ts-type(ts2) ⟶ ts-type(ts1)].  (ts-refinement(ts1;ts2;f) ∈ ℙ)

Definition: rel-preserving
λx.f[x]:T2->T1 takes R2 into R1*) ==  ∀x,y:T2.  ((x R2 y)  (f[x] (R1^*) f[y]))

Lemma: rel-preserving_wf
[T1,T2:Type]. ∀[R1:T1 ⟶ T1 ⟶ Type]. ∀[R2:T2 ⟶ T2 ⟶ Type]. ∀[f:T2 ⟶ T1].  x.f[x]:T2->T1 takes R2 into R1*) ∈ ℙ)

Lemma: rel-preserving-star
[T1,T2:Type]. ∀[R1:T1 ⟶ T1 ⟶ Type]. ∀[R2:T2 ⟶ T2 ⟶ Type].
  ∀f:T2 ⟶ T1. x.f[x]:T2->T1 takes R2 into R1*)  λx.f[x]:T2->T1 takes R2^* into R1*))

Lemma: rel-preserving-star-reachable
[T1,T2:Type]. ∀[i2:T2]. ∀[R1:T1 ⟶ T1 ⟶ Type]. ∀[R2:T2 ⟶ T2 ⟶ Type].
  ∀f:T2 ⟶ T1
    ((∀x,y:{s:T2| i2 (R2^*) s} .  ((x R2 y)  ((f x) (R1^*) (f y))))
     {∀x,y:{s:T2| i2 (R2^*) s} .  ((x (R2^*) y)  ((f x) (R1^*) (f y)))})

Lemma: rel-preserving-composes
[T1,T2,T3:Type]. ∀[R1:T1 ⟶ T1 ⟶ Type]. ∀[R2:T2 ⟶ T2 ⟶ Type]. ∀[R3:T3 ⟶ T3 ⟶ ℙ].
  ∀f:T2 ⟶ T1. ∀g:T3 ⟶ T2.
    x.f[x]:T2->T1 takes R2 into R1*)  λx.g[x]:T3->T2 takes R3 into R2*)  λx.f[g[x]]:T3->T1 takes R3 into R1*))

Lemma: ts-init_wf_reachable
[ts:transition-system{i:l}]. (ts-init(ts) ∈ ts-reachable(ts))

Lemma: ts-refinement-composition
ts1,ts2,ts3:transition-system{i:l}. ∀g:ts-type(ts3) ⟶ ts-type(ts2). ∀f:ts-type(ts2) ⟶ ts-type(ts1).
  (ts-refinement(ts1;ts2;f)  ts-refinement(ts2;ts3;g)  ts-refinement(ts1;ts3;f g))

Lemma: ts-reachable-induction
[ts:transition-system{i:l}]. ∀[P:ts-reachable(ts) ⟶ ℙ].
  ((∀x:ts-reachable(ts). SqStable(P[x]))
   (∀x,y:ts-reachable(ts).  (P[x]  (x ts-rel(ts) y)  P[y]))  {∀x:ts-reachable(ts). P[x]} 
     supposing P[ts-init(ts)])

Lemma: ts-reachable-induction3
ts:transition-system{i:l}
  ∀[P:ts-reachable(ts) ⟶ ℙ]
    (P[ts-init(ts)]
     (∀x:ts-type(ts). ((ts-init(ts) (ts-rel(ts)^*) x)  (∀y:ts-reachable(ts). (P[x]  (x ts-rel(ts) y)  P[y]))))
     {∀x:ts-type(ts). ((ts-init(ts) (ts-rel(ts)^*) x)  P[x])})

Lemma: ts-reachable-induction2
ts:transition-system{i:l}
  ∀[P:ts-reachable(ts) ⟶ ℙ]
    (P[ts-init(ts)]
     (∀x,y:ts-reachable(ts).  (P[x]  (x ts-rel(ts) y)  P[y]))
     {∀x:ts-type(ts). ((ts-init(ts) (ts-rel(ts)^*) x)  P[x])})

Lemma: ts-refinement-reachable2
ts1,ts2:transition-system{i:l}. ∀f:ts-type(ts2) ⟶ ts-type(ts1).
  (ts-refinement(ts1;ts2;f)
   (∀s1,s2:ts-reachable(ts2).  ((s1 (ts-rel(ts2)^*) s2)  ((f s1) (ts-rel(ts1)^*) (f s2)))))

Lemma: ts-refinement-reachable
[ts1,ts2:transition-system{i:l}]. ∀[f:ts-type(ts2) ⟶ ts-type(ts1)].
  ∀[s:ts-reachable(ts2)]. (f s ∈ ts-reachable(ts1)) supposing ts-refinement(ts1;ts2;f)

Definition: ts-stable
ts-stable(ts;x.P[x]) ==  ∀x,y:ts-type(ts).  (P[x]  (x ts-rel(ts) y)  P[y])

Lemma: ts-stable_wf
[ts:transition-system{i:l}]. ∀[P:ts-type(ts) ⟶ ℙ].  (ts-stable(ts;x.P[x]) ∈ ℙ)

Lemma: ts-stable-star
ts:transition-system{i:l}
  ∀[P:ts-type(ts) ⟶ ℙ]. (ts-stable(ts;x.P[x])  (∀x,y:ts-type(ts).  (P[x]  (x (ts-rel(ts)^*) y)  P[y])))

Definition: ts-stable-rel
ts-stable-rel(ts;x,y.R[x; y]) ==  ∀x,y:ts-type(ts).  ((x (ts-rel(ts)^*) y)  R[x; y])

Lemma: ts-stable-rel_wf
[ts:transition-system{i:l}]. ∀[R:ts-type(ts) ⟶ ts-type(ts) ⟶ ℙ].  (ts-stable-rel(ts;x,y.R[x;y]) ∈ ℙ)

Lemma: ts-transitive-stable
ts:transition-system{i:l}
  ∀[R:ts-type(ts) ⟶ ts-type(ts) ⟶ ℙ]
    (Refl(ts-type(ts);x,y.R[x;y])
     Trans(ts-type(ts);x,y.R[x;y])
     ts-rel(ts) => λx,y. R[x;y]
     ts-stable-rel(ts;x,y.R[x;y]))

Definition: uncurry
uncurry(n;f) ==  λa.primrec(n;f;λi,z. (z (a i)))

Lemma: uncurry_wf
[T:Type]. ∀[n:ℕ]. ∀[A:ℕn ⟶ Type]. ∀[f:funtype(n;A;T)].  (uncurry(n;f) ∈ (i:ℕn ⟶ A[i]) ⟶ T)

Lemma: uncurry1_lemma
F:Top. (uncurry(1;λa.F[a]) ~ λa.F[a 0])

Lemma: uncurry2_lemma
F:Top. (uncurry(2;λx,y. F[x;y]) ~ λa.F[a 0;a 1])

Definition: apply-nth
apply-nth(n; f; x) ==  primrec(n;λf.(f x);λi,H,f. (H f)) f

Lemma: apply-nth_wf
[T:Type]. ∀[n:ℕ]. ∀[A:ℕ1 ⟶ Type]. ∀[f:funtype(n 1;A;T)]. ∀[x:A[n]].  (apply-nth(n; f; x) ∈ funtype(n;A;T))

Comment: I-S-K combinators
The I, and combinators are defined in Error :core_2
as icombkcomb and scomb

Definition: C-comb
C-comb() ==  λf,x,y. (f x)

Lemma: C-comb_wf
[A,B,C:Type].  (C-comb() ∈ (A ⟶ B ⟶ C) ⟶ B ⟶ A ⟶ C)

Lemma: C-comb_wf_funtype
[T:Type]. ∀[n:ℕ]. ∀[A:ℕn ⟶ Type].
  C-comb() ∈ funtype(n;A;T) ⟶ funtype(n;λk.if (k =z 0) then if (k =z 1) then else fi ;T) supposing 2 ≤ n

Definition: Cn-comb
Cn-comb(n) ==  primrec(n;λf.f;λi,F,f,x. (F (C-comb() x)))

Lemma: Cn-comb_wf
[T:Type]. ∀[n,m:ℕ]. ∀[A:ℕm ⟶ Type].
  Cn-comb(n) ∈ funtype(m;A;T) ⟶ funtype(m;λk.if k <then (k 1)
                                              if (k =z n) then 0
                                              else k
                                              fi ;T) 
  supposing n < m

Lemma: finite-function-type-equality
[k:ℕ]. ∀[F:ℕk ⟶ Type].  ((i:ℕk ⟶ F[i]) (i:ℕk ⟶ map(λi.F[i];upto(k))[i]) ∈ Type)

Definition: type-recoding
TypeRecoding ==  T:Type ⟶ (L:Type List × h:T ⟶ tuple-type(L) × {j:tuple-type(L) ⟶ T| ∀s:T. ((j (h s)) s ∈ T)} )

Lemma: type-recoding_wf
TypeRecoding ∈ 𝕌'

Definition: tuple-recoding
TupleRecoding ==
  L:(Type List) ⟶ (L':Type List
                   × h:tuple-type(L) ⟶ tuple-type(L')
                   × {j:tuple-type(L') ⟶ tuple-type(L)| ∀s:tuple-type(L). ((j (h s)) s ∈ tuple-type(L))} )

Lemma: tuple-recoding_wf
TupleRecoding ∈ 𝕌'

Definition: recode-tuple
recode-tuple(f) ==
  λL.rec-case(L) of
     [] => <[], λx.x, λx.x>
     T::Ts =>
      r.let l2,h2,j2 in 
     let l1,h1,j1 in 
     if null(Ts)
     then <l1, h1, j1>
     else <l1 l2
          , λx.let a,b 
               in append-tuple(||l1||;||l2||;h1 a;h2 b)
          if null(l2) then λx.<j1 x, j2 ⋅> else λx.let a,b split-tuple(x;||l1||) in <j1 a, j2 b> fi >
     fi 

Lemma: recode-tuple_wf
[f:TypeRecoding]. (recode-tuple(f) ∈ TupleRecoding)

Lemma: fst-recode-tuple
[T:Type]. ∀[f:T ⟶ (T List × Top × Top)]. ∀[L:T List].
  ((fst((recode-tuple(f) L))) reduce(λT,X. ((fst((f T))) X);[];L) ∈ (T List))

Definition: compose-tuple-recodings
r2 r1 ==  λL1.let L2,h1,j1 r1 L1 in let L3,h2,j2 r2 L2 in <L3, h2 h1, j1 j2>

Lemma: compose-tuple-recodings_wf
[r1,r2:TupleRecoding].  (r2 r1 ∈ TupleRecoding)

Lemma: finite-function-equipollent-tuple
n:ℕ. ∀F:ℕn ⟶ Type.  i:ℕn ⟶ F[i] tuple-type(map(λi.F[i];upto(n)))

Lemma: product-equipollent-tuple
[A,B:Type].  A × tuple-type([A; B])

Lemma: product-equipollent-tuple2
[A:Type]. ∀L:Type List. A × tuple-type(L) tuple-type([A L])

Lemma: product-equipollent-tuple3
L:Type List. ∀[A:Type]. tuple-type(L) × tuple-type(L [A])

Definition: is_list_splitting
is_list_splitting(T;L;LL;L2;f) ==
  (L (concat(LL) L2) ∈ (T List))
  ∧ (∀L'∈LL.(↑(f L')) ∧ (¬↑null(L')) ∧ (∀S:T List. ((¬↑null(S))  S ≤ L'  (S L' ∈ (T List)))  (¬↑(f S)))))
  ∧ ((¬↑null(L))  (¬↑null(L2)))
  ∧ (∀S:T List. ((¬↑null(S))  S ≤ L2  (S L2 ∈ (T List)))  (¬↑(f S))))

Lemma: is_list_splitting_wf
[T:Type]. ∀[L:T List]. ∀[LL:T List List]. ∀[L2:T List]. ∀[f:(T List) ⟶ 𝔹].  (is_list_splitting(T;L;LL;L2;f) ∈ ℙ)

Lemma: sq_stable__is_list_splitting
[T:Type]. ∀[L:T List]. ∀[LL:T List List]. ∀[L2:T List]. ∀[f:(T List) ⟶ 𝔹].  SqStable(is_list_splitting(T;L;LL;L2;f))

Definition: list_split
list_split(f;L) ==
  accumulate (with value and list item v):
   let LL,L2 
   in if null(L2) then <LL, [v]>
      if L2 then <LL [L2], [v]>
      else <LL, L2 [v]>
      fi 
  over list:
    L
  with starting value:
   <[], []>)

Lemma: list_split_wf
[T:Type]. ∀[f:(T List) ⟶ 𝔹]. ∀[L:T List].
  (list_split(f;L) ∈ {p:T List List × (T List)| let LL,L2 in is_list_splitting(T;L;LL;L2;f)} )

Lemma: list_split_iseg
[T:Type]
  ∀f:(T List) ⟶ 𝔹. ∀L1,L2:T List.
    (L1 ≤ L2
     let LL1,X list_split(f;L1) 
       in let LL2,Y list_split(f;L2) 
          in ((LL1 LL2 ∈ (T List List)) ∧ X ≤ Y)
             ∨ (∃Z:T List. ∃ZZ:T List List. (((LL1 [Z ZZ]) LL2 ∈ (T List List)) ∧ X ≤ Z)))

Lemma: list_split_iseg2
[T:Type]
  ∀f:(T List) ⟶ 𝔹. ∀L1,L2:T List.
    (L1 ≤ L2
     (∀LL1,LL2:T List List. ∀X,Y:T List.
          (((LL1 LL2 ∈ (T List List)) ∧ X ≤ Y)
             ∨ (∃Z:T List. ∃ZZ:T List List. (((LL1 [Z ZZ]) LL2 ∈ (T List List)) ∧ X ≤ Z))) supposing 
             ((list_split(f;L2) = <LL2, Y> ∈ (T List List × (T List))) and 
             (list_split(f;L1) = <LL1, X> ∈ (T List List × (T List))))))

Lemma: list_split_inverse
[T:Type]. ∀[f:(T List) ⟶ 𝔹]. ∀[L:T List]. ∀[LL:T List List]. ∀[X:T List].
  (concat(LL) X) ∈ (T List) supposing list_split(f;L) = <LL, X> ∈ (T List List × (T List))

Lemma: list_split_one_one
[T:Type]. ∀[f:(T List) ⟶ 𝔹]. ∀[X,Y:T List].
  Y ∈ (T List) supposing list_split(f;X) list_split(f;Y) ∈ (T List List × (T List))

Lemma: list_split_prefix
[T:Type]. ∀[L:T List]. ∀[g:(T List) ⟶ 𝔹].
  ↑(g (snd(list_split(g;concat(fst(list_split(g;L))))))) supposing ¬↑null(fst(list_split(g;L)))

Definition: is_accum_splitting
is_accum_splitting(T;A;L;LL;L2;f;g;x) ==
  (L (concat(map(λp.(fst(p));LL)) (fst(L2))) ∈ (T List))
  ∧ (∀L'∈LL.(↑(f L'))
        ∧ (¬↑null(fst(L')))
        ∧ (∀S:T List. ((¬↑null(S))  S ≤ fst(L')  (S (fst(L')) ∈ (T List)))  (¬↑(f <S, snd(L')>)))))
  ∧ ((¬↑null(L))  (¬↑null(fst(L2))))
  ∧ (∀S:T List. ((¬↑null(S))  S ≤ fst(L2)  (S (fst(L2)) ∈ (T List)))  (¬↑(f <S, snd(L2)>))))
  ∧ ((snd(hd(LL [L2]))) x ∈ A)
  ∧ (∀i:ℕ||LL||. ((snd(LL [L2][i 1])) (g LL[i]) ∈ A))

Lemma: is_accum_splitting_wf
[T,A:Type]. ∀[L:T List]. ∀[LL:(T List × A) List]. ∀[L2:T List × A]. ∀[f:(T List × A) ⟶ 𝔹]. ∀[x:A].
[g:(T List × A) ⟶ A].
  (is_accum_splitting(T;A;L;LL;L2;f;g;x) ∈ ℙ)

Lemma: sq_stable__is_accum_splitting
[T,A:Type]. ∀[L:T List]. ∀[LL:(T List × A) List]. ∀[L2:T List × A]. ∀[f:(T List × A) ⟶ 𝔹]. ∀[x:A].
[g:(T List × A) ⟶ A].
  SqStable(is_accum_splitting(T;A;L;LL;L2;f;g;x))

Definition: accum_split
accum_split(g;x;f;L) ==
  accumulate (with value and list item v):
   let LL,L2,z in 
   if null(L2) then <LL, [v], z>
   if f <L2, z> then <LL [<L2, z>], [v], g <L2, z>>
   else <LL, L2 [v], z>
   fi 
  over list:
    L
  with starting value:
   <[], [], x>)

Lemma: accum_split_wf
[A,T:Type]. ∀[f:(T List × A) ⟶ 𝔹]. ∀[g:(T List × A) ⟶ A]. ∀[x:A]. ∀[L:T List].
  (accum_split(g;x;f;L) ∈ {p:(T List × A) List × List × A| let LL,L2 in is_accum_splitting(T;A;L;LL;L2;f;g;x)} )

Lemma: accum_split_iseg
[T,A:Type].
  ∀x:A. ∀g:(T List × A) ⟶ A. ∀f:(T List × A) ⟶ 𝔹. ∀L1,L2:T List.
    (L1 ≤ L2
     let LL1,X,z1 accum_split(g;x;f;L1) in 
       let LL2,Y,z2 accum_split(g;x;f;L2) in 
       ((LL1 LL2 ∈ ((T List × A) List)) ∧ X ≤ Y ∧ (z1 z2 ∈ A))
       ∨ (∃Z:T List. ∃ZZ:(T List × A) List. (((LL1 [<Z, z1> ZZ]) LL2 ∈ ((T List × A) List)) ∧ X ≤ Z)))

Lemma: accum_split_iseg2
[T,A:Type].
  ∀x:A. ∀g:(T List × A) ⟶ A. ∀f:(T List × A) ⟶ 𝔹. ∀L1,L2:T List.
    (L1 ≤ L2
     (∀LL1,LL2:(T List × A) List. ∀X,Y:T List. ∀z1,z2:A.
          (((LL1 LL2 ∈ ((T List × A) List)) ∧ X ≤ Y ∧ (z1 z2 ∈ A))
             ∨ (∃Z:T List
                 ∃ZZ:(T List × A) List. (((LL1 [<Z, z1> ZZ]) LL2 ∈ ((T List × A) List)) ∧ X ≤ Z))) supposing 
             ((accum_split(g;x;f;L2) = <LL2, Y, z2> ∈ ((T List × A) List × List × A)) and 
             (accum_split(g;x;f;L1) = <LL1, X, z1> ∈ ((T List × A) List × List × A)))))

Lemma: accum_split_inverse
[A,T:Type]. ∀[x:A]. ∀[g:(T List × A) ⟶ A]. ∀[f:(T List × A) ⟶ 𝔹]. ∀[L:T List]. ∀[LL:(T List × A) List]. ∀[X:T List].
[z:A].
  (concat(map(λp.(fst(p));LL)) X) ∈ (T List) 
  supposing accum_split(g;x;f;L) = <LL, X, z> ∈ ((T List × A) List × List × A)

Lemma: accum_split_one_one
[A,T:Type]. ∀[x:A]. ∀[g:(T List × A) ⟶ A]. ∀[f:(T List × A) ⟶ 𝔹]. ∀[X,Y:T List].
  Y ∈ (T List) supposing accum_split(g;x;f;X) accum_split(g;x;f;Y) ∈ ((T List × A) List × List × A)

Lemma: accum_split_prefix
[A,T:Type]. ∀[x:A]. ∀[g:(T List × A) ⟶ A]. ∀[f:(T List × A) ⟶ 𝔹]. ∀[L:T List].
  ↑(f (snd(accum_split(g;x;f;concat(map(λp.(fst(p));fst(accum_split(g;x;f;L)))))))) 
  supposing ¬↑null(fst(accum_split(g;x;f;L)))

Lemma: accum_split_prefix2
[A,T:Type]. ∀[x:A]. ∀[g:(T List × A) ⟶ A]. ∀[f:(T List × A) ⟶ 𝔹]. ∀[L:T List]. ∀[ZZ:(T List × A) List].
[Z,X:T List × A].
  accum_split(g;x;f;concat(map(λp.(fst(p));ZZ [Z]))) = <ZZ, Z> ∈ ((T List × A) List × List × A) 
  supposing accum_split(g;x;f;L) = <ZZ [Z], X> ∈ ((T List × A) List × List × A)

Definition: combination
Combination(n;T) ==  {L:T List| no_repeats(T;L) ∧ (||L|| n ∈ ℤ)} 

Lemma: combination_wf
[T:Type]. ∀[n:ℤ].  (Combination(n;T) ∈ Type)

Lemma: map_wf_combination
[A,B:Type]. ∀[f:A ⟶ B].  ∀[n:ℤ]. ∀[L:Combination(n;A)].  (map(f;L) ∈ Combination(n;B)) supposing Inj(A;B;f)

Lemma: combination_functionality
[A,B:Type].  ∀n,m:ℤ.  (A  Combination(n;A) Combination(m;B) supposing m ∈ ℤ)

Lemma: combination-decomp
[A:Type]. ∀[n:ℕ+]. ∀[L:Combination(n;A)].  {(hd(L) ∈ A) ∧ (tl(L) ∈ Combination(n 1;{a:A| ¬(a hd(L) ∈ A)} ))}

Definition: combinations_aux
combinations_aux(b;n;m) ==
  fix((λcombinations_aux,b,n,m. if (n =z 0)
                               then b
                               else eval b2 in
                                    eval n2 in
                                    eval m2 in
                                      combinations_aux b2 n2 m2
                               fi )) 
  
  
  m

Lemma: combinations_aux_wf_int
[n:ℕ]. ∀[b,m:ℤ].  (combinations_aux(b;n;m) ∈ ℤ)

Lemma: combinations_aux_wf
[n,b,m:ℕ].  (combinations_aux(b;n;m) ∈ ℕ)

Lemma: combinations_aux_linear
[n:ℕ]. ∀[b,m:ℤ].  (combinations_aux(b;n;m) (b combinations_aux(1;n;m)) ∈ ℤ)

Definition: combinations
C(n;m) ==  combinations_aux(1;n;m)

Lemma: combinations_wf_int
[n:ℕ]. ∀[m:ℤ].  (C(n;m) ∈ ℤ)

Lemma: combinations_wf
[n,m:ℕ].  (C(n;m) ∈ ℕ)

Lemma: combinations-step
[n,m:ℕ].  (C(n;m) if (n =z 0) then else C(n 1;m 1) fi )

Lemma: count-combinations
n,m:ℕ.  Combination(n;ℕm) ~ ℕC(n;m)

Lemma: combinations-positive
[n:ℕ]. ∀[m:ℤ].  0 < C(n;m) supposing n ≤ m

Lemma: divides-combinations
n:ℕ. ∀m:ℤ. ∀k:ℕ.  (k C(n;m)) supposing ((k ≤ m) and n < k)

Lemma: combinations-split
[m,n,k:ℕ].  C(n k;m) (C(k;m) C(n;m k)) ∈ ℤ supposing (n k) ≤ m

Lemma: combinations-list
[A:Type]. ∀n:ℕ(A ~ ℕ (∀k:ℕ. ∃C:A List List. ∀L:A List. ((L ∈ C) ⇐⇒ (||L|| k ∈ ℤ) ∧ no_repeats(A;L))))

Definition: combinations_aux_rem
combinations_aux_rem(b;n;m;k) ==
  fix((λcombinations_aux_rem,b,n,m. if (n =z 0)
                                   then b
                                   else eval b2 rem in
                                        eval n2 in
                                        eval m2 in
                                          combinations_aux_rem b2 n2 m2
                                   fi )) 
  
  
  m

Lemma: combinations_aux_rem_wf
[k:ℕ+]. ∀[n,b,m:ℕ].  (combinations_aux_rem(b;n;m;k) ∈ ℕ)

Lemma: combinations_aux_rem_property
[k:ℕ+]. ∀[n,b,m:ℕ].  (combinations_aux_rem(b rem k;n;m;k) (combinations_aux(b;n;m) rem k) ∈ ℤ)

Definition: iseg_product
iseg_product(i;j) ==  C((j i) 1;j)

Lemma: iseg_product_wf
[j:ℕ]. ∀[i:ℕ1].  (iseg_product(i;j) ∈ ℕ)

Lemma: iseg_product-property
i,j:ℤ. ∀k:ℕ.  (k iseg_product(i;j)) supposing ((k ≤ j) and (i ≤ k))

Lemma: iseg_product-positive
[i,j:ℕ].  (0 < iseg_product(i;j)) supposing ((i ≤ j) and 0 < i)

Lemma: iseg_product-split
[i,j,k:ℤ].  (iseg_product(i;j) iseg_product(i;k) iseg_product(k 1;j)) supposing (k < and (i ≤ k) and (1 ≤ i))

Definition: iseg_product_rem
iseg_product_rem(i;j;k) ==  combinations_aux_rem(1;(j i) 1;j;k)

Lemma: iseg_product_rem_wf
[k:ℕ+]. ∀[j:ℕ]. ∀[i:ℕ1].  (iseg_product_rem(i;j;k) ∈ ℕ)

Lemma: iseg_product_rem_property
[k,j:ℕ]. ∀[i:ℕ1].  iseg_product_rem(i;j;k) (iseg_product(i;j) rem k) ∈ ℤ supposing 1 < k

Lemma: eqmod-test
m:ℤ(((m 1) (m 1)) ≡ mod m)

Lemma: chinese-remainder1
r:ℤ. ∀s:{s':ℤCoPrime(r,s')} . ∀a,b:ℤ.  (∃x:ℤ [((x ≡ mod r) ∧ (x ≡ mod s))])

Lemma: chinese-remainder2
r,s,a,b:ℤ.  Dec(∃x:ℤ [((x ≡ mod r) ∧ (x ≡ mod s))])

Lemma: chinese-remainder2-extract
r,s,a,b:ℤ.  Dec(∃x:ℤ [((x ≡ mod r) ∧ (x ≡ mod s))])

Definition: chrem
chrem(a;r;b;s) ==
  let x,y letrec rec(p)=λq.if p=0
                             then <q, 0, 1, 1, 1, Ax, Ax, Ax>
                             else let x,y divrem(q; p) 
                                  in eval in
                                     let g,%13 rec 
                                     in let a,b,x1,y1,%18,%20,%21 %13 in 
                                         eval q' in
                                         eval b' (b q') in
                                         eval x' y1 x1 q' in
                                           <g, b, b', x', x1, Ax, Ax, Ax> in
             rec(|r|) 
            |s| 
  in let x1,x2,x3,x4,x5,x5,y in 
      if x=0
      then if a=b then inl else (inr x.Ax) )
      else let x4,r1 divrem(b a; x) 
           in if r1=0
              then eval if (r) < (0)  then ((-1) x3) x4 r  else ((1 x3) x4 r) in
                   inl z
              else (inr x.Ax) )

Lemma: chrem_wf
r,s,a,b:ℤ.  (chrem(a;r;b;s) ∈ {x:ℤ(x ≡ mod r) ∧ (x ≡ mod s)}  (∃x:ℤ [((x ≡ mod r) ∧ (x ≡ mod s))])))

Lemma: divisor-in-range
n:{2...}
  ∀[k:ℕ]
    ∀i:{1...}. ∀j:{i..i k-}.
      (∃m:ℤ [(m < n ∧ (2 ≤ m) ∧ (m n))]) ∨ (∃m:ℤ [((2 ≤ m) ∧ (i ≤ m) ∧ (m ≤ j) ∧ (m n))])) supposing j < n

Lemma: decidable__proper_divisor
n:{2...}. Dec(∃n1:ℤ [(n1 < n ∧ (2 ≤ n1) ∧ (n1 n))])

Lemma: prime-factors2
n:{2...}. (∃factors:{m:{2...}| prime(m)}  List [(n = Π(factors)  ∈ ℤ)])

Lemma: prime-factors3
n:{2...}. (∃factors:{m:{2...}| prime(m)}  List [(n = Π(factors)  ∈ ℤ)])

Definition: primefactors
primefactors(n) ==  TERMOF{prime-factors3:o, 1:l} n

Lemma: primefactors_wf
n:{2...}. (primefactors(n) ∈ {factors:{m:{2...}| prime(m)}  List| = Π(factors)  ∈ ℤ)

Definition: power-type
(T^k) ==  fix((λpower-type,k. if (k =z 0) then Unit else T × (power-type (k 1)) fi )) k

Lemma: power-type_wf
[T:Type]. ∀[k:ℕ].  ((T^k) ∈ Type)

Lemma: list-subtype-power-type
T:Type. ∀L:T List.  (L ∈ (T^||L||))

Lemma: power-type-subtype-list
T:Type. ∀k:ℕ.  ((T^k) ⊆(T List))

Lemma: power-type-length
T:Type. ∀k:ℕ. ∀x:(T^k).  (||x|| k ∈ ℤ)

Lemma: equipollent-type-unit-pair
[T:Type]. T × Unit

Lemma: equipollent-nat-powered
n:ℕ. ℕ (ℕ^n 1)

Lemma: equipollent-nat-powered2
f:n:ℕ ⟶ ℕ ⟶ (ℕ^n 1). ∀n:ℕBij(ℕ;(ℕ^n 1);f n)

Lemma: equipollent-nat-powered3
f:n:ℕ ⟶ ℕ ⟶ (ℕ^n 1). ∀n:ℕ. ∃g:(ℕ^n 1) ⟶ ℕInvFuns(ℕ;(ℕ^n 1);f n;g)

Lemma: equipollent-list-as-product
[T:Type]. List k:ℕ × (T^k)

Lemma: equipollent-nat-list-as-product
ℕ k:ℕ × (ℕ^k)

Lemma: equipollent-nat-list-nat
ℕ List ~ ℕ

Definition: divisor-test
divisor-test(n;i;j) ==
  fix((λdivisor-test,i,j. eval iseg_product_rem(i;j;n) in
                          eval better-gcd(n;p) in
                            if 1 <g
                            then if g <n
                                 then inl g
                                 else eval ((j i) ÷ 2) in
                                      case divisor-test k
                                       of inl(x) =>
                                       inl x
                                       inr(x) =>
                                       eval k' in
                                       divisor-test k' j
                                 fi 
                            else inr ⋅ 
                            fi )) 
  
  j

Lemma: divisor-test_wf
[n:ℕ]. ∀[i:ℕ+]. ∀[j:ℤ].
  (divisor-test(n;i;j) ∈ {n1:ℤn1 < n ∧ (2 ≤ n1) ∧ (n1 n)}  ∨ (gcd(n;iseg_product(i;j)) 1 ∈ ℤ)) supposing ((i ≤ j) \000Cand j < n)

Definition: proper-divisor-aux
proper-divisor-aux(n;m;b;i;j) ==
  fix((λproper-divisor-aux,b,i,j. if (b =z 0)
                                 then inr x.any x) 
                                 else case divisor-test(n;i;j)
                                       of inl(x) =>
                                       inl x
                                       inr(x) =>
                                       eval j' in
                                       eval i' in
                                       eval b' in
                                         proper-divisor-aux b' i' j'
                                 fi )) 
  
  
  j

Lemma: proper-divisor-aux_wf
[n:ℕ+]. ∀[m:ℕ].
  (proper-divisor-aux(n;m;m;1;m) ∈ Dec(∃n1:ℤ [(n1 < n ∧ (2 ≤ n1) ∧ (n1 n))])) supposing 
     (m m < and 
     n < (m m) m)

Definition: trial-division
trial-division(n;L) ==
  rec-case(L) of
  [] => inr ⋅ 
  a::as =>
   r.case r
   of inl(x) =>
   inl x
   inr(x) =>
   if a <then eval better-gcd(a;n) in if 1 <then inl else inr x  fi  else inr x  fi 

Lemma: trial-division_wf
[n:ℕ+]. ∀[L:{2...} List].  (trial-division(n;L) ∈ ∃n1:ℤ [(n1 < n ∧ (2 ≤ n1) ∧ (n1 n))] Top)

Definition: proper-divisor
proper-divisor(n) ==
  case trial-division(n;[510510; 14535931; 75186702453419])
   of inl(x) =>
   inl x
   inr(x) =>
   if n <then if (n =z 4) then inl else inr x.any x)  fi 
   if n <16 then proper-divisor-aux(n;2;2;1;2)
   if n <81 then proper-divisor-aux(n;3;3;1;3)
   else eval iroot(4;n) in
        proper-divisor-aux(n;m;m;1;m)
   fi 

Lemma: proper-divisor_wf
[n:ℕ+]. (proper-divisor(n) ∈ Dec(∃n1:ℤ [(n1 < n ∧ (2 ≤ n1) ∧ (n1 n))]))

Definition: injection
A →⟶ ==  {f:A ⟶ B| Inj(A;B;f)} 

Lemma: injection_wf
[A,B:Type].  (A →⟶ B ∈ Type)

Lemma: injection-bijection
n:ℕ. ∀f:ℕn →⟶ ℕn.  Bij(ℕn;ℕn;f)

Lemma: injection-inverse
[A,B:Type].  ∀f:A →⟶ B. (A  finite-type(A)  (∀x,y:B.  Dec(x y ∈ B))  (∃g:B ⟶ A. ∀a:A. ((g (f a)) a ∈ A)))

Lemma: injection-inverse2
n:ℕ. ∀f:ℕn →⟶ ℕn.  ∃g:ℕn →⟶ ℕn. ((∀a:ℕn. ((g (f a)) a ∈ ℕn)) ∧ (∀a:ℕn. ((f (g a)) a ∈ ℕn)))

Lemma: funinv_wf3
[n:ℕ]. ∀[f:ℕn →⟶ ℕn].  (inv(f) ∈ ℕn →⟶ ℕn)

Lemma: compose_wf-injection
[A,B,C:Type]. ∀[f:A →⟶ B]. ∀[g:B →⟶ C].  (g f ∈ A →⟶ C)

Lemma: funinv-ap-equals
[n:ℕ]. ∀[f:ℕn →⟶ ℕn]. ∀[a,b:ℕn].  (inv(f) b) a ∈ ℤ supposing (f a) b ∈ ℤ

Definition: extend-injection
extend-injection(a;f) ==  λx.if x <then else fi 

Lemma: extend-injection_wf
[a:ℕ]. ∀[f:ℕa →⟶ ℕa]. ∀[b:{a...}].  (extend-injection(a;f) ∈ ℕb →⟶ ℕb)

Lemma: decidable__equal_injection
n:ℕ. ∀T:Type.  ((∀x,y:T.  Dec(x y ∈ T))  (∀f,g:ℕn →⟶ T.  Dec(f g ∈ ℕn →⟶ T)))

Definition: conjugate
conjugate(f;g) ==  inv(g) (f g)

Lemma: conjugate_wf
[n:ℕ]. ∀[f,g:ℕn →⟶ ℕn].  (conjugate(f;g) ∈ ℕn →⟶ ℕn)

Lemma: iterated-conjugate
[T:Type]. ∀[f,g,h:T ⟶ T]. ∀[m:ℕ].  ((λf.(g (f h))^m f) (g^m (f h^m)) ∈ (T ⟶ T))

Lemma: iterated-conjugate2
[T:Type]. ∀[f,g,h:T ⟶ T].
  (∀[n:ℕ]. (g (f h)^n (g (f^n h)) ∈ (T ⟶ T))) supposing 
     ((∀b:T. ((h (g b)) b ∈ T)) and 
     (∀a:T. ((g (h a)) a ∈ T)))

Lemma: rem-one
[x:ℤ]. (x rem 0)

Lemma: div-one
[x:ℤ]. (x ÷ x)

Definition: rotate-by
rotate-by(n;i) ==  λx.(x rem n)

Lemma: rotate-by_wf
[n,i:ℕ].  (rotate-by(n;i) ∈ ℕn ⟶ ℕn)

Lemma: rotate-by-id
[n,i:ℕ].  uiff(rotate-by(n;i) x.x) ∈ (ℕn ⟶ ℕn);(i rem n) 0 ∈ ℤ supposing 0 < n)

Lemma: rotate-by-is-id
[n,i:ℕ].  rotate-by(n;i) x.x) ∈ (ℕn ⟶ ℕn) supposing i

Lemma: rotate-by-zero
[n:ℕ]. (rotate-by(n;0) x.x) ∈ (ℕn ⟶ ℕn))

Lemma: rotate-by-trivial
[n:ℕ]. ∀[x:ℕn].  (((rotate-by(n;0) x) x ∈ ℕn) ∧ ((rotate-by(n;n) x) x ∈ ℕn))

Lemma: rotate-by-trivial-test
[n:ℕ]. ∀[x:ℕn].  (((rotate-by(n;0) x) x ∈ ℕn) ∧ ((rotate-by(n;n) x) x ∈ ℕn))

Lemma: iterate-rotate-rotate-by
[n,i:ℕ].  (rot(n)^i rotate-by(n;i) ∈ (ℕn ⟶ ℕn))

Lemma: compose-rotate-by
[n,i,j:ℕ].  ((rotate-by(n;i) rotate-by(n;j)) rotate-by(n;i j) ∈ (ℕn ⟶ ℕn))

Lemma: rotate-by-rotate-by
[n,i,j:ℕ]. ∀[x:ℕn].  ((rotate-by(n;i) (rotate-by(n;j) x)) (rotate-by(n;i j) x) ∈ ℤ)

Lemma: iterate-rotate-by
[n,i,k:ℕ].  (rotate-by(n;i)^k rotate-by(n;k i) ∈ (ℕn ⟶ ℕn))

Lemma: rotate-by-injection
[n,i:ℕ].  Inj(ℕn;ℕn;rotate-by(n;i))

Lemma: rotate-by-transitive
n,b:ℕ.  (gcd(b;n) 1 ∈ ℤ supposing 0 < ⇐⇒ ∀x,y:ℕn.  ∃k:ℕ((rotate-by(n;b)^k x) y ∈ ℤ))

Lemma: injections-combinations
n:ℕ. ∀[T:Type]. ℕn →⟶ Combination(n;T)

Lemma: combinations-formula
[n,m:ℕ].  ((C(n;m) (m n)!) (m)! ∈ ℤ supposing n ≤ m ∧ (C(n;m) if n ≤then (m)! ÷ (m n)! else fi  ∈ ℤ))

Lemma: choose-formula
[n,m:ℕ].  (choose(n;m) (m)! (n m)!) (n)! ∈ ℤ supposing m ≤ n

Lemma: choose-inequality1
n:ℕ. ∀i:ℕn.  (choose(n;i) ≤ (n choose(n 1;i)))

Lemma: combinations-choose
[m,n:ℕ].  C(m;n) (choose(n;m) (m)!) ∈ ℤ supposing m ≤ n

Lemma: equipollent-factorial
n:ℕ. ℕn →⟶ ℕ~ ℕ(n)!

Lemma: list-permutations
n:ℕ(∃P:ℕn →⟶ ℕList [(no_repeats(ℕn →⟶ ℕn;P) ∧ (∀f:ℕn →⟶ ℕn. (f ∈ P)))])

Definition: permutations-list
permutations-list(n) ==  TERMOF{list-permutations:o, 1:l} n

Lemma: permutations-list_wf
[n:ℕ]. (permutations-list(n) ∈ {P:ℕn →⟶ ℕList| no_repeats(ℕn →⟶ ℕn;P) ∧ (∀f:ℕn →⟶ ℕn. (f ∈ P))} )

Lemma: permutations-list-0
permutations-list(0) x.⊥]

Lemma: permutations-list-1
permutations-list(1) x.[0][x]]

Lemma: member-permutations-list
n:ℕ. ∀f:ℕn →⟶ ℕn.  (f ∈ permutations-list(n))

Lemma: no_repeats-permutations-list
n:ℕno_repeats(ℕn →⟶ ℕn;permutations-list(n))

Lemma: permutation-of-permutations-list
n:ℕ. ∀f:{p:ℕn ⟶ ℕn| Inj(ℕn;ℕn;p)}  ⟶ {p:ℕn ⟶ ℕn| Inj(ℕn;ℕn;p)} .
  (Bij({p:ℕn ⟶ ℕn| Inj(ℕn;ℕn;p)} ;{p:ℕn ⟶ ℕn| Inj(ℕn;ℕn;p)} ;f)
   permutation({p:ℕn ⟶ ℕn| Inj(ℕn;ℕn;p)} ;permutations-list(n);map(f;permutations-list(n))))

Lemma: funinv-permutes-permutations-list
n:ℕpermutation({p:ℕn ⟶ ℕn| Inj(ℕn;ℕn;p)} ;permutations-list(n);map(λf.inv(f);permutations-list(n)))

Lemma: flip-permutes-permutations-list
n:ℕ. ∀i,j:ℕn.  permutation({p:ℕn ⟶ ℕn| Inj(ℕn;ℕn;p)} ;permutations-list(n);map(λf.(f (i, j));permutations-list(n)))

Lemma: flip-permutes-permutations-list2
n:ℕ. ∀i,j:ℕn.  permutation({p:ℕn ⟶ ℕn| Inj(ℕn;ℕn;p)} ;permutations-list(n);map(λf.((i, j) f);permutations-list(n)))

Definition: partial-permutations-list
partial-permutations-list(n;i) ==  filter(λa.(a (n 1) =z i);permutations-list(n))

Lemma: partial-permutations-list_wf
[n:ℕ+]. ∀[i:ℤ].  (partial-permutations-list(n;i) ∈ ℕn →⟶ ℕList)

Lemma: no_repeats-partial-permutations-list
[n:ℕ+]. ∀[i:ℤ].  no_repeats(ℕn →⟶ ℕn;partial-permutations-list(n;i))

Lemma: member-partial-permutations-list
n:ℕ+. ∀i:ℤ. ∀f:ℕn →⟶ ℕn.  (f ∈ partial-permutations-list(n;i)) supposing (f (n 1)) i ∈ ℤ

Lemma: list-functions
n,b:ℕ.  (∃P:(ℕn ⟶ ℕb) List [(no_repeats(ℕn ⟶ ℕb;P) ∧ (∀f:ℕn ⟶ ℕb. (f ∈ P)))])

Definition: functions-list
functions-list(n;b) ==  TERMOF{list-functions:o, 1:l} b

Lemma: functions-list_wf
[n,b:ℕ].  (functions-list(n;b) ∈ {P:(ℕn ⟶ ℕb) List| no_repeats(ℕn ⟶ ℕb;P) ∧ (∀f:ℕn ⟶ ℕb. (f ∈ P))} )

Lemma: member-functions-list
n,b:ℕ. ∀f:ℕn ⟶ ℕb.  (f ∈ functions-list(n;b))

Lemma: no_repeats-functions-list
[n,b:ℕ].  no_repeats(ℕn ⟶ ℕb;functions-list(n;b))

Lemma: functions-list-0
[b:ℕ]. (functions-list(0;b) x.x] ∈ ((ℕ0 ⟶ ℕb) List))

Definition: cyclic-map
cyclic-map(T) ==  {f:T →⟶ T| ∀x,y:T.  ∃n:ℕ((f^n x) y ∈ T)} 

Lemma: cyclic-map_wf
[T:Type]. (cyclic-map(T) ∈ Type)

Lemma: cyclic-map-transitive
n:ℕ. ∀f:cyclic-map(ℕn). ∀x,y:ℕn.  ∃m:ℕn. ((f^m x) y ∈ ℕn)

Lemma: cyclic-map-conjugate
[T:Type]. ∀[g,h:T ⟶ T].
  (∀[f:cyclic-map(T)]. (g (f h) ∈ cyclic-map(T))) supposing 
     ((∀a:T. ((g (h a)) a ∈ T)) and 
     (∀b:T. ((h (g b)) b ∈ T)))

Lemma: cycle_wf2
[n:ℕ]. ∀[L:Combination(n;ℕn)].  (cycle(L) ∈ cyclic-map(ℕn))

Lemma: cyclic-map-equipollent
n:ℕ+Combination(n 1;ℕ1) cyclic-map(ℕn)

Lemma: count-cyclic-map
n:ℕ+cyclic-map(ℕn) ~ ℕ(n 1)!

Lemma: rotate-by-cyclic-map
[n,i:ℕ].  rotate-by(n;i) ∈ cyclic-map(ℕn) supposing gcd(i;n) 1 ∈ ℤ

Definition: ring-as-list
ring-as-list(T;L;f) ==  Inj({i:T| (i ∈ L)} ;{i:T| (i ∈ L)} ;f) ∧ (∀i,j:{i:T| (i ∈ L)} .  ∃n:ℕ(j (f^n i) ∈ {i:T| (i ∈\000C L)} ))

Lemma: ring-as-list_wf
[T:Type]. ∀[L:T List]. ∀[f:{i:T| (i ∈ L)}  ⟶ {i:T| (i ∈ L)} ].  (ring-as-list(T;L;f) ∈ ℙ)

Lemma: gcd-prime
[p:ℕ]. ∀[n:ℕ+p]. (gcd(n;p) 1 ∈ ℤsupposing prime(p)

Lemma: wilson-theorem
n:{i:ℤ1 < i} (prime(n) ⇐⇒ (n 1)! ≡ (-1) mod n)

Lemma: rng_sum-int
[a,b:ℤ]. ∀[f:{a..b-} ⟶ ℤ].  (ℤ-rng) a ≤ i < b. f[i]) = Σ(f[a i] i < a) ∈ ℤ supposing a ≤ b

Lemma: geom-sum-int
[a:ℤ]. ∀[n:ℕ].  (((1 a) * Σ(a^i i < n)) (1 a^n) ∈ ℤ)

Lemma: binomial-int
[a,b:ℤ]. ∀[n:ℕ].  ((a b)^n = Σ(choose(n;i) a^i b^(n i) i < 1) ∈ ℤ)

Lemma: binomial-inequality1
[a,b:ℕ]. ∀[n:ℕ+].  ((a^n b^n) ≤ (a b)^n)

Lemma: exp-difference-inequality
[n:ℕ+]. ∀[a,b:ℕ].  (((a b)^n a^n) ≤ (n (a b)^(n 1)))

Lemma: iroot-lemma
a:ℕ. ∀n,b,k:ℕ+.  ∃x:ℕ. ∃y:ℕ+(a y^n < (x b)^n ∧ ((x b)^n ≤ ((a k) y^n)))

Lemma: iroot-lemma2
a:ℕ. ∀n,b,k:ℕ+.  (∃p:ℕ × ℕ+ [let x,y in y^n < (x b)^n ∧ ((x b)^n ≤ ((a k) y^n))])

Definition: n-intersecting
n-intersecting(A;T;n) ==  ∀Ls:T List. ((||Ls|| n ∈ ℤ (∃a:A. (∀L∈Ls.(a ∈ L))))

Lemma: n-intersecting_wf
[A,T:Type].  ∀[n:ℤ]. (n-intersecting(A;T;n) ∈ ℙsupposing T ⊆(A List)

Lemma: combinations-n-intersecting
n,t:ℕ.  ∀[A:Type]. (A ~ ℕ(n t)  n-intersecting(A;Combination(((n 1) t) 1;A);n))

Definition: sum-map
Σf[x] for x ∈ ==  Σ(f[L[i]] i < ||L||)

Lemma: sum-map_wf
[T:Type]. ∀[f:T ⟶ ℤ]. ∀[L:T List].  f[x] for x ∈ L ∈ ℤ)

Lemma: sum_map_nil_lemma
f:Top. f[x] for x ∈ [] 0)

Lemma: sum-map-append1
[f:Top]. ∀[L:Top List]. ∀[x:Top].  f[x] for x ∈ [x] ~ Σf[x] for x ∈ f[x])

Lemma: sum-map-cons
[T:Type]. ∀[f:T ⟶ ℤ]. ∀[L:T List]. ∀[x:T].  f[x] for x ∈ [x L] f[x] + Σf[x] for x ∈ L)

Lemma: sum-map-append
[T:Type]. ∀[f:T ⟶ ℤ]. ∀[L1,L2:T List].  f[x] for x ∈ L1 L2 ~ Σf[x] for x ∈ L1 + Σf[x] for x ∈ L2)

Lemma: non_neg_sum-map
[T:Type]. ∀[f:T ⟶ ℤ]. ∀[L:T List].  0 ≤ Σf[x] for x ∈ supposing (∀x∈L.0 ≤ f[x])

Definition: lift guard
lift guard(x) ==  x

Lemma: lift guard_wf
[T:Type]. ∀[x:T].  (lift guard(x) ∈ T)

Lemma: sum-equal-terms
[n:ℕ]. ∀[a:ℕn ⟶ ℤ]. ∀[m:ℕ]. ∀[b:ℕm ⟶ ℤ].
  Σ(a[i] i < n) = Σ(b[j] j < m) ∈ ℤ 
  supposing permutation(ℤ;filter(λx.(¬b(x =z 0));map(λi.a[i];upto(n)));filter(λx.(¬b(x =z 0));map(λj.b[j];upto(m))))

Definition: power-sum
Σi<n.a[i]*x^i ==  Σ(a[i] x^i i < n)

Lemma: power-sum_wf
[n:ℕ]. ∀[x:ℤ]. ∀[a:ℕn ⟶ ℤ].  i<n.a[i]*x^i ∈ ℤ)

Lemma: power-sum-split
[n:ℕ]. ∀[k:ℕ1]. ∀[x:ℤ]. ∀[a:ℕn ⟶ ℤ].  i<n.a[i]*x^i i<k.a[i]*x^i (x^k * Σi<k.a[k i]*x^i)) ∈ ℤ)

Lemma: power-sum-linear
[n:ℕ]. ∀[x:ℤ]. ∀[a,b:ℕn ⟶ ℤ]. ∀[c,d:ℤ].
  i<n.(c a[i]) (d b[i])*x^i ((c * Σi<n.a[i]*x^i) (d * Σi<n.b[i]*x^i)) ∈ ℤ)

Lemma: power-sum-product
[x:ℤ]. ∀[n:ℕ]. ∀[a:ℕ1 ⟶ ℤ]. ∀[m:ℕ]. ∀[b:ℕ1 ⟶ ℤ].
  ((Σi<1.a[i]*x^i * Σi<1.b[i]*x^i)
  = Σi<(n m) 1.Σ(if j ≤then a[j] else fi  if j ≤then b[i j] else fi  j < 1)*x^i
  ∈ ℤ)

Lemma: power-sum_functionality_wrt_eqmod
m:ℤ. ∀n:ℕ. ∀x,y:ℤ. ∀a,b:ℕn ⟶ ℤ.
  ((x ≡ mod m)  (∀i:ℕn. (a[i] ≡ b[i] mod m))  i<n.a[i]*x^i ≡ Σi<n.b[i]*y^i mod m))

Lemma: power-sum_functionality_wrt_le
m:ℤ. ∀n,x:ℕ. ∀a,b:ℕn ⟶ ℤ.  ((∀i:ℕn. (a[i] ≤ b[i]))  i<n.a[i]*x^i ≤ Σi<n.b[i]*x^i))

Lemma: divisibility-by-3-rule
n:ℕ. ∀a:ℕn ⟶ ℤ.  (3 | Σi<n.a[i]*10^i ⇐⇒ | Σ(a[i] i < n))

Lemma: divisibility-by-9-rule
n:ℕ. ∀a:ℕn ⟶ ℤ.  (9 | Σi<n.a[i]*10^i ⇐⇒ | Σ(a[i] i < n))

Lemma: divisibility-by-5-rule
n:ℕ+. ∀a:ℕn ⟶ ℤ.  (5 | Σi<n.a[i]*10^i ⇐⇒ a[0])

Lemma: divisibility-by-2-rule
n:ℕ+. ∀a:ℕn ⟶ ℤ.  (2 | Σi<n.a[i]*10^i ⇐⇒ a[0])

Lemma: sparse-signed-rep-lemma1
m:ℤ(∃p:ℤ × {-2..3-[let k,b in (m ((4 k) b) ∈ ℤ) ∧ ((|b| 2 ∈ ℤ (↑isEven(k)))])

Lemma: sparse-signed-rep-lemma1-ext
m:ℤ(∃p:ℤ × {-2..3-[let k,b in (m ((4 k) b) ∈ ℤ) ∧ ((|b| 2 ∈ ℤ (↑isEven(k)))])

Definition: special-mod4-decomp
special-mod4-decomp(m) ==  TERMOF{sparse-signed-rep-lemma1-ext:o, 1:l} m

Lemma: special-mod4-decomp_wf
[m:ℤ]
  (special-mod4-decomp(m) ∈ {p:ℤ × {-2..3-}| let k,b in (m ((4 k) b) ∈ ℤ) ∧ ((|b| 2 ∈ ℤ (↑isEven(k)))} )

Lemma: special-mod4-decomp-unique
m:ℤ. ∃!k:ℤ. ∃b:{-2..3-}. ((m ((4 k) b) ∈ ℤ) ∧ ((|b| 2 ∈ ℤ (↑isEven(k))))

Lemma: small-sparse-rep
r:{-2..3-}
  (∃L:{-1..2-List [((r = Σi<||L||.L[i]*2^i ∈ ℤ)
                    ∧ (||L|| 2 ∈ ℤ)
                    ∧ (∀i:ℕ||L|| 1. ((L[i] 0 ∈ ℤ) ∨ (L[i 1] 0 ∈ ℤ))))])

Lemma: small-sparse-rep-ext
r:{-2..3-}
  (∃L:{-1..2-List [((r = Σi<||L||.L[i]*2^i ∈ ℤ)
                    ∧ (||L|| 2 ∈ ℤ)
                    ∧ (∀i:ℕ||L|| 1. ((L[i] 0 ∈ ℤ) ∨ (L[i 1] 0 ∈ ℤ))))])

Definition: sparse-rep-base
sparse-rep-base(r) ==  TERMOF{small-sparse-rep-ext:o, 1:l} r

Lemma: sparse-rep-base_wf
[r:{-2..3-}]
  (sparse-rep-base(r) ∈ {L:{-1..2-List| 
                         (r = Σi<||L||.L[i]*2^i ∈ ℤ)
                         ∧ (||L|| 2 ∈ ℤ)
                         ∧ (∀i:ℕ||L|| 1. ((L[i] 0 ∈ ℤ) ∨ (L[i 1] 0 ∈ ℤ)))} )

Lemma: sparse-signed-rep-exists
m:ℤ
  (∃L:{-1..2-List [((m = Σi<||L||.L[i]*2^i ∈ ℤ)
                    ∧ (0 < ||L||  (last(L) 0 ∈ ℤ)))
                    ∧ (∀i:ℕ||L|| 1. ((L[i] 0 ∈ ℤ) ∨ (L[i 1] 0 ∈ ℤ))))])

Lemma: sparse-signed-rep-exists-ext
m:ℤ
  (∃L:{-1..2-List [((m = Σi<||L||.L[i]*2^i ∈ ℤ)
                    ∧ (0 < ||L||  (last(L) 0 ∈ ℤ)))
                    ∧ (∀i:ℕ||L|| 1. ((L[i] 0 ∈ ℤ) ∨ (L[i 1] 0 ∈ ℤ))))])

Definition: sparse-signed-rep
sparse-signed-rep(m) ==  evalall(TERMOF{sparse-signed-rep-exists-ext:o, 1:l} m)

Lemma: sparse-signed-rep_wf
[m:ℤ]
  (sparse-signed-rep(m) ∈ {L:{-1..2-List| 
                           (m = Σi<||L||.L[i]*2^i ∈ ℤ)
                           ∧ (0 < ||L||  (last(L) 0 ∈ ℤ)))
                           ∧ (∀i:ℕ||L|| 1. ((L[i] 0 ∈ ℤ) ∨ (L[i 1] 0 ∈ ℤ)))} )

Definition: spsr
spsr(m) ==  let answer ⟵ <"display-as", "signed-rep", m, sparse-signed-rep(m)> in answer

Lemma: cantor-theorem-on-power-set-prop
[T:Type]. ∀f:T ⟶ T ⟶ ℙ. ∃P:T ⟶ ℙ. ∀x:T. (∀y:T. (P ⇐⇒ y)))

Lemma: cantor-theorem-on-power-set
[T:Type]. powerset(T))

Lemma: lift-test
[t:Top]
  ((let x,y let a,b 
              in <b, a> 
    in let a,b 
               in a)
  ∧ (let x,y case of inl(a) => <1, 2> inr(b) => <3, 4> 
     in case of inl(a) => inr(b) => 7)
  ∧ (let x,y if t=33 then <1, 2> else <3, 4> 
     in if t=33 then else 7))

Lemma: lift-test2
[r,g,a:Top].  (let a,b let a,b in <g, a> in let a,b in let g,u = <g, a> in u)

Lemma: map_functionality
[T,A:Type]. ∀[L1,L2:T List]. ∀[f,g:{x:T| (x ∈ L1)}  ⟶ A].
  (map(f;L1) map(g;L2) ∈ (A List)) supposing ((f g ∈ ({x:T| (x ∈ L1)}  ⟶ A)) and (L1 L2 ∈ (T List)))

Definition: lifted-rel
lifted-rel(R) ==  λx,y. ((↑isl(y)) ∧ ((↑isl(x))  (R outl(x) outl(y))))

Lemma: lifted-rel_wf
[A,B:Type]. ∀[R:A ⟶ A ⟶ ℙ].  (lifted-rel(R) ∈ (A B) ⟶ (A B) ⟶ ℙ)

Definition: lexico
lexico(T; a,b.lt[a; b]) ==
  λas,bs. (||as|| < ||bs|| ∨ ((||as|| ||bs|| ∈ ℤ) ∧ (∃i:ℕ||as||. (lt[as[i]; bs[i]] ∧ (∀j:ℕi. (as[j] bs[j] ∈ T))))))

Lemma: lexico_wf
[T:Type]. ∀[lt:T ⟶ T ⟶ ℙ].  (lexico(T; a,b.lt[a;b]) ∈ (T List) ⟶ (T List) ⟶ ℙ)

Lemma: lexico_well_fnd
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  (WellFnd{i}(T;a,b.R[a;b])  WellFnd{i}(T List;as,bs.as lexico(T; a,b.R[a;b]) bs))

Lemma: lexico_induction
[T:Type]. ∀[lt:T ⟶ T ⟶ ℙ].
  ((∀[Q:T ⟶ ℙ]. ((∀b:T. ((∀a:T. (lt[a;b]  Q[a]))  Q[b]))  (∀b:T. Q[b])))
   (∀[P:(T List) ⟶ ℙ]
        ((∀L:T List. ((∀L':T List. ((L' lexico(T; a,b.lt[a;b]) L)  P[L']))  P[L]))  (∀L:T List. P[L]))))

Definition: mul-initial-seg
mul-initial-seg(f) ==  λn.reduce(λx,y. (x y);1;map(f;upto(n)))

Lemma: mul-initial-seg_wf
[f:ℕ ⟶ ℕ]. (mul-initial-seg(f) ∈ ℕ ⟶ ℕ)

Lemma: mul-initial-seg-step
[f:ℕ ⟶ ℕ]. ∀[m:ℕ+].  ((mul-initial-seg(f) m) ((mul-initial-seg(f) (m 1)) (f (m 1))) ∈ ℤ)

Lemma: mul-initial-seg-property
f:ℕ ⟶ ℕ. ∀m:ℕ.  (∃n:ℕ(n < m ∧ ((f n) 0 ∈ ℤ)) ⇐⇒ (mul-initial-seg(f) m) 0 ∈ ℤ)

Lemma: mul-initial-seg-property2
f:ℕ ⟶ ℕ(∃n:ℕ((f n) 0 ∈ ℤ⇐⇒ ∃n:ℕ. ∀m:ℕ(mul-initial-seg(f) m) 0 ∈ ℤ supposing n ≤ m)

Lemma: sq_exists_subtype_rel
[A,B:Type]. ∀[P:A ⟶ ℙ]. ∀[Q:B ⟶ ℙ].
  ((∃a:A [P[a]]) ⊆(∃b:B [Q[b]])) supposing ((∀a:A. (P[a]  Q[a])) and (A ⊆B))

Lemma: not_subtype_rel
[A,B:Type].  A) ⊆B) supposing B ⊆A

Definition: record
record(x.T[x]) ==  x:Atom ⟶ T[x]

Lemma: record_wf
[T:Atom ⟶ 𝕌']. (record(x.T[x]) ∈ 𝕌')

Lemma: top-subtype-record
Top ⊆record(x.Top)

Definition: record+
Tz:B[self] ==  self:T ⋂ x:Atom ⟶ if =a then B[self] else Top fi 

Lemma: record+_wf
[T:𝕌']. ∀[B:T ⟶ 𝕌']. ∀[z:Atom].  (Tz:B[self] ∈ 𝕌')

Lemma: subtype_rel_record+
[T1,T2:𝕌']. ∀[B1:T1 ⟶ 𝕌']. ∀[B2:T2 ⟶ 𝕌'].
  (∀[z:Atom]. (T1z:B1[self] ⊆T2z:B2[self])) supposing ((∀x:T1. (B1[x] ⊆B2[x])) and (T1 ⊆T2))

Lemma: record+_subtype_rel
[T:Type]. ∀[B:T ⟶ Type]. ∀[z:Atom].  (T z:B[self] ⊆T)

Lemma: subtype_rel_record+_easy
[T1,T2:𝕌']. ∀[B:T2 ⟶ 𝕌'].  ∀[z:Atom]. (T1z:B[self] ⊆T2z:B[self]) supposing T1 ⊆T2

Lemma: record+_record
[T:Atom ⟶ 𝕌']. ∀[B:record(x.T[x]) ⟶ 𝕌']. ∀[z:Atom]. ∀[r:record(x.T[x])
                                                           z:B[self]].
  (r ∈ record(x.if =a then B[r] else T[x] fi ))

Definition: record-select
r.x ==  x

Lemma: record-select_wf
[T:Atom ⟶ 𝕌']. ∀[z:Atom]. ∀[r:record(x.T[x])].  (r.z ∈ T[z])

Lemma: record-select_wf2
[T:𝕌']. ∀[z:Atom]. ∀[B:T ⟶ 𝕌']. ∀[r:Tz:B[self]].  (r.z ∈ B[r])

Lemma: record_extensionality
[T:Atom ⟶ 𝕌']. ∀[r1,r2:record(x.T[x])].  uiff(r1 r2 ∈ record(x.T[x]);∀a:Atom. (r1.a r2.a ∈ T[a]))

Lemma: record+_extensionality
[T:Atom ⟶ 𝕌']. ∀[B:record(x.T[x]) ⟶ 𝕌']. ∀[z:Atom]. ∀[r1,r2:record(x.T[x])
                                                               z:B[self]].
  uiff(r1 r2 ∈ record(x.T[x])z:B[self];(r1 r2 ∈ record(x.T[x])) ∧ (r1.z r2.z ∈ B[r1]))

Definition: record-update
r[x := v] ==  λz.if =a then else fi 

Lemma: record-update_wf
[T:Atom ⟶ 𝕌']. ∀[z:Atom]. ∀[r:record(x.T[x])]. ∀[v:T[z]].  (r[z := v] ∈ record(x.T[x]))

Lemma: record-update-update
[z,r,v1,v2:Top].  (r[z := v1][z := v2] r[z := v2])

Lemma: trivial-record-update
[r,z:Top].  (r[z := r.z] ~ λx.if =a then else fi )

Lemma: identity-record-update
[T:Atom ⟶ 𝕌']. ∀[z:Atom]. ∀[r:record(x.T[x])].  (r[z := r.z] r ∈ record(x.T[x]))

Lemma: identity-record+-update
[T:Atom ⟶ 𝕌']. ∀[B:record(x.T[x]) ⟶ 𝕌']. ∀[z:Atom]. ∀[r:record(x.T[x])
                                                           z:B[self]]. ∀[x:Atom].
  (r[x := r.x] r ∈ record(x.T[x])z:B[self])

Lemma: rec_select_update_lemma
y,v,x,r:Top.  (r[x := v].y if =a then else r.y fi )

Lemma: record-select-update
[r,x,y,v:Top].  (r[x := v].y if =a then else r.y fi )

Lemma: subtype_rel_record
[T1,T2:Atom ⟶ Type].  uiff(record(x.T1[x]) ⊆record(x.T2[x]);∀[x:Atom]. (T1[x] ⊆T2[x]) supposing record(x.T1[x]))

Lemma: isect2-record
[T1,T2:Atom ⟶ Type].
  ((record(x.T1[x]) ⋂ record(x.T2[x]) ⊆record(x.T1[x] ⋂ T2[x]))
  ∧ (record(x.T1[x] ⋂ T2[x]) ⊆record(x.T1[x]) ⋂ record(x.T2[x])))

Definition: tag-case
z:T ==  x:Atom × if =a then else Top fi 

Lemma: tag-case_wf
[T:Type]. ∀[z:Atom].  (z:T ∈ Type)

Lemma: subtype_rel-tag-case
[T1,T2:Type]. ∀[z:Atom].  z:T1 ⊆z:T2 supposing T1 ⊆T2

Definition: tag-by
==  {tg:Atom| tg z ∈ Atom}  × T

Lemma: tag-by_wf
[T:Type]. ∀[z:Atom].  (z×T ∈ Type)

Lemma: tag-by-subtype-tag-case
[T,S:Type].  ∀z,x:Atom.  (z×T ⊆x:S ⇐⇒ (¬↑=a x) ∨ (T ⊆S))

Lemma: subtype_rel-tag-by
[T,S:Type]. ∀[z:Atom].  T ⊆supposing T ⊆S

Lemma: strong-continuous-tag-case
[z:Atom]. ∀[F:Type ⟶ Type].  Continuous+(T.z:F[T]) supposing Continuous+(T.F[T])

Definition: tagged
tagged(x.T[x]) ==  x:Atom × T[x]

Lemma: tagged_wf
[T:Atom ⟶ Type]. (tagged(x.T[x]) ∈ Type)

Definition: tagged+
|+ z:B ==  T ⋂ z:B

Lemma: tagged+_wf
[T,B:Type]. ∀[z:Atom].  (T |+ z:B ∈ Type)

Lemma: tagged+_subtype_rel
[T,B:Type]. ∀[z:Atom].  (T |+ z:B ⊆T)

Lemma: subtype_rel_tagged+_general
[T,A,B:Type]. ∀[z:Atom].  (T ⊆|+ z:B) supposing ((T ⊆z:B) and (T ⊆A))

Lemma: subtype_rel_tagged+
[T1,T2,B1,B2:Type].  (∀[z:Atom]. (T1 |+ z:B1 ⊆T2 |+ z:B2)) supposing ((B1 ⊆B2) and (T1 ⊆T2))

Definition: tagged-tag
x.tag ==  fst(x)

Lemma: tagged-tag_wf
[T:Atom ⟶ Type]. ∀[x:tagged(x.T[x])].  (x.tag ∈ Atom)

Lemma: tagged-tag_wf2
[T,B:Type]. ∀[z:Atom]. ∀[x:T |+ z:B].  (x.tag ∈ Atom)

Definition: tagged-val
x.val ==  snd(x)

Lemma: tagged-val_wf
[T:Atom ⟶ Type]. ∀[x:tagged(x.T[x])].  (x.val ∈ T[x.tag])

Lemma: tagged-val_wf2
[B:Type]. ∀[z:Atom]. ∀[x:z:B].  x.val ∈ supposing x.tag z ∈ Atom

Definition: mk-tagged
mk-tagged(tg;x) ==  <tg, x>

Lemma: mk-tagged_wf_unequal
[B:Type]. ∀[tg,a:Atom].  ∀[x:Top]. (mk-tagged(a;x) ∈ tg:B) supposing ¬(tg a ∈ Atom)

Lemma: mk-tagged_wf
[B:Type]. ∀[tg:Atom]. ∀[x:B].  (mk-tagged(tg;x) ∈ tg:B)

Lemma: mk-tagged_wf2
[B:Type]. ∀[tg:Atom]. ∀[x:B].  (mk-tagged(tg;x) ∈ tg×B)

Definition: tagged-test
tagged-test(x) ==  if x.tag =a "a" then mk-tagged("a";x.val 1) else fi 

Lemma: tagged-test_wf
[T:Type]. ∀[x:"b":𝔹 |+ "a":ℤ].  (tagged-test(x) ∈ "b":𝔹 |+ "a":ℤ)

Lemma: member-tagged+-right
[T,B:Type]. ∀[a:Atom]. ∀[p:T |+ a:B].  p ∈ supposing ¬(p.tag a ∈ Atom)

Lemma: member-of-tagged+1
[T,B:Type]. ∀[tg,a:Atom]. ∀[x:T].  mk-tagged(tg;x) ∈ tg:T |+ a:B supposing ¬(tg a ∈ Atom)

Lemma: member-of-tagged+2
[T,B:Type]. ∀[tg,a:Atom]. ∀[x:T].  mk-tagged(tg;x) ∈ a:B |+ tg:T supposing ¬(tg a ∈ Atom)

Lemma: base-member-of-tagged+
[T,B:Type]. ∀[tg,a:Atom]. ∀[x:Base].
  (mk-tagged(tg;x) ∈ |+ a:B) supposing ((¬(tg a ∈ Atom)) and (mk-tagged(tg;x) ∈ T))

Definition: urec
urec(F) ==  ⋃n:ℕ.(F^n Void)

Lemma: urec_wf
[F:Type ⟶ Type]. (urec(F) ∈ Type)

Lemma: urec_subtype
[F:Type ⟶ Type]. urec(F) ⊆(F urec(F)) supposing Monotone(T.F[T])

Lemma: urec_subtype_base
[F:Type ⟶ Type]. urec(F) ⊆Base supposing ∀T:Type. ((T ⊆Base)  ((F T) ⊆Base))

Definition: type-incr-chain
type-incr-chain{i:l}() ==  {X:ℕ ⟶ Type| ∀n:ℕ((X n) ⊆(X (n 1)))} 

Lemma: type-incr-chain_wf
type-incr-chain{i:l}() ∈ 𝕌'

Lemma: type-incr-chain-subtype
[X:type-incr-chain{i:l}()]. ∀[n,m:ℕ].  (X n) ⊆(X m) supposing n ≤ m

Lemma: monotone-incr-chain
F:Type ⟶ Type. (Monotone(T.F T)  n.(F^n Void) ∈ type-incr-chain{i:l}()))

Definition: type-continuous'
semi-continuous(λT.F[T]) ==  ∀[X:type-incr-chain{i:l}()]. (F[⋃n:ℕ.(X n)] ⊆r ⋃n:ℕ.F[X n])

Lemma: type-continuous'_wf
[F:Type ⟶ Type]. (semi-continuous(λT.F[T]) ∈ ℙ')

Definition: continuous'-monotone
continuous'-monotone{i:l}(T.F[T]) ==  Monotone(T.F[T]) ∧ semi-continuous(λT.F[T])

Lemma: continuous'-monotone_wf
[F:Type ⟶ Type]. (continuous'-monotone{i:l}(T.F[T]) ∈ ℙ')

Lemma: continuous'-monotone-product
[F,G:Type ⟶ Type].
  (continuous'-monotone{i:l}(T.F[T] × G[T])) supposing 
     (continuous'-monotone{i:l}(T.G[T]) and 
     continuous'-monotone{i:l}(T.F[T]))

Lemma: continuous'-monotone-sum
[F,G:Type ⟶ Type].
  (continuous'-monotone{i:l}(T.F[T] G[T])) supposing 
     (continuous'-monotone{i:l}(T.G[T]) and 
     continuous'-monotone{i:l}(T.F[T]))

Lemma: continuous'-monotone-constant
[A:Type]. continuous'-monotone{i:l}(T.A)

Lemma: continuous'-monotone-identity
continuous'-monotone{i:l}(T.T)

Lemma: continuous'-monotone-bunion
[F,G:Type ⟶ Type].
  (continuous'-monotone{i:l}(T.F[T] ⋃ G[T])) supposing 
     (continuous'-monotone{i:l}(T.G[T]) and 
     continuous'-monotone{i:l}(T.F[T]))

Lemma: subtype_urec
[F:Type ⟶ Type]. (F urec(F)) ⊆urec(F) supposing continuous'-monotone{i:l}(T.F T)

Lemma: urec-is-fixedpoint
[F:Type ⟶ Type]. urec(F) ≡ urec(F) supposing continuous'-monotone{i:l}(T.F T)

Lemma: urec-is-least-fixedpoint
[F:Type ⟶ Type]. ∀T:Type. urec(F) ⊆supposing T ≡ supposing Monotone(T.F T)

Definition: constructor
Constr(T.F[T]) ==  ⋂T:{T:Type| T ⊆Base} ((T List) ⟶ F[T])

Lemma: constructor_wf
[F:Type ⟶ Type]. (Constr(T.F[T]) ∈ 𝕌')

Definition: ap-con
ap-con(con;L) ==  con L

Lemma: ap-con_wf
[F:Type ⟶ Type]. ∀[T:{T:Type| T ⊆Base} ]. ∀[con:Constr(T.F[T])]. ∀[L:T List].  (ap-con(con;L) ∈ F[T])

Definition: decomp
decomp{i:l}(S.F[S];T;x) ==  con:Constr(T.F[T]) × {L:T List| ap-con(con;L) x ∈ F[T]} 

Lemma: decomp_wf
[F:Type ⟶ Type]. ∀[T:{T:Type| T ⊆Base} ]. ∀[x:F[T]].  (decomp{i:l}(T.F[T];T;x) ∈ 𝕌')

Definition: destructor
destructor{i:l}(T.F[T]) ==  ⋂T:{T:Type| T ⊆Base} (x:F[T] ⟶ decomp{i:l}(T.F[T];T;x))

Lemma: destructor_wf
[F:Type ⟶ Type]. (destructor{i:l}(T.F[T]) ∈ 𝕌')

Lemma: destructor-product
[F,G:Type ⟶ Type].  (destructor{i:l}(T.F[T])  destructor{i:l}(T.G[T])  destructor{i:l}(T.F[T] × G[T]))

Lemma: destructor-sum
[F,G:Type ⟶ Type].  (destructor{i:l}(T.F[T])  destructor{i:l}(T.G[T])  destructor{i:l}(T.F[T] G[T]))

Lemma: destructor-const
[A:Type]. destructor{i:l}(T.A)

Definition: urec-level
urec-level(f;x) ==
  fix((λurec-level,x. let con,L in if null(L) then else imax-list(map(λt.(urec-level t);L)) fi )) x

Lemma: urec-level_wf
[F:Type ⟶ Type]
  ∀[f:destructor{i:l}(T.F[T])]. ∀[x:urec(F)].  (urec-level(f;x) ∈ ℕsupposing ∀T:Type. ((T ⊆Base)  (F[T] ⊆Base))

Lemma: urec-level-property
[F:Type ⟶ Type]
  (∀[f:destructor{i:l}(T.F[T])]. ∀[x:urec(F)].  (x ∈ F^urec-level(f;x) Void)) supposing 
     ((∀T:Type. ((T ⊆Base)  (F[T] ⊆Base))) and 
     Monotone(T.F[T]))

Lemma: urec_induction
[F:Type ⟶ Type]
  (destructor{i:l}(T.F[T])
      (∀[P:urec(F) ⟶ ℙ]
           ((∀[T:Type]. ((∀x:T ⋂ urec(F). P[x])  (∀x:F T ⋂ urec(F). P[x])))  (∀x:urec(F). P[x])))) supposing 
     ((∀T:Type. ((T ⊆Base)  ((F T) ⊆Base))) and 
     Monotone(T.F[T]))

Lemma: urec-lfp
[f:Type ⟶ Type]. ⋂r:{r:Type| (f r) ⊆r} r ≡ urec(f) supposing continuous'-monotone{i:l}(T.f T)

Definition: ulist
ulist(T) ==  urec(λA.(Unit ⋃ (T × A)))

Lemma: ulist_wf
[T:Type]. (ulist(T) ∈ Type)

Lemma: ulist-ext
[T:Type]. ulist(T) ≡ List

Definition: formulaco
formulaco() ==
  corec(X.lbl:Atom × if lbl =a "var" then Atom
                     if lbl =a "not" then X
                     if lbl =a "and" then left:X × X
                     if lbl =a "or" then left:X × X
                     if lbl =a "imp" then left:X × X
                     else Void
                     fi )

Lemma: formulaco_wf
formulaco() ∈ Type

Lemma: formulaco-ext
formulaco() ≡ lbl:Atom × if lbl =a "var" then Atom
                         if lbl =a "not" then formulaco()
                         if lbl =a "and" then left:formulaco() × formulaco()
                         if lbl =a "or" then left:formulaco() × formulaco()
                         if lbl =a "imp" then left:formulaco() × formulaco()
                         else Void
                         fi 

Definition: formulaco_size
formulaco_size(p) ==
  fix((λsize,p. let lbl,x 
                in if lbl =a "var" then 0
                   if lbl =a "not" then (size x)
                   if lbl =a "and" then let left,right in (1 (size left)) (size right)
                   if lbl =a "or" then let left,right in (1 (size left)) (size right)
                   if lbl =a "imp" then let left,right in (1 (size left)) (size right)
                   else 0
                   fi )) 
  p

Lemma: formulaco_size_wf
[p:formulaco()]. (formulaco_size(p) ∈ partial(ℕ))

Definition: formula
formula() ==  {p:formulaco()| (formulaco_size(p))↓

Lemma: formula_wf
formula() ∈ Type

Definition: formula_size
formula_size(p) ==
  fix((λsize,p. let lbl,x 
                in if lbl =a "var" then 0
                   if lbl =a "not" then (size x)
                   if lbl =a "and" then let left,right in (1 (size left)) (size right)
                   if lbl =a "or" then let left,right in (1 (size left)) (size right)
                   if lbl =a "imp" then let left,right in (1 (size left)) (size right)
                   else 0
                   fi )) 
  p

Lemma: formula_size_wf
[p:formula()]. (formula_size(p) ∈ ℕ)

Lemma: formula-ext
formula() ≡ lbl:Atom × if lbl =a "var" then Atom
                       if lbl =a "not" then formula()
                       if lbl =a "and" then left:formula() × formula()
                       if lbl =a "or" then left:formula() × formula()
                       if lbl =a "imp" then left:formula() × formula()
                       else Void
                       fi 

Definition: pvar
pvar(name) ==  <"var", name>

Lemma: pvar_wf
[name:Atom]. (pvar(name) ∈ formula())

Definition: pnot
pnot(sub) ==  <"not", sub>

Lemma: pnot_wf
[sub:formula()]. (pnot(sub) ∈ formula())

Definition: pand
pand(left;right) ==  <"and", left, right>

Lemma: pand_wf
[left,right:formula()].  (pand(left;right) ∈ formula())

Definition: por
por(left;right) ==  <"or", left, right>

Lemma: por_wf
[left,right:formula()].  (por(left;right) ∈ formula())

Definition: pimp
pimp(left;right) ==  <"imp", left, right>

Lemma: pimp_wf
[left,right:formula()].  (pimp(left;right) ∈ formula())

Definition: pvar?
pvar?(v) ==  fst(v) =a "var"

Lemma: pvar?_wf
[v:formula()]. (pvar?(v) ∈ 𝔹)

Definition: pvar-name
pvar-name(v) ==  snd(v)

Lemma: pvar-name_wf
[v:formula()]. pvar-name(v) ∈ Atom supposing ↑pvar?(v)

Definition: pnot?
pnot?(v) ==  fst(v) =a "not"

Lemma: pnot?_wf
[v:formula()]. (pnot?(v) ∈ 𝔹)

Definition: pnot-sub
pnot-sub(v) ==  snd(v)

Lemma: pnot-sub_wf
[v:formula()]. pnot-sub(v) ∈ formula() supposing ↑pnot?(v)

Definition: pand?
pand?(v) ==  fst(v) =a "and"

Lemma: pand?_wf
[v:formula()]. (pand?(v) ∈ 𝔹)

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

Lemma: pand-left_wf
[v:formula()]. pand-left(v) ∈ formula() supposing ↑pand?(v)

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

Lemma: pand-right_wf
[v:formula()]. pand-right(v) ∈ formula() supposing ↑pand?(v)

Definition: por?
por?(v) ==  fst(v) =a "or"

Lemma: por?_wf
[v:formula()]. (por?(v) ∈ 𝔹)

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

Lemma: por-left_wf
[v:formula()]. por-left(v) ∈ formula() supposing ↑por?(v)

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

Lemma: por-right_wf
[v:formula()]. por-right(v) ∈ formula() supposing ↑por?(v)

Definition: pimp?
pimp?(v) ==  fst(v) =a "imp"

Lemma: pimp?_wf
[v:formula()]. (pimp?(v) ∈ 𝔹)

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

Lemma: pimp-left_wf
[v:formula()]. pimp-left(v) ∈ formula() supposing ↑pimp?(v)

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

Lemma: pimp-right_wf
[v:formula()]. pimp-right(v) ∈ formula() supposing ↑pimp?(v)

Lemma: formula-induction
[P:formula() ⟶ ℙ]
  ((∀name:Atom. P[pvar(name)])
   (∀sub:formula(). (P[sub]  P[pnot(sub)]))
   (∀left,right:formula().  (P[left]  P[right]  P[pand(left;right)]))
   (∀left,right:formula().  (P[left]  P[right]  P[por(left;right)]))
   (∀left,right:formula().  (P[left]  P[right]  P[pimp(left;right)]))
   {∀v:formula(). P[v]})

Lemma: formula-definition
[A:Type]. ∀[R:A ⟶ formula() ⟶ ℙ].
  ((∀name:Atom. {x:A| R[x;pvar(name)]} )
   (∀sub:formula(). ({x:A| R[x;sub]}   {x:A| R[x;pnot(sub)]} ))
   (∀left,right:formula().  ({x:A| R[x;left]}   {x:A| R[x;right]}   {x:A| R[x;pand(left;right)]} ))
   (∀left,right:formula().  ({x:A| R[x;left]}   {x:A| R[x;right]}   {x:A| R[x;por(left;right)]} ))
   (∀left,right:formula().  ({x:A| R[x;left]}   {x:A| R[x;right]}   {x:A| R[x;pimp(left;right)]} ))
   {∀v:formula(). {x:A| R[x;v]} })

Definition: formula_ind
formula_ind(v;
            pvar(name) var[name];
            pnot(sub) rec1.not[sub; rec1];
            pand(left,right) rec2,rec3.and[left;
                                             right;
                                             rec2;
                                             rec3];
            por(left,right) rec4,rec5.or[left;
                                           right;
                                           rec4;
                                           rec5];
            pimp(left,right) rec6,rec7.imp[left;
                                             right;
                                             rec6;
                                             rec7])  ==
  fix((λformula_ind,v. let lbl,v1 
                       in if lbl="var" then var[v1]
                          if lbl="not" then not[v1; formula_ind v1]
                          if lbl="and"
                            then let left,v2 v1 
                                 in and[left;
                                        v2;
                                        formula_ind left;
                                        formula_ind v2]
                          if lbl="or"
                            then let left,v2 v1 
                                 in or[left;
                                       v2;
                                       formula_ind left;
                                       formula_ind v2]
                          if lbl="imp"
                            then let left,v2 v1 
                                 in imp[left;
                                        v2;
                                        formula_ind left;
                                        formula_ind v2]
                          else Ax
                          fi )) 
  v

Lemma: formula_ind_wf
[A:Type]. ∀[R:A ⟶ formula() ⟶ ℙ]. ∀[v:formula()]. ∀[var:name:Atom ⟶ {x:A| R[x;pvar(name)]} ].
[not:sub:formula() ⟶ {x:A| R[x;sub]}  ⟶ {x:A| R[x;pnot(sub)]} ]. ∀[and:left:formula()
                                                                        ⟶ right:formula()
                                                                        ⟶ {x:A| R[x;left]} 
                                                                        ⟶ {x:A| R[x;right]} 
                                                                        ⟶ {x:A| R[x;pand(left;right)]} ].
[or:left:formula() ⟶ right:formula() ⟶ {x:A| R[x;left]}  ⟶ {x:A| R[x;right]}  ⟶ {x:A| R[x;por(left;right)]} ].
[imp:left:formula() ⟶ right:formula() ⟶ {x:A| R[x;left]}  ⟶ {x:A| R[x;right]}  ⟶ {x:A| R[x;pimp(left;right)]} ].
  (formula_ind(v;
               pvar(name) var[name];
               pnot(sub) rec1.not[sub;rec1];
               pand(left,right) rec2,rec3.and[left;right;rec2;rec3];
               por(left,right) rec4,rec5.or[left;right;rec4;rec5];
               pimp(left,right) rec6,rec7.imp[left;right;rec6;rec7])  ∈ {x:A| R[x;v]} )

Lemma: formula_ind_wf_simple
[A:Type]. ∀[v:formula()]. ∀[var:name:Atom ⟶ A]. ∀[not:sub:formula() ⟶ A ⟶ A]. ∀[and,or,imp:left:formula()
                                                                                               ⟶ right:formula()
                                                                                               ⟶ A
                                                                                               ⟶ A
                                                                                               ⟶ A].
  (formula_ind(v;
               pvar(name) var[name];
               pnot(sub) rec1.not[sub;rec1];
               pand(left,right) rec2,rec3.and[left;right;rec2;rec3];
               por(left,right) rec4,rec5.or[left;right;rec4;rec5];
               pimp(left,right) rec6,rec7.imp[left;right;rec6;rec7])  ∈ A)

Definition: prank
prank(x) ==
  formula_ind(x;
              pvar(v) 0;
              pnot(a) r.r 1;
              pand(a,b) r1,r2.imax(r1;r2) 1;
              por(a,b) r1,r2.imax(r1;r2) 1;
              pimp(a,b) r1,r2.imax(r1;r2) 1) 

Lemma: prank_wf
[x:formula()]. (prank(x) ∈ ℕ)

Definition: psub
a ⊆ ==
  formula_ind(b;
              pvar(v) pvar(v) ∈ formula();
              pnot(c) r.(a pnot(c) ∈ formula()) ∨ r;
              pand(c,d) r1,r2.(a pand(c;d) ∈ formula()) ∨ r1 ∨ r2;
              por(c,d) r1,r2.(a por(c;d) ∈ formula()) ∨ r1 ∨ r2;
              pimp(c,d) r1,r2.(a pimp(c;d) ∈ formula()) ∨ r1 ∨ r2) 

Lemma: psub_wf
[a,b:formula()].  (a ⊆ b ∈ ℙ)

Lemma: psub_weakening
a,b:formula().  a ⊆ supposing b ∈ formula()

Lemma: psub-same
a:formula(). (a ⊆ ⇐⇒ True)

Lemma: psub_transitivity
a,b,c:formula().  (a ⊆  b ⊆  a ⊆ c)

Lemma: prank-psub
a,b:formula().  (a ⊆  ((a b ∈ formula()) ∨ prank(a) < prank(b)))

Lemma: psub_antisymmetry
[P,Q:formula()].  (P Q ∈ formula()) supposing (Q ⊆ and P ⊆ Q)

Lemma: prank_functionality
[P,Q:formula()].  prank(P) ≤ prank(Q) supposing P ⊆ Q

Lemma: prank-pvar
[a:formula()]. uiff(prank(a) 0 ∈ ℤ;↑pvar?(a))

Definition: extend-val
extend-val(v0;g;x) ==
  formula_ind(x;
              pvar(v) v0 pvar(v);
              pnot(a) r.¬b(g a);
              pand(a,b) r1,r2.(g a) ∧b (g b);
              por(a,b) r1,r2.(g a) ∨b(g b);
              pimp(a,b) r1,r2.(¬b(g a)) ∨b(g b)) 

Lemma: extend-val_wf
[a:formula()]. ∀[v0:{b:formula()| b ⊆ a ∧ (↑pvar?(b))}  ⟶ 𝔹]. ∀[g:{b:formula()| b ⊆ a}  ⟶ 𝔹].  (extend-val(v0;g;a) ∈ \000C𝔹)

Definition: valuation
valuation(v0;x;f) ==  ∀a:{a:formula()| a ⊆ x} extend-val(v0;f;a)

Lemma: valuation_wf
[x:formula()]. ∀[v0:{a:formula()| a ⊆ x ∧ (↑pvar?(a))}  ⟶ 𝔹]. ∀[f:{a:formula()| a ⊆ x}  ⟶ 𝔹].  (valuation(v0;x;f) ∈ ℙ\000C)

Definition: bdd-val
bdd-val(v0;x;n) ==
  {f:{a:formula()| a ⊆ x ∧ prank(a) < n}  ⟶ 𝔹| ∀a:{a:formula()| a ⊆ x ∧ prank(a) < n} extend-val(v0;f;a)} 

Lemma: bdd-val_wf
[x:formula()]. ∀[v0:{a:formula()| a ⊆ x ∧ (↑pvar?(a))}  ⟶ 𝔹]. ∀[n:ℕ].  (bdd-val(v0;x;n) ∈ Type)

Lemma: valuation-exists
x:formula(). ∀v0:{a:formula()| a ⊆ x ∧ (↑pvar?(a))}  ⟶ 𝔹.  (∃f:{a:formula()| a ⊆ x}  ⟶ 𝔹 [valuation(v0;x;f)])

Lemma: valuation-exists-ext
x:formula(). ∀v0:{a:formula()| a ⊆ x ∧ (↑pvar?(a))}  ⟶ 𝔹.  (∃f:{a:formula()| a ⊆ x}  ⟶ 𝔹 [valuation(v0;x;f)])

Lemma: valuation-val
[z:formula()]. ∀[v0:{a:formula()| a ⊆ z ∧ (↑pvar?(a))}  ⟶ 𝔹]. ∀[f:{f:{a:formula()| a ⊆ z}  ⟶ 𝔹valuation(v0;z;f)} ].
  extend-val(v0;f;z)

Definition: peval
peval(v0;x) ==  TERMOF{valuation-exists-ext:o, 1:l} v0 x

Lemma: peval_wf
[x:formula()]. ∀[v0:{a:formula()| a ⊆ x ∧ (↑pvar?(a))}  ⟶ 𝔹].  (peval(v0;x) ∈ 𝔹)

Lemma: peval-unroll
[z,v0:Top].  (peval(v0;z) extend-val(v0;fix((λ%,a. extend-val(v0;%;a)));z))

Definition: pnegate
pnegate(x) ==
  formula_ind(x;
              pvar(v) pnot(pvar(v));
              pnot(A) nA.A;
              pand(A,B) nA,nB.por(nA;nB);
              por(A,B) nA,nB.pand(nA;nB);
              pimp(A,B) nA,nB.pand(A;nB)) 

Lemma: pnegate_wf
[x:formula()]. (pnegate(x) ∈ formula())

Lemma: peval-pnegate
[x:formula()]. ∀[v:{a:formula()| a ⊆ x ∧ (↑pvar?(a))}  ⟶ 𝔹].  peval(v;pnegate(x)) = ¬bpeval(v;x)

Lemma: xmiddle-elim
[P,Q:ℙ].  (((P ∨ P))  Q))  Q))

Lemma: xmiddle-elim2
[P:ℙ]. (((P ∨ P))  False)  False)

Lemma: example1
[A,B,C:Type].  (((A  B)  A)  (B  C)  (A  B)  (¬¬C))

Lemma: example1-ext
[A,B,C:Type].  (((A  B)  A)  (B  C)  (A  B)  (¬¬C))

Lemma: example2
[A,B,C:Type].  (((A  B)  A)  (B  C)  (A  B)  (¬¬C))

Lemma: example2-ext
[A,B,C:Type].  (((A  B)  A)  (B  C)  (A  B)  (¬¬C))

Definition: anyprop
anyprop() ==  False

Lemma: anyprop_wf
anyprop() ∈ ℙ

Lemma: example3
[A,B,C:Type].  ((A  (B ∨ (A  C)))   (B ∨ C))

Lemma: infinite-domain-example
[A,D:Type]. ∀[R,Eq:D ⟶ D ⟶ ℙ].
  ((∀x,y,z:D.  (R[x;y]  (R[y;z] ∨ Eq[y;z])  R[x;z]))
   (∀x:D. (R[x;x]  A))
   (∀x:D. ∃y:D. R[x;y])
   (∃m:D. ∀x:D. ((Eq[x;m]  A)  R[x;m]))
   A)

Lemma: infinite-domain-example-ext
[A,D:Type]. ∀[R,Eq:D ⟶ D ⟶ ℙ].
  ((∀x,y,z:D.  (R[x;y]  (R[y;z] ∨ Eq[y;z])  R[x;z]))
   (∀x:D. (R[x;x]  A))
   (∀x:D. ∃y:D. R[x;y])
   (∃m:D. ∀x:D. ((Eq[x;m]  A)  R[x;m]))
   A)

Lemma: red-black-example
[A,D:Type]. ∀[Red,Black:D ⟶ ℙ]. ∀[R:D ⟶ D ⟶ ℙ].
  ((∀x:D. (Red[x] ∨ Black[x]))
   (∀x,y,z:D.  (R[x;y]  R[y;z]  R[x;z]))
   (∀x:D. (R[x;x]  A))
   (∀x:D. (Red[x]  (∃y:D. (Black[y] ∧ R[x;y]))))
   (∀x:D. (Black[x]  (∃y:D. (Red[y] ∧ R[x;y]))))
   (∃m:D. ((∀x:D. (Red[x]  R[x;m])) ∨ (∀x:D. (Black[x]  R[x;m]))))
   A)

Lemma: red-black-example-ext
[A,D:Type]. ∀[Red,Black:D ⟶ ℙ]. ∀[R:D ⟶ D ⟶ ℙ].
  ((∀x:D. (Red[x] ∨ Black[x]))
   (∀x,y,z:D.  (R[x;y]  R[y;z]  R[x;z]))
   (∀x:D. (R[x;x]  A))
   (∀x:D. (Red[x]  (∃y:D. (Black[y] ∧ R[x;y]))))
   (∀x:D. (Black[x]  (∃y:D. (Red[y] ∧ R[x;y]))))
   (∃m:D. ((∀x:D. (Red[x]  R[x;m])) ∨ (∀x:D. (Black[x]  R[x;m]))))
   A)

Lemma: sum-reindex
[n:ℕ]. ∀[a:ℕn ⟶ ℤ]. ∀[m:ℕ]. ∀[b:ℕm ⟶ ℤ].
  Σ(a[i] i < n) = Σ(b[j] j < m) ∈ ℤ 
  supposing ∃f:{i:ℕn| ¬(a[i] 0 ∈ ℤ)}  ⟶ {j:ℕm| ¬(b[j] 0 ∈ ℤ)} (Bij({i:ℕn| ¬(a[i] 0 ∈ ℤ)} ;{j:ℕm| ¬(b[j] 0 ∈ ℤ\000C)} ;f) ∧ (∀i:{i:ℕn| ¬(a[i] 0 ∈ ℤ)} (a[i] b[f i] ∈ ℤ)))

Lemma: sum-reindex2
[n:ℕ]. ∀[a:ℕn ⟶ ℤ]. ∀[m:ℕ]. ∀[b:ℕm ⟶ ℤ]. ∀[f:{i:ℕn| ¬(a[i] 0 ∈ ℤ)}  ⟶ {j:ℕm| ¬(b[j] 0 ∈ ℤ)} ].
  (a[i] i < n) = Σ(b[j] j < m) ∈ ℤsupposing 
     ((∀i:{i:ℕn| ¬(a[i] 0 ∈ ℤ)} (a[i] b[f i] ∈ ℤ)) and 
     Bij({i:ℕn| ¬(a[i] 0 ∈ ℤ)} ;{j:ℕm| ¬(b[j] 0 ∈ ℤ)} ;f))

Definition: awf
awf(T) ==  corec(S.T (S List))

Lemma: awf_wf
[T:Type]. (awf(T) ∈ Type)

Lemma: awf-subtype
[T:Type]. ∀[s:awf(T)].  (s ∈ (awf(T) List))

Definition: awf-leaf
awf-leaf(x) ==  inl x

Lemma: awf-leaf_wf
[A:Type]. ∀[x:A].  (awf-leaf(x) ∈ awf(A))

Definition: awf-system
awf-system{i:l}(I;A) ==
  ⋂T:{T:Type| ((A (T List)) ⊆T) ∧ (awf(A) ⊆T)} ((I ⟶ T) ⟶ I ⟶ (A (T List))) ⋂ Top ⟶ I ⟶ Top

Lemma: awf-system_wf
[I,A:Type].  (awf-system{i:l}(I;A) ∈ 𝕌')

Lemma: awf-system-subtype
[I,A:Type].  (awf-system{i:l}(I;A) ⊆((I ⟶ awf(A)) ⟶ I ⟶ awf(A)))

Lemma: unique-awf
[A,I:Type].
  ∀G:awf-system{i:l}(I;A)
    (∃s:I ⟶ awf(A) [((∀i:I. ((s i) (G i) ∈ awf(A)))
                    ∧ (∀s':I ⟶ awf(A). ((∀i:I. ((s' i) (G s' i) ∈ awf(A)))  (s' s ∈ (I ⟶ awf(A))))))])

Definition: awf-solution
awf-solution(G) ==  fix(G)

Lemma: awf-solution_wf
[A,I:Type].
  ∀G:awf-system{i:l}(I;A)
    (awf-solution(G) ∈ {s:I ⟶ awf(A)| 
                        (∀i:I. ((s i) (G i) ∈ awf(A)))
                        ∧ (∀s':I ⟶ awf(A). ((∀i:I. ((s' i) (G s' i) ∈ awf(A)))  (s' s ∈ (I ⟶ awf(A)))))} )

Definition: awf-example1
awf-example1() ==  λg,i. [inr [awf-leaf(1); 2] inr [g 0; 2] inr [g 0; 1; 2] ][i]

Lemma: awf-example1_wf
awf-example1() ∈ awf-system{i:l}(ℕ3;ℤ)

Definition: awf_sum
awf_sum(n;s) ==
  fix((λawf_sum,n,s. case s
                     of inl(z) =>
                     z
                     inr(L) =>
                     if (n =z 0)
                     then 0
                     else eval in
                          accumulate (with value and list item s):
                           (awf_sum s)
                          over list:
                            L
                          with starting value:
                           0)
                     fi )) 
  
  s

Lemma: awf_sum_wf
[n:ℕ]. ∀[s:awf(ℤ)].  (awf_sum(n;s) ∈ ℤ)

Lemma: gcd-mod
x:ℕ+. ∀y:ℕ.  (gcd(x;y mod x) gcd(x;y) ∈ ℤ)

Definition: residue
residue(n) ==  {k:ℕn| CoPrime(n,k)} 

Lemma: residue_wf
n:ℕ(residue(n) ∈ Type)

Lemma: test-sq-stuff
L,L2:ℤ List.  ((L L2 ∈ (ℤ List))  (||L|| ||L2|| ∈ ℤ))

Definition: residue-mul
(ai mod n) ==  (a i) mod n

Lemma: residue-mul_wf
[n:ℕ+]. ∀[a,i:ℤ].  (ai mod n) ∈ residue(n) supposing CoPrime(n,a) ∧ CoPrime(n,i)

Lemma: residue-mul-assoc
[n:ℕ+]. ∀[a,b,c:{i:ℤCoPrime(n,i)} ].  (((ab mod n)c mod n) (a(bc mod n) mod n) ∈ ℤ)

Lemma: residue-mul-inverse
n:{2...}. ∀a:ℕ.
  (CoPrime(n,a)
   (∃b:ℤ
       (CoPrime(n,b)
       ∧ (∀i:residue(n)
            ((((ba mod n) 1 ∈ residue(n)) ∧ ((ab mod n) 1 ∈ residue(n)))
            ∧ ((b(ai mod n) mod n) i ∈ residue(n))
            ∧ ((a(bi mod n) mod n) i ∈ residue(n)))))))

Definition: residues-mod
residues-mod(n) ==  filter(λi.(gcd(n;i) =z 1);upto(n))

Lemma: residues-mod_wf
[n:ℕ]. (residues-mod(n) ∈ residue(n) List)

Lemma: no_repeats-residues-mod
[n:ℕ]. no_repeats(residue(n);residues-mod(n))

Lemma: member-residues-mod
n:ℕ. ∀x:residue(n).  (x ∈ residues-mod(n))

Lemma: residues-permute
n:{2...}. ∀a:ℕ+.  (CoPrime(n,a)  permutation(residue(n);map(λi.(ai mod n);residues-mod(n));residues-mod(n)))

Lemma: residues-equal-bags
n:{2...}. ∀a:ℕ+.  (CoPrime(n,a)  (map(λi.(ai mod n);residues-mod(n)) residues-mod(n) ∈ bag(residue(n))))

Definition: totient
totient(n) ==  ||residues-mod(n)||

Lemma: totient_wf
[n:ℕ]. (totient(n) ∈ ℕ)

Lemma: Euler-Fermat
Euler's Generalization Of Fermat's 
 (we also have different, combinatorial, proof of
  fermat-littlefermat-little2)⋅

n:{2...}. ∀a:ℕ+.  (CoPrime(n,a)  (a^totient(n) ≡ mod n))

Definition: oob-getleft?
oob-getleft?(x) ==  if oob-hasleft(x) then {oob-getleft(x)} else {} fi 

Lemma: oob-getleft?_wf
[B,A:Type]. ∀[x:one_or_both(A;B)].  (oob-getleft?(x) ∈ bag(A))

Definition: oob-getright?
oob-getright?(x) ==  if oob-hasright(x) then {oob-getright(x)} else {} fi 

Lemma: oob-getright?_wf
[B,A:Type]. ∀[x:one_or_both(A;B)].  (oob-getright?(x) ∈ bag(B))

Definition: oob-apply
oob-apply(xs;ys) ==
  if (#(xs) =z 1) then if (#(ys) =z 1) then {oobboth(<only(xs), only(ys)>)} else {oobleft(only(xs))} fi 
  if (#(ys) =z 1) then {oobright(only(ys))}
  else {}
  fi 

Lemma: oob-apply_wf
[A,B:Type]. ∀[X:bag(A)]. ∀[Y:bag(B)].  (oob-apply(X;Y) ∈ bag(one_or_both(A;B)))

Definition: cnv-taba
cnv-taba(xs;ys) ==  fst(rec-case(xs) of [] => <[], ys> x::xs' => p.let a,ys in let h,t ys in <[<x, h> a], t>)

Lemma: cnv-taba_wf
[A,B:Type].  ∀xs:A List. ∀ys:B List.  ((||xs|| ≤ ||ys||)  (cnv-taba(xs;ys) ∈ (A × B) List))

Lemma: cnv-taba-property
[A,B:Type].
  ∀xs:A List. ∀ys:B List.  ((||xs|| ≤ ||ys||)  (cnv-taba(xs;ys) zip(xs;rev(firstn(||xs||;ys))) ∈ ((A × B) List)))

Definition: taba
taba(init;x,x',a.F[x;
                   x';
                   a];l) ==
  fst(rec-case(l) of
      [] => <init, l>
      x::xs' =>
       p.let a,ys 
         in let h,t ys 
            in <F[x;
                  h;
                  a]
               t
               >)

Lemma: taba_wf
[A,B:Type]. ∀[init:B]. ∀[F:A ⟶ A ⟶ B ⟶ B].  ∀xs:A List. (taba(init;x,x',a.F[x;x';a];xs) ∈ B)

Lemma: taba-property
[A,B:Type]. ∀[init:B]. ∀[F:A ⟶ A ⟶ B ⟶ B]. ∀[xs:A List].
  (taba(init;x,x',a.F[x;x';a];xs)
  accumulate (with value and list item p):
     let x,x' 
     in F[x;x';a]
    over list:
      zip(rev(xs);xs)
    with starting value:
     init)
  ∈ B)

Comment: --- three int list palindrome tests ----
Example of some Nuprl terms that implement the same
function in different ways.
The function has type ⌜(ℤ List) ⟶ 𝔹⌝ and returns ⌜tt⌝ if and only if the list is
palindrome.
One algorithm is slow-int-palindrome-test(L). It computes the reverse of and then
tests and reverse(L) for equality.

Another algorithm is palindrome-test(IntDeq;L).
It is an instance of Danvy Goldberg's "there and back again" functional program
and does the test without copying the list or traversing it many times.

Finally, int-palindrome-test(L) is what get by doing the 
"compile-time" symbolic (partial) evaluation of palindrome-test(IntDeq;L)
to get an even faster version.

These three terms are building the list [0;1;...;9] and then checking that it is
not palindrome. 

a) ⌜let L ⟵ upto(10)
    in slow-palindrome-test(L)⌝ 
b) ⌜let L ⟵ upto(10)
    in palindrome-test(IntDeq;L)⌝
c) ⌜let L ⟵ upto(10)
    in int-palindrome-test(L)⌝ 
a) takes 57621 steps  (seems to be exponential in for upto(n))
b) takes 418  steps  (38n+38 steps for upto(n) n>)
c) takes 370 steps   (34n+30 steps for upto(n) n>0)⋅

Definition: palindrome-test
palindrome-test(eq;L) ==  taba(tt;x,x',a.eval a ∧b (eq x') in b;L)

Lemma: palindrome-test_wf
[T:Type]. ∀[eq:EqDecider(T)]. ∀[L:T List].  (palindrome-test(eq;L) ∈ 𝔹)

Lemma: assert-palindrome-test
[T:Type]. ∀[eq:EqDecider(T)]. ∀[L:T List].  uiff(↑palindrome-test(eq;L);rev(L) L ∈ (T List))

Definition: int-palindrome-test
int-palindrome-test(L) ==
  let x,y fix((λlist_ind,L@0. eval L@0 in
                                if is pair then let x,xs' 
                                                    in let a,ys list_ind xs' 
                                                       in let h,t ys 
                                                          in <case a
                                                               of inl() =>
                                                               if x=h then inl Ax else (inr Ax )
                                                               inr() =>
                                                               inr Ax 
                                                             t
                                                             > otherwise if Ax then <inl Ax, L> otherwise ⊥)) 
            
  in x

Lemma: int-palindrome-test-sq
[L:ℤ List]. (int-palindrome-test(L) palindrome-test(IntDeq;L))

Lemma: int-palindrome-test_wf
[L:ℤ List]. (int-palindrome-test(L) ∈ 𝔹)

Lemma: assert-int-palindrome-test
[L:ℤ List]. uiff(↑int-palindrome-test(L);rev(L) L ∈ (ℤ List))

Definition: slow-int-palindrome-test
slow-int-palindrome-test(L) ==  list-deq(IntDeq) rev(L)

Lemma: slow-int-palindrome-test_wf
[L:ℤ List]. (slow-int-palindrome-test(L) ∈ 𝔹)

Lemma: assert-slow-int-palindrome-test
[L:ℤ List]. uiff(↑slow-int-palindrome-test(L);rev(L) L ∈ (ℤ List))

Definition: streamless
streamless(T) ==  ∀f:ℕ ⟶ T. ∃i,j:ℕ((¬(i j ∈ ℕ)) ∧ ((f i) (f j) ∈ T))

Lemma: streamless_wf
[T:Type]. (streamless(T) ∈ ℙ)

Lemma: streamless-dec-equal
[T:Type]. (streamless(T)  (∀x,y:T.  Dec(x y ∈ T)))

Lemma: streamless-implies-not-not-enum
[T:Type]. (streamless(T)  (¬¬(∃L:T List. ∀x:T. (x ∈ L))))

Lemma: markov-streamless-iff-not-not-enum
(∀P:ℕ ⟶ ℙ((∀m:ℕ((P m) ∨ (P m))))  (∀m:ℕ(P m))))  (∃m:ℕ(P m))))
 (∀T:Type. (streamless(T) ⇐⇒ (∀x,y:T.  Dec(x y ∈ T)) ∧ (¬¬(∃L:T List. ∀x:T. (x ∈ L)))))

Lemma: markov-streamless-iff
(∀P:ℕ ⟶ ℙ((∀m:ℕ((P m) ∨ (P m))))  (∀m:ℕ(P m))))  (∃m:ℕ(P m))))
 (∀T:Type. (streamless(T) ⇐⇒ (∀x,y:T.  Dec(x y ∈ T)) ∧ (¬¬(∃n:ℕ~ ℕn))))

Lemma: markov-streamless-product
(∀P:ℕ ⟶ ℙ((∀m:ℕ((P m) ∨ (P m))))  (∀m:ℕ(P m))))  (∃m:ℕ(P m))))
 (∀A,B:Type.  (streamless(A)  streamless(B)  streamless(A × B)))

Lemma: markov-streamless-function
(∀P:ℕ ⟶ ℙ((∀m:ℕ((P m) ∨ (P m))))  (∀m:ℕ(P m))))  (∃m:ℕ(P m))))
 (∀A,B:Type.  ((∃a:ℕ~ ℕa)  streamless(B)  streamless(A ⟶ B)))

Lemma: testsq
λx.(x 2) ~ λx.(x 1)

Lemma: page55
[U:Type]. ∀[P,Q:U ⟶ ℙ].  ((∀x:U. (P[x]  Q[x]))  (∀x:U. P[x])  (∀x:U. Q[x]))

Lemma: page55-again
U:Type. ∀P,Q:U ⟶ ℙ.  ((∀x:U. ((P x)  (Q x)))  (∀x:U. (P x))  (∀x:U. (Q x)))

Lemma: page55'
[U:Type]. ∀[P,Q:U ⟶ ℙ].  ∀f:∀x:U. (P[x]  Q[x]). ∀g:∀x:U. P[x]. ∀x:U.  Q[x]

Lemma: page55_witness
[U:Type]. ∀[P,Q:U ⟶ ℙ].  ((∀x:U. (P[x]  Q[x]))  (∀x:U. P[x])  (∀x:U. Q[x]))

Lemma: prob2.1
[U:Type]. ∀[P:U ⟶ ℙ].  ∀x:U. ((P x)  (∃x:U. (P x)))

Lemma: prob2.2
[U:Type]. ∀[P,Q:U ⟶ ℙ].  ((∃x:U. ((P x) ∨ (Q x)))  ((∃x:U. (P x)) ∨ (∃x:U. (Q x))))

Lemma: prob2.3
[U:Type]. ∀[P,Q:U ⟶ ℙ].  (((∃x:U. (P x)) ∨ (∃x:U. (Q x)))  (∃x:U. ((P x) ∨ (Q x))))

Lemma: prob2.4
[U:Type]. ∀[P,Q:U ⟶ ℙ].  ((∃x:U. ((P x) ∧ (Q x)))  ((∃x:U. (P x)) ∧ (∃x:U. (Q x))))

Lemma: comp-nat-ind-ext
[P:ℕ ⟶ ℙ]. ((∀i:ℕ((∀j:ℕi. P[j])  P[i]))  (∀i:ℕP[i]))

Lemma: ppcc-test6
[A,B:Type]. ∀[f:A ⟶ B]. ∀[as,L:A List]. ∀[n:ℤ].
  ||L as|| (n n) ∈ ℤ ⇐⇒ ||map(f;as)|| n ∈ ℤ supposing ||L|| ||as|| ∈ ℤ

Lemma: map_wf_bag
[A,B:Type]. ∀[f:A ⟶ B]. ∀[bs:bag(A)].  (map(f;bs) ∈ bag(B))

Lemma: bag_qinc
A:Type. ((A List) ⊆bag(A))

Definition: testrec1
testrec1(x;L) ==  fix((λtestrec1,x,L. if null(L) then else testrec1 (x x) tl(L) fi )) L

Lemma: testrec1_wf
[A:Type]. ∀[x:ℤ]. ∀[L:A List].  (testrec1(x;L) ∈ ℤ)

Definition: testrec2
testrec2(x;L) ==  letrec f(p)=let x,L in if null(L) then else f <x, tl(L)> fi  in f <x, L>

Definition: testrec3
testrec3(x;f;L) ==  letrec g(p)=let x,L in if null(L) then else g <x, map(f;tl(L))> fi  in g <x, L>

Lemma: test1
p:ℕ × ℕ. ∀bs:ℕ List.  (let x,y in y ∈ ℤ)

Lemma: test2
p:ℕ (ℕ ⟶ ℕ). ∀bs:ℕ List.  (case of inl(x) => inr(y) => 3 ∈ ℤ)

Lemma: test3
λloc,L. loc ∈ ⋂A1,B1:Type.  (B1 ⟶ A1 ⟶ B1)

Lemma: test4
p:ℕ × ℕ × 𝔹. ∀bs:ℕ List.  (let x,y,ok in if ok then else fi  ∈ ℤ)

Lemma: test5
[A:Type]. ∀[B:A ⟶ Type]. ∀[f:a:A ⟶ B[a]]. ∀[x:A].  (f x ∈ B[x])

Lemma: test6
[x:ℤ]. (if 0 <then else -x fi  ∈ ℕ)

Lemma: test7
[A:Type]. ∀[B:A ⟶ Type]. ∀[f:a:A ⟶ B[a]]. ∀[x:A].  (f x ∈ B[x])

Lemma: test34
p:ℕ × ℕ. ∀bs:ℕ List.  (let x,y in y ∈ ℤ)

Lemma: test35
p:ℕ (ℕ ⟶ ℕ). ∀bs:ℕ List.  (case of inl(x) => inr(y) => 3 ∈ ℤ)

Definition: foo
foo(x) ==  False

Lemma: subtype_neg_polymorphism_test
((⋂T:Type. (T ⟶ T ⟶ ℙ)) ⊆(Top ⟶ Top ⟶ ℙ)) ∧ ((Top ⟶ Top ⟶ ℙ) ⊆(⋂T:Type. (T ⟶ T ⟶ ℙ)))

Lemma: subtype_pos_polymorphism_test
((⋂T:Type. (ℙ ⟶ ℤ ⟶ (T × T))) ⊆(ℙ ⟶ ℤ ⟶ (Void × Void Void)))
∧ ((ℙ ⟶ ℤ ⟶ (Void × Void Void)) ⊆(⋂T:Type. (ℙ ⟶ ℤ ⟶ (T × T))))

Lemma: test-squash-simp
[B:ℤ ⟶ ℤ ⟶ ℙ]. ∀[A:ℤ ⟶ ℙ].
  (↓∃x:ℤ
     ((↓∃v:ℤ(↓A[x] ∨ (↓A[v]))) ∧ (↓(∃u:ℤ(↓((↓B[x;u]) ∧ (↓B[u;u])) ∨ ((↓B[u;x]) ∧ B[x;x]))) ∨ (↓∃n:ℕ(↓4 ≤ n))))
  ⇐⇒ ↓∃x:ℤ((∃v:ℤ(A[x] ∨ A[v])) ∧ ((∃u:ℤ((B[x;u] ∧ B[u;u]) ∨ (B[u;x] ∧ B[x;x]))) ∨ (∃n:ℕ(4 ≤ n)))))

Lemma: test-add-member-elim
x,y:Base. ∀d:ℤ.  ((x d)  (0 < x ∨ (x ≤ 0)))

Definition: unique-minimal
unique-minimal(T;x,y.R[x; y];m) ==  (∀x:T. R[x; m])) ∧ (∀y:T. ((∀x:T. R[x; y]))  (y m ∈ T)))

Lemma: unique-minimal_wf
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ]. ∀[m:T].  (unique-minimal(T;x,y.R[x;y];m) ∈ ℙ)

Definition: decidable-non-minimal
decidable-non-minimal(T;x,y.R[x; y]) ==  ∀y:T. Dec(↓∃x:T. R[x; y])

Lemma: decidable-non-minimal_wf
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  (decidable-non-minimal(T;x,y.R[x;y]) ∈ ℙ)

Lemma: unique-minimal-wellfounded-implies
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (decidable-non-minimal(T;x,y.R[x;y])
   WellFnd{i}(T;x,y.R[x;y])
   (∀m:T. (unique-minimal(T;x,y.R[x;y];m)  (∀y:T. (↓((λx,y. R[x;y])^*) y)))))

Definition: non-forking
non-forking(T;x,y.R[x; y]) ==  ∀x,y,z:T.  (R[x; y]  R[x; z]  (y z ∈ T))

Lemma: non-forking_wf
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  (non-forking(T;x,y.R[x;y]) ∈ ℙ)

Lemma: non-forking-rel_exp
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].  (non-forking(T;x,y.R[x;y])  (∀n:ℕnon-forking(T;x,y.x λx,y. R[x;y]^n y)))

Lemma: non-forking-wellfounded-linorder
[T:Type]. ∀[R:T ⟶ T ⟶ ℙ].
  (decidable-non-minimal(T;x,y.R[x;y])
   WellFnd{i}(T;x,y.R[x;y])
   (∀m:T. (unique-minimal(T;x,y.R[x;y];m)  non-forking(T;x,y.R[x;y])  WeakLinorder(T;x,y.x (R^*) y))))

Definition: two-factorizations
two-factorizations(n) ==  mapfilter(λa.<a, n ÷ a>a.(n rem =z 0);[1, 1))

Lemma: two-factorizations_wf
[n:ℕ]. (two-factorizations(n) ∈ {p:ℤ × ℤ((1 ≤ (fst(p))) ∧ ((fst(p)) ≤ n)) ∧ (((fst(p)) (snd(p))) n ∈ ℤ)}  List)

Lemma: two-factorizations-no-repeats
[n:ℕ]. no_repeats(ℤ × ℤ;two-factorizations(n))

Lemma: bag-member-two-factorizations
[n:ℕ]. ∀[a,b:ℤ].  uiff(<a, b> ↓∈ two-factorizations(n);(1 ≤ a) ∧ (a ≤ n) ∧ ((a b) n ∈ ℤ))

Definition: d-conv
d-conv(r;f;g) ==  λn.Σ(p∈two-factorizations(n)). (f (fst(p))) (g (snd(p)))

Lemma: d-conv_wf
[r:CRng]. ∀[f,g:ℕ ⟶ |r|].  (d-conv(r;f;g) ∈ ℕ ⟶ |r|)

Definition: andrew
andrew{i:l}(T) ==  {c:Type| (c ⊆T) ∧ (⋂X:c. X)} 

Lemma: andrew_wf
[T:𝕌']. andrew{i:l}(T) ∈ 𝕌supposing T ⊆Type

Definition: letrec
let rec f(x) F[f; x] in f(a) ==  fix((λf,x. F[f; x])) a

Definition: HofstadterF
This is an example of definition that uses  mutual recursion
(see HofstadterM).

The two functions are called "married" functions by Douglas Hofstadter.

We prove joint well-formedness lemma Hofstadter_wf
from which we then derive the typing lemma HofstadterF_wf.

We use explict callbyvalue in the definition because otherwise the
computation is extremely inefficient.
Even with callbyvalue, the recursion is inefficent without memoization.

The memoized version is given in HofstadterL (see HofstadterL_wf).⋅

HofstadterF(n) ==
  if (n) < (1)  then 1  else eval in eval HofstadterF(p) in eval HofstadterM(f) in   m

Definition: HofstadterM
This is an example of definition that uses  mutual recursion
(see HofstadterF).

The two functions are called "married" functions by Douglas Hofstadter.

We prove joint well-formedness lemma Hofstadter_wf
from which we then derive the typing lemma HofstadterM_wf.

We use explict callbyvalue in the definition because otherwise the
computation is extremely inefficient.
Even with callbyvalue, the recursion is inefficent without memoization.

The memoized version is given in HofstadterL (see HofstadterL_wf)⋅

HofstadterM(n) ==
  if (n) < (1)  then 0  else eval in eval HofstadterM(p) in eval HofstadterF(m) in   f

Lemma: Hofstadter_wf
n:ℤ((HofstadterF(n) ∈ if 0 <then ℕelse ℕ+fi ) ∧ (HofstadterM(n) ∈ if 0 ≤then ℕelse ℕfi ))

Lemma: HofstadterM_wf
n:ℤ(HofstadterM(n) ∈ if 0 ≤then ℕelse ℕfi )

Lemma: HofstadterF_wf
n:ℤ(HofstadterF(n) ∈ if 0 <then ℕelse ℕ+fi )

Definition: HofstadterL
HofstadterL(n) ==
  if (n =z 0)
  then [<0, 1>]
  else eval in
       eval HofstadterL(p) in
         let m,f hd(L) 
         in eval m' snd(L[p m]) in
            eval f' fst(L[p f]) in
              [<m', f'> L]
  fi 

Lemma: HofstadterL_wf
n:ℕ
  (HofstadterL(n) ∈ {L:(ℤ × ℤList| 
                     (||L|| (n 1) ∈ ℤ) ∧ (∀i:ℕ1. (L[i] = <HofstadterM(n i), HofstadterF(n i)> ∈ (ℤ × ℤ)))} )

Definition: int_mod_ring
Integers mod n ⌜ℤ_n⌝ form "discrete" commutative ring.

We use the form defined by Paul Jackson. It is dependent tuple
and includes an equality test, and ordering, and division operator.
Since ⌜ℤ_n⌝  isn't ordered and isn't field (unless is prime)
we supply "dummy" operations ⌜λu,v. ff⌝ and ⌜λu,v. (inr ⋅ )⌝ for the
ordering and division.

The equality test ⌜λu,v. (u mod =z mod n)⌝ does provide decidable
equality, so this structure is a ⌜CDRng⌝  (see int_mod_ring_wf) ⋅

int_mod_ring(n) ==  <ℤ_n, λu,v. (u mod =z mod n), λu,v. ff, λu,v. (u v), 0, λu.(-u), λu,v. (u v), 1, λu,v. (inr \000C⋅ )>

Lemma: int_mod_ring_wf
[n:ℕ+]. (int_mod_ring(n) ∈ CDRng)

Lemma: int_mod_ring_wf_field
[p:ℕ+]. int_mod_ring(p) ∈ Field{i} supposing prime(p)

Lemma: vr_test_foo_bar345566
7 ⋂ {3..9-} ≡ {3..7-}

Lemma: vr_test_foo_bar345566546
+ 𝔹 ∈ Void ⟶ 1

Definition: nim-sum
nim-sum(x;y) ==
  if x=0
  then y
  else if y=0
       then x
       else eval rx rem in
            eval qx x ÷ in
            eval ry rem in
            eval qy y ÷ in
            eval nim-sum(qx;qy) in
            eval if rx=ry then else in
              (2 s) d

Lemma: nim-sum_wf
[x,y:ℕ].  (nim-sum(x;y) ∈ ℕ)

Lemma: nim_sum0_lemma
y:Top. (nim-sum(0;y) y)

Lemma: nim-sum-com
[x,y:ℕ].  (nim-sum(x;y) nim-sum(y;x) ∈ ℤ)

Lemma: nim-sum-0
[x:ℕ]. (nim-sum(x;0) x ∈ ℤ)

Lemma: nim-sum-div2
[x,y:ℕ].  (nim-sum(x;y) ÷ nim-sum(x ÷ 2;y ÷ 2))

Lemma: nim-sum-rem2
[x,y:ℕ].  (nim-sum(x;y) rem if rem 2=y rem then else 1)

Lemma: nim-sum-rec
[x,y:ℕ].  (nim-sum(x;y) (2 nim-sum(x ÷ 2;y ÷ 2)) if rem 2=y rem then else 1)

Lemma: nim-sum-assoc
[x,y,z:ℕ].  (nim-sum(x;nim-sum(y;z)) nim-sum(nim-sum(x;y);z) ∈ ℤ)

Lemma: nim-sum-same
[x:ℕ]. (nim-sum(x;x) 0 ∈ ℤ)

Definition: nim-grp
nim-grp() ==  <ℕ, λx,y. (x =z y), λx,y. x ≤y, λx,y. nim-sum(x;y), 0, λx.x>

Lemma: nim-grp_wf
nim-grp() ∈ AbDGrp

Definition: ut1-test
ut1-test{$aaa,$bbb}(x) ==  case of inl(a) => '$aaa'1 inr(b) => '$bbb'1

Lemma: ut1-test_wf
[x:Top Top]. (ut1-test{$aaa,$bbb}(x) ∈ Atom1)



Home Index