feat (hott/init): move nat.of_num to init.num and make it reducible outside namespace nat
This is so that init.trunc can already use nat.of_num. Also make nat.of_num reducible in the standard library Also make gt and ge abbreviations
This commit is contained in:
parent
88e287df5a
commit
9201bd7ca6
5 changed files with 31 additions and 21 deletions
|
@ -1,6 +1,8 @@
|
|||
/-
|
||||
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Module init.nat
|
||||
Authors: Floris van Doorn, Leonardo de Moura
|
||||
-/
|
||||
|
||||
|
@ -274,18 +276,15 @@ namespace nat
|
|||
have aux : a = max a b, from max.left_eq (lt.asymm h),
|
||||
eq.rec_on aux (le.of_lt h)))
|
||||
|
||||
definition gt a b := lt b a
|
||||
abbreviation gt a b := lt b a
|
||||
|
||||
notation a > b := gt a b
|
||||
|
||||
definition ge a b := le b a
|
||||
abbreviation ge a b := le b a
|
||||
|
||||
notation a ≥ b := ge a b
|
||||
|
||||
definition add (a b : nat) : nat :=
|
||||
nat.rec_on b a (λ b₁ r, succ r)
|
||||
|
||||
notation a + b := add a b
|
||||
-- add is defined in init.num
|
||||
|
||||
definition sub (a b : nat) : nat :=
|
||||
nat.rec_on b a (λ b₁ r, pred r)
|
||||
|
@ -343,7 +342,4 @@ namespace nat
|
|||
(le.refl a)
|
||||
(λ b₁ ih, le.trans !pred_le ih)
|
||||
|
||||
definition of_num [coercion] [reducible] (n : num) : ℕ :=
|
||||
num.rec zero
|
||||
(λ n, pos_num.rec (succ zero) (λ n r, r + r + (succ zero)) (λ n r, r + r) n) n
|
||||
end nat
|
||||
|
|
|
@ -2,6 +2,7 @@
|
|||
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Module init.num
|
||||
Authors: Leonardo de Moura
|
||||
-/
|
||||
prelude
|
||||
|
@ -72,3 +73,17 @@ namespace num
|
|||
notation a + b := add a b
|
||||
notation a * b := mul a b
|
||||
end num
|
||||
|
||||
-- the coercion from num to nat is defined here, so that it can already be used in init.trunc
|
||||
namespace nat
|
||||
|
||||
definition add (a b : nat) : nat :=
|
||||
nat.rec_on b a (λ b₁ r, succ r)
|
||||
|
||||
notation a + b := add a b
|
||||
|
||||
definition of_num [coercion] (n : num) : nat :=
|
||||
num.rec zero
|
||||
(λ n, pos_num.rec (succ zero) (λ n r, r + r + (succ zero)) (λ n r, r + r) n) n
|
||||
end nat
|
||||
attribute nat.of_num [reducible] -- of_num is also reducible if namespace "nat" is not opened
|
||||
|
|
|
@ -6,20 +6,18 @@ Module: init.trunc
|
|||
Authors: Jeremy Avigad, Floris van Doorn
|
||||
|
||||
Ported from Coq HoTT.
|
||||
|
||||
TODO: can we replace some definitions with a hprop as codomain by theorems?
|
||||
-/
|
||||
|
||||
prelude
|
||||
import .path .logic .datatypes .equiv .types.empty .types.sigma
|
||||
import .logic .equiv .types.empty .types.sigma
|
||||
open eq nat sigma unit
|
||||
|
||||
/- Truncation levels -/
|
||||
|
||||
-- TODO: can we replace some definitions with a hprop as codomain by theorems?
|
||||
|
||||
/- truncation indices -/
|
||||
|
||||
namespace is_trunc
|
||||
|
||||
/- Truncation levels -/
|
||||
|
||||
inductive trunc_index : Type₁ :=
|
||||
| minus_two : trunc_index
|
||||
| succ : trunc_index → trunc_index
|
||||
|
@ -32,7 +30,7 @@ namespace is_trunc
|
|||
postfix `.+2`:(max+1) := λn, (n .+1 .+1)
|
||||
notation `-2` := trunc_index.minus_two
|
||||
notation `-1` := -2.+1
|
||||
export [coercions] nat -- does this export
|
||||
export [coercions] nat
|
||||
|
||||
namespace trunc_index
|
||||
definition add (n m : trunc_index) : trunc_index :=
|
||||
|
@ -81,7 +79,7 @@ namespace is_trunc
|
|||
|
||||
abbreviation is_contr := is_trunc -2
|
||||
abbreviation is_hprop := is_trunc -1
|
||||
abbreviation is_hset := is_trunc nat.zero
|
||||
abbreviation is_hset := is_trunc 0
|
||||
|
||||
variables {A B : Type}
|
||||
|
||||
|
|
|
@ -11,7 +11,6 @@ Properties of is_trunc
|
|||
import types.pi types.path
|
||||
|
||||
open sigma sigma.ops pi function eq equiv path funext
|
||||
|
||||
namespace is_trunc
|
||||
|
||||
definition is_contr.sigma_char (A : Type) :
|
||||
|
@ -108,8 +107,9 @@ namespace is_trunc
|
|||
|
||||
section
|
||||
open decidable
|
||||
--this is proven differently in init.hedberg
|
||||
definition is_hset_of_decidable_eq (A : Type)
|
||||
[H : Π(a b : A), decidable (a = b)] : is_hset A :=
|
||||
[H : decidable_eq A] : is_hset A :=
|
||||
is_hset_of_double_neg_elim (λa b, by_contradiction)
|
||||
end
|
||||
|
||||
|
|
|
@ -334,7 +334,8 @@ namespace nat
|
|||
(le.refl a)
|
||||
(λ b₁ ih, le.trans !pred_le ih)
|
||||
|
||||
definition of_num [coercion] [reducible] (n : num) : ℕ :=
|
||||
definition of_num [coercion] (n : num) : ℕ :=
|
||||
num.rec zero
|
||||
(λ n, pos_num.rec (succ zero) (λ n r, r + r + (succ zero)) (λ n r, r + r) n) n
|
||||
end nat
|
||||
attribute nat.of_num [reducible] -- of_num is also reducible if namespace "nat" is not opened
|
||||
|
|
Loading…
Reference in a new issue