Theory ComputeNumeral

Up to index of Isabelle/HOL/HOL-Algebra/example_Bicomplex

theory ComputeNumeral
imports ComputeHOL Float
begin

theory ComputeNumeral
imports ComputeHOL "~~/src/HOL/Real/Float"
begin

(* normalization of bit strings *)
lemmas bitnorm = normalize_bin_simps

(* neg for bit strings *)
lemma neg1: "neg Int.Pls = False" by (simp add: Int.Pls_def)
lemma neg2: "neg Int.Min = True" apply (subst Int.Min_def) by auto
lemma neg3: "neg (Int.Bit0 x) = neg x" apply (simp add: neg_def) apply (subst Bit0_def) by auto
lemma neg4: "neg (Int.Bit1 x) = neg x" apply (simp add: neg_def) apply (subst Bit1_def) by auto  
lemmas bitneg = neg1 neg2 neg3 neg4

(* iszero for bit strings *)
lemma iszero1: "iszero Int.Pls = True" by (simp add: Int.Pls_def iszero_def)
lemma iszero2: "iszero Int.Min = False" apply (subst Int.Min_def) apply (subst iszero_def) by simp
lemma iszero3: "iszero (Int.Bit0 x) = iszero x" apply (subst Int.Bit0_def) apply (subst iszero_def)+ by auto
lemma iszero4: "iszero (Int.Bit1 x) = False" apply (subst Int.Bit1_def) apply (subst iszero_def)+  apply simp by arith
lemmas bitiszero = iszero1 iszero2 iszero3 iszero4

(* lezero for bit strings *)
constdefs
  "lezero x == (x ≤ 0)"
lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto
lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto
lemma lezero3: "lezero (Int.Bit0 x) = lezero x" unfolding Int.Bit0_def lezero_def by auto
lemma lezero4: "lezero (Int.Bit1 x) = neg x" unfolding Int.Bit1_def lezero_def neg_def by auto
lemmas bitlezero = lezero1 lezero2 lezero3 lezero4

(* equality for bit strings *)
lemma biteq1: "(Int.Pls = Int.Pls) = True" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq2: "(Int.Min = Int.Min) = True" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq3: "(Int.Pls = Int.Min) = False" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq4: "(Int.Min = Int.Pls) = False" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq5: "(Int.Bit0 x = Int.Bit0 y) = (x = y)" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq6: "(Int.Bit1 x = Int.Bit1 y) = (x = y)" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq7: "(Int.Bit0 x = Int.Bit1 y) = False" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq8: "(Int.Bit1 x = Int.Bit0 y) = False" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq9: "(Int.Pls = Int.Bit0 x) = (Int.Pls = x)" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq10: "(Int.Pls = Int.Bit1 x) = False" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq11: "(Int.Min = Int.Bit0 x) = False" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq12: "(Int.Min = Int.Bit1 x) = (Int.Min = x)" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq13: "(Int.Bit0 x = Int.Pls) = (x = Int.Pls)" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq14: "(Int.Bit1 x = Int.Pls) = False" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq15: "(Int.Bit0 x = Int.Min) = False" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma biteq16: "(Int.Bit1 x = Int.Min) = (x = Int.Min)" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemmas biteq = biteq1 biteq2 biteq3 biteq4 biteq5 biteq6 biteq7 biteq8 biteq9 biteq10 biteq11 biteq12 biteq13 biteq14 biteq15 biteq16

(* x < y for bit strings *)
lemma bitless1: "(Int.Pls < Int.Min) = False" by (simp add: less_int_code)
lemma bitless2: "(Int.Pls < Int.Pls) = False" by (simp add: less_int_code)
lemma bitless3: "(Int.Min < Int.Pls) = True" by (simp add: less_int_code)
lemma bitless4: "(Int.Min < Int.Min) = False" by (simp add: less_int_code)
lemma bitless5: "(Int.Bit0 x < Int.Bit0 y) = (x < y)" by (simp add: less_int_code)
lemma bitless6: "(Int.Bit1 x < Int.Bit1 y) = (x < y)" by (simp add: less_int_code)
lemma bitless7: "(Int.Bit0 x < Int.Bit1 y) = (x ≤ y)" by (simp add: less_int_code)
lemma bitless8: "(Int.Bit1 x < Int.Bit0 y) = (x < y)" by (simp add: less_int_code)
lemma bitless9: "(Int.Pls < Int.Bit0 x) = (Int.Pls < x)" by (simp add: less_int_code)
lemma bitless10: "(Int.Pls < Int.Bit1 x) = (Int.Pls ≤ x)" by (simp add: less_int_code)
lemma bitless11: "(Int.Min < Int.Bit0 x) = (Int.Pls ≤ x)" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma bitless12: "(Int.Min < Int.Bit1 x) = (Int.Min < x)" by (simp add: less_int_code)
lemma bitless13: "(Int.Bit0 x < Int.Pls) = (x < Int.Pls)" by (simp add: less_int_code)
lemma bitless14: "(Int.Bit1 x < Int.Pls) = (x < Int.Pls)" by (simp add: less_int_code)
lemma bitless15: "(Int.Bit0 x < Int.Min) = (x < Int.Pls)" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma bitless16: "(Int.Bit1 x < Int.Min) = (x < Int.Min)" by (simp add: less_int_code)
lemmas bitless = bitless1 bitless2 bitless3 bitless4 bitless5 bitless6 bitless7 bitless8 
                 bitless9 bitless10 bitless11 bitless12 bitless13 bitless14 bitless15 bitless16

(* x ≤ y for bit strings *)
lemma bitle1: "(Int.Pls ≤ Int.Min) = False" by (simp add: less_eq_int_code)
lemma bitle2: "(Int.Pls ≤ Int.Pls) = True" by (simp add: less_eq_int_code)
lemma bitle3: "(Int.Min ≤ Int.Pls) = True" by (simp add: less_eq_int_code)
lemma bitle4: "(Int.Min ≤ Int.Min) = True" by (simp add: less_eq_int_code)
lemma bitle5: "(Int.Bit0 x ≤ Int.Bit0 y) = (x ≤ y)" by (simp add: less_eq_int_code)
lemma bitle6: "(Int.Bit1 x ≤ Int.Bit1 y) = (x ≤ y)" by (simp add: less_eq_int_code)
lemma bitle7: "(Int.Bit0 x ≤ Int.Bit1 y) = (x ≤ y)" by (simp add: less_eq_int_code)
lemma bitle8: "(Int.Bit1 x ≤ Int.Bit0 y) = (x < y)" by (simp add: less_eq_int_code)
lemma bitle9: "(Int.Pls ≤ Int.Bit0 x) = (Int.Pls ≤ x)" by (simp add: less_eq_int_code)
lemma bitle10: "(Int.Pls ≤ Int.Bit1 x) = (Int.Pls ≤ x)" by (simp add: less_eq_int_code)
lemma bitle11: "(Int.Min ≤ Int.Bit0 x) = (Int.Pls ≤ x)" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma bitle12: "(Int.Min ≤ Int.Bit1 x) = (Int.Min ≤ x)" by (simp add: less_eq_int_code)
lemma bitle13: "(Int.Bit0 x ≤ Int.Pls) = (x ≤ Int.Pls)" by (simp add: less_eq_int_code)
lemma bitle14: "(Int.Bit1 x ≤ Int.Pls) = (x < Int.Pls)" by (simp add: less_eq_int_code)
lemma bitle15: "(Int.Bit0 x ≤ Int.Min) = (x < Int.Pls)" unfolding Pls_def Min_def Bit0_def Bit1_def by presburger
lemma bitle16: "(Int.Bit1 x ≤ Int.Min) = (x ≤ Int.Min)" by (simp add: less_eq_int_code)
lemmas bitle = bitle1 bitle2 bitle3 bitle4 bitle5 bitle6 bitle7 bitle8 
                 bitle9 bitle10 bitle11 bitle12 bitle13 bitle14 bitle15 bitle16

(* succ for bit strings *)
lemmas bitsucc = succ_bin_simps

(* pred for bit strings *)
lemmas bitpred = pred_bin_simps

(* unary minus for bit strings *)
lemmas bituminus = minus_bin_simps

(* addition for bit strings *)
lemmas bitadd = add_bin_simps

(* multiplication for bit strings *) 
lemma mult_Pls_right: "x * Int.Pls = Int.Pls" by (simp add: Pls_def)
lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute, simp add: mult_Min)
lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0)
lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp
lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)"
  unfolding Bit0_def Bit1_def by (simp add: ring_simps)
lemmas bitmul = mult_Pls mult_Min mult_Pls_right mult_Min_right multb0x multxb0 multb1

lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul 

constdefs 
  "nat_norm_number_of (x::nat) == x"

lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)"
  apply (simp add: nat_norm_number_of_def)
  unfolding lezero_def iszero_def neg_def
  apply (simp add: number_of_is_id)
  done

(* Normalization of nat literals *)
lemma natnorm0: "(0::nat) = number_of (Int.Pls)" by auto
lemma natnorm1: "(1 :: nat) = number_of (Int.Bit1 Int.Pls)"  by auto 
lemmas natnorm = natnorm0 natnorm1 nat_norm_number_of

(* Suc *)
lemma natsuc: "Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))" by (auto simp add: number_of_is_id)

(* Addition for nat *)
lemma natadd: "number_of x + ((number_of y)::nat) = (if neg x then (number_of y) else (if neg y then number_of x else (number_of (x + y))))"
  by (auto simp add: number_of_is_id)

(* Subtraction for nat *)
lemma natsub: "(number_of x) - ((number_of y)::nat) = 
  (if neg x then 0 else (if neg y then number_of x else (nat_norm_number_of (number_of (x + (- y))))))"
  unfolding nat_norm_number_of
  by (auto simp add: number_of_is_id neg_def lezero_def iszero_def Let_def nat_number_of_def)

(* Multiplication for nat *)
lemma natmul: "(number_of x) * ((number_of y)::nat) = 
  (if neg x then 0 else (if neg y then 0 else number_of (x * y)))"
  apply (auto simp add: number_of_is_id neg_def iszero_def)
  apply (case_tac "x > 0")
  apply auto
  apply (simp add: mult_strict_left_mono[where a=y and b=0 and c=x, simplified])
  done

lemma nateq: "(((number_of x)::nat) = (number_of y)) = ((lezero x ∧ lezero y) ∨ (x = y))"
  by (auto simp add: iszero_def lezero_def neg_def number_of_is_id)

lemma natless: "(((number_of x)::nat) < (number_of y)) = ((x < y) ∧ (¬ (lezero y)))"
  by (auto simp add: number_of_is_id neg_def lezero_def)

lemma natle: "(((number_of x)::nat) ≤ (number_of y)) = (y < x --> lezero x)"
  by (auto simp add: number_of_is_id lezero_def nat_number_of_def)

fun natfac :: "nat => nat"
where
   "natfac n = (if n = 0 then 1 else n * (natfac (n - 1)))"

lemmas compute_natarith = bitarith natnorm natsuc natadd natsub natmul nateq natless natle natfac.simps

lemma number_eq: "(((number_of x)::'a::{number_ring, ordered_idom}) = (number_of y)) = (x = y)"
  unfolding number_of_eq
  apply simp
  done

lemma number_le: "(((number_of x)::'a::{number_ring, ordered_idom}) ≤  (number_of y)) = (x ≤ y)"
  unfolding number_of_eq
  apply simp
  done

lemma number_less: "(((number_of x)::'a::{number_ring, ordered_idom}) <  (number_of y)) = (x < y)"
  unfolding number_of_eq 
  apply simp
  done

lemma number_diff: "((number_of x)::'a::{number_ring, ordered_idom}) - number_of y = number_of (x + (- y))"
  apply (subst diff_number_of_eq)
  apply simp
  done

lemmas number_norm = number_of_Pls[symmetric] numeral_1_eq_1[symmetric]

lemmas compute_numberarith = number_of_minus[symmetric] number_of_add[symmetric] number_diff number_of_mult[symmetric] number_norm number_eq number_le number_less

lemma compute_real_of_nat_number_of: "real ((number_of v)::nat) = (if neg v then 0 else number_of v)"
  by (simp only: real_of_nat_number_of number_of_is_id)

lemma compute_nat_of_int_number_of: "nat ((number_of v)::int) = (number_of v)"
  by simp

lemmas compute_num_conversions = compute_real_of_nat_number_of compute_nat_of_int_number_of real_number_of

lemmas zpowerarith = zpower_number_of_even
  zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring]
  zpower_Pls zpower_Min

(* div, mod *)

lemma adjust: "adjust b (q, r) = (if 0 ≤ r - b then (2 * q + 1, r - b) else (2 * q, r))"
  by (auto simp only: adjust_def)

lemma negateSnd: "negateSnd (q, r) = (q, -r)" 
  by (auto simp only: negateSnd_def)

lemma divAlg: "divAlg (a, b) = (if 0≤a then
                  if 0≤b then posDivAlg a b
                  else if a=0 then (0, 0)
                       else negateSnd (negDivAlg (-a) (-b))
               else 
                  if 0<b then negDivAlg a b
                  else negateSnd (posDivAlg (-a) (-b)))"
  by (auto simp only: divAlg_def)

lemmas compute_div_mod = div_def mod_def divAlg adjust negateSnd posDivAlg.simps negDivAlg.simps



(* collecting all the theorems *)

lemma even_Pls: "even (Int.Pls) = True"
  apply (unfold Pls_def even_def)
  by simp

lemma even_Min: "even (Int.Min) = False"
  apply (unfold Min_def even_def)
  by simp

lemma even_B0: "even (Int.Bit0 x) = True"
  apply (unfold Bit0_def)
  by simp

lemma even_B1: "even (Int.Bit1 x) = False"
  apply (unfold Bit1_def)
  by simp

lemma even_number_of: "even ((number_of w)::int) = even w"
  by (simp only: number_of_is_id)

lemmas compute_even = even_Pls even_Min even_B0 even_B1 even_number_of

lemmas compute_numeral = compute_if compute_let compute_pair compute_bool 
                         compute_natarith compute_numberarith max_def min_def compute_num_conversions zpowerarith compute_div_mod compute_even

end




lemma bitnorm:

  Int.Bit0 Int.Pls = Int.Pls
  Int.Bit1 Int.Min = Int.Min

lemma neg1:

  neg Int.Pls = False

lemma neg2:

  neg Int.Min = True

lemma neg3:

  neg (Int.Bit0 x) = neg x

lemma neg4:

  neg (Int.Bit1 x) = neg x

lemma bitneg:

  neg Int.Pls = False
  neg Int.Min = True
  neg (Int.Bit0 x) = neg x
  neg (Int.Bit1 x) = neg x

lemma iszero1:

  iszero Int.Pls = True

lemma iszero2:

  iszero Int.Min = False

lemma iszero3:

  iszero (Int.Bit0 x) = iszero x

lemma iszero4:

  iszero (Int.Bit1 x) = False

lemma bitiszero:

  iszero Int.Pls = True
  iszero Int.Min = False
  iszero (Int.Bit0 x) = iszero x
  iszero (Int.Bit1 x) = False

lemma lezero1:

  lezero Int.Pls = True

lemma lezero2:

  lezero Int.Min = True

lemma lezero3:

  lezero (Int.Bit0 x) = lezero x

lemma lezero4:

  lezero (Int.Bit1 x) = neg x

lemma bitlezero:

  lezero Int.Pls = True
  lezero Int.Min = True
  lezero (Int.Bit0 x) = lezero x
  lezero (Int.Bit1 x) = neg x

lemma biteq1:

  (Int.Pls = Int.Pls) = True

lemma biteq2:

  (Int.Min = Int.Min) = True

lemma biteq3:

  (Int.Pls = Int.Min) = False

lemma biteq4:

  (Int.Min = Int.Pls) = False

lemma biteq5:

  (Int.Bit0 x = Int.Bit0 y) = (x = y)

lemma biteq6:

  (Int.Bit1 x = Int.Bit1 y) = (x = y)

lemma biteq7:

  (Int.Bit0 x = Int.Bit1 y) = False

lemma biteq8:

  (Int.Bit1 x = Int.Bit0 y) = False

lemma biteq9:

  (Int.Pls = Int.Bit0 x) = (Int.Pls = x)

lemma biteq10:

  (Int.Pls = Int.Bit1 x) = False

lemma biteq11:

  (Int.Min = Int.Bit0 x) = False

lemma biteq12:

  (Int.Min = Int.Bit1 x) = (Int.Min = x)

lemma biteq13:

  (Int.Bit0 x = Int.Pls) = (x = Int.Pls)

lemma biteq14:

  (Int.Bit1 x = Int.Pls) = False

lemma biteq15:

  (Int.Bit0 x = Int.Min) = False

lemma biteq16:

  (Int.Bit1 x = Int.Min) = (x = Int.Min)

lemma biteq:

  (Int.Pls = Int.Pls) = True
  (Int.Min = Int.Min) = True
  (Int.Pls = Int.Min) = False
  (Int.Min = Int.Pls) = False
  (Int.Bit0 x = Int.Bit0 y) = (x = y)
  (Int.Bit1 x = Int.Bit1 y) = (x = y)
  (Int.Bit0 x = Int.Bit1 y) = False
  (Int.Bit1 x = Int.Bit0 y) = False
  (Int.Pls = Int.Bit0 x) = (Int.Pls = x)
  (Int.Pls = Int.Bit1 x) = False
  (Int.Min = Int.Bit0 x) = False
  (Int.Min = Int.Bit1 x) = (Int.Min = x)
  (Int.Bit0 x = Int.Pls) = (x = Int.Pls)
  (Int.Bit1 x = Int.Pls) = False
  (Int.Bit0 x = Int.Min) = False
  (Int.Bit1 x = Int.Min) = (x = Int.Min)

lemma bitless1:

  (Int.Pls < Int.Min) = False

lemma bitless2:

  (Int.Pls < Int.Pls) = False

lemma bitless3:

  (Int.Min < Int.Pls) = True

lemma bitless4:

  (Int.Min < Int.Min) = False

lemma bitless5:

  (Int.Bit0 x < Int.Bit0 y) = (x < y)

lemma bitless6:

  (Int.Bit1 x < Int.Bit1 y) = (x < y)

lemma bitless7:

  (Int.Bit0 x < Int.Bit1 y) = (x  y)

lemma bitless8:

  (Int.Bit1 x < Int.Bit0 y) = (x < y)

lemma bitless9:

  (Int.Pls < Int.Bit0 x) = (Int.Pls < x)

lemma bitless10:

  (Int.Pls < Int.Bit1 x) = (Int.Pls  x)

lemma bitless11:

  (Int.Min < Int.Bit0 x) = (Int.Pls  x)

lemma bitless12:

  (Int.Min < Int.Bit1 x) = (Int.Min < x)

lemma bitless13:

  (Int.Bit0 x < Int.Pls) = (x < Int.Pls)

lemma bitless14:

  (Int.Bit1 x < Int.Pls) = (x < Int.Pls)

lemma bitless15:

  (Int.Bit0 x < Int.Min) = (x < Int.Pls)

lemma bitless16:

  (Int.Bit1 x < Int.Min) = (x < Int.Min)

lemma bitless:

  (Int.Pls < Int.Min) = False
  (Int.Pls < Int.Pls) = False
  (Int.Min < Int.Pls) = True
  (Int.Min < Int.Min) = False
  (Int.Bit0 x < Int.Bit0 y) = (x < y)
  (Int.Bit1 x < Int.Bit1 y) = (x < y)
  (Int.Bit0 x < Int.Bit1 y) = (x  y)
  (Int.Bit1 x < Int.Bit0 y) = (x < y)
  (Int.Pls < Int.Bit0 x) = (Int.Pls < x)
  (Int.Pls < Int.Bit1 x) = (Int.Pls  x)
  (Int.Min < Int.Bit0 x) = (Int.Pls  x)
  (Int.Min < Int.Bit1 x) = (Int.Min < x)
  (Int.Bit0 x < Int.Pls) = (x < Int.Pls)
  (Int.Bit1 x < Int.Pls) = (x < Int.Pls)
  (Int.Bit0 x < Int.Min) = (x < Int.Pls)
  (Int.Bit1 x < Int.Min) = (x < Int.Min)

lemma bitle1:

  (Int.Pls  Int.Min) = False

lemma bitle2:

  (Int.Pls  Int.Pls) = True

lemma bitle3:

  (Int.Min  Int.Pls) = True

lemma bitle4:

  (Int.Min  Int.Min) = True

lemma bitle5:

  (Int.Bit0 x  Int.Bit0 y) = (x  y)

lemma bitle6:

  (Int.Bit1 x  Int.Bit1 y) = (x  y)

lemma bitle7:

  (Int.Bit0 x  Int.Bit1 y) = (x  y)

lemma bitle8:

  (Int.Bit1 x  Int.Bit0 y) = (x < y)

lemma bitle9:

  (Int.Pls  Int.Bit0 x) = (Int.Pls  x)

lemma bitle10:

  (Int.Pls  Int.Bit1 x) = (Int.Pls  x)

lemma bitle11:

  (Int.Min  Int.Bit0 x) = (Int.Pls  x)

lemma bitle12:

  (Int.Min  Int.Bit1 x) = (Int.Min  x)

lemma bitle13:

  (Int.Bit0 x  Int.Pls) = (x  Int.Pls)

lemma bitle14:

  (Int.Bit1 x  Int.Pls) = (x < Int.Pls)

lemma bitle15:

  (Int.Bit0 x  Int.Min) = (x < Int.Pls)

lemma bitle16:

  (Int.Bit1 x  Int.Min) = (x  Int.Min)

lemma bitle:

  (Int.Pls  Int.Min) = False
  (Int.Pls  Int.Pls) = True
  (Int.Min  Int.Pls) = True
  (Int.Min  Int.Min) = True
  (Int.Bit0 x  Int.Bit0 y) = (x  y)
  (Int.Bit1 x  Int.Bit1 y) = (x  y)
  (Int.Bit0 x  Int.Bit1 y) = (x  y)
  (Int.Bit1 x  Int.Bit0 y) = (x < y)
  (Int.Pls  Int.Bit0 x) = (Int.Pls  x)
  (Int.Pls  Int.Bit1 x) = (Int.Pls  x)
  (Int.Min  Int.Bit0 x) = (Int.Pls  x)
  (Int.Min  Int.Bit1 x) = (Int.Min  x)
  (Int.Bit0 x  Int.Pls) = (x  Int.Pls)
  (Int.Bit1 x  Int.Pls) = (x < Int.Pls)
  (Int.Bit0 x  Int.Min) = (x < Int.Pls)
  (Int.Bit1 x  Int.Min) = (x  Int.Min)

lemma bitsucc:

  Int.succ Int.Pls = Int.Bit1 Int.Pls
  Int.succ Int.Min = Int.Pls
  Int.succ (Int.Bit0 k) = Int.Bit1 k
  Int.succ (Int.Bit1 k) = Int.Bit0 (Int.succ k)

lemma bitpred:

  Int.pred Int.Pls = Int.Min
  Int.pred Int.Min = Int.Bit0 Int.Min
  Int.pred (Int.Bit0 k) = Int.Bit1 (Int.pred k)
  Int.pred (Int.Bit1 k) = Int.Bit0 k

lemma bituminus:

  - Int.Pls = Int.Pls
  - Int.Min = Int.Bit1 Int.Pls
  - Int.Bit0 k = Int.Bit0 (- k)
  - Int.Bit1 k = Int.Bit1 (Int.pred (- k))

lemma bitadd:

  Int.Pls + k = k
  Int.Min + k = Int.pred k
  k + Int.Pls = k
  k + Int.Min = Int.pred k
  Int.Bit0 k + Int.Bit0 l = Int.Bit0 (k + l)
  Int.Bit0 k + Int.Bit1 l = Int.Bit1 (k + l)
  Int.Bit1 k + Int.Bit0 l = Int.Bit1 (k + l)
  Int.Bit1 k + Int.Bit1 l = Int.Bit0 (k + Int.succ l)

lemma mult_Pls_right:

  x * Int.Pls = Int.Pls

lemma mult_Min_right:

  x * Int.Min = - x

lemma multb0x:

  Int.Bit0 x * y = Int.Bit0 (x * y)

lemma multxb0:

  x * Int.Bit0 y = Int.Bit0 (x * y)

lemma multb1:

  Int.Bit1 x * Int.Bit1 y = Int.Bit1 (Int.Bit0 (x * y) + x + y)

lemma bitmul:

  Int.Pls * w = Int.Pls
  Int.Min * k = - k
  x * Int.Pls = Int.Pls
  x * Int.Min = - x
  Int.Bit0 x * y = Int.Bit0 (x * y)
  x * Int.Bit0 y = Int.Bit0 (x * y)
  Int.Bit1 x * Int.Bit1 y = Int.Bit1 (Int.Bit0 (x * y) + x + y)

lemma bitarith:

  Int.Bit0 Int.Pls = Int.Pls
  Int.Bit1 Int.Min = Int.Min
  iszero Int.Pls = True
  iszero Int.Min = False
  iszero (Int.Bit0 x) = iszero x
  iszero (Int.Bit1 x) = False
  neg Int.Pls = False
  neg Int.Min = True
  neg (Int.Bit0 x) = neg x
  neg (Int.Bit1 x) = neg x
  lezero Int.Pls = True
  lezero Int.Min = True
  lezero (Int.Bit0 x) = lezero x
  lezero (Int.Bit1 x) = neg x
  (Int.Pls = Int.Pls) = True
  (Int.Min = Int.Min) = True
  (Int.Pls = Int.Min) = False
  (Int.Min = Int.Pls) = False
  (Int.Bit0 x = Int.Bit0 y) = (x = y)
  (Int.Bit1 x = Int.Bit1 y) = (x = y)
  (Int.Bit0 x = Int.Bit1 y) = False
  (Int.Bit1 x = Int.Bit0 y) = False
  (Int.Pls = Int.Bit0 x) = (Int.Pls = x)
  (Int.Pls = Int.Bit1 x) = False
  (Int.Min = Int.Bit0 x) = False
  (Int.Min = Int.Bit1 x) = (Int.Min = x)
  (Int.Bit0 x = Int.Pls) = (x = Int.Pls)
  (Int.Bit1 x = Int.Pls) = False
  (Int.Bit0 x = Int.Min) = False
  (Int.Bit1 x = Int.Min) = (x = Int.Min)
  (Int.Pls < Int.Min) = False
  (Int.Pls < Int.Pls) = False
  (Int.Min < Int.Pls) = True
  (Int.Min < Int.Min) = False
  (Int.Bit0 x < Int.Bit0 y) = (x < y)
  (Int.Bit1 x < Int.Bit1 y) = (x < y)
  (Int.Bit0 x < Int.Bit1 y) = (x  y)
  (Int.Bit1 x < Int.Bit0 y) = (x < y)
  (Int.Pls < Int.Bit0 x) = (Int.Pls < x)
  (Int.Pls < Int.Bit1 x) = (Int.Pls  x)
  (Int.Min < Int.Bit0 x) = (Int.Pls  x)
  (Int.Min < Int.Bit1 x) = (Int.Min < x)
  (Int.Bit0 x < Int.Pls) = (x < Int.Pls)
  (Int.Bit1 x < Int.Pls) = (x < Int.Pls)
  (Int.Bit0 x < Int.Min) = (x < Int.Pls)
  (Int.Bit1 x < Int.Min) = (x < Int.Min)
  (Int.Pls  Int.Min) = False
  (Int.Pls  Int.Pls) = True
  (Int.Min  Int.Pls) = True
  (Int.Min  Int.Min) = True
  (Int.Bit0 x  Int.Bit0 y) = (x  y)
  (Int.Bit1 x  Int.Bit1 y) = (x  y)
  (Int.Bit0 x  Int.Bit1 y) = (x  y)
  (Int.Bit1 x  Int.Bit0 y) = (x < y)
  (Int.Pls  Int.Bit0 x) = (Int.Pls  x)
  (Int.Pls  Int.Bit1 x) = (Int.Pls  x)
  (Int.Min  Int.Bit0 x) = (Int.Pls  x)
  (Int.Min  Int.Bit1 x) = (Int.Min  x)
  (Int.Bit0 x  Int.Pls) = (x  Int.Pls)
  (Int.Bit1 x  Int.Pls) = (x < Int.Pls)
  (Int.Bit0 x  Int.Min) = (x < Int.Pls)
  (Int.Bit1 x  Int.Min) = (x  Int.Min)
  Int.succ Int.Pls = Int.Bit1 Int.Pls
  Int.succ Int.Min = Int.Pls
  Int.succ (Int.Bit0 k) = Int.Bit1 k
  Int.succ (Int.Bit1 k) = Int.Bit0 (Int.succ k)
  Int.pred Int.Pls = Int.Min
  Int.pred Int.Min = Int.Bit0 Int.Min
  Int.pred (Int.Bit0 k) = Int.Bit1 (Int.pred k)
  Int.pred (Int.Bit1 k) = Int.Bit0 k
  - Int.Pls = Int.Pls
  - Int.Min = Int.Bit1 Int.Pls
  - Int.Bit0 k = Int.Bit0 (- k)
  - Int.Bit1 k = Int.Bit1 (Int.pred (- k))
  Int.Pls + k = k
  Int.Min + k = Int.pred k
  k + Int.Pls = k
  k + Int.Min = Int.pred k
  Int.Bit0 k + Int.Bit0 l = Int.Bit0 (k + l)
  Int.Bit0 k + Int.Bit1 l = Int.Bit1 (k + l)
  Int.Bit1 k + Int.Bit0 l = Int.Bit1 (k + l)
  Int.Bit1 k + Int.Bit1 l = Int.Bit0 (k + Int.succ l)
  Int.Pls * w = Int.Pls
  Int.Min * k = - k
  x * Int.Pls = Int.Pls
  x * Int.Min = - x
  Int.Bit0 x * y = Int.Bit0 (x * y)
  x * Int.Bit0 y = Int.Bit0 (x * y)
  Int.Bit1 x * Int.Bit1 y = Int.Bit1 (Int.Bit0 (x * y) + x + y)

lemma nat_norm_number_of:

  nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)

lemma natnorm0:

  0 = Numeral0

lemma natnorm1:

  1 = Numeral1

lemma natnorm:

  0 = Numeral0
  1 = Numeral1
  nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)

lemma natsuc:

  Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))

lemma natadd:

  number_of x + number_of y =
  (if neg x then number_of y
   else if neg y then number_of x else number_of (x + y))

lemma natsub:

  number_of x - number_of y =
  (if neg x then 0
   else if neg y then number_of x else nat_norm_number_of (number_of (x + - y)))

lemma natmul:

  number_of x * number_of y =
  (if neg x then 0 else if neg y then 0 else number_of (x * y))

lemma nateq:

  (number_of x = number_of y) = (lezero x ∧ lezero yx = y)

lemma natless:

  (number_of x < number_of y) = (x < y ∧ ¬ lezero y)

lemma natle:

  (number_of x  number_of y) = (y < x --> lezero x)

lemma compute_natarith:

  Int.Bit0 Int.Pls = Int.Pls
  Int.Bit1 Int.Min = Int.Min
  iszero Int.Pls = True
  iszero Int.Min = False
  iszero (Int.Bit0 x) = iszero x
  iszero (Int.Bit1 x) = False
  neg Int.Pls = False
  neg Int.Min = True
  neg (Int.Bit0 x) = neg x
  neg (Int.Bit1 x) = neg x
  lezero Int.Pls = True
  lezero Int.Min = True
  lezero (Int.Bit0 x) = lezero x
  lezero (Int.Bit1 x) = neg x
  (Int.Pls = Int.Pls) = True
  (Int.Min = Int.Min) = True
  (Int.Pls = Int.Min) = False
  (Int.Min = Int.Pls) = False
  (Int.Bit0 x = Int.Bit0 y) = (x = y)
  (Int.Bit1 x = Int.Bit1 y) = (x = y)
  (Int.Bit0 x = Int.Bit1 y) = False
  (Int.Bit1 x = Int.Bit0 y) = False
  (Int.Pls = Int.Bit0 x) = (Int.Pls = x)
  (Int.Pls = Int.Bit1 x) = False
  (Int.Min = Int.Bit0 x) = False
  (Int.Min = Int.Bit1 x) = (Int.Min = x)
  (Int.Bit0 x = Int.Pls) = (x = Int.Pls)
  (Int.Bit1 x = Int.Pls) = False
  (Int.Bit0 x = Int.Min) = False
  (Int.Bit1 x = Int.Min) = (x = Int.Min)
  (Int.Pls < Int.Min) = False
  (Int.Pls < Int.Pls) = False
  (Int.Min < Int.Pls) = True
  (Int.Min < Int.Min) = False
  (Int.Bit0 x < Int.Bit0 y) = (x < y)
  (Int.Bit1 x < Int.Bit1 y) = (x < y)
  (Int.Bit0 x < Int.Bit1 y) = (x  y)
  (Int.Bit1 x < Int.Bit0 y) = (x < y)
  (Int.Pls < Int.Bit0 x) = (Int.Pls < x)
  (Int.Pls < Int.Bit1 x) = (Int.Pls  x)
  (Int.Min < Int.Bit0 x) = (Int.Pls  x)
  (Int.Min < Int.Bit1 x) = (Int.Min < x)
  (Int.Bit0 x < Int.Pls) = (x < Int.Pls)
  (Int.Bit1 x < Int.Pls) = (x < Int.Pls)
  (Int.Bit0 x < Int.Min) = (x < Int.Pls)
  (Int.Bit1 x < Int.Min) = (x < Int.Min)
  (Int.Pls  Int.Min) = False
  (Int.Pls  Int.Pls) = True
  (Int.Min  Int.Pls) = True
  (Int.Min  Int.Min) = True
  (Int.Bit0 x  Int.Bit0 y) = (x  y)
  (Int.Bit1 x  Int.Bit1 y) = (x  y)
  (Int.Bit0 x  Int.Bit1 y) = (x  y)
  (Int.Bit1 x  Int.Bit0 y) = (x < y)
  (Int.Pls  Int.Bit0 x) = (Int.Pls  x)
  (Int.Pls  Int.Bit1 x) = (Int.Pls  x)
  (Int.Min  Int.Bit0 x) = (Int.Pls  x)
  (Int.Min  Int.Bit1 x) = (Int.Min  x)
  (Int.Bit0 x  Int.Pls) = (x  Int.Pls)
  (Int.Bit1 x  Int.Pls) = (x < Int.Pls)
  (Int.Bit0 x  Int.Min) = (x < Int.Pls)
  (Int.Bit1 x  Int.Min) = (x  Int.Min)
  Int.succ Int.Pls = Int.Bit1 Int.Pls
  Int.succ Int.Min = Int.Pls
  Int.succ (Int.Bit0 k) = Int.Bit1 k
  Int.succ (Int.Bit1 k) = Int.Bit0 (Int.succ k)
  Int.pred Int.Pls = Int.Min
  Int.pred Int.Min = Int.Bit0 Int.Min
  Int.pred (Int.Bit0 k) = Int.Bit1 (Int.pred k)
  Int.pred (Int.Bit1 k) = Int.Bit0 k
  - Int.Pls = Int.Pls
  - Int.Min = Int.Bit1 Int.Pls
  - Int.Bit0 k = Int.Bit0 (- k)
  - Int.Bit1 k = Int.Bit1 (Int.pred (- k))
  Int.Pls + k = k
  Int.Min + k = Int.pred k
  k + Int.Pls = k
  k + Int.Min = Int.pred k
  Int.Bit0 k + Int.Bit0 l = Int.Bit0 (k + l)
  Int.Bit0 k + Int.Bit1 l = Int.Bit1 (k + l)
  Int.Bit1 k + Int.Bit0 l = Int.Bit1 (k + l)
  Int.Bit1 k + Int.Bit1 l = Int.Bit0 (k + Int.succ l)
  Int.Pls * w = Int.Pls
  Int.Min * k = - k
  x * Int.Pls = Int.Pls
  x * Int.Min = - x
  Int.Bit0 x * y = Int.Bit0 (x * y)
  x * Int.Bit0 y = Int.Bit0 (x * y)
  Int.Bit1 x * Int.Bit1 y = Int.Bit1 (Int.Bit0 (x * y) + x + y)
  0 = Numeral0
  1 = Numeral1
  nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)
  Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))
  number_of x + number_of y =
  (if neg x then number_of y
   else if neg y then number_of x else number_of (x + y))
  number_of x - number_of y =
  (if neg x then 0
   else if neg y then number_of x else nat_norm_number_of (number_of (x + - y)))
  number_of x * number_of y =
  (if neg x then 0 else if neg y then 0 else number_of (x * y))
  (number_of x = number_of y) = (lezero x ∧ lezero yx = y)
  (number_of x < number_of y) = (x < y ∧ ¬ lezero y)
  (number_of x  number_of y) = (y < x --> lezero x)
  natfac n = (if n = 0 then 1 else n * natfac (n - 1))

lemma number_eq:

  (number_of x = number_of y) = (x = y)

lemma number_le:

  (number_of x  number_of y) = (x  y)

lemma number_less:

  (number_of x < number_of y) = (x < y)

lemma number_diff:

  number_of x - number_of y = number_of (x + - y)

lemma number_norm:

  (0::'a) = Numeral0
  (1::'a) = Numeral1

lemma compute_numberarith:

  - number_of w = number_of (- w)
  number_of v + number_of w = number_of (v + w)
  number_of x - number_of y = number_of (x + - y)
  number_of v * number_of w = number_of (v * w)
  (0::'a) = Numeral0
  (1::'a) = Numeral1
  (number_of x = number_of y) = (x = y)
  (number_of x  number_of y) = (x  y)
  (number_of x < number_of y) = (x < y)

lemma compute_real_of_nat_number_of:

  real (number_of v) = (if neg v then 0 else number_of v)

lemma compute_nat_of_int_number_of:

  nat (number_of v) = number_of v

lemma compute_num_conversions:

  real (number_of v) = (if neg v then 0 else number_of v)
  nat (number_of v) = number_of v
  real (number_of v) = number_of v

lemma zpowerarith:

  z ^ number_of (Int.Bit0 w) = (let w = z ^ number_of w in w * w)
  z ^ number_of (Int.Bit1 w) =
  (if Numeral0  number_of w then let w = z ^ number_of w in z * w * w
   else Numeral1)
  z ^ Numeral0 = Numeral1
  z ^ -1 = Numeral1

lemma adjust:

  adjust b (q, r) = (if 0  r - b then (2 * q + 1, r - b) else (2 * q, r))

lemma negateSnd:

  negateSnd (q, r) = (q, - r)

lemma divAlg:

  divAlg (a, b) =
  (if 0  a
   then if 0  b then posDivAlg a b
        else if a = 0 then 0N else negateSnd (negDivAlg (- a) (- b))
   else if 0 < b then negDivAlg a b else negateSnd (posDivAlg (- a) (- b)))

lemma compute_div_mod:

  a div b = fst (divAlg (a, b))
  a mod b = snd (divAlg (a, b))
  divAlg (a, b) =
  (if 0  a
   then if 0  b then posDivAlg a b
        else if a = 0 then 0N else negateSnd (negDivAlg (- a) (- b))
   else if 0 < b then negDivAlg a b else negateSnd (posDivAlg (- a) (- b)))
  adjust b (q, r) = (if 0  r - b then (2 * q + 1, r - b) else (2 * q, r))
  negateSnd (q, r) = (q, - r)
  posDivAlg a b =
  (if a < bb  0 then (0, a) else adjust b (posDivAlg a (2 * b)))
  negDivAlg a b =
  (if 0  a + bb  0 then (-1, a + b) else adjust b (negDivAlg a (2 * b)))

lemma even_Pls:

  even Int.Pls = True

lemma even_Min:

  even Int.Min = False

lemma even_B0:

  even (Int.Bit0 x) = True

lemma even_B1:

  even (Int.Bit1 x) = False

lemma even_number_of:

  even (number_of w) = even w

lemma compute_even:

  even Int.Pls = True
  even Int.Min = False
  even (Int.Bit0 x) = True
  even (Int.Bit1 x) = False
  even (number_of w) = even w

lemma compute_numeral:

  If True = (λx y. x)
  If False = (λx y. y)
  Let s f == f s
  fst (x, y) = x
  snd (x, y) = y
  ((a, b) = (c, d)) = (a = cb = d)
  prod_case f (x, y) = f x y
  (¬ True) = False
  (¬ False) = True
  (P ∧ True) = P
  (True ∧ P) = P
  (P ∧ False) = False
  (False ∧ P) = False
  (P ∨ True) = True
  (True ∨ P) = True
  (P ∨ False) = P
  (False ∨ P) = P
  (True --> P) = P
  (P --> True) = True
  (True --> P) = P
  (P --> False) = (¬ P)
  (False --> P) = True
  (False = False) = True
  (True = True) = True
  (False = True) = False
  (True = False) = False
  Int.Bit0 Int.Pls = Int.Pls
  Int.Bit1 Int.Min = Int.Min
  iszero Int.Pls = True
  iszero Int.Min = False
  iszero (Int.Bit0 x) = iszero x
  iszero (Int.Bit1 x) = False
  neg Int.Pls = False
  neg Int.Min = True
  neg (Int.Bit0 x) = neg x
  neg (Int.Bit1 x) = neg x
  lezero Int.Pls = True
  lezero Int.Min = True
  lezero (Int.Bit0 x) = lezero x
  lezero (Int.Bit1 x) = neg x
  (Int.Pls = Int.Pls) = True
  (Int.Min = Int.Min) = True
  (Int.Pls = Int.Min) = False
  (Int.Min = Int.Pls) = False
  (Int.Bit0 x = Int.Bit0 y) = (x = y)
  (Int.Bit1 x = Int.Bit1 y) = (x = y)
  (Int.Bit0 x = Int.Bit1 y) = False
  (Int.Bit1 x = Int.Bit0 y) = False
  (Int.Pls = Int.Bit0 x) = (Int.Pls = x)
  (Int.Pls = Int.Bit1 x) = False
  (Int.Min = Int.Bit0 x) = False
  (Int.Min = Int.Bit1 x) = (Int.Min = x)
  (Int.Bit0 x = Int.Pls) = (x = Int.Pls)
  (Int.Bit1 x = Int.Pls) = False
  (Int.Bit0 x = Int.Min) = False
  (Int.Bit1 x = Int.Min) = (x = Int.Min)
  (Int.Pls < Int.Min) = False
  (Int.Pls < Int.Pls) = False
  (Int.Min < Int.Pls) = True
  (Int.Min < Int.Min) = False
  (Int.Bit0 x < Int.Bit0 y) = (x < y)
  (Int.Bit1 x < Int.Bit1 y) = (x < y)
  (Int.Bit0 x < Int.Bit1 y) = (x  y)
  (Int.Bit1 x < Int.Bit0 y) = (x < y)
  (Int.Pls < Int.Bit0 x) = (Int.Pls < x)
  (Int.Pls < Int.Bit1 x) = (Int.Pls  x)
  (Int.Min < Int.Bit0 x) = (Int.Pls  x)
  (Int.Min < Int.Bit1 x) = (Int.Min < x)
  (Int.Bit0 x < Int.Pls) = (x < Int.Pls)
  (Int.Bit1 x < Int.Pls) = (x < Int.Pls)
  (Int.Bit0 x < Int.Min) = (x < Int.Pls)
  (Int.Bit1 x < Int.Min) = (x < Int.Min)
  (Int.Pls  Int.Min) = False
  (Int.Pls  Int.Pls) = True
  (Int.Min  Int.Pls) = True
  (Int.Min  Int.Min) = True
  (Int.Bit0 x  Int.Bit0 y) = (x  y)
  (Int.Bit1 x  Int.Bit1 y) = (x  y)
  (Int.Bit0 x  Int.Bit1 y) = (x  y)
  (Int.Bit1 x  Int.Bit0 y) = (x < y)
  (Int.Pls  Int.Bit0 x) = (Int.Pls  x)
  (Int.Pls  Int.Bit1 x) = (Int.Pls  x)
  (Int.Min  Int.Bit0 x) = (Int.Pls  x)
  (Int.Min  Int.Bit1 x) = (Int.Min  x)
  (Int.Bit0 x  Int.Pls) = (x  Int.Pls)
  (Int.Bit1 x  Int.Pls) = (x < Int.Pls)
  (Int.Bit0 x  Int.Min) = (x < Int.Pls)
  (Int.Bit1 x  Int.Min) = (x  Int.Min)
  Int.succ Int.Pls = Int.Bit1 Int.Pls
  Int.succ Int.Min = Int.Pls
  Int.succ (Int.Bit0 k) = Int.Bit1 k
  Int.succ (Int.Bit1 k) = Int.Bit0 (Int.succ k)
  Int.pred Int.Pls = Int.Min
  Int.pred Int.Min = Int.Bit0 Int.Min
  Int.pred (Int.Bit0 k) = Int.Bit1 (Int.pred k)
  Int.pred (Int.Bit1 k) = Int.Bit0 k
  - Int.Pls = Int.Pls
  - Int.Min = Int.Bit1 Int.Pls
  - Int.Bit0 k = Int.Bit0 (- k)
  - Int.Bit1 k = Int.Bit1 (Int.pred (- k))
  Int.Pls + k = k
  Int.Min + k = Int.pred k
  k + Int.Pls = k
  k + Int.Min = Int.pred k
  Int.Bit0 k + Int.Bit0 l = Int.Bit0 (k + l)
  Int.Bit0 k + Int.Bit1 l = Int.Bit1 (k + l)
  Int.Bit1 k + Int.Bit0 l = Int.Bit1 (k + l)
  Int.Bit1 k + Int.Bit1 l = Int.Bit0 (k + Int.succ l)
  Int.Pls * w = Int.Pls
  Int.Min * k = - k
  x * Int.Pls = Int.Pls
  x * Int.Min = - x
  Int.Bit0 x * y = Int.Bit0 (x * y)
  x * Int.Bit0 y = Int.Bit0 (x * y)
  Int.Bit1 x * Int.Bit1 y = Int.Bit1 (Int.Bit0 (x * y) + x + y)
  0 = Numeral0
  1 = Numeral1
  nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)
  Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))
  number_of x + number_of y =
  (if neg x then number_of y
   else if neg y then number_of x else number_of (x + y))
  number_of x - number_of y =
  (if neg x then 0
   else if neg y then number_of x else nat_norm_number_of (number_of (x + - y)))
  number_of x * number_of y =
  (if neg x then 0 else if neg y then 0 else number_of (x * y))
  (number_of x = number_of y) = (lezero x ∧ lezero yx = y)
  (number_of x < number_of y) = (x < y ∧ ¬ lezero y)
  (number_of x  number_of y) = (y < x --> lezero x)
  natfac n = (if n = 0 then 1 else n * natfac (n - 1))
  - number_of w = number_of (- w)
  number_of v + number_of w = number_of (v + w)
  number_of x - number_of y = number_of (x + - y)
  number_of v * number_of w = number_of (v * w)
  (0::'a) = Numeral0
  (1::'a) = Numeral1
  (number_of x = number_of y) = (x = y)
  (number_of x  number_of y) = (x  y)
  (number_of x < number_of y) = (x < y)
  max a b = (if a  b then b else a)
  min a b = (if a  b then a else b)
  real (number_of v) = (if neg v then 0 else number_of v)
  nat (number_of v) = number_of v
  real (number_of v) = number_of v
  z ^ number_of (Int.Bit0 w) = (let w = z ^ number_of w in w * w)
  z ^ number_of (Int.Bit1 w) =
  (if Numeral0  number_of w then let w = z ^ number_of w in z * w * w
   else Numeral1)
  z ^ Numeral0 = Numeral1
  z ^ -1 = Numeral1
  a div b = fst (divAlg (a, b))
  a mod b = snd (divAlg (a, b))
  divAlg (a, b) =
  (if 0  a
   then if 0  b then posDivAlg a b
        else if a = 0 then 0N else negateSnd (negDivAlg (- a) (- b))
   else if 0 < b then negDivAlg a b else negateSnd (posDivAlg (- a) (- b)))
  adjust b (q, r) = (if 0  r - b then (2 * q + 1, r - b) else (2 * q, r))
  negateSnd (q, r) = (q, - r)
  posDivAlg a b =
  (if a < bb  0 then (0, a) else adjust b (posDivAlg a (2 * b)))
  negDivAlg a b =
  (if 0  a + bb  0 then (-1, a + b) else adjust b (negDivAlg a (2 * b)))
  even Int.Pls = True
  even Int.Min = False
  even (Int.Bit0 x) = True
  even (Int.Bit1 x) = False
  even (number_of w) = even w