checkpoint: new numeral encoding
This commit is contained in:
parent
f6d22c0002
commit
8ee214f133
24 changed files with 430 additions and 242 deletions
|
@ -3,7 +3,6 @@ Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Leonardo de Moura
|
||||
-/
|
||||
|
||||
prelude
|
||||
import init.datatypes
|
||||
definition std.priority.default : num := 1000
|
||||
|
|
|
@ -16,15 +16,6 @@ namespace algebra
|
|||
|
||||
variable {A : Type}
|
||||
|
||||
structure has_one [class] (A : Type) :=
|
||||
(one : A)
|
||||
|
||||
structure has_zero [class] (A : Type) :=
|
||||
(zero : A)
|
||||
|
||||
notation 1 := !has_one.one
|
||||
notation 0 := !has_zero.zero
|
||||
|
||||
/- semigroup -/
|
||||
|
||||
structure semigroup [class] (A : Type) extends has_mul A :=
|
||||
|
|
|
@ -76,14 +76,14 @@ namespace vector
|
|||
|
||||
definition tabulate : Π {n : nat}, (fin n → A) → vector A n
|
||||
| 0 f := []
|
||||
| (n+1) f := f (@zero n) :: tabulate (λ i : fin n, f (succ i))
|
||||
| (n+1) f := f (fin.zero n) :: tabulate (λ i : fin n, f (succ i))
|
||||
|
||||
theorem nth_tabulate : ∀ {n : nat} (f : fin n → A) (i : fin n), nth (tabulate f) i = f i
|
||||
| 0 f i := elim0 i
|
||||
| (n+1) f (mk 0 h) := by reflexivity
|
||||
| (n+1) f (mk (succ i) h) :=
|
||||
begin
|
||||
change nth (f (@zero n) :: tabulate (λ i : fin n, f (succ i))) (mk (succ i) h) = f (mk (succ i) h),
|
||||
change nth (f (fin.zero n) :: tabulate (λ i : fin n, f (succ i))) (mk (succ i) h) = f (mk (succ i) h),
|
||||
rewrite nth_succ,
|
||||
rewrite nth_tabulate
|
||||
end
|
||||
|
|
|
@ -86,9 +86,12 @@ assume Pex, absurd Pmltn (not_lt_of_ge
|
|||
|
||||
end pigeonhole
|
||||
|
||||
definition zero (n : nat) : fin (succ n) :=
|
||||
protected definition zero (n : nat) : fin (succ n) :=
|
||||
mk 0 !zero_lt_succ
|
||||
|
||||
definition fin_has_zero [instance] [reducible] (n : nat) : has_zero (fin (succ n)) :=
|
||||
has_zero.mk (fin.zero n)
|
||||
|
||||
definition mk_mod [reducible] (n i : nat) : fin (succ n) :=
|
||||
mk (i mod (succ n)) (mod_lt _ !zero_lt_succ)
|
||||
|
||||
|
@ -112,7 +115,7 @@ mk n !lt_succ_self
|
|||
theorem val_lift : ∀ (i : fin n) (m : nat), val i = val (lift i m)
|
||||
| (mk v h) m := rfl
|
||||
|
||||
lemma mk_succ_ne_zero {i : nat} : ∀ {P}, mk (succ i) P ≠ zero n :=
|
||||
lemma mk_succ_ne_zero {i : nat} : ∀ {P}, mk (succ i) P ≠ fin.zero n :=
|
||||
assume P Pe, absurd (veq_of_eq Pe) !succ_ne_zero
|
||||
|
||||
lemma mk_mod_eq {i : fin (succ n)} : i = mk_mod n i :=
|
||||
|
@ -123,7 +126,7 @@ begin esimp [mk_mod], congruence, exact mod_eq_of_lt Plt end
|
|||
|
||||
section lift_lower
|
||||
|
||||
lemma lift_zero : lift_succ (zero n) = zero (succ n) := rfl
|
||||
lemma lift_zero : lift_succ (fin.zero n) = fin.zero (succ n) := rfl
|
||||
|
||||
lemma ne_max_of_lt_max {i : fin (succ n)} : i < n → i ≠ maxi :=
|
||||
by intro hlt he; substvars; exact absurd hlt (lt.irrefl n)
|
||||
|
@ -235,23 +238,23 @@ lemma val_mod : ∀ i : fin (succ n), (val i) mod (succ n) = val i
|
|||
lemma madd_comm (i j : fin (succ n)) : madd i j = madd j i :=
|
||||
by apply eq_of_veq; rewrite [*val_madd, add.comm (val i)]
|
||||
|
||||
lemma zero_madd (i : fin (succ n)) : madd (zero n) i = i :=
|
||||
by apply eq_of_veq; rewrite [val_madd, ↑zero, nat.zero_add, mod_eq_of_lt (is_lt i)]
|
||||
lemma zero_madd (i : fin (succ n)) : madd (fin.zero n) i = i :=
|
||||
by apply eq_of_veq; rewrite [val_madd, ↑fin.zero, nat.zero_add, mod_eq_of_lt (is_lt i)]
|
||||
|
||||
lemma madd_zero (i : fin (succ n)) : madd i (zero n) = i :=
|
||||
lemma madd_zero (i : fin (succ n)) : madd i (fin.zero n) = i :=
|
||||
!madd_comm ▸ zero_madd i
|
||||
|
||||
lemma madd_assoc (i j k : fin (succ n)) : madd (madd i j) k = madd i (madd j k) :=
|
||||
by apply eq_of_veq; rewrite [*val_madd, mod_add_mod, add_mod_mod, add.assoc (val i)]
|
||||
|
||||
lemma madd_left_inv : ∀ i : fin (succ n), madd (minv i) i = zero n
|
||||
lemma madd_left_inv : ∀ i : fin (succ n), madd (minv i) i = fin.zero n
|
||||
| (mk iv ilt) := eq_of_veq (by
|
||||
rewrite [val_madd, ↑minv, ↑zero, mod_add_mod, sub_add_cancel (le_of_lt ilt), mod_self])
|
||||
rewrite [val_madd, ↑minv, ↑fin.zero, mod_add_mod, sub_add_cancel (le_of_lt ilt), mod_self])
|
||||
|
||||
open algebra
|
||||
|
||||
definition madd_is_comm_group [instance] : add_comm_group (fin (succ n)) :=
|
||||
add_comm_group.mk madd madd_assoc (zero n) zero_madd madd_zero minv madd_left_inv madd_comm
|
||||
add_comm_group.mk madd madd_assoc (fin.zero n) zero_madd madd_zero minv madd_left_inv madd_comm
|
||||
|
||||
end madd
|
||||
|
||||
|
@ -261,7 +264,7 @@ definition pred : fin n → fin n
|
|||
lemma val_pred : ∀ (i : fin n), val (pred i) = nat.pred (val i)
|
||||
| (mk v h) := rfl
|
||||
|
||||
lemma pred_zero : pred (zero n) = zero n :=
|
||||
lemma pred_zero : pred (fin.zero n) = fin.zero n :=
|
||||
rfl
|
||||
|
||||
definition mk_pred (i : nat) (h : succ i < succ n) : fin n :=
|
||||
|
@ -282,7 +285,7 @@ definition elim0 {C : fin 0 → Type} : Π i : fin 0, C i
|
|||
| (mk v h) := absurd h !not_lt_zero
|
||||
|
||||
definition zero_succ_cases {C : fin (nat.succ n) → Type} :
|
||||
C (zero n) → (Π j : fin n, C (succ j)) → (Π k : fin (nat.succ n), C k) :=
|
||||
C (fin.zero n) → (Π j : fin n, C (succ j)) → (Π k : fin (nat.succ n), C k) :=
|
||||
begin
|
||||
intros CO CS k,
|
||||
induction k with [vk, pk],
|
||||
|
@ -299,7 +302,7 @@ begin
|
|||
{ show C (mk vk pk), from
|
||||
have vk = 0, from
|
||||
eq_zero_of_le_zero (le_of_not_gt HF),
|
||||
have zero n = mk vk pk, from
|
||||
have fin.zero n = mk vk pk, from
|
||||
val_inj (eq.symm this),
|
||||
eq.rec_on this CO }
|
||||
end
|
||||
|
@ -323,7 +326,7 @@ begin
|
|||
end
|
||||
|
||||
definition foldr {A B : Type} (m : A → B → B) (b : B) : ∀ {n : nat}, (fin n → A) → B :=
|
||||
nat.rec (λ f, b) (λ n IH f, m (f (zero n)) (IH (λ i : fin n, f (succ i))))
|
||||
nat.rec (λ f, b) (λ n IH f, m (f (fin.zero n)) (IH (λ i : fin n, f (succ i))))
|
||||
|
||||
definition foldl {A B : Type} (m : B → A → B) (b : B) : ∀ {n : nat}, (fin n → A) → B :=
|
||||
nat.rec (λ f, b) (λ n IH f, m (IH (λ i : fin n, f (lift_succ i))) (f maxi))
|
||||
|
@ -337,7 +340,7 @@ begin
|
|||
apply nonempty.intro,
|
||||
exact elim0 },
|
||||
{ intros C H,
|
||||
fapply nonempty.elim (H (zero n)),
|
||||
fapply nonempty.elim (H (fin.zero n)),
|
||||
intro CO,
|
||||
fapply nonempty.elim (IH (λ i, C (succ i)) (λ i, H (succ i))),
|
||||
intro CS,
|
||||
|
@ -364,10 +367,10 @@ begin
|
|||
apply dmap_map_lift, apply @list.lt_of_mem_upto
|
||||
end
|
||||
|
||||
definition upto_step : ∀ {n : nat}, fin.upto (n +1) = (map succ (upto n))++[zero n]
|
||||
definition upto_step : ∀ {n : nat}, fin.upto (n +1) = (map succ (upto n))++[fin.zero n]
|
||||
| 0 := rfl
|
||||
| (i +1) := begin rewrite [upto_succ i, map_cons, append_cons, succ_max, upto_succ, -lift_zero],
|
||||
congruence, rewrite [map_map, -lift_succ.comm, -map_map, -(map_singleton _ (zero i)), -map_append, -upto_step] end
|
||||
congruence, rewrite [map_map, -lift_succ.comm, -map_map, -(map_singleton _ (fin.zero i)), -map_append, -upto_step] end
|
||||
end
|
||||
|
||||
open sum equiv decidable
|
||||
|
|
|
@ -227,8 +227,8 @@ namespace tuple
|
|||
lemma ith_zero {n : nat} (a : A) (v : tuple A n) (h : 0 < succ n) : ith (a::v) (mk 0 h) = a :=
|
||||
by induction v; reflexivity
|
||||
|
||||
lemma ith_fin_zero {n : nat} (a : A) (v : tuple A n) : ith (a::v) (zero n) = a :=
|
||||
by unfold zero; apply ith_zero
|
||||
lemma ith_fin_zero {n : nat} (a : A) (v : tuple A n) : ith (a::v) (fin.zero n) = a :=
|
||||
by unfold fin.zero; apply ith_zero
|
||||
|
||||
lemma ith_succ {n : nat} (a : A) (v : tuple A n) (i : nat) (h : succ i < succ n)
|
||||
: ith (a::v) (mk (succ i) h) = ith v (mk_pred i h) :=
|
||||
|
@ -238,7 +238,7 @@ namespace tuple
|
|||
: ith (a::v) (succ i) = ith v i :=
|
||||
begin cases i, unfold fin.succ, rewrite ith_succ end
|
||||
|
||||
lemma ith_zero_eq_head {n : nat} (v : tuple A (nat.succ n)) : ith v (zero n) = head v :=
|
||||
lemma ith_zero_eq_head {n : nat} (v : tuple A (nat.succ n)) : ith v (fin.zero n) = head v :=
|
||||
by rewrite [-eta v, ith_fin_zero, head_cons]
|
||||
|
||||
lemma ith_succ_eq_ith_tail {n : nat} (v : tuple A (nat.succ n)) (i : fin n) : ith v (succ i) = ith (tail v) i :=
|
||||
|
@ -259,7 +259,7 @@ namespace tuple
|
|||
|
||||
definition tabulate : Π {n : nat}, (fin n → A) → tuple A n
|
||||
| 0 f := nil
|
||||
| (n+1) f := f (@zero n) :: tabulate (λ i : fin n, f (succ i))
|
||||
| (n+1) f := f (fin.zero n) :: tabulate (λ i : fin n, f (succ i))
|
||||
|
||||
theorem ith_tabulate {n : nat} (f : fin n → A) (i : fin n) : ith (tabulate f) i = f i :=
|
||||
begin
|
||||
|
|
|
@ -12,38 +12,12 @@ definition pos_num.is_inhabited [instance] : inhabited pos_num :=
|
|||
inhabited.mk pos_num.one
|
||||
|
||||
namespace pos_num
|
||||
definition is_one (a : pos_num) : bool :=
|
||||
pos_num.rec_on a tt (λn r, ff) (λn r, ff)
|
||||
|
||||
definition pred (a : pos_num) : pos_num :=
|
||||
pos_num.rec_on a one (λn r, bit0 n) (λn r, cond (is_one n) one (bit1 r))
|
||||
|
||||
definition size (a : pos_num) : pos_num :=
|
||||
pos_num.rec_on a one (λn r, succ r) (λn r, succ r)
|
||||
|
||||
definition add (a b : pos_num) : pos_num :=
|
||||
pos_num.rec_on a
|
||||
succ
|
||||
(λn f b, pos_num.rec_on b
|
||||
(succ (bit1 n))
|
||||
(λm r, succ (bit1 (f m)))
|
||||
(λm r, bit1 (f m)))
|
||||
(λn f b, pos_num.rec_on b
|
||||
(bit1 n)
|
||||
(λm r, bit1 (f m))
|
||||
(λm r, bit0 (f m)))
|
||||
b
|
||||
|
||||
notation a + b := add a b
|
||||
|
||||
definition mul (a b : pos_num) : pos_num :=
|
||||
protected definition mul (a b : pos_num) : pos_num :=
|
||||
pos_num.rec_on a
|
||||
b
|
||||
(λn r, bit0 r + b)
|
||||
(λn r, bit0 r)
|
||||
|
||||
notation a * b := mul a b
|
||||
|
||||
definition lt (a b : pos_num) : bool :=
|
||||
pos_num.rec_on a
|
||||
(λ b, pos_num.cases_on b
|
||||
|
@ -61,9 +35,12 @@ namespace pos_num
|
|||
b
|
||||
|
||||
definition le (a b : pos_num) : bool :=
|
||||
lt a (succ b)
|
||||
pos_num.lt a (succ b)
|
||||
end pos_num
|
||||
|
||||
definition pos_num_has_mul [instance] [reducible] : has_mul pos_num :=
|
||||
has_mul.mk pos_num.mul
|
||||
|
||||
definition num.is_inhabited [instance] : inhabited num :=
|
||||
inhabited.mk num.zero
|
||||
|
||||
|
@ -76,16 +53,15 @@ namespace num
|
|||
definition size (a : num) : num :=
|
||||
num.rec_on a (pos one) (λp, pos (size p))
|
||||
|
||||
definition add (a b : num) : num :=
|
||||
num.rec_on a b (λpa, num.rec_on b (pos pa) (λpb, pos (pos_num.add pa pb)))
|
||||
|
||||
definition mul (a b : num) : num :=
|
||||
protected definition mul (a b : num) : num :=
|
||||
num.rec_on a zero (λpa, num.rec_on b zero (λpb, pos (pos_num.mul pa pb)))
|
||||
end num
|
||||
|
||||
notation a + b := add a b
|
||||
notation a * b := mul a b
|
||||
definition num_has_mul [instance] [reducible] : has_mul num :=
|
||||
has_mul.mk num.mul
|
||||
|
||||
definition le (a b : num) : bool :=
|
||||
namespace num
|
||||
protected definition le (a b : num) : bool :=
|
||||
num.rec_on a tt (λpa, num.rec_on b ff (λpb, pos_num.le pa pb))
|
||||
|
||||
private definition psub (a b : pos_num) : num :=
|
||||
|
@ -107,13 +83,13 @@ namespace num
|
|||
(λm, 2 * f m)))
|
||||
b
|
||||
|
||||
definition sub (a b : num) : num :=
|
||||
protected definition sub (a b : num) : num :=
|
||||
num.rec_on a zero (λpa, num.rec_on b a (λpb, psub pa pb))
|
||||
|
||||
notation a ≤ b := le a b
|
||||
notation a - b := sub a b
|
||||
end num
|
||||
|
||||
definition num_has_sub [instance] [reducible] : has_sub num :=
|
||||
has_sub.mk num.sub
|
||||
|
||||
-- the coercion from num to nat is defined here,
|
||||
-- so that it can already be used in init.tactic
|
||||
namespace nat
|
||||
|
@ -122,7 +98,14 @@ namespace nat
|
|||
protected definition add (a b : nat) : nat :=
|
||||
nat.rec_on b a (λ b₁ r, succ r)
|
||||
|
||||
definition nat_has_add [reducible] [instance] [priority nat.prio] : has_add nat := has_add.mk nat.add
|
||||
definition nat_has_zero [reducible] [instance] : has_zero nat :=
|
||||
has_zero.mk nat.zero
|
||||
|
||||
definition nat_has_one [reducible] [instance] : has_one nat :=
|
||||
has_one.mk (nat.succ (nat.zero))
|
||||
|
||||
definition nat_has_add [reducible] [instance] [priority nat.prio] : has_add nat :=
|
||||
has_add.mk nat.add
|
||||
|
||||
definition of_num [coercion] (n : num) : nat :=
|
||||
num.rec zero
|
||||
|
|
|
@ -4,6 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Leonardo de Moura
|
||||
-/
|
||||
prelude
|
||||
import init.datatypes
|
||||
import init.reserved_notation
|
||||
definition std.priority.default : num := 1000
|
||||
definition std.priority.max : num := 4294967295
|
||||
|
|
|
@ -9,6 +9,85 @@ import init.datatypes
|
|||
notation `assume` binders `,` r:(scoped f, f) := r
|
||||
notation `take` binders `,` r:(scoped f, f) := r
|
||||
|
||||
structure has_zero [class] (A : Type) := (zero : A)
|
||||
structure has_one [class] (A : Type) := (one : A)
|
||||
structure has_add [class] (A : Type) := (add : A → A → A)
|
||||
structure has_mul [class] (A : Type) := (mul : A → A → A)
|
||||
structure has_inv [class] (A : Type) := (inv : A → A)
|
||||
structure has_neg [class] (A : Type) := (neg : A → A)
|
||||
structure has_sub [class] (A : Type) := (sub : A → A → A)
|
||||
structure has_division [class] (A : Type) := (division : A → A → A)
|
||||
structure has_divide [class] (A : Type) := (divide : A → A → A)
|
||||
structure has_modulo [class] (A : Type) := (modulo : A → A → A)
|
||||
structure has_dvd [class] (A : Type) := (dvd : A → A → Prop)
|
||||
structure has_le [class] (A : Type) := (le : A → A → Prop)
|
||||
structure has_lt [class] (A : Type) := (lt : A → A → Prop)
|
||||
|
||||
definition zero {A : Type} [s : has_zero A] : A := has_zero.zero A
|
||||
definition one {A : Type} [s : has_one A] : A := has_one.one A
|
||||
definition add {A : Type} [s : has_add A] : A → A → A := has_add.add
|
||||
definition mul {A : Type} [s : has_mul A] : A → A → A := has_mul.mul
|
||||
definition sub {A : Type} [s : has_sub A] : A → A → A := has_sub.sub
|
||||
definition division {A : Type} [s : has_division A] : A → A → A := has_division.division
|
||||
definition divide {A : Type} [s : has_divide A] : A → A → A := has_divide.divide
|
||||
definition modulo {A : Type} [s : has_modulo A] : A → A → A := has_modulo.modulo
|
||||
definition dvd {A : Type} [s : has_dvd A] : A → A → Prop := has_dvd.dvd
|
||||
definition neg {A : Type} [s : has_neg A] : A → A := has_neg.neg
|
||||
definition inv {A : Type} [s : has_inv A] : A → A := has_inv.inv
|
||||
definition le {A : Type} [s : has_le A] : A → A → Prop := has_le.le
|
||||
definition lt {A : Type} [s : has_lt A] : A → A → Prop := has_lt.lt
|
||||
definition ge [reducible] {A : Type} [s : has_le A] (a b : A) : Prop := le b a
|
||||
definition gt [reducible] {A : Type} [s : has_lt A] (a b : A) : Prop := lt b a
|
||||
definition bit0 {A : Type} [s : has_add A] (a : A) : A := add a a
|
||||
definition bit1 {A : Type} [s₁ : has_one A] [s₂ : has_add A] (a : A) : A := add (bit0 a) one
|
||||
|
||||
definition num_has_zero [reducible] [instance] : has_zero num :=
|
||||
has_zero.mk num.zero
|
||||
|
||||
definition num_has_one [reducible] [instance] : has_one num :=
|
||||
has_one.mk (num.pos pos_num.one)
|
||||
|
||||
definition pos_num_has_one [reducible] [instance] : has_one pos_num :=
|
||||
has_one.mk (pos_num.one)
|
||||
|
||||
namespace pos_num
|
||||
open bool
|
||||
definition is_one (a : pos_num) : bool :=
|
||||
pos_num.rec_on a tt (λn r, ff) (λn r, ff)
|
||||
|
||||
definition pred (a : pos_num) : pos_num :=
|
||||
pos_num.rec_on a one (λn r, bit0 n) (λn r, bool.rec_on (is_one n) one (bit1 r))
|
||||
|
||||
definition size (a : pos_num) : pos_num :=
|
||||
pos_num.rec_on a one (λn r, succ r) (λn r, succ r)
|
||||
|
||||
definition add (a b : pos_num) : pos_num :=
|
||||
pos_num.rec_on a
|
||||
succ
|
||||
(λn f b, pos_num.rec_on b
|
||||
(succ (bit1 n))
|
||||
(λm r, succ (bit1 (f m)))
|
||||
(λm r, bit1 (f m)))
|
||||
(λn f b, pos_num.rec_on b
|
||||
(bit1 n)
|
||||
(λm r, bit1 (f m))
|
||||
(λm r, bit0 (f m)))
|
||||
b
|
||||
end pos_num
|
||||
|
||||
definition pos_num_has_add [reducible] [instance] : has_add pos_num :=
|
||||
has_add.mk pos_num.add
|
||||
|
||||
namespace num
|
||||
open pos_num
|
||||
|
||||
definition add (a b : num) : num :=
|
||||
num.rec_on a b (λpa, num.rec_on b (pos pa) (λpb, pos (pos_num.add pa pb)))
|
||||
end num
|
||||
|
||||
definition num_has_add [reducible] [instance] : has_add num :=
|
||||
has_add.mk num.add
|
||||
|
||||
/-
|
||||
Global declarations of right binding strength
|
||||
|
||||
|
@ -17,7 +96,6 @@ notation `take` binders `,` r:(scoped f, f) := r
|
|||
|
||||
When hovering over a symbol, use "C-c C-k" to see how to input it.
|
||||
-/
|
||||
|
||||
definition std.prec.max : num := 1024 -- the strength of application, identifiers, (, [, etc.
|
||||
definition std.prec.arrow : num := 25
|
||||
|
||||
|
@ -96,32 +174,6 @@ reserve infix ` ∣ `:50
|
|||
reserve infixl ` ++ `:65
|
||||
reserve infixr ` :: `:65
|
||||
|
||||
structure has_add [class] (A : Type) := (add : A → A → A)
|
||||
structure has_mul [class] (A : Type) := (mul : A → A → A)
|
||||
structure has_inv [class] (A : Type) := (inv : A → A)
|
||||
structure has_neg [class] (A : Type) := (neg : A → A)
|
||||
structure has_sub [class] (A : Type) := (sub : A → A → A)
|
||||
structure has_division [class] (A : Type) := (division : A → A → A)
|
||||
structure has_divide [class] (A : Type) := (divide : A → A → A)
|
||||
structure has_modulo [class] (A : Type) := (modulo : A → A → A)
|
||||
structure has_dvd [class] (A : Type) := (dvd : A → A → Prop)
|
||||
structure has_le [class] (A : Type) := (le : A → A → Prop)
|
||||
structure has_lt [class] (A : Type) := (lt : A → A → Prop)
|
||||
|
||||
definition add {A : Type} [s : has_add A] : A → A → A := has_add.add
|
||||
definition mul {A : Type} [s : has_mul A] : A → A → A := has_mul.mul
|
||||
definition sub {A : Type} [s : has_sub A] : A → A → A := has_sub.sub
|
||||
definition division {A : Type} [s : has_division A] : A → A → A := has_division.division
|
||||
definition divide {A : Type} [s : has_divide A] : A → A → A := has_divide.divide
|
||||
definition modulo {A : Type} [s : has_modulo A] : A → A → A := has_modulo.modulo
|
||||
definition dvd {A : Type} [s : has_dvd A] : A → A → Prop := has_dvd.dvd
|
||||
definition neg {A : Type} [s : has_neg A] : A → A := has_neg.neg
|
||||
definition inv {A : Type} [s : has_inv A] : A → A := has_inv.inv
|
||||
definition le {A : Type} [s : has_le A] : A → A → Prop := has_le.le
|
||||
definition lt {A : Type} [s : has_lt A] : A → A → Prop := has_lt.lt
|
||||
definition ge [reducible] {A : Type} [s : has_le A] (a b : A) : Prop := le b a
|
||||
definition gt [reducible] {A : Type} [s : has_lt A] (a b : A) : Prop := lt b a
|
||||
|
||||
infix + := add
|
||||
infix * := mul
|
||||
infix - := sub
|
||||
|
|
|
@ -10,4 +10,4 @@ init_module.cpp elaborator_context.cpp calc_proof_elaborator.cpp
|
|||
parse_tactic_location.cpp parse_rewrite_tactic.cpp builtin_tactics.cpp
|
||||
type_util.cpp elaborator_exception.cpp migrate_cmd.cpp local_ref_info.cpp
|
||||
obtain_expr.cpp decl_attributes.cpp nested_declaration.cpp
|
||||
parse_with_options_tactic.cpp parse_simp_tactic.cpp opt_cmd.cpp)
|
||||
parse_with_options_tactic.cpp parse_simp_tactic.cpp opt_cmd.cpp prenum.cpp)
|
||||
|
|
|
@ -57,6 +57,7 @@ Author: Leonardo de Moura
|
|||
#include "frontends/lean/calc.h"
|
||||
#include "frontends/lean/decl_cmds.h"
|
||||
#include "frontends/lean/opt_cmd.h"
|
||||
#include "frontends/lean/prenum.h"
|
||||
|
||||
namespace lean {
|
||||
type_checker_ptr mk_coercion_from_type_checker(environment const & env, name_generator && ngen) {
|
||||
|
@ -1579,8 +1580,53 @@ expr elaborator::visit_obtain_expr(expr const & e, constraint_seq & cs) {
|
|||
return process_obtain_expr(to_list(s), to_list(from), decls_goal, true, cs, e);
|
||||
}
|
||||
|
||||
expr elaborator::visit_prenum(expr const & e, constraint_seq & cs) {
|
||||
lean_assert(is_prenum(e));
|
||||
mpz const & v = prenum_value(e);
|
||||
tag e_tag = e.get_tag();
|
||||
levels ls = levels(mk_meta_univ(m_ngen.next()));
|
||||
expr A = m_full_context.mk_meta(m_ngen, none_expr(), e_tag);
|
||||
bool is_strict = true;
|
||||
bool inst_imp = true;
|
||||
if (v.is_neg())
|
||||
throw_elaborator_exception("invalid pre-numeral, it must be a non-negative value", e);
|
||||
if (v.is_zero()) {
|
||||
expr has_zero_A = mk_app(mk_constant(get_has_zero_name(), ls), A, e_tag);
|
||||
expr S = mk_placeholder_meta(optional<name>(), some_expr(has_zero_A),
|
||||
e_tag, is_strict, inst_imp, cs);
|
||||
return mk_app(mk_app(mk_constant(get_zero_name(), ls), A, e_tag), S, e_tag);
|
||||
} else {
|
||||
expr has_one_A = mk_app(mk_constant(get_has_one_name(), ls), A, e_tag);
|
||||
expr S_one = mk_placeholder_meta(optional<name>(), some_expr(has_one_A),
|
||||
e_tag, is_strict, inst_imp, cs);
|
||||
expr one = mk_app(mk_app(mk_constant(get_one_name(), ls), A, e_tag), S_one, e_tag);
|
||||
if (v == 1) {
|
||||
return one;
|
||||
} else {
|
||||
expr has_add_A = mk_app(mk_constant(get_has_add_name(), ls), A, e_tag);
|
||||
expr S_add = mk_placeholder_meta(optional<name>(), some_expr(has_add_A),
|
||||
e_tag, is_strict, inst_imp, cs);
|
||||
std::function<expr(mpz const & v)> convert = [&](mpz const & v) {
|
||||
lean_assert(v > 0);
|
||||
if (v == 1)
|
||||
return one;
|
||||
else if (v % mpz(2) == 0) {
|
||||
expr r = convert(v / 2);
|
||||
return mk_app(mk_app(mk_app(mk_constant(get_bit0_name(), ls), A, e_tag), S_add, e_tag), r, e_tag);
|
||||
} else {
|
||||
expr r = convert(v / 2);
|
||||
return mk_app(mk_app(mk_app(mk_app(mk_constant(get_bit1_name(), ls), A, e_tag), S_one, e_tag), S_add, e_tag), r, e_tag);
|
||||
}
|
||||
};
|
||||
return convert(v);
|
||||
}
|
||||
}
|
||||
}
|
||||
|
||||
expr elaborator::visit_core(expr const & e, constraint_seq & cs) {
|
||||
if (is_placeholder(e)) {
|
||||
if (is_prenum(e)) {
|
||||
return visit_prenum(e, cs);
|
||||
} else if (is_placeholder(e)) {
|
||||
return visit_placeholder(e, cs);
|
||||
} else if (is_choice(e)) {
|
||||
return visit_choice(e, none_expr(), cs);
|
||||
|
|
|
@ -188,6 +188,8 @@ class elaborator : public coercion_info_manager {
|
|||
expr const & goal, bool first, constraint_seq & cs, expr const & src);
|
||||
expr visit_obtain_expr(expr const & e, constraint_seq & cs);
|
||||
|
||||
expr visit_prenum(expr const & e, constraint_seq & cs);
|
||||
|
||||
void check_used_local_tactic_hints();
|
||||
|
||||
void show_goal(proof_state const & ps, expr const & start, expr const & end, expr const & curr);
|
||||
|
|
|
@ -32,9 +32,11 @@ Author: Leonardo de Moura
|
|||
#include "frontends/lean/obtain_expr.h"
|
||||
#include "frontends/lean/decl_cmds.h"
|
||||
#include "frontends/lean/nested_declaration.h"
|
||||
#include "frontends/lean/prenum.h"
|
||||
|
||||
namespace lean {
|
||||
void initialize_frontend_lean_module() {
|
||||
initialize_prenum();
|
||||
initialize_info_annotation();
|
||||
initialize_tokens();
|
||||
initialize_token_table();
|
||||
|
@ -93,5 +95,6 @@ void finalize_frontend_lean_module() {
|
|||
finalize_token_table();
|
||||
finalize_tokens();
|
||||
finalize_info_annotation();
|
||||
finalize_prenum();
|
||||
}
|
||||
}
|
||||
|
|
|
@ -16,6 +16,8 @@ Author: Leonardo de Moura
|
|||
#include "library/num.h"
|
||||
#include "library/normalize.h"
|
||||
#include "library/aliases.h"
|
||||
#include "library/constants.h"
|
||||
#include "library/typed_expr.h"
|
||||
#include "frontends/lean/parser.h"
|
||||
#include "frontends/lean/tokens.h"
|
||||
#include "frontends/lean/util.h"
|
||||
|
@ -39,12 +41,13 @@ static unsigned parse_precedence_core(parser & p) {
|
|||
if (p.curr_is_numeral()) {
|
||||
return p.parse_small_nat();
|
||||
} else {
|
||||
environment env = open_prec_aliases(open_num_notation(p.env()));
|
||||
environment env = open_prec_aliases(p.env());
|
||||
parser::local_scope scope(p, env);
|
||||
expr val = p.parse_expr(get_max_prec());
|
||||
val = elim_choice_num(val);
|
||||
expr pre_val = p.parse_expr(get_max_prec());
|
||||
pre_val = mk_typed_expr(mk_constant(get_num_name()), pre_val);
|
||||
expr val = std::get<0>(p.elaborate(pre_val, list<expr>()));
|
||||
val = normalize(p.env(), val);
|
||||
if (optional<mpz> mpz_val = to_num(val)) {
|
||||
if (optional<mpz> mpz_val = to_num_core(val)) {
|
||||
if (!mpz_val->is_unsigned_int())
|
||||
throw parser_error("invalid 'precedence', argument does not fit in a machine integer", pos);
|
||||
return mpz_val->get_unsigned_int();
|
||||
|
|
|
@ -54,6 +54,7 @@ Author: Leonardo de Moura
|
|||
#include "frontends/lean/local_ref_info.h"
|
||||
#include "frontends/lean/opt_cmd.h"
|
||||
#include "frontends/lean/builtin_cmds.h"
|
||||
#include "frontends/lean/prenum.h"
|
||||
|
||||
#ifndef LEAN_DEFAULT_PARSER_SHOW_ERRORS
|
||||
#define LEAN_DEFAULT_PARSER_SHOW_ERRORS true
|
||||
|
@ -1536,52 +1537,30 @@ expr parser::parse_id() {
|
|||
return id_to_expr(id, p);
|
||||
}
|
||||
|
||||
expr parser::parse_numeral_expr(bool user_notation) {
|
||||
expr parser::parse_numeral_expr() {
|
||||
auto p = pos();
|
||||
mpz n = get_num_val().get_numerator();
|
||||
next();
|
||||
if (!m_has_num)
|
||||
m_has_num = has_num_decls(m_env);
|
||||
list<expr> vals;
|
||||
if (user_notation)
|
||||
vals = get_mpz_notation(m_env, n);
|
||||
if (!*m_has_num && !vals) {
|
||||
throw parser_error("numeral cannot be encoded as expression, environment does not contain the type 'num' "
|
||||
"nor notation was defined for the given numeral "
|
||||
"(solution: use 'import data.num', or define notation for the given numeral)", p);
|
||||
if (!*m_has_num) {
|
||||
throw parser_error("numeral cannot be encoded as expression, environment does not contain "
|
||||
"the auxiliary declarations zero, one, bit0 and bit1", p);
|
||||
}
|
||||
buffer<expr> cs;
|
||||
for (expr const & c : vals)
|
||||
cs.push_back(copy_with_new_pos(c, p));
|
||||
if (*m_has_num)
|
||||
cs.push_back(save_pos(copy(from_num(n)), p));
|
||||
// Remark: choices are processed from right to left.
|
||||
// We want to process user provided notation before the default 'num'.
|
||||
lean_assert(!cs.empty());
|
||||
if (cs.size() == 1)
|
||||
return cs[0];
|
||||
else
|
||||
return save_pos(mk_choice(cs.size(), cs.data()), p);
|
||||
return mk_prenum(n);
|
||||
}
|
||||
|
||||
expr parser::parse_decimal_expr() {
|
||||
auto p = pos();
|
||||
mpq val = get_num_val();
|
||||
next();
|
||||
if (!m_has_rat_of_num) {
|
||||
m_has_rat_of_num = static_cast<bool>(m_env.find(get_rat_of_num_name()));
|
||||
}
|
||||
if (!*m_has_rat_of_num) {
|
||||
throw parser_error("invalid decimal number, environment does not contain 'rat.of_num' "
|
||||
"(solution: use 'import data.rat')", p);
|
||||
}
|
||||
expr of_num = save_pos(mk_constant(get_rat_of_num_name()), p);
|
||||
expr num = mk_app(of_num, save_pos(copy(from_num(val.get_numerator())), p), p);
|
||||
expr num = save_pos(mk_prenum(val.get_numerator()), p);
|
||||
if (val.get_denominator() == 1) {
|
||||
return num;
|
||||
} else {
|
||||
expr den = mk_app(of_num, save_pos(copy(from_num(val.get_denominator())), p), p);
|
||||
return mk_app({save_pos(mk_constant(get_rat_divide_name()), p), num, den}, p);
|
||||
expr den = save_pos(mk_prenum(val.get_denominator()), p);
|
||||
expr div = save_pos(mk_constant(get_div_name()), p);
|
||||
return save_pos(lean::mk_app(div, num, den), p);
|
||||
}
|
||||
}
|
||||
|
||||
|
@ -1851,7 +1830,7 @@ expr parser::parse_tactic_opt_id_list() {
|
|||
expr parser::parse_tactic_option_num() {
|
||||
auto p = pos();
|
||||
if (curr_is_numeral()) {
|
||||
expr n = parse_numeral_expr(false);
|
||||
expr n = parse_numeral_expr(); // TODO(Leo): it should be a num
|
||||
return mk_app(save_pos(mk_constant(get_option_some_name()), p), n, p);
|
||||
} else {
|
||||
return save_pos(mk_constant(get_option_none_name()), p);
|
||||
|
|
|
@ -119,7 +119,6 @@ class parser {
|
|||
undef_id_behavior m_undef_id_behavior;
|
||||
optional<bool> m_has_num;
|
||||
optional<bool> m_has_string;
|
||||
optional<bool> m_has_rat_of_num;
|
||||
optional<bool> m_has_tactic_decls;
|
||||
// We process theorems in parallel
|
||||
theorem_queue m_theorem_queue;
|
||||
|
@ -215,7 +214,7 @@ class parser {
|
|||
expr parse_led_notation(expr left);
|
||||
expr parse_nud();
|
||||
bool curr_starts_expr();
|
||||
expr parse_numeral_expr(bool user_notation = true);
|
||||
expr parse_numeral_expr();
|
||||
expr parse_decimal_expr();
|
||||
expr parse_string_expr();
|
||||
expr parse_binder_core(binder_info const & bi, unsigned rbp);
|
||||
|
|
68
src/frontends/lean/prenum.cpp
Normal file
68
src/frontends/lean/prenum.cpp
Normal file
|
@ -0,0 +1,68 @@
|
|||
/*
|
||||
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Author: Leonardo de Moura
|
||||
*/
|
||||
#include "library/kernel_serializer.h"
|
||||
|
||||
namespace lean {
|
||||
static name * g_prenum_name = nullptr;
|
||||
static std::string * g_prenum_opcode = nullptr;
|
||||
[[ noreturn ]] static void throw_ex() { throw exception("unexpected occurrence of 'pre-numeral' expression"); }
|
||||
|
||||
class prenum_macro_definition_cell : public macro_definition_cell {
|
||||
mpz m_value;
|
||||
public:
|
||||
prenum_macro_definition_cell(mpz const & v):m_value(v) {}
|
||||
mpz const & get_value() const { return m_value; }
|
||||
virtual name get_name() const { return *g_prenum_name; }
|
||||
virtual format pp(formatter const &) const { return format(m_value); }
|
||||
virtual void display(std::ostream & out) const { out << m_value; }
|
||||
virtual pair<expr, constraint_seq> check_type(expr const &, extension_context &, bool) const { throw_ex(); }
|
||||
virtual optional<expr> expand(expr const &, extension_context &) const { throw_ex(); }
|
||||
virtual void write(serializer & s) const {
|
||||
s.write_string(*g_prenum_opcode);
|
||||
s << m_value;
|
||||
}
|
||||
virtual bool operator==(macro_definition_cell const & other) const {
|
||||
if (auto other_ptr = dynamic_cast<prenum_macro_definition_cell const *>(&other)) {
|
||||
return m_value == other_ptr->m_value;
|
||||
} else {
|
||||
return false;
|
||||
}
|
||||
}
|
||||
virtual unsigned hash() const {
|
||||
return ::lean::hash(m_value.hash(), g_prenum_name->hash());
|
||||
}
|
||||
};
|
||||
|
||||
expr mk_prenum(mpz const & v) {
|
||||
return mk_macro(macro_definition(new prenum_macro_definition_cell(v)), 0, nullptr);
|
||||
}
|
||||
|
||||
bool is_prenum(expr const & e) {
|
||||
return is_macro(e) && dynamic_cast<prenum_macro_definition_cell const*>(macro_def(e).raw()) != nullptr;
|
||||
}
|
||||
|
||||
mpz const & prenum_value(expr const & e) {
|
||||
lean_assert(is_prenum(e));
|
||||
return static_cast<prenum_macro_definition_cell const *>(macro_def(e).raw())->get_value();
|
||||
}
|
||||
|
||||
void initialize_prenum() {
|
||||
g_prenum_name = new name("prenum");
|
||||
g_prenum_opcode = new std::string("Prenum");
|
||||
register_macro_deserializer(*g_prenum_opcode,
|
||||
[](deserializer & d, unsigned, expr const *) {
|
||||
mpz v;
|
||||
d >> v;
|
||||
return mk_prenum(v);
|
||||
});
|
||||
}
|
||||
|
||||
void finalize_prenum() {
|
||||
delete g_prenum_name;
|
||||
delete g_prenum_opcode;
|
||||
}
|
||||
}
|
22
src/frontends/lean/prenum.h
Normal file
22
src/frontends/lean/prenum.h
Normal file
|
@ -0,0 +1,22 @@
|
|||
/*
|
||||
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Author: Leonardo de Moura
|
||||
*/
|
||||
#pragma once
|
||||
#include "kernel/expr.h"
|
||||
|
||||
namespace lean {
|
||||
/** \brief Create a pre-numeral. We create pre-numerals at parsing time. The elaborator is responsible for
|
||||
encoding them using the polymorphic operations zero, one, bit0 and bit1
|
||||
\remark Fully elaborated terms do not contain pre-numerals */
|
||||
expr mk_prenum(mpz const & v);
|
||||
/** \brief Return true iff \c e is a pre-numeral */
|
||||
bool is_prenum(expr const & e);
|
||||
/** \brief Return the value stored in the pre-numeral
|
||||
\pre is_prenum(e) */
|
||||
mpz const & prenum_value(expr const & e);
|
||||
void initialize_prenum();
|
||||
void finalize_prenum();
|
||||
}
|
|
@ -16,6 +16,8 @@ Author: Leonardo de Moura
|
|||
#include "library/locals.h"
|
||||
#include "library/explicit.h"
|
||||
#include "library/aliases.h"
|
||||
#include "library/constants.h"
|
||||
#include "library/typed_expr.h"
|
||||
#include "library/normalize.h"
|
||||
#include "library/placeholder.h"
|
||||
#include "library/abbreviation.h"
|
||||
|
@ -475,14 +477,11 @@ optional<unsigned> parse_priority(parser & p) {
|
|||
if (p.curr_is_token(get_priority_tk())) {
|
||||
p.next();
|
||||
auto pos = p.pos();
|
||||
environment env = open_num_notation(p.env());
|
||||
parser::local_scope scope(p, env);
|
||||
expr val = p.parse_expr();
|
||||
// Remark: open_num_notation will not override numeral overloading.
|
||||
// So, we use the following helper class for eliminating it.
|
||||
val = elim_choice_num(val);
|
||||
expr pre_val = p.parse_expr(get_max_prec());
|
||||
pre_val = mk_typed_expr(mk_constant(get_num_name()), pre_val);
|
||||
expr val = std::get<0>(p.elaborate(pre_val, list<expr>()));
|
||||
val = normalize(p.env(), val);
|
||||
if (optional<mpz> mpz_val = to_num(val)) {
|
||||
if (optional<mpz> mpz_val = to_num_core(val)) {
|
||||
if (!mpz_val->is_unsigned_int())
|
||||
throw parser_error("invalid 'priority', argument does not fit in a machine integer", pos);
|
||||
p.check_token_next(get_rbracket_tk(), "invalid 'priority', ']' expected");
|
||||
|
|
|
@ -8,6 +8,8 @@ name const * g_and = nullptr;
|
|||
name const * g_and_elim_left = nullptr;
|
||||
name const * g_and_elim_right = nullptr;
|
||||
name const * g_and_intro = nullptr;
|
||||
name const * g_bit0 = nullptr;
|
||||
name const * g_bit1 = nullptr;
|
||||
name const * g_bool = nullptr;
|
||||
name const * g_bool_ff = nullptr;
|
||||
name const * g_bool_tt = nullptr;
|
||||
|
@ -15,6 +17,7 @@ name const * g_char = nullptr;
|
|||
name const * g_char_mk = nullptr;
|
||||
name const * g_congr = nullptr;
|
||||
name const * g_dite = nullptr;
|
||||
name const * g_div = nullptr;
|
||||
name const * g_empty = nullptr;
|
||||
name const * g_empty_rec = nullptr;
|
||||
name const * g_eq = nullptr;
|
||||
|
@ -31,6 +34,9 @@ name const * g_exists_elim = nullptr;
|
|||
name const * g_false = nullptr;
|
||||
name const * g_false_rec = nullptr;
|
||||
name const * g_funext = nullptr;
|
||||
name const * g_has_zero = nullptr;
|
||||
name const * g_has_one = nullptr;
|
||||
name const * g_has_add = nullptr;
|
||||
name const * g_heq = nullptr;
|
||||
name const * g_heq_refl = nullptr;
|
||||
name const * g_heq_to_eq = nullptr;
|
||||
|
@ -54,6 +60,7 @@ name const * g_not = nullptr;
|
|||
name const * g_num = nullptr;
|
||||
name const * g_num_zero = nullptr;
|
||||
name const * g_num_pos = nullptr;
|
||||
name const * g_one = nullptr;
|
||||
name const * g_option = nullptr;
|
||||
name const * g_option_some = nullptr;
|
||||
name const * g_option_none = nullptr;
|
||||
|
@ -146,12 +153,15 @@ name const * g_true_intro = nullptr;
|
|||
name const * g_is_trunc_is_hset = nullptr;
|
||||
name const * g_is_trunc_is_hprop = nullptr;
|
||||
name const * g_well_founded = nullptr;
|
||||
name const * g_zero = nullptr;
|
||||
void initialize_constants() {
|
||||
g_absurd = new name{"absurd"};
|
||||
g_and = new name{"and"};
|
||||
g_and_elim_left = new name{"and", "elim_left"};
|
||||
g_and_elim_right = new name{"and", "elim_right"};
|
||||
g_and_intro = new name{"and", "intro"};
|
||||
g_bit0 = new name{"bit0"};
|
||||
g_bit1 = new name{"bit1"};
|
||||
g_bool = new name{"bool"};
|
||||
g_bool_ff = new name{"bool", "ff"};
|
||||
g_bool_tt = new name{"bool", "tt"};
|
||||
|
@ -159,6 +169,7 @@ void initialize_constants() {
|
|||
g_char_mk = new name{"char", "mk"};
|
||||
g_congr = new name{"congr"};
|
||||
g_dite = new name{"dite"};
|
||||
g_div = new name{"div"};
|
||||
g_empty = new name{"empty"};
|
||||
g_empty_rec = new name{"empty", "rec"};
|
||||
g_eq = new name{"eq"};
|
||||
|
@ -175,6 +186,9 @@ void initialize_constants() {
|
|||
g_false = new name{"false"};
|
||||
g_false_rec = new name{"false", "rec"};
|
||||
g_funext = new name{"funext"};
|
||||
g_has_zero = new name{"has_zero"};
|
||||
g_has_one = new name{"has_one"};
|
||||
g_has_add = new name{"has_add"};
|
||||
g_heq = new name{"heq"};
|
||||
g_heq_refl = new name{"heq", "refl"};
|
||||
g_heq_to_eq = new name{"heq", "to_eq"};
|
||||
|
@ -198,6 +212,7 @@ void initialize_constants() {
|
|||
g_num = new name{"num"};
|
||||
g_num_zero = new name{"num", "zero"};
|
||||
g_num_pos = new name{"num", "pos"};
|
||||
g_one = new name{"one"};
|
||||
g_option = new name{"option"};
|
||||
g_option_some = new name{"option", "some"};
|
||||
g_option_none = new name{"option", "none"};
|
||||
|
@ -290,6 +305,7 @@ void initialize_constants() {
|
|||
g_is_trunc_is_hset = new name{"is_trunc", "is_hset"};
|
||||
g_is_trunc_is_hprop = new name{"is_trunc", "is_hprop"};
|
||||
g_well_founded = new name{"well_founded"};
|
||||
g_zero = new name{"zero"};
|
||||
}
|
||||
void finalize_constants() {
|
||||
delete g_absurd;
|
||||
|
@ -297,6 +313,8 @@ void finalize_constants() {
|
|||
delete g_and_elim_left;
|
||||
delete g_and_elim_right;
|
||||
delete g_and_intro;
|
||||
delete g_bit0;
|
||||
delete g_bit1;
|
||||
delete g_bool;
|
||||
delete g_bool_ff;
|
||||
delete g_bool_tt;
|
||||
|
@ -304,6 +322,7 @@ void finalize_constants() {
|
|||
delete g_char_mk;
|
||||
delete g_congr;
|
||||
delete g_dite;
|
||||
delete g_div;
|
||||
delete g_empty;
|
||||
delete g_empty_rec;
|
||||
delete g_eq;
|
||||
|
@ -320,6 +339,9 @@ void finalize_constants() {
|
|||
delete g_false;
|
||||
delete g_false_rec;
|
||||
delete g_funext;
|
||||
delete g_has_zero;
|
||||
delete g_has_one;
|
||||
delete g_has_add;
|
||||
delete g_heq;
|
||||
delete g_heq_refl;
|
||||
delete g_heq_to_eq;
|
||||
|
@ -343,6 +365,7 @@ void finalize_constants() {
|
|||
delete g_num;
|
||||
delete g_num_zero;
|
||||
delete g_num_pos;
|
||||
delete g_one;
|
||||
delete g_option;
|
||||
delete g_option_some;
|
||||
delete g_option_none;
|
||||
|
@ -435,12 +458,15 @@ void finalize_constants() {
|
|||
delete g_is_trunc_is_hset;
|
||||
delete g_is_trunc_is_hprop;
|
||||
delete g_well_founded;
|
||||
delete g_zero;
|
||||
}
|
||||
name const & get_absurd_name() { return *g_absurd; }
|
||||
name const & get_and_name() { return *g_and; }
|
||||
name const & get_and_elim_left_name() { return *g_and_elim_left; }
|
||||
name const & get_and_elim_right_name() { return *g_and_elim_right; }
|
||||
name const & get_and_intro_name() { return *g_and_intro; }
|
||||
name const & get_bit0_name() { return *g_bit0; }
|
||||
name const & get_bit1_name() { return *g_bit1; }
|
||||
name const & get_bool_name() { return *g_bool; }
|
||||
name const & get_bool_ff_name() { return *g_bool_ff; }
|
||||
name const & get_bool_tt_name() { return *g_bool_tt; }
|
||||
|
@ -448,6 +474,7 @@ name const & get_char_name() { return *g_char; }
|
|||
name const & get_char_mk_name() { return *g_char_mk; }
|
||||
name const & get_congr_name() { return *g_congr; }
|
||||
name const & get_dite_name() { return *g_dite; }
|
||||
name const & get_div_name() { return *g_div; }
|
||||
name const & get_empty_name() { return *g_empty; }
|
||||
name const & get_empty_rec_name() { return *g_empty_rec; }
|
||||
name const & get_eq_name() { return *g_eq; }
|
||||
|
@ -464,6 +491,9 @@ name const & get_exists_elim_name() { return *g_exists_elim; }
|
|||
name const & get_false_name() { return *g_false; }
|
||||
name const & get_false_rec_name() { return *g_false_rec; }
|
||||
name const & get_funext_name() { return *g_funext; }
|
||||
name const & get_has_zero_name() { return *g_has_zero; }
|
||||
name const & get_has_one_name() { return *g_has_one; }
|
||||
name const & get_has_add_name() { return *g_has_add; }
|
||||
name const & get_heq_name() { return *g_heq; }
|
||||
name const & get_heq_refl_name() { return *g_heq_refl; }
|
||||
name const & get_heq_to_eq_name() { return *g_heq_to_eq; }
|
||||
|
@ -487,6 +517,7 @@ name const & get_not_name() { return *g_not; }
|
|||
name const & get_num_name() { return *g_num; }
|
||||
name const & get_num_zero_name() { return *g_num_zero; }
|
||||
name const & get_num_pos_name() { return *g_num_pos; }
|
||||
name const & get_one_name() { return *g_one; }
|
||||
name const & get_option_name() { return *g_option; }
|
||||
name const & get_option_some_name() { return *g_option_some; }
|
||||
name const & get_option_none_name() { return *g_option_none; }
|
||||
|
@ -579,4 +610,5 @@ name const & get_true_intro_name() { return *g_true_intro; }
|
|||
name const & get_is_trunc_is_hset_name() { return *g_is_trunc_is_hset; }
|
||||
name const & get_is_trunc_is_hprop_name() { return *g_is_trunc_is_hprop; }
|
||||
name const & get_well_founded_name() { return *g_well_founded; }
|
||||
name const & get_zero_name() { return *g_zero; }
|
||||
}
|
||||
|
|
|
@ -10,6 +10,8 @@ name const & get_and_name();
|
|||
name const & get_and_elim_left_name();
|
||||
name const & get_and_elim_right_name();
|
||||
name const & get_and_intro_name();
|
||||
name const & get_bit0_name();
|
||||
name const & get_bit1_name();
|
||||
name const & get_bool_name();
|
||||
name const & get_bool_ff_name();
|
||||
name const & get_bool_tt_name();
|
||||
|
@ -17,6 +19,7 @@ name const & get_char_name();
|
|||
name const & get_char_mk_name();
|
||||
name const & get_congr_name();
|
||||
name const & get_dite_name();
|
||||
name const & get_div_name();
|
||||
name const & get_empty_name();
|
||||
name const & get_empty_rec_name();
|
||||
name const & get_eq_name();
|
||||
|
@ -33,6 +36,9 @@ name const & get_exists_elim_name();
|
|||
name const & get_false_name();
|
||||
name const & get_false_rec_name();
|
||||
name const & get_funext_name();
|
||||
name const & get_has_zero_name();
|
||||
name const & get_has_one_name();
|
||||
name const & get_has_add_name();
|
||||
name const & get_heq_name();
|
||||
name const & get_heq_refl_name();
|
||||
name const & get_heq_to_eq_name();
|
||||
|
@ -56,6 +62,7 @@ name const & get_not_name();
|
|||
name const & get_num_name();
|
||||
name const & get_num_zero_name();
|
||||
name const & get_num_pos_name();
|
||||
name const & get_one_name();
|
||||
name const & get_option_name();
|
||||
name const & get_option_some_name();
|
||||
name const & get_option_none_name();
|
||||
|
@ -148,4 +155,5 @@ name const & get_true_intro_name();
|
|||
name const & get_is_trunc_is_hset_name();
|
||||
name const & get_is_trunc_is_hprop_name();
|
||||
name const & get_well_founded_name();
|
||||
name const & get_zero_name();
|
||||
}
|
||||
|
|
|
@ -3,6 +3,8 @@ and
|
|||
and.elim_left
|
||||
and.elim_right
|
||||
and.intro
|
||||
bit0
|
||||
bit1
|
||||
bool
|
||||
bool.ff
|
||||
bool.tt
|
||||
|
@ -10,6 +12,7 @@ char
|
|||
char.mk
|
||||
congr
|
||||
dite
|
||||
div
|
||||
empty
|
||||
empty.rec
|
||||
eq
|
||||
|
@ -26,6 +29,9 @@ exists.elim
|
|||
false
|
||||
false.rec
|
||||
funext
|
||||
has_zero
|
||||
has_one
|
||||
has_add
|
||||
heq
|
||||
heq.refl
|
||||
heq.to_eq
|
||||
|
@ -49,6 +55,7 @@ not
|
|||
num
|
||||
num.zero
|
||||
num.pos
|
||||
one
|
||||
option
|
||||
option.some
|
||||
option.none
|
||||
|
@ -141,3 +148,4 @@ true.intro
|
|||
is_trunc.is_hset
|
||||
is_trunc.is_hprop
|
||||
well_founded
|
||||
zero
|
||||
|
|
|
@ -12,7 +12,6 @@ Author: Leonardo de Moura
|
|||
#include "library/choice.h"
|
||||
#include "library/class.h"
|
||||
#include "library/string.h"
|
||||
#include "library/num.h"
|
||||
#include "library/resolve_macro.h"
|
||||
#include "library/annotation.h"
|
||||
#include "library/explicit.h"
|
||||
|
@ -57,7 +56,6 @@ void initialize_library_module() {
|
|||
initialize_let();
|
||||
initialize_typed_expr();
|
||||
initialize_choice();
|
||||
initialize_num();
|
||||
initialize_string();
|
||||
initialize_resolve_macro();
|
||||
initialize_annotation();
|
||||
|
@ -111,7 +109,6 @@ void finalize_library_module() {
|
|||
finalize_annotation();
|
||||
finalize_resolve_macro();
|
||||
finalize_string();
|
||||
finalize_num();
|
||||
finalize_choice();
|
||||
finalize_typed_expr();
|
||||
finalize_let();
|
||||
|
|
|
@ -9,88 +9,92 @@ Author: Leonardo de Moura
|
|||
#include "library/constants.h"
|
||||
|
||||
namespace lean {
|
||||
static expr * g_num = nullptr;
|
||||
static expr * g_pos_num = nullptr;
|
||||
static expr * g_zero = nullptr;
|
||||
static expr * g_pos = nullptr;
|
||||
static expr * g_one = nullptr;
|
||||
static expr * g_bit0 = nullptr;
|
||||
static expr * g_bit1 = nullptr;
|
||||
|
||||
void initialize_num() {
|
||||
g_num = new expr(Const(get_num_name()));
|
||||
g_pos_num = new expr(Const(get_pos_num_name()));
|
||||
g_zero = new expr(Const(get_num_zero_name()));
|
||||
g_pos = new expr(Const(get_num_pos_name()));
|
||||
g_one = new expr(Const(get_pos_num_one_name()));
|
||||
g_bit0 = new expr(Const(get_pos_num_bit0_name()));
|
||||
g_bit1 = new expr(Const(get_pos_num_bit1_name()));
|
||||
}
|
||||
|
||||
void finalize_num() {
|
||||
delete g_num;
|
||||
delete g_pos_num;
|
||||
delete g_zero;
|
||||
delete g_pos;
|
||||
delete g_one;
|
||||
delete g_bit0;
|
||||
delete g_bit1;
|
||||
}
|
||||
|
||||
bool has_num_decls(environment const & env) {
|
||||
try {
|
||||
type_checker tc(env);
|
||||
return
|
||||
tc.infer(*g_zero).first == *g_num &&
|
||||
tc.infer(*g_pos).first == mk_arrow(*g_pos_num, *g_num) &&
|
||||
tc.infer(*g_one).first == *g_pos_num &&
|
||||
tc.infer(*g_bit0).first == mk_arrow(*g_pos_num, *g_pos_num) &&
|
||||
tc.infer(*g_bit1).first == mk_arrow(*g_pos_num, *g_pos_num);
|
||||
} catch (exception&) {
|
||||
return
|
||||
env.find(get_zero_name()) &&
|
||||
env.find(get_one_name()) &&
|
||||
env.find(get_bit0_name()) &&
|
||||
env.find(get_bit1_name());
|
||||
}
|
||||
|
||||
static bool is_const_app(expr const & e, name const & n, unsigned nargs) {
|
||||
expr const & f = get_app_fn(e);
|
||||
return is_constant(f) && const_name(f) == n && get_app_num_args(e) == nargs;
|
||||
}
|
||||
|
||||
bool is_zero(expr const & e) {
|
||||
return is_const_app(e, get_zero_name(), 2);
|
||||
}
|
||||
|
||||
bool is_one(expr const & e) {
|
||||
return is_const_app(e, get_one_name(), 2);
|
||||
}
|
||||
|
||||
optional<expr> is_bit0(expr const & e) {
|
||||
if (!is_const_app(e, get_bit0_name(), 3))
|
||||
return none_expr();
|
||||
return some_expr(app_arg(e));
|
||||
}
|
||||
|
||||
optional<expr> is_bit1(expr const & e) {
|
||||
if (!is_const_app(e, get_bit1_name(), 4))
|
||||
return none_expr();
|
||||
return some_expr(app_arg(e));
|
||||
}
|
||||
|
||||
static bool is_num(expr const & e, bool first) {
|
||||
if (is_zero(e))
|
||||
return first;
|
||||
else if (is_one(e))
|
||||
return true;
|
||||
else if (auto a = is_bit0(e))
|
||||
return is_num(*a, false);
|
||||
else if (auto a = is_bit1(e))
|
||||
return is_num(*a, false);
|
||||
else
|
||||
return false;
|
||||
}
|
||||
}
|
||||
|
||||
expr from_pos_num(mpz const & n) {
|
||||
lean_assert(n > 0);
|
||||
if (n == 1)
|
||||
return *g_one;
|
||||
if (n % mpz(2) == 1)
|
||||
return mk_app(*g_bit1, from_pos_num(n / 2));
|
||||
else
|
||||
return mk_app(*g_bit0, from_pos_num(n / 2));
|
||||
bool is_num(expr const & e) {
|
||||
return is_num(e, true);
|
||||
}
|
||||
|
||||
expr from_num(mpz const & n) {
|
||||
expr r;
|
||||
lean_assert(n >= 0);
|
||||
if (n == 0)
|
||||
r = *g_zero;
|
||||
else
|
||||
r = mk_app(*g_pos, from_pos_num(n));
|
||||
lean_assert(*to_num(r) == n);
|
||||
return r;
|
||||
}
|
||||
|
||||
optional<mpz> to_pos_num(expr const & e) {
|
||||
if (e == *g_one) {
|
||||
static optional<mpz> to_num(expr const & e, bool first) {
|
||||
if (is_zero(e)) {
|
||||
return first ? some(mpz(0)) : optional<mpz>();
|
||||
} else if (is_one(e)) {
|
||||
return some(mpz(1));
|
||||
} else if (is_app(e)) {
|
||||
if (app_fn(e) == *g_bit0) {
|
||||
if (auto r = to_pos_num(app_arg(e)))
|
||||
return some(2*(*r));
|
||||
} else if (app_fn(e) == *g_bit1) {
|
||||
if (auto r = to_pos_num(app_arg(e)))
|
||||
return some(2*(*r) + 1);
|
||||
}
|
||||
} else if (auto a = is_bit0(e)) {
|
||||
if (auto r = to_num(*a, false))
|
||||
return some(2*(*r));
|
||||
} else if (auto a = is_bit1(e)) {
|
||||
if (auto r = to_num(*a, false))
|
||||
return some(2*(*r)+1);
|
||||
}
|
||||
return optional<mpz>();
|
||||
}
|
||||
|
||||
optional<mpz> to_num(expr const & e) {
|
||||
if (e == *g_zero)
|
||||
return to_num(e, true);
|
||||
}
|
||||
|
||||
optional<mpz> to_pos_num(expr const & e) {
|
||||
if (is_constant(e, get_pos_num_one_name())) {
|
||||
return some(mpz(1));
|
||||
} else if (is_const_app(e, get_pos_num_bit0_name(), 1)) {
|
||||
if (auto r = to_pos_num(app_arg(e)))
|
||||
return some(2*(*r));
|
||||
} else if (is_const_app(e, get_pos_num_bit1_name(), 1)) {
|
||||
if (auto r = to_pos_num(app_arg(e)))
|
||||
return some(2*(*r) + 1);
|
||||
}
|
||||
return optional<mpz>();
|
||||
}
|
||||
|
||||
optional<mpz> to_num_core(expr const & e) {
|
||||
if (is_constant(e, get_num_zero_name()))
|
||||
return some(mpz(0));
|
||||
else if (is_app(e) && app_fn(e) == *g_pos)
|
||||
else if (is_const_app(e, get_num_pos_name(), 1))
|
||||
return to_pos_num(app_arg(e));
|
||||
else
|
||||
return optional<mpz>();
|
||||
|
|
|
@ -8,35 +8,25 @@ Author: Leonardo de Moura
|
|||
#include "kernel/environment.h"
|
||||
|
||||
namespace lean {
|
||||
/**
|
||||
\brief Return true iff the environment \c env contains the following declarations
|
||||
in the namespace 'num'
|
||||
one : pos_num
|
||||
bit0 : pos_num -> pos_num
|
||||
bit1 : pos_num -> pos_num
|
||||
zero : num
|
||||
pos : pos_num -> num
|
||||
*/
|
||||
/** \brief Return true iff the given environment contains the declarations needed to encode numerals:
|
||||
zero, one, bit0, bit1 */
|
||||
bool has_num_decls(environment const & env);
|
||||
|
||||
/**
|
||||
\brief Return an expression that encodes the given numeral in binary using
|
||||
the declarations one, bit0, bit1, zero, pos.
|
||||
/** \brief Return true iff the given expression encodes a numeral. */
|
||||
bool is_num(expr const & e);
|
||||
|
||||
\see has_num_decls
|
||||
bool is_zero(expr const & e);
|
||||
bool is_one(expr const & e);
|
||||
optional<expr> is_bit0(expr const & e);
|
||||
optional<expr> is_bit1(expr const & e);
|
||||
|
||||
\pre n >= 0
|
||||
\post *to_num(from_num(n)) == n
|
||||
*/
|
||||
expr from_num(mpz const & n);
|
||||
|
||||
/**
|
||||
\brief If the given expression encodes a numeral, then convert it back to mpz numeral.
|
||||
|
||||
\see from_num
|
||||
*/
|
||||
/** \brief If the given expression encodes a numeral, then convert it back to mpz numeral.
|
||||
\see from_num */
|
||||
optional<mpz> to_num(expr const & e);
|
||||
|
||||
/** \brief If the given expression is a numeral encode the num and pos_num types, return the encoded numeral */
|
||||
optional<mpz> to_num_core(expr const & e);
|
||||
|
||||
void initialize_num();
|
||||
void finalize_num();
|
||||
}
|
||||
|
|
Loading…
Reference in a new issue