feat(*): change inductive datatype syntax

Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
Leonardo de Moura 2014-08-22 15:46:10 -07:00
parent 626cd952e7
commit a5f0593df1
82 changed files with 292 additions and 300 deletions

View file

@ -7,8 +7,8 @@ import logic.connectives.basic logic.classes.decidable logic.classes.inhabited
using eq_ops decidable
inductive bool : Type :=
| ff : bool
| tt : bool
ff : bool,
tt : bool
namespace bool

View file

@ -24,8 +24,8 @@ namespace list
-- ----
inductive list (T : Type) : Type :=
| nil {} : list T
| cons : T → list T → list T
nil {} : list T,
cons : T → list T → list T
infix `::` : 65 := cons

View file

@ -25,8 +25,8 @@ using helper_tactics
namespace nat
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
notation ``:max := nat

View file

@ -12,13 +12,13 @@ namespace num
-- The parser will generate the terms (pos (bit1 (bit1 (bit0 one)))), zero, and (pos (bit0 (bit1 (bit1 one)))).
-- This representation can be coerced in whatever we want (e.g., naturals, integers, reals, etc).
inductive pos_num : Type :=
| one : pos_num
| bit1 : pos_num → pos_num
| bit0 : pos_num → pos_num
one : pos_num,
bit1 : pos_num → pos_num,
bit0 : pos_num → pos_num
inductive num : Type :=
| zero : num
| pos : pos_num → num
zero : num,
pos : pos_num → num
theorem inhabited_pos_num [instance] : inhabited pos_num :=
inhabited_mk one

View file

@ -8,8 +8,8 @@ using eq_ops decidable
namespace option
inductive option (A : Type) : Type :=
| none {} : option A
| some : A → option A
none {} : option A,
some : A → option A
theorem induction_on {A : Type} {p : option A → Prop} (o : option A)
(H1 : p none) (H2 : ∀a, p (some a)) : p o :=

View file

@ -7,7 +7,7 @@ import logic.classes.inhabited logic.connectives.eq logic.classes.decidable
using inhabited decidable
inductive prod (A B : Type) : Type :=
| pair : A → B → prod A B
pair : A → B → prod A B
precedence `×`:30
infixr × := prod
@ -47,8 +47,8 @@ section
(H2 : decidable (pr2 u = pr2 v)) : decidable (u = v) :=
have H3 : u = v ↔ (pr1 u = pr1 v) ∧ (pr2 u = pr2 v), from
iff_intro
(assume H, subst H (and_intro (refl _) (refl _)))
(assume H, and_elim H (assume H4 H5, prod_eq H4 H5)),
(assume H, subst H (and_intro (refl _) (refl _)))
(assume H, and_elim H (assume H4 H5, prod_eq H4 H5)),
decidable_iff_equiv _ (iff_symm H3)
end

View file

@ -7,7 +7,7 @@ import logic.classes.inhabited logic.connectives.eq
using inhabited
inductive sigma {A : Type} (B : A → Type) : Type :=
| dpair : Πx : A, B x → sigma B
dpair : Πx : A, B x → sigma B
notation `Σ` binders `,` r:(scoped P, sigma P) := r
@ -36,12 +36,12 @@ section
(show ∀(b2 : B a2) (H1 : a1 = a2) (H2 : eq_rec_on H1 b1 = b2), dpair a1 b1 = dpair a2 b2, from
eq_rec
(take (b2' : B a1),
assume (H1' : a1 = a1),
assume (H2' : eq_rec_on H1' b1 = b2'),
show dpair a1 b1 = dpair a1 b2', from
calc
dpair a1 b1 = dpair a1 (eq_rec_on H1' b1) : {symm (eq_rec_on_id H1' b1)}
... = dpair a1 b2' : {H2'}) H1)
assume (H1' : a1 = a1),
assume (H2' : eq_rec_on H1' b1 = b2'),
show dpair a1 b1 = dpair a1 b2', from
calc
dpair a1 b1 = dpair a1 (eq_rec_on H1' b1) : {symm (eq_rec_on_id H1' b1)}
... = dpair a1 b2' : {H2'}) H1)
b2 H1 H2
theorem sigma_eq {p1 p2 : Σx : A, B x} :

View file

@ -9,11 +9,11 @@ using bool inhabited
namespace string
inductive char : Type :=
| ascii : bool → bool → bool → bool → bool → bool → bool → bool → char
ascii : bool → bool → bool → bool → bool → bool → bool → bool → char
inductive string : Type :=
| empty : string
| str : char → string → string
empty : string,
str : char → string → string
theorem char_inhabited [instance] : inhabited char :=
inhabited_mk (ascii ff ff ff ff ff ff ff ff)

View file

@ -7,7 +7,7 @@ import logic.classes.inhabited logic.connectives.eq logic.classes.decidable
using decidable
inductive subtype {A : Type} (P : A → Prop) : Type :=
| tag : Πx : A, P x → subtype P
tag : Πx : A, P x → subtype P
notation `{` binders `|` r:(scoped P, subtype P) `}` := r

View file

@ -10,8 +10,8 @@ namespace sum
-- TODO: take this outside the namespace when the inductive package handles it better
inductive sum (A B : Type) : Type :=
| inl : A → sum A B
| inr : B → sum A B
inl : A → sum A B,
inr : B → sum A B
infixr `+`:25 := sum

View file

@ -9,7 +9,7 @@ using decidable
namespace unit
inductive unit : Type :=
| star : unit
star : unit
notation `⋆`:max := star

View file

@ -16,7 +16,7 @@ abbreviation Sect {A B : Type} (s : A → B) (r : B → A) := Πx : A, r (s x)
-- Structure IsEquiv
inductive IsEquiv {A B : Type} (f : A → B) :=
| IsEquiv_mk : Π
IsEquiv_mk : Π
(equiv_inv : B → A)
(eisretr : Sect equiv_inv f)
(eissect : Sect f equiv_inv)
@ -40,7 +40,7 @@ IsEquiv_rec (λequiv_inv eisretr eissect eisadj, eisadj) H
-- Structure Equiv
inductive Equiv (A B : Type) : Type :=
| Equiv_mk : Π
Equiv_mk : Π
(equiv_fun : A → B)
(equiv_isequiv : IsEquiv equiv_fun),
Equiv A B

View file

@ -8,7 +8,7 @@ import data.prod data.sum data.sigma
using unit bool nat prod sum sigma
inductive fibrant (T : Type) : Type :=
| fibrant_mk : fibrant T
fibrant_mk : fibrant T
namespace fibrant

View file

@ -16,7 +16,7 @@ using function
-- ----
inductive path {A : Type} (a : A) : A → Type :=
| idpath : path a a
idpath : path a a
infix `≈`:50 := path
notation x `≈` y:50 `:>`:0 A:0 := @path A x y -- TODO: is this right?

View file

@ -10,7 +10,7 @@ import .path
-- -----------------
inductive Contr (A : Type) : Type :=
| Contr_mk : Π
Contr_mk : Π
(center : A)
(contr : Πy : A, center ≈ y),
Contr A
@ -21,8 +21,8 @@ definition contr {A : Type} (C : Contr A) : Πy : A, center C ≈ y :=
Contr_rec (λcenter contr, contr) C
inductive trunc_index : Type :=
| minus_two : trunc_index
| trunc_S : trunc_index → trunc_index
minus_two : trunc_index,
trunc_S : trunc_index → trunc_index
-- TODO: add coercions to / from nat

View file

@ -7,8 +7,8 @@ import logic.connectives.basic logic.connectives.eq
namespace decidable
inductive decidable (p : Prop) : Type :=
| inl : p → decidable p
| inr : ¬p → decidable p
inl : p → decidable p,
inr : ¬p → decidable p
theorem true_decidable [instance] : decidable true :=
inl trivial
@ -76,7 +76,7 @@ rec_on Ha
(assume Hna, rec_on Hb
(assume Hb : b, inr (assume H : a ↔ b, absurd (iff_elim_right H Hb) Hna))
(assume Hnb : ¬b, inl
(iff_intro (assume Ha, absurd_elim b Ha Hna) (assume Hb, absurd_elim a Hb Hnb))))
(iff_intro (assume Ha, absurd_elim b Ha Hna) (assume Hb, absurd_elim a Hb Hnb))))
theorem implies_decidable [instance] {a b : Prop} (Ha : decidable a) (Hb : decidable b) :
decidable (a → b) :=

View file

@ -5,7 +5,7 @@
import logic.connectives.basic
inductive inhabited (A : Type) : Type :=
| inhabited_mk : A → inhabited A
inhabited_mk : A → inhabited A
namespace inhabited

View file

@ -9,7 +9,7 @@ using inhabited
namespace nonempty
inductive nonempty (A : Type) : Prop :=
| nonempty_intro : A → nonempty A
nonempty_intro : A → nonempty A
definition nonempty_elim {A : Type} {B : Type} (H1 : nonempty A) (H2 : A → B) : B :=
nonempty_rec H2 H1

View file

@ -19,7 +19,7 @@ theorem false_elim (c : Prop) (H : false) : c :=
false_rec c H
inductive true : Prop :=
| trivial : true
trivial : true
abbreviation not (a : Prop) := a → false
prefix `¬`:40 := not
@ -63,7 +63,7 @@ assume Hnb Ha, Hnb (Hab Ha)
-- ---
inductive and (a b : Prop) : Prop :=
| and_intro : a → b → and a b
and_intro : a → b → and a b
infixr `/\`:35 := and
infixr `∧`:35 := and
@ -100,8 +100,8 @@ and_elim H1 (assume Hc : c, assume Ha : a, and_intro Hc (H Ha))
-- --
inductive or (a b : Prop) : Prop :=
| or_intro_left : a → or a b
| or_intro_right : b → or a b
or_intro_left : a → or a b,
or_intro_right : b → or a b
infixr `\/`:30 := or
infixr ``:30 := or

View file

@ -10,7 +10,7 @@ import .basic
-- --
inductive eq {A : Type} (a : A) : A → Prop :=
| refl : eq a a
refl : eq a a
infix `=`:50 := eq

View file

@ -7,7 +7,7 @@ import .basic .eq ..classes.nonempty
using inhabited nonempty
inductive Exists {A : Type} (P : A → Prop) : Prop :=
| exists_intro : ∀ (a : A), P a → Exists P
exists_intro : ∀ (a : A), P a → Exists P
notation `exists` binders `,` r:(scoped P, Exists P) := r
notation `∃` binders `,` r:(scoped P, Exists P) := r

View file

@ -16,7 +16,7 @@ abbreviation transitive {T : Type} (R : T → T → Type) : Type := ∀⦃x y z
inductive is_reflexive {T : Type} (R : T → T → Type) : Prop :=
| is_reflexive_mk : reflexive R → is_reflexive R
is_reflexive_mk : reflexive R → is_reflexive R
namespace is_reflexive
@ -30,7 +30,7 @@ end is_reflexive
inductive is_symmetric {T : Type} (R : T → T → Type) : Prop :=
| is_symmetric_mk : symmetric R → is_symmetric R
is_symmetric_mk : symmetric R → is_symmetric R
namespace is_symmetric
@ -44,7 +44,7 @@ end is_symmetric
inductive is_transitive {T : Type} (R : T → T → Type) : Prop :=
| is_transitive_mk : transitive R → is_transitive R
is_transitive_mk : transitive R → is_transitive R
namespace is_transitive
@ -58,7 +58,7 @@ end is_transitive
inductive is_equivalence {T : Type} (R : T → T → Type) : Prop :=
| is_equivalence_mk : is_reflexive R → is_symmetric R → is_transitive R → is_equivalence R
is_equivalence_mk : is_reflexive R → is_symmetric R → is_transitive R → is_equivalence R
namespace is_equivalence
@ -80,7 +80,7 @@ instance is_equivalence.is_transitive
-- partial equivalence relation
inductive is_PER {T : Type} (R : T → T → Type) : Prop :=
| is_PER_mk : is_symmetric R → is_transitive R → is_PER R
is_PER_mk : is_symmetric R → is_transitive R → is_PER R
namespace is_PER
@ -101,12 +101,12 @@ instance is_PER.is_transitive
inductive congr {T1 : Type} (R1 : T1 → T1 → Prop) {T2 : Type} (R2 : T2 → T2 → Prop)
(f : T1 → T2) : Prop :=
| congr_mk : (∀x y, R1 x y → R2 (f x) (f y)) → congr R1 R2 f
congr_mk : (∀x y, R1 x y → R2 (f x) (f y)) → congr R1 R2 f
-- for binary functions
inductive congr2 {T1 : Type} (R1 : T1 → T1 → Prop) {T2 : Type} (R2 : T2 → T2 → Prop)
{T3 : Type} (R3 : T3 → T3 → Prop) (f : T1 → T2 → T3) : Prop :=
| congr2_mk : (∀(x1 y1 : T1) (x2 y2 : T2), R1 x1 y1 → R2 x2 y2 → R3 (f x1 x2) (f y1 y2)) →
congr2_mk : (∀(x1 y1 : T1) (x2 y2 : T2), R1 x1 y1 → R2 x2 y2 → R3 (f x1 x2) (f y1 y2)) →
congr2 R1 R2 R3 f
namespace congr
@ -171,7 +171,7 @@ congr_mk (λx y H, H)
-- ---------------------------------------------------------
inductive mp_like {R : Type → Type → Prop} {a b : Type} (H : R a b) : Prop :=
| mp_like_mk {} : (a → b) → @mp_like R a b H
mp_like_mk {} : (a → b) → @mp_like R a b H
namespace mp_like

View file

@ -15,7 +15,7 @@ namespace tactic
-- builtin_tactic is just a "dummy" for creating the
-- definitions that are actually implemented in C++
inductive tactic : Type :=
| builtin_tactic : tactic
builtin_tactic : tactic
-- Remark the following names are not arbitrary, the tactic module
-- uses them when converting Lean expressions into actual tactic objects.
-- The bultin 'by' construct triggers the process of converting a

View file

@ -27,7 +27,7 @@ namespace lean {
static name g_assign(":=");
static name g_with("with");
static name g_colon(":");
static name g_bar("|");
static name g_comma(",");
static name g_lcurly("{");
static name g_rcurly("}");
@ -258,8 +258,7 @@ struct inductive_cmd_fn {
*/
list<intro_rule> parse_intro_rules(buffer<expr> & params) {
buffer<intro_rule> intros;
while (m_p.curr_is_token(g_bar)) {
m_p.next();
while (true) {
name intro_name = parse_intro_decl_name();
bool strict = true;
if (m_p.curr_is_token(g_lcurly)) {
@ -273,6 +272,9 @@ struct inductive_cmd_fn {
intro_type = Pi(params, intro_type, m_p);
intro_type = infer_implicit(intro_type, params.size(), strict);
intros.push_back(intro_rule(intro_name, intro_type));
if (!m_p.curr_is_token(g_comma))
break;
m_p.next();
}
return to_list(intros.begin(), intros.end());
}

View file

@ -17,10 +17,10 @@ end
namespace algebra
inductive mul_struct (A : Type) : Type :=
| mk_mul_struct : (A → A → A) → mul_struct A
mk_mul_struct : (A → A → A) → mul_struct A
inductive add_struct (A : Type) : Type :=
| mk_add_struct : (A → A → A) → add_struct A
mk_add_struct : (A → A → A) → add_struct A
definition mul [inline] {A : Type} {s : mul_struct A} (a b : A)
:= mul_struct_rec (fun f, f) s a b
@ -35,8 +35,8 @@ end algebra
namespace nat
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
variable add : nat → nat → nat
variable mul : nat → nat → nat
@ -55,7 +55,7 @@ end nat
namespace algebra
namespace semigroup
inductive semigroup_struct (A : Type) : Type :=
| mk_semigroup_struct : Π (mul : A → A → A), is_assoc mul → semigroup_struct A
mk_semigroup_struct : Π (mul : A → A → A), is_assoc mul → semigroup_struct A
definition mul [inline] {A : Type} (s : semigroup_struct A) (a b : A)
:= semigroup_struct_rec (fun f h, f) s a b
@ -67,7 +67,7 @@ namespace semigroup
:= mk_mul_struct (mul s)
inductive semigroup : Type :=
| mk_semigroup : Π (A : Type), semigroup_struct A → semigroup
mk_semigroup : Π (A : Type), semigroup_struct A → semigroup
definition carrier [inline] [coercion] (g : semigroup)
:= semigroup_rec (fun c s, c) g
@ -80,7 +80,7 @@ namespace monoid
check semigroup.mul
inductive monoid_struct (A : Type) : Type :=
| mk_monoid_struct : Π (mul : A → A → A) (id : A), is_assoc mul → is_id mul id → monoid_struct A
mk_monoid_struct : Π (mul : A → A → A) (id : A), is_assoc mul → is_id mul id → monoid_struct A
definition mul [inline] {A : Type} (s : monoid_struct A) (a b : A)
:= monoid_struct_rec (fun mul id a i, mul) s a b
@ -93,7 +93,7 @@ namespace monoid
:= mk_semigroup_struct (mul s) (assoc s)
inductive monoid : Type :=
| mk_monoid : Π (A : Type), monoid_struct A → monoid
mk_monoid : Π (A : Type), monoid_struct A → monoid
definition carrier [inline] [coercion] (m : monoid)
:= monoid_rec (fun c s, c) m

View file

@ -17,8 +17,8 @@ namespace N2
section
parameter A : Type
inductive list : Type :=
| nil {} : list
| cons : A → list → list
nil {} : list,
cons : A → list → list
check list
end
check list

View file

@ -1,8 +1,8 @@
import standard
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
definition add (x y : nat)
:= nat_rec x (λ n r, succ r) y
@ -54,7 +54,7 @@ theorem not_zero_add (x y : nat) (H : ¬ is_zero y) : ¬ is_zero (x + y)
subst (symm H1) H2)
inductive not_zero (x : nat) : Prop :=
| not_zero_intro : ¬ is_zero x → not_zero x
not_zero_intro : ¬ is_zero x → not_zero x
theorem not_zero_not_is_zero {x : nat} (H : not_zero x) : ¬ is_zero x
:= not_zero_rec (λ H1, H1) H

View file

@ -2,7 +2,7 @@ import standard
namespace algebra
inductive mul_struct (A : Type) : Type :=
| mk_mul_struct : (A → A → A) → mul_struct A
mk_mul_struct : (A → A → A) → mul_struct A
definition mul [inline] {A : Type} {s : mul_struct A} (a b : A)
:= mul_struct_rec (λ f, f) s a b
@ -12,8 +12,8 @@ end algebra
namespace nat
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
variable mul : nat → nat → nat
variable add : nat → nat → nat

View file

@ -2,10 +2,10 @@ import standard
using prod
inductive t1 : Type :=
| mk1 : t1
mk1 : t1
inductive t2 : Type :=
| mk2 : t2
mk2 : t2
theorem inhabited_t1 : inhabited t1
:= inhabited_mk mk1
@ -17,4 +17,3 @@ instance inhabited_t1 inhabited_t2
theorem T : inhabited (t1 × t2)
:= _

View file

@ -2,7 +2,7 @@ import standard
using num tactic
inductive inh (A : Type) : Type :=
| inh_intro : A -> inh A
inh_intro : A -> inh A
instance inh_intro

View file

@ -2,7 +2,7 @@ import standard
using num tactic prod
inductive inh (A : Type) : Prop :=
| inh_intro : A -> inh A
inh_intro : A -> inh A
instance inh_intro

View file

@ -7,7 +7,7 @@ variable int_add : int → int → int
variable real_add : real → real → real
inductive add_struct (A : Type) :=
| mk : (A → A → A) → add_struct A
mk : (A → A → A) → add_struct A
definition add {A : Type} {S : add_struct A} (a b : A) : A :=
add_struct_rec (λ m, m) S a b
@ -51,4 +51,3 @@ infixl `=`:50 := eq
abbreviation id (A : Type) (a : A) := a
notation A `=` B `:` C := @eq C A B
check nat_to_int n + nat_to_int m = (n + m) : int

View file

@ -18,7 +18,7 @@ set_option pp.universes true
check eq a1 b1
inductive pair (A : Type) (B: Type) : Type :=
| mk_pair : A → B → pair A B
mk_pair : A → B → pair A B
check mk_pair a1 b2
check B

View file

@ -2,7 +2,7 @@ import standard
namespace setoid
inductive setoid : Type :=
| mk_setoid: Π (A : Type), (A → A → Prop) → setoid
mk_setoid: Π (A : Type), (A → A → Prop) → setoid
definition carrier (s : setoid)
:= setoid_rec (λ a eq, a) s
@ -15,6 +15,6 @@ namespace setoid
coercion carrier
inductive morphism (s1 s2 : setoid) : Type :=
| mk_morphism : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism s1 s2
mk_morphism : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism s1 s2
end setoid

View file

@ -2,7 +2,7 @@ import standard
namespace setoid
inductive setoid : Type :=
| mk_setoid: Π (A : Type), (A → A → Prop) → setoid
mk_setoid: Π (A : Type), (A → A → Prop) → setoid
definition carrier (s : setoid)
:= setoid_rec (λ a eq, a) s
@ -15,7 +15,7 @@ namespace setoid
coercion carrier
inductive morphism (s1 s2 : setoid) : Type :=
| mk_morphism : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism s1 s2
mk_morphism : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism s1 s2
set_option pp.universes true
@ -24,7 +24,7 @@ namespace setoid
check λ (s1 s2 : Type), s1
inductive morphism2 (s1 : setoid) (s2 : setoid) : Type :=
| mk_morphism2 : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism2 s1 s2
mk_morphism2 : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism2 s1 s2
check mk_morphism2
end setoid

View file

@ -2,7 +2,7 @@ import standard
namespace setoid
inductive setoid : Type :=
| mk_setoid: Π (A : Type'), (A → A → Prop) → setoid
mk_setoid: Π (A : Type'), (A → A → Prop) → setoid
set_option pp.universes true
@ -20,14 +20,14 @@ namespace setoid
coercion carrier
inductive morphism (s1 s2 : setoid) : Type :=
| mk_morphism : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism s1 s2
mk_morphism : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism s1 s2
check mk_morphism
check λ (s1 s2 : setoid), s1
check λ (s1 s2 : Type), s1
inductive morphism2 (s1 : setoid) (s2 : setoid) : Type :=
| mk_morphism2 : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism2 s1 s2
mk_morphism2 : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism2 s1 s2
check morphism2
check mk_morphism2

View file

@ -2,7 +2,7 @@ import standard
namespace setoid
inductive setoid : Type :=
| mk_setoid: Π (A : Type'), (A → A → Prop) → setoid
mk_setoid: Π (A : Type'), (A → A → Prop) → setoid
set_option pp.universes true
@ -20,20 +20,20 @@ namespace setoid
coercion carrier
inductive morphism (s1 s2 : setoid) : Type :=
| mk_morphism : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism s1 s2
mk_morphism : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism s1 s2
check mk_morphism
check λ (s1 s2 : setoid), s1
check λ (s1 s2 : Type), s1
inductive morphism2 (s1 : setoid) (s2 : setoid) : Type :=
| mk_morphism2 : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism2 s1 s2
mk_morphism2 : Π (f : s1 → s2), (∀ x y, x ≈ y → f x ≈ f y) → morphism2 s1 s2
check morphism2
check mk_morphism2
inductive my_struct : Type :=
| mk_foo : Π (s1 s2 : setoid) (s3 s4 : setoid), morphism2 s1 s2 → morphism2 s3 s4 → my_struct
mk_foo : Π (s1 s2 : setoid) (s3 s4 : setoid), morphism2 s1 s2 → morphism2 s3 s4 → my_struct
check my_struct
definition tst2 : Type.{4} := my_struct.{1 2 1 2}

View file

@ -2,10 +2,9 @@ import data.nat
using nat
inductive list (T : Type) : Type :=
| nil {} : list T
| cons : T → list T → list T
nil {} : list T,
cons : T → list T → list T
definition length {T : Type} : list T → nat := list_rec 0 (fun x l m, succ m)
theorem length_nil {T : Type} : length (@nil T) = 0
:= refl _

View file

@ -10,7 +10,7 @@ namespace congr
inductive struc {T1 : Type} {T2 : Type} (R1 : T1 → T1 → Prop) (R2 : T2 → T2 → Prop)
(f : T1 → T2) : Prop :=
| mk : (∀x y : T1, R1 x y → R2 (f x) (f y)) → struc R1 R2 f
mk : (∀x y : T1, R1 x y → R2 (f x) (f y)) → struc R1 R2 f
abbreviation app {T1 : Type} {T2 : Type} {R1 : T1 → T1 → Prop} {R2 : T2 → T2 → Prop}
{f : T1 → T2} (C : struc R1 R2 f) {x y : T1} : R1 x y → R2 (f x) (f y) :=
@ -18,7 +18,7 @@ struc_rec id C x y
inductive struc2 {T1 : Type} {T2 : Type} {T3 : Type} (R1 : T1 → T1 → Prop)
(R2 : T2 → T2 → Prop) (R3 : T3 → T3 → Prop) (f : T1 → T2 → T3) : Prop :=
| mk2 : (∀(x1 y1 : T1) (x2 y2 : T2), R1 x1 y1 → R2 x2 y2 → R3 (f x1 x2) (f y1 y2)) →
mk2 : (∀(x1 y1 : T1) (x2 y2 : T2), R1 x1 y1 → R2 x2 y2 → R3 (f x1 x2) (f y1 y2)) →
struc2 R1 R2 R3 f
abbreviation app2 {T1 : Type} {T2 : Type} {T3 : Type} {R1 : T1 → T1 → Prop}

View file

@ -1,10 +1,10 @@
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
inductive list (A : Type) : Type :=
| nil {} : list A
| cons : A → list A → list A
nil {} : list A,
cons : A → list A → list A
check nil
check nil.{1}
@ -14,8 +14,8 @@ check @nil nat
check cons zero nil
inductive vector (A : Type) : nat → Type :=
| vnil {} : vector A zero
| vcons : forall {n : nat}, A → vector A n → vector A (succ n)
vnil {} : vector A zero,
vcons : forall {n : nat}, A → vector A n → vector A (succ n)
check vcons zero vnil
variable n : nat

View file

@ -1,10 +1,10 @@
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
inductive list (A : Type) : Type :=
| nil {} : list A
| cons : A → list A → list A
nil {} : list A,
cons : A → list A → list A
check nil
check nil.{1}
@ -14,8 +14,8 @@ check @nil nat
check cons zero nil
inductive vector (A : Type) : nat → Type :=
| vnil {} : vector A zero
| vcons : forall {n : nat}, A → vector A n → vector A (succ n)
vnil {} : vector A zero,
vcons : forall {n : nat}, A → vector A n → vector A (succ n)
check vcons zero vnil
variable n : nat

View file

@ -1,10 +1,10 @@
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
inductive list (A : Type) : Type :=
| nil {} : list A
| cons : A → list A → list A
nil {} : list A,
cons : A → list A → list A
check nil
check nil.{1}
@ -14,8 +14,8 @@ check @nil nat
check cons zero nil
inductive vector (A : Type) : nat → Type :=
| vnil {} : vector A zero
| vcons : forall {n : nat}, A → vector A n → vector A (succ n)
vnil {} : vector A zero,
vcons : forall {n : nat}, A → vector A n → vector A (succ n)
check vcons zero vnil
variable n : nat

View file

@ -1,14 +1,14 @@
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
inductive list (A : Type) : Type :=
| nil {} : list A
| cons : A → list A → list A
nil {} : list A,
cons : A → list A → list A
inductive int : Type :=
| of_nat : nat → int
| neg : nat → int
of_nat : nat → int,
neg : nat → int
coercion of_nat

View file

@ -1,14 +1,14 @@
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
inductive list (A : Type) : Type :=
| nil {} : list A
| cons : A → list A → list A
nil {} : list A,
cons : A → list A → list A
inductive int : Type :=
| of_nat : nat → int
| neg : nat → int
of_nat : nat → int,
neg : nat → int
coercion of_nat

View file

@ -32,8 +32,8 @@ theorem trans {A : Type} {a b c : A} (H1 : a = b) (H2 : b = c) : a = c
:= subst H2 H1
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
print "using strict implicit arguments"
abbreviation symmetric {A : Type} (R : A → A → Prop) := ∀ ⦃a b⦄, R a b → R b a
@ -66,4 +66,3 @@ axiom H5 : symmetric3 p
axiom H6 : p zero (succ zero)
check H5
check H5 H6

View file

@ -18,7 +18,7 @@ abbreviation reflexive {T : Type} (R : T → T → Type) : Prop := ∀x, R x x
inductive congruence {T1 : Type} {T2 : Type} (R1 : T1 → T1 → Prop) (R2 : T2 → T2 → Prop)
(f : T1 → T2) : Prop :=
| mk : (∀x y : T1, R1 x y → R2 (f x) (f y)) → congruence R1 R2 f
mk : (∀x y : T1, R1 x y → R2 (f x) (f y)) → congruence R1 R2 f
-- to trigger class inference
theorem congr_app {T1 : Type} {T2 : Type} (R1 : T1 → T1 → Prop) (R2 : T2 → T2 → Prop)

View file

@ -14,10 +14,10 @@ section
end
inductive group_struct (A : Type) : Type :=
| mk_group_struct : Π (mul : A → A → A) (one : A) (inv : A → A), is_assoc mul → is_id mul one → is_inv mul one inv → group_struct A
mk_group_struct : Π (mul : A → A → A) (one : A) (inv : A → A), is_assoc mul → is_id mul one → is_inv mul one inv → group_struct A
inductive group : Type :=
| mk_group : Π (A : Type), group_struct A → group
mk_group : Π (A : Type), group_struct A → group
definition carrier (g : group) : Type
:= group_rec (λ c s, c) g

View file

@ -14,10 +14,10 @@ section
end
inductive group_struct (A : Type) : Type :=
| mk_group_struct : Π (mul : A → A → A) (one : A) (inv : A → A), is_assoc mul → is_id mul one → is_inv mul one inv → group_struct A
mk_group_struct : Π (mul : A → A → A) (one : A) (inv : A → A), is_assoc mul → is_id mul one → is_inv mul one inv → group_struct A
inductive group : Type :=
| mk_group : Π (A : Type), group_struct A → group
mk_group : Π (A : Type), group_struct A → group
definition carrier (g : group) : Type
:= group_rec (λ c s, c) g

View file

@ -1,6 +1,6 @@
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
check nat
check nat_rec.{1}

View file

@ -1,6 +1,6 @@
inductive list (A : Type) : Type :=
| nil : list A
| cons : A → list A → list A
nil : list A,
cons : A → list A → list A
check list.{1}
check cons.{1}

View file

@ -1,10 +1,10 @@
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
inductive vector (A : Type) : nat → Type :=
| vnil : vector A zero
| vcons : Π {n : nat}, A → vector A n → vector A (succ n)
vnil : vector A zero,
vcons : Π {n : nat}, A → vector A n → vector A (succ n)
check vector.{1}
check vnil.{1}

View file

@ -1,8 +1,8 @@
inductive tree (A : Type) : Type :=
| node : A → forest A → tree A
node : A → forest A → tree A
with forest (A : Type) : Type :=
| nil : forest A
| cons : tree A → forest A → forest A
nil : forest A,
cons : tree A → forest A → forest A
check tree.{1}
check forest.{1}
@ -12,7 +12,7 @@ check forest_rec.{1 1}
print "==============="
inductive group : Type :=
| mk_group : Π (carrier : Type) (mul : carrier → carrier → carrier) (one : carrier), group
mk_group : Π (carrier : Type) (mul : carrier → carrier → carrier) (one : carrier), group
check group.{1}
check group.{2}

View file

@ -1,8 +1,8 @@
section
parameter A : Type
inductive list : Type :=
| nil : list
| cons : A → list → list
nil : list,
cons : A → list → list
end
check list.{1}
@ -11,10 +11,10 @@ check cons.{1}
section
parameter A : Type
inductive tree : Type :=
| node : A → forest → tree
node : A → forest → tree
with forest : Type :=
| fnil : forest
| fcons : tree → forest → forest
fnil : forest,
fcons : tree → forest → forest
check tree
check forest
end

View file

@ -1,8 +1,8 @@
definition Prop [inline] : Type.{1} := Type.{0}
inductive or (A B : Prop) : Prop :=
| or_intro_left : A → or A B
| or_intro_right : B → or A B
or_intro_left : A → or A B,
or_intro_right : B → or A B
check or
check or_intro_left

View file

@ -1,8 +1,8 @@
inductive tree.{u} (A : Type.{u}) : Type.{max u 1} :=
| node : A → forest.{u} A → tree.{u} A
node : A → forest.{u} A → tree.{u} A
with forest.{u} (A : Type.{u}) : Type.{max u 1} :=
| nil : forest.{u} A
| cons : tree.{u} A → forest.{u} A → forest.{u} A
nil : forest.{u} A,
cons : tree.{u} A → forest.{u} A → forest.{u} A
check tree.{1}
check forest.{1}

View file

@ -1,7 +1,7 @@
namespace list
inductive list (A : Type) : Type :=
| nil : list A
| cons : A → list A → list A
nil : list A,
cons : A → list A → list A
check list.{1}
check cons.{1}

View file

@ -1,21 +1,19 @@
abbreviation Prop := Type.{0}
inductive nat :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
inductive list (A : Type) :=
| nil {} : list A
| cons : A → list A → list A
nil {} : list A,
cons : A → list A → list A
inductive list2 (A : Type) : Type :=
| nil2 {} : list2 A
| cons2 : A → list2 A → list2 A
nil2 {} : list2 A,
cons2 : A → list2 A → list2 A
inductive and (A B : Prop) : Prop :=
| and_intro : A → B → and A B
and_intro : A → B → and A B
inductive class {T1 : Type} (R1 : T1 → T1 → Prop) {T2 : Type} (R2 : T2 → T2 → Prop) (f : T1 → T2) :=
| mk : (∀x y : T1, R1 x y → R2 (f x) (f y)) → class R1 R2 f
mk : (∀x y : T1, R1 x y → R2 (f x) (f y)) → class R1 R2 f

View file

@ -1,45 +1,45 @@
inductive list (A : Type) : Type :=
| nil {} : list A
| cons : A → list A → list A
nil {} : list A,
cons : A → list A → list A
section
parameter A : Type
inductive list2 : Type :=
| nil2 {} : list2
| cons2 : A → list2 → list2
nil2 {} : list2,
cons2 : A → list2 → list2
end
variable num : Type.{1}
namespace Tree
inductive tree (A : Type) : Type :=
| node : A → forest A → tree A
node : A → forest A → tree A
with forest (A : Type) : Type :=
| nil : forest A
| cons : tree A → forest A → forest A
nil : forest A,
cons : tree A → forest A → forest A
end Tree
inductive group_struct (A : Type) : Type :=
| mk_group_struct : (A → A → A) → A → group_struct A
mk_group_struct : (A → A → A) → A → group_struct A
inductive group : Type :=
| mk_group : Π (A : Type), (A → A → A) → A → group
mk_group : Π (A : Type), (A → A → A) → A → group
section
parameter A : Type
parameter B : Type
inductive pair : Type :=
| mk_pair : A → B → pair
mk_pair : A → B → pair
end
definition Prop [inline] := Type.{0}
inductive eq {A : Type} (a : A) : A → Prop :=
| refl : eq a a
refl : eq a a
section
parameter {A : Type}
inductive eq2 (a : A) : A → Prop :=
| refl2 : eq2 a a
refl2 : eq2 a a
end
@ -47,5 +47,5 @@ section
parameter A : Type
parameter B : Type
inductive triple (C : Type) : Type :=
| mk_triple : A → B → C → triple C
mk_triple : A → B → C → triple C
end

View file

@ -2,8 +2,8 @@ import standard
using tactic
inductive list (A : Type) : Type :=
| nil {} : list A
| cons : A → list A → list A
nil {} : list A,
cons : A → list A → list A
definition is_nil {A : Type} (l : list A) : Prop
:= list_rec true (fun h t r, false) l

View file

@ -12,8 +12,8 @@
import data.nat
using nat eq_ops
inductive list (T : Type) : Type :=
| nil {} : list T
| cons : T → list T → list T
nil {} : list T,
cons : T → list T → list T
theorem list_induction_on {T : Type} {P : list T → Prop} (l : list T) (Hnil : P nil)
(Hind : forall x : T, forall l : list T, forall H : P l, P (cons x l)) : P l :=

View file

@ -2,8 +2,8 @@ import logic
using decidable
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
theorem induction_on {P : nat → Prop} (a : nat) (H1 : P zero) (H2 : ∀ (n : nat) (IH : P n), P (succ n)) : P a
:= nat_rec H1 H2 a

View file

@ -2,8 +2,8 @@ import standard
using num eq_ops
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
abbreviation plus (x y : nat) : nat
:= nat_rec x (λn r, succ r) y

View file

@ -2,8 +2,8 @@ import standard
using num eq_ops
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
abbreviation plus (x y : nat) : nat
:= nat_rec x (λn r, succ r) y

View file

@ -2,8 +2,8 @@ import standard
using num eq_ops
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
definition add (x y : nat) : nat := nat_rec x (λn r, succ r) y
infixl `+`:65 := add
@ -14,4 +14,3 @@ axiom mul_succ_right (n m : nat) : n * succ m = n * m + n
theorem small2 (n m l : nat) : n * succ l + m = n * l + n + m
:= subst (mul_succ_right _ _) (refl (n * succ l + m))

View file

@ -2,8 +2,8 @@ import standard
using num eq_ops
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
definition add (x y : nat) : nat := nat_rec x (λn r, succ r) y
infixl `+`:65 := add

View file

@ -2,8 +2,8 @@ import standard
using eq_ops
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
definition add (x y : nat) : nat := nat_rec x (λn r, succ r) y
infixl `+`:65 := add

View file

@ -1,8 +1,8 @@
import logic
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
definition add (x y : nat) : nat := nat_rec x (λn r, succ r) y
infixl `+`:65 := add
@ -12,4 +12,3 @@ axiom add_right_comm (n m k : nat) : n + m + k = n + k + m
print "==========================="
theorem bug (a b c d : nat) : a + b + c + d = a + c + b + d
:= subst (add_right_comm _ _ _) (refl (a + b + c + d))

View file

@ -1,8 +1,6 @@
inductive list (T : Type) : Type :=
| nil {} : list T
| cons : T → list T → list T
nil {} : list T,
cons : T → list T → list T
section

View file

@ -4,7 +4,7 @@ using relation
namespace is_equivalence
inductive class {T : Type} (R : T → T → Type) : Prop :=
| mk : is_reflexive R → is_symmetric R → is_transitive R → class R
mk : is_reflexive R → is_symmetric R → is_transitive R → class R
theorem is_reflexive {T : Type} {R : T → T → Type} {C : class R} : is_reflexive R :=
class_rec (λx y z, x) C

View file

@ -10,8 +10,8 @@ namespace sum
-- TODO: take this outside the namespace when the inductive package handles it better
inductive sum (A B : Type) : Type :=
| inl : A → sum A B
| inr : B → sum A B
inl : A → sum A B,
inr : B → sum A B
infixr `+`:25 := sum

View file

@ -3,8 +3,8 @@ using num (num pos_num num_rec pos_num_rec)
using tactic
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
definition add [inline] (a b : nat) : nat
:= nat_rec a (λ n r, succ r) b

View file

@ -2,8 +2,8 @@ import standard
using tactic inhabited
inductive sum (A : Type) (B : Type) : Type :=
| inl : A → sum A B
| inr : B → sum A B
inl : A → sum A B,
inr : B → sum A B
theorem inl_inhabited {A : Type} (B : Type) (H : inhabited A) : inhabited (sum A B)
:= inhabited_destruct H (λ a, inhabited_mk (inl B a))

View file

@ -2,8 +2,8 @@ import standard
using tactic inhabited
inductive sum (A : Type) (B : Type) : Type :=
| inl : A → sum A B
| inr : B → sum A B
inl : A → sum A B,
inr : B → sum A B
theorem inl_inhabited {A : Type} (B : Type) (H : inhabited A) : inhabited (sum A B)
:= inhabited_destruct H (λ a, inhabited_mk (inl B a))

View file

@ -1,8 +1,8 @@
import logic
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
check @nat_rec

View file

@ -1,8 +1,8 @@
import logic
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
variable f : nat → nat

View file

@ -1,8 +1,8 @@
import standard
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
definition is_zero (n : nat)
:= nat_rec true (λ n r, false) n

View file

@ -13,7 +13,7 @@ namespace simp
-- first define a class of homogeneous equality
inductive simplifies_to {T : Type} (t1 t2 : T) : Prop :=
| mk : t1 = t2 → simplifies_to t1 t2
mk : t1 = t2 → simplifies_to t1 t2
theorem get_eq {T : Type} {t1 t2 : T} (C : simplifies_to t1 t2) : t1 = t2 :=
simplifies_to_rec (λx, x) C

View file

@ -10,7 +10,7 @@ using nat
namespace simp
-- first define a class of homogeneous equality
inductive simplifies_to {T : Type} (t1 t2 : T) : Prop :=
| mk : t1 = t2 → simplifies_to t1 t2
mk : t1 = t2 → simplifies_to t1 t2
theorem get_eq {T : Type} {t1 t2 : T} (C : simplifies_to t1 t2) : t1 = t2 :=
simplifies_to_rec (λx, x) C

View file

@ -4,17 +4,17 @@ notation `take` binders `,` r:(scoped f, f) := r
inductive empty : Type
inductive unit : Type :=
| tt : unit
tt : unit
inductive nat : Type :=
| O : nat
| S : nat → nat
O : nat,
S : nat → nat
inductive paths {A : Type} (a : A) : A → Type :=
| idpath : paths a a
idpath : paths a a
inductive sum (A : Type) (B : Type) : Type :=
| inl : A -> sum A B
| inr : B -> sum A B
inl : A -> sum A B,
inr : B -> sum A B
definition coprod := sum
definition ii1fun {A : Type} (B : Type) (a : A) := inl B a
@ -23,7 +23,7 @@ definition ii1 {A : Type} {B : Type} (a : A) := inl B a
definition ii2 {A : Type} {B : Type} (b : B) := inl A b
inductive total2 {T: Type} (P: T → Type) : Type :=
| tpair : Π (t : T) (tp : P t), total2 P
tpair : Π (t : T) (tp : P t), total2 P
definition pr1 {T : Type} {P : T → Type} (tp : total2 P) : T
:= total2_rec (λ a b, a) tp
@ -31,7 +31,7 @@ definition pr2 {T : Type} {P : T → Type} (tp : total2 P) : P (pr1 tp)
:= total2_rec (λ a b, b) tp
inductive Phant (T : Type) : Type :=
| phant : Phant T
phant : Phant T
definition fromempty {X : Type} : empty → X
:= λe, empty_rec (λe, X) e

View file

@ -22,8 +22,8 @@ namespace list
-- ----
inductive list (T : Type) : Type :=
| nil {} : list T
| cons : T → list T → list T
nil {} : list T,
cons : T → list T → list T
infix `::` : 65 := cons

View file

@ -9,8 +9,8 @@ using decidable
namespace nat
inductive nat : Type :=
| zero : nat
| succ : nat → nat
zero : nat,
succ : nat → nat
notation ``:max := nat

View file

@ -14,7 +14,7 @@ infixl `∘`:60 := compose
-- ----
inductive path {A : Type} (a : A) : A → Type :=
| idpath : path a a
idpath : path a a
infix `≈`:50 := path
-- TODO: is this right?