coq-barebone/theories/Prim/Int63.v

51 lines
1.3 KiB
Coq

Require Import Barebone.Prelude.
Require Barebone.Prim.Bool.
Register Bool.bool as kernel.ind_bool.
Require Barebone.Prim.Prod.
Register Prod.prod as kernel.ind_pair.
Variant carry (A : Type) : Type := C0 (a : A) | C1 (a : A).
Register carry as kernel.ind_carry.
Variant cmp : Set := Eq | Lt | Gt.
Register cmp as kernel.ind_cmp.
Primitive int := #int63_type.
Primitive lsl := #int63_lsl.
Primitive lsr := #int63_lsr.
Primitive land := #int63_land.
Primitive lor := #int63_lor.
Primitive lxor := #int63_lxor.
Primitive asr := #int63_asr.
Primitive add := #int63_add.
Primitive sub := #int63_sub.
Primitive mul := #int63_mul.
Primitive mulc := #int63_mulc.
Primitive div := #int63_div.
Primitive mod := #int63_mod.
Primitive divs := #int63_divs.
Primitive mods := #int63_mods.
Primitive eqb := #int63_eq.
Primitive ltb := #int63_lt.
Primitive leb := #int63_le.
Primitive ltsb := #int63_lts.
Primitive lesb := #int63_les.
Primitive compare := #int63_compare.
Primitive compares := #int63_compares.
Primitive addc := #int63_addc.
Primitive addcarryc := #int63_addcarryc.
Primitive subc := #int63_subc.
Primitive subcarryc := #int63_subcarryc.
Primitive diveucl := #int63_diveucl.
Primitive diveucl_21 := #int63_div21.
Primitive addmuldiv := #int63_addmuldiv.
Primitive head0 := #int63_head0.
Primitive tail0 := #int63_tail0.