Built with Alectryon, running Coq+SerAPI v8.19.0+0.19.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
[Loading ML file number_string_notation_plugin.cmxs (using legacy method) ... done]
Require Import Spaces.Pos.Core.Local Set Universe Minimization ToSet.(** * Binary Integers *)LocalClose Scope trunc_scope.LocalClose Scope nat_scope.LocalOpen Scope positive_scope.(** ** Definition of the Integers *)(** We define an integer as being a positive number labelled negative, zero or a positive number labelled positive. *)InductiveBinInt : Type0 :=
| neg : Pos -> BinInt
| zero : BinInt
| pos : Pos -> BinInt.Arguments pos p%_pos.Declare Scope binint_scope.LocalOpen Scope binint_scope.Delimit Scope binint_scope with binint.(** The integers are a pointed type *)Instanceispointed_BinInt : IsPointed BinInt := zero.(** Properties of constructors *)Definitionneg_inj {zw : Pos} (p : neg z = neg w) : z = w
:= transport (funs => z =
(match s with neg a => a | zero => w | pos a => w end)) p (idpath z).Definitionpos_inj {zw : Pos} (p : pos z = pos w) : z = w
:= transport (funs => z =
(match s with neg a => w | zero => w | pos a => a end)) p (idpath z).Definitionneg_neq_zero {z : Pos} : ~ (neg z = zero)
:= funp => transport (funs =>
match s with neg a => z = a| zero => Empty | pos _ => Empty end) p (idpath z).Definitionpos_neq_zero {z : Pos} : ~ (pos z = zero)
:= funp => transport (funs =>
match s with pos a => z = a | zero => Empty | neg _ => Empty end) p (idpath z).Definitionneg_neq_pos {zw : Pos} : ~ (neg z = pos w)
:= funp => transport (funs =>
match s with neg a => z = a | zero => Empty | pos _ => Empty end) p (idpath z).Definitionzero_neq_neg {z : Pos} := @neg_neq_zero z o symmetry _ _.Definitionzero_neq_pos {z : Pos} := @pos_neq_zero z o symmetry _ _.Definitionpos_neq_neg {zw : Pos} := @neg_neq_pos z w o symmetry _ _.(** ** Conversion with a decimal representation for printing/parsing *)Definitionbinint_to_decimal_binint (n : BinInt) : Decimal.int :=
match n with
| neg m => Decimal.Neg (pos_to_uint m)
| zero => Decimal.Pos Decimal.Nil
| pos m => Decimal.Pos (pos_to_uint m)
end.Definitionbinint_to_number_binint (n : BinInt) : Numeral.int :=
Numeral.IntDec (binint_to_decimal_binint n).Fixpointbinint_of_decimal_uint (d : Decimal.uint) : BinInt :=
match d with
| Decimal.Nil => zero
| Decimal.D0 l => binint_of_decimal_uint l
| Decimal.D1 l => pos (pos_of_uint_acc l 1)
| Decimal.D2 l => pos (pos_of_uint_acc l 1~0)
| Decimal.D3 l => pos (pos_of_uint_acc l 1~1)
| Decimal.D4 l => pos (pos_of_uint_acc l 1~0~0)
| Decimal.D5 l => pos (pos_of_uint_acc l 1~0~1)
| Decimal.D6 l => pos (pos_of_uint_acc l 1~1~0)
| Decimal.D7 l => pos (pos_of_uint_acc l 1~1~1)
| Decimal.D8 l => pos (pos_of_uint_acc l 1~0~0~0)
| Decimal.D9 l => pos (pos_of_uint_acc l 1~0~0~1)
end.Definitionbinint_of_decimal_binint (d : Decimal.int) : BinInt :=
match d with
| Decimal.Pos u => binint_of_decimal_uint u
| Decimal.Neg u => lett := binint_of_decimal_uint u inmatch t with
| pos v => neg v
| _ => zero
endend.Definitionbinint_of_number_binint (d:Numeral.int) :=
match d with
| Numeral.IntDec d => Some (binint_of_decimal_binint d)
| Numeral.IntHex _ => None
end.Number NotationBinInt binint_of_number_binint binint_to_number_binint : binint_scope.(* For some reason 0 can be parsed as an integer, but is printed as [zero]. This notation fixes that. *)Notation"0" := zero : binint_scope.(** ** Doubling and variants *)Definitionbinint_doublex :=
match x with
| 0 => 0
| pos p => pos p~0
| neg p => neg p~0end.Definitionbinint_succ_doublex :=
match x with
| 0 => 1
| pos p => pos p~1
| neg p => neg (pos_pred_double p)
end.Definitionbinint_pred_doublex :=
match x with
| 0 => neg 1%pos
| neg p => neg p~1
| pos p => pos (pos_pred_double p)
end.(** ** Subtraction of positive into [BinInt] *)Fixpointbinint_pos_sub (xy : Pos) {structy} : BinInt :=
match x, y with
| p~1, q~1 => binint_double (binint_pos_sub p q)
| p~1, q~0 => binint_succ_double (binint_pos_sub p q)
| p~1, 1 => pos p~0
| p~0, q~1 => binint_pred_double (binint_pos_sub p q)
| p~0, q~0 => binint_double (binint_pos_sub p q)
| p~0, 1 => pos (pos_pred_double p)
| 1, q~1 => neg q~0
| 1, q~0 => neg (pos_pred_double q)
| 1, 1 => zero
end%pos.(** ** Negation *)Definitionbinint_negationx :=
match x with
| zero => zero
| pos x => neg x
| neg x => pos x
end.Notation"- x" := (binint_negation x) : binint_scope.
n: BinInt
- - n = n
n: BinInt
- - n = n
bydestruct n.Qed.(** ** Addition *)Definitionbinint_addxy :=
match x, y with
| 0, y => y
| x, 0 => x
| pos x', pos y' => pos (x' + y')
| pos x', neg y' => binint_pos_sub x' y'
| neg x', pos y' => binint_pos_sub y' x'
| neg x', neg y' => neg (x' + y')
end.Infix"+" := binint_add : binint_scope.(** ** Successor *)Definitionbinint_succx := x + 1.(** ** Predecessor *)Definitionbinint_predx := x + neg 1%pos.(** ** Subtraction *)Definitionbinint_submn := m + -n.Infix"-" := binint_sub : binint_scope.(** ** Multiplication *)Definitionbinint_mulxy :=
match x, y with
| 0, _ => 0
| _, 0 => 0
| pos x', pos y' => pos (x' * y')
| pos x', neg y' => neg (x' * y')
| neg x', pos y' => neg (x' * y')
| neg x', neg y' => pos (x' * y')
end.Infix"*" := binint_mul : binint_scope.(** ** Power function *)Definitionbinint_powxy :=
match y with
| pos p => pos_iter (binint_mul x) p 1
| 0 => 1
| neg _ => 0end.(** ** Square *)Definitionbinint_squarex :=
match x with
| 0 => 0
| pos p => pos (pos_square p)
| neg p => pos (pos_square p)
end.(** ** Sign function *)Definitionbinint_sgnz :=
match z with
| 0 => 0
| pos p => 1
| neg p => neg 1%pos
end.(* ** Decidable paths and truncation. *)
DecidablePaths BinInt
DecidablePaths BinInt
n, m: Pos
Decidable (neg n = neg m)
n: Pos
Decidable (neg n = 0)
n, m: Pos
Decidable (neg n = pos m)
m: Pos
Decidable (0 = neg m)
Decidable (0 = 0)
m: Pos
Decidable (0 = pos m)
n, m: Pos
Decidable (pos n = neg m)
n: Pos
Decidable (pos n = 0)
n, m: Pos
Decidable (pos n = pos m)
n, m: Pos
Decidable (neg n = neg m)
n, m: Pos p: n = m
Decidable (neg n = neg m)
n, m: Pos q: n <> m
Decidable (neg n = neg m)
n, m: Pos p: n = m
Decidable (neg n = neg m)
apply inl, ap, p.
n, m: Pos q: n <> m
Decidable (neg n = neg m)
byapply inr; intro; apply q, neg_inj.
n: Pos
Decidable (neg n = 0)
exact (inr neg_neq_zero).
n, m: Pos
Decidable (neg n = pos m)
exact (inr neg_neq_pos).
m: Pos
Decidable (0 = neg m)
exact (inr zero_neq_neg).
Decidable (0 = 0)
exact (inl idpath).
m: Pos
Decidable (0 = pos m)
exact (inr zero_neq_pos).
n, m: Pos
Decidable (pos n = neg m)
exact (inr pos_neq_neg).
n: Pos
Decidable (pos n = 0)
exact (inr pos_neq_zero).
n, m: Pos
Decidable (pos n = pos m)
n, m: Pos p: n = m
Decidable (pos n = pos m)
n, m: Pos q: n <> m
Decidable (pos n = pos m)
n, m: Pos p: n = m
Decidable (pos n = pos m)
apply inl, ap, p.
n, m: Pos q: n <> m
Decidable (pos n = pos m)
byapply inr; intro; apply q, pos_inj.Defined.(** Since integers have decidable paths they are a hset *)Instancehset_binint : IsHSet BinInt | 0 := _.