Library Paux

Require Export Digit.
Require Export Option.
Require Export Inverse_Image.
Require Export Wf_nat.

Fixpoint exp (n m : nat) {struct m} : nat :=
  match m with
  | O => 1
  | S m' => n * exp n m'
  end.

Theorem expPlus : forall n p q : nat, exp n (p + q) = exp n p * exp n q.

Fixpoint positive_exp (p n : positive) {struct n} : positive :=
  match n with
  | xH => p
  | xO n1 =>
      match positive_exp p n1 with
      | r =>
          (fun (x : positive) (_ : positive -> positive)
             (y : positive) => (x * y)%positive) r (
            fun x => x) r
      end
  | xI n1 =>
      match positive_exp p n1 with
      | r =>
          (fun (x : positive) (_ : positive -> positive)
             (y : positive) => (x * y)%positive) p (
            fun x => x)
            ((fun (x : positive)
                (_ : positive -> positive)
                (y : positive) => (x * y)%positive) r (
               fun x => x) r)
      end
  end.

Theorem positive_exp_correct :
 forall p n : positive,
 nat_of_P (positive_exp p n) =
 exp (nat_of_P p) (nat_of_P n).

Fixpoint Pdiv (p q : positive) {struct p} :
 Option positive * Option positive :=
  match p with
  | xH =>
      match q with
      | xH => (Some _ 1%positive, None _)
      | xO r => (None _, Some _ p)
      | xI r => (None _, Some _ p)
      end
  | xI p' =>
      match Pdiv p' q with
      | (None, None) =>
          match (1 - Zpos q)%Z with
          | Z0 => (Some _ 1%positive, None _)
          | Zpos r' => (Some _ 1%positive, Some _ r')
          | Zneg r' => (None _, Some _ 1%positive)
          end
      | (None, Some r1) =>
          match (Zpos (xI r1) - Zpos q)%Z with
          | Z0 => (Some _ 1%positive, None _)
          | Zpos r' => (Some _ 1%positive, Some _ r')
          | Zneg r' => (None _, Some _ (xI r1))
          end
      | (Some q1, None) =>
          match (1 - Zpos q)%Z with
          | Z0 => (Some _ (xI q1), None _)
          | Zpos r' => (Some _ (xI q1), Some _ r')
          | Zneg r' => (Some _ (xO q1), Some _ 1%positive)
          end
      | (Some q1, Some r1) =>
          match (Zpos (xI r1) - Zpos q)%Z with
          | Z0 => (Some _ (xI q1), None _)
          | Zpos r' => (Some _ (xI q1), Some _ r')
          | Zneg r' => (Some _ (xO q1), Some _ (xI r1))
          end
      end
  | xO p' =>
      match Pdiv p' q with
      | (None, None) => (None _, None _)
      | (None, Some r1) =>
          match (Zpos (xO r1) - Zpos q)%Z with
          | Z0 => (Some _ 1%positive, None _)
          | Zpos r' => (Some _ 1%positive, Some _ r')
          | Zneg r' => (None _, Some _ (xO r1))
          end
      | (Some q1, None) => (Some _ (xO q1), None _)
      | (Some q1, Some r1) =>
          match (Zpos (xO r1) - Zpos q)%Z with
          | Z0 => (Some _ (xI q1), None _)
          | Zpos r' => (Some _ (xI q1), Some _ r')
          | Zneg r' => (Some _ (xO q1), Some _ (xO r1))
          end
      end
  end.

Definition oZ h := match h with
                   | None => 0
                   | Some p => nat_of_P p
                   end.

Theorem Pdiv_correct :
 forall p q,
 nat_of_P p =
 oZ (fst (Pdiv p q)) * nat_of_P q + oZ (snd (Pdiv p q)) /\
 oZ (snd (Pdiv p q)) < nat_of_P q.

Section bugFix.
Variable
  PdivAux :
    positive ->
    positive -> Option positive * Option positive.

Fixpoint PdivlessAux (bound p base length : positive) {struct length}
 : Option positive * Option positive * nat :=
  match (bound ?= p)%positive Datatypes.Eq with
  | Datatypes.Gt => (Some _ p, None _, 0)
  | _ =>
      match PdivAux p base with
      | (None, None) => (None _, None _, 1)
      | (None, Some r1) => (None _, Some _ r1, 1)
      | (Some q1, None) =>
          match length with
          | xH => (Some _ q1, None _, 0)
          | xO length' =>
              match PdivlessAux bound q1 base length' with
              | (s2, None, n) => (s2, None _, S n)
              | (s2, Some r2, n) =>
                  (s2,
                  Some _
                    ((fun (x : positive)
                        (_ : positive -> positive)
                        (y : positive) => (x * y)%positive) r2
                       (fun x => x) base), S n)
              end
          | xI length' =>
              match PdivlessAux bound q1 base length' with
              | (s2, None, n) => (s2, None _, S n)
              | (s2, Some r2, n) =>
                  (s2,
                  Some _
                    ((fun (x : positive)
                        (_ : positive -> positive)
                        (y : positive) => (x * y)%positive) r2
                       (fun x => x) base), S n)
              end
          end
      | (Some q1, Some r1) =>
          match length with
          | xH => (Some _ q1, None _, 0)
          | xO length' =>
              match PdivlessAux bound q1 base length' with
              | (s2, None, n) => (s2, Some _ r1, S n)
              | (s2, Some r2, n) =>
                  (s2,
                  Some _
                    ((fun (x : positive)
                        (_ : positive -> positive)
                        (y : positive) => x * y) r2 (
                       fun x => x) base + r1)%positive,
                  S n)
              end
          | xI length' =>
              match PdivlessAux bound q1 base length' with
              | (s2, None, n) => (s2, Some _ r1, S n)
              | (s2, Some r2, n) =>
                  (s2,
                  Some _
                    ((fun (x : positive)
                        (_ : positive -> positive)
                        (y : positive) => x * y) r2 (
                       fun x => x) base + r1)%positive,
                  S n)
              end
          end
      end
  end.
End bugFix.

Definition Pdivless := PdivlessAux Pdiv.

Theorem Pdivless1 :
 forall bound p q base,
 (bound ?= p)%positive Datatypes.Eq = Datatypes.Gt ->
 Pdivless bound p base q = (Some _ p, None _, 0).

Theorem Pdivless2 :
 forall bound p length base,
 (bound ?= p)%positive Datatypes.Eq <> Datatypes.Gt ->
 Pdivless bound p base length =
 match Pdiv p base with
 | (None, None) => (None _, None _, 1)
 | (None, Some r1) => (None _, Some _ r1, 1)
 | (Some q1, None) =>
     match length with
     | xH => (Some _ q1, None _, 0)
     | xO length' =>
         match Pdivless bound q1 base length' with
         | (s2, None, n) => (s2, None _, S n)
         | (s2, Some r2, n) =>
             (s2,
             Some _
               ((fun (x : positive)
                   (_ : positive -> positive)
                   (y : positive) => (x * y)%positive) r2 (
                  fun x => x) base), S n)
         end
     | xI length' =>
         match Pdivless bound q1 base length' with
         | (s2, None, n) => (s2, None _, S n)
         | (s2, Some r2, n) =>
             (s2,
             Some _
               ((fun (x : positive)
                   (_ : positive -> positive)
                   (y : positive) => (x * y)%positive) r2 (
                  fun x => x) base), S n)
         end
     end
 | (Some q1, Some r1) =>
     match length with
     | xH => (Some _ q1, None _, 0)
     | xO length' =>
         match Pdivless bound q1 base length' with
         | (s2, None, n) => (s2, Some _ r1, S n)
         | (s2, Some r2, n) =>
             (s2,
             Some _
               ((fun (x : positive)
                   (_ : positive -> positive)
                   (y : positive) => x * y) r2 (
                  fun x => x) base + r1)%positive,
             S n)
         end
     | xI length' =>
         match Pdivless bound q1 base length' with
         | (s2, None, n) => (s2, Some _ r1, S n)
         | (s2, Some r2, n) =>
             (s2,
             Some _
               ((fun (x : positive)
                   (_ : positive -> positive)
                   (y : positive) => x * y) r2 (
                  fun x => x) base + r1)%positive,
             S n)
         end
     end
 end.

Theorem compare_SUP_dec :
 forall p q : positive,
 (p ?= q)%positive Datatypes.Eq = Datatypes.Gt \/
 (p ?= q)%positive Datatypes.Eq <> Datatypes.Gt.
Hint Resolve lt_O_nat_of_P: arith.

Theorem odd_even_lem : forall p q, 2 * p + 1 <> 2 * q.

Theorem Pdivless_correct :
 forall bound p q base,
 1 < nat_of_P base ->
 nat_of_P p <= nat_of_P q ->
 nat_of_P p =
 oZ (fst (fst (Pdivless bound p base q))) *
 exp (nat_of_P base) (snd (Pdivless bound p base q)) +
 oZ (snd (fst (Pdivless bound p base q))) /\
 (oZ (fst (fst (Pdivless bound p base q))) < nat_of_P bound /\
  oZ (snd (fst (Pdivless bound p base q))) <
  exp (nat_of_P base) (snd (Pdivless bound p base q))) /\
 (forall bound',
  nat_of_P bound = nat_of_P base * bound' ->
  nat_of_P bound <= nat_of_P p ->
  nat_of_P bound <=
  nat_of_P base * oZ (fst (fst (Pdivless bound p base q)))).

Definition PdivBound bound p base := Pdivless bound p base p.

Theorem PdivBound_correct :
 forall bound p base,
 1 < nat_of_P base ->
 nat_of_P p =
 oZ (fst (fst (PdivBound bound p base))) *
 exp (nat_of_P base) (snd (PdivBound bound p base)) +
 oZ (snd (fst (PdivBound bound p base))) /\
 (oZ (fst (fst (PdivBound bound p base))) < nat_of_P bound /\
  oZ (snd (fst (PdivBound bound p base))) <
  exp (nat_of_P base) (snd (PdivBound bound p base))) /\
 (forall bound',
  nat_of_P bound = nat_of_P base * bound' ->
  nat_of_P bound <= nat_of_P p ->
  nat_of_P bound <=
  nat_of_P base * oZ (fst (fst (PdivBound bound p base)))).

Theorem PdivBound_correct1 :
 forall bound p base,
 1 < nat_of_P base ->
 nat_of_P p =
 oZ (fst (fst (PdivBound bound p base))) *
 exp (nat_of_P base) (snd (PdivBound bound p base)) +
 oZ (snd (fst (PdivBound bound p base))).

Theorem PdivBound_correct2 :
 forall bound p base,
 1 < nat_of_P base ->
 oZ (fst (fst (PdivBound bound p base))) < nat_of_P bound.

Theorem PdivBound_correct3 :
 forall bound p base,
 nat_of_P p < nat_of_P bound ->
 PdivBound bound p base = (Some _ p, None _, 0).

Theorem PdivBound_correct4 :
 forall bound p base bound',
 1 < nat_of_P base ->
 nat_of_P bound = nat_of_P base * bound' ->
 nat_of_P bound <= nat_of_P p ->
 nat_of_P bound <=
 nat_of_P base * oZ (fst (fst (PdivBound bound p base))).
Transparent Pdiv.