feat(frontends/lean/inductive_cmd): allow '|' in inductive datatype declarations

This commit is contained in:
Leonardo de Moura 2015-02-25 17:00:10 -08:00
parent 5ca52d81ec
commit 68110faa4d
127 changed files with 345 additions and 339 deletions

View file

@ -36,8 +36,8 @@ structure prod (A B : Type) :=
mk :: (pr1 : A) (pr2 : B) mk :: (pr1 : A) (pr2 : B)
inductive sum (A B : Type) : Type := inductive sum (A B : Type) : Type :=
inl {} : A → sum A B, | inl {} : A → sum A B
inr {} : B → sum A B | inr {} : B → sum A B
definition sum.intro_left [reducible] {A : Type} (B : Type) (a : A) : sum A B := definition sum.intro_left [reducible] {A : Type} (B : Type) (a : A) : sum A B :=
sum.inl a sum.inl a
@ -49,29 +49,29 @@ sum.inr b
-- The parser will generate the terms (pos (bit1 (bit1 (bit0 one)))), zero, and (pos (bit0 (bit1 (bit1 one)))). -- 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). -- This representation can be coerced in whatever we want (e.g., naturals, integers, reals, etc).
inductive pos_num : Type := inductive pos_num : Type :=
one : pos_num, | one : pos_num
bit1 : pos_num → pos_num, | bit1 : pos_num → pos_num
bit0 : pos_num → pos_num | bit0 : pos_num → pos_num
inductive num : Type := inductive num : Type :=
zero : num, | zero : num
pos : pos_num → num | pos : pos_num → num
inductive bool : Type := inductive bool : Type :=
ff : bool, | ff : bool
tt : bool | tt : bool
inductive char : Type := inductive char : Type :=
mk : bool → bool → bool → bool → bool → bool → bool → bool → char mk : bool → bool → bool → bool → bool → bool → bool → bool → char
inductive string : Type := inductive string : Type :=
empty : string, | empty : string
str : char → string → string | str : char → string → string
inductive nat := inductive nat :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
inductive option (A : Type) : Type := inductive option (A : Type) : Type :=
none {} : option A, | none {} : option A
some : A → option A | some : A → option A

View file

@ -199,8 +199,8 @@ end inhabited
/- decidable -/ /- decidable -/
inductive decidable.{l} [class] (p : Type.{l}) : Type.{l} := inductive decidable.{l} [class] (p : Type.{l}) : Type.{l} :=
inl : p → decidable p, | inl : p → decidable p
inr : ¬p → decidable p | inr : ¬p → decidable p
namespace decidable namespace decidable
variables {p q : Type} variables {p q : Type}

View file

@ -14,8 +14,8 @@ namespace nat
notation `` := nat notation `` := nat
inductive lt (a : nat) : nat → Type := inductive lt (a : nat) : nat → Type :=
base : lt a (succ a), | base : lt a (succ a)
step : Π {b}, lt a b → lt a (succ b) | step : Π {b}, lt a b → lt a (succ b)
notation a < b := lt a b notation a < b := lt a b

View file

@ -38,5 +38,5 @@ definition inv_image.irreflexive (f : A → B) (H : irreflexive R) : irreflexive
λ (a : A) (H₁ : inv_image R f a a), H (f a) H₁ λ (a : A) (H₁ : inv_image R f a a), H (f a) H₁
inductive tc {A : Type} (R : A → A → Type) : A → A → Type := inductive tc {A : Type} (R : A → A → Type) : A → A → Type :=
base : ∀a b, R a b → tc R a b, | base : ∀a b, R a b → tc R a b
trans : ∀a b c, tc R a b → tc R b c → tc R a c | trans : ∀a b c, tc R a b → tc R b c → tc R a c

View file

@ -67,8 +67,8 @@ opaque definition inversion (id : expr) : tactic := builtin
notation a `↦` b:max := rename a b notation a `↦` b:max := rename a b
inductive expr_list : Type := inductive expr_list : Type :=
nil : expr_list, | nil : expr_list
cons : expr → expr_list → expr_list | cons : expr → expr_list → expr_list
-- rewrite_tac is just a marker for the builtin 'rewrite' notation -- rewrite_tac is just a marker for the builtin 'rewrite' notation
-- used to create instances of this tactic. -- used to create instances of this tactic.

View file

@ -21,8 +21,8 @@ open eq nat sigma unit
namespace is_trunc namespace is_trunc
inductive trunc_index : Type₁ := inductive trunc_index : Type₁ :=
minus_two : trunc_index, | minus_two : trunc_index
succ : trunc_index → trunc_index | succ : trunc_index → trunc_index
/- /-
notation for trunc_index is -2, -1, 0, 1, ... notation for trunc_index is -2, -1, 0, 1, ...

View file

@ -47,8 +47,8 @@ namespace prod
-- Lexicographical order based on Ra and Rb -- Lexicographical order based on Ra and Rb
inductive lex : A × B → A × B → Type := inductive lex : A × B → A × B → Type :=
left : ∀{a₁ b₁} a₂ b₂, Ra a₁ a₂ → lex (a₁, b₁) (a₂, b₂), | left : ∀{a₁ b₁} a₂ b₂, Ra a₁ a₂ → lex (a₁, b₁) (a₂, b₂)
right : ∀a {b₁ b₂}, Rb b₁ b₂ → lex (a, b₁) (a, b₂) | right : ∀a {b₁ b₂}, Rb b₁ b₂ → lex (a, b₁) (a, b₂)
-- Relational product based on Ra and Rb -- Relational product based on Ra and Rb
inductive rprod : A × B → A × B → Type := inductive rprod : A × B → A × B → Type :=

View file

@ -11,8 +11,8 @@ import data.nat
open nat open nat
inductive fin : nat → Type := inductive fin : nat → Type :=
fz : Π n, fin (succ n), | fz : Π n, fin (succ n)
fs : Π {n}, fin n → fin (succ n) | fs : Π {n}, fin n → fin (succ n)
namespace fin namespace fin
definition to_nat : Π {n}, fin n → nat definition to_nat : Π {n}, fin n → nat

View file

@ -37,8 +37,8 @@ open decidable binary fake_simplifier
/- the type of integers -/ /- the type of integers -/
inductive int : Type := inductive int : Type :=
of_nat : nat → int, | of_nat : nat → int
neg_succ_of_nat : nat → int | neg_succ_of_nat : nat → int
notation `` := int notation `` := int
attribute int.of_nat [coercion] attribute int.of_nat [coercion]

View file

@ -13,8 +13,8 @@ import logic tools.helper_tactics data.nat.basic
open eq.ops helper_tactics nat open eq.ops helper_tactics nat
inductive list (T : Type) : Type := inductive list (T : Type) : Type :=
nil {} : list T, | nil {} : list T
cons : T → list T → list T | cons : T → list T → list T
namespace list namespace list
notation h :: t := cons h t notation h :: t := cons h t

View file

@ -9,8 +9,8 @@ import data.nat.basic
open nat prod open nat prod
inductive vector (A : Type) : nat → Type := inductive vector (A : Type) : nat → Type :=
nil {} : vector A zero, | nil {} : vector A zero
cons : Π {n}, A → vector A n → vector A (succ n) | cons : Π {n}, A → vector A n → vector A (succ n)
namespace vector namespace vector
notation a :: b := cons a b notation a :: b := cons a b

View file

@ -50,8 +50,8 @@ and.rec (λa b, b) H
definition and.right := @and.elim_right definition and.right := @and.elim_right
inductive sum (A B : Type) : Type := inductive sum (A B : Type) : Type :=
inl {} : A → sum A B, | inl {} : A → sum A B
inr {} : B → sum A B | inr {} : B → sum A B
definition sum.intro_left [reducible] {A : Type} (B : Type) (a : A) : sum A B := definition sum.intro_left [reducible] {A : Type} (B : Type) (a : A) : sum A B :=
sum.inl a sum.inl a
@ -60,8 +60,8 @@ definition sum.intro_right [reducible] (A : Type) {B : Type} (b : B) : sum A B :
sum.inr b sum.inr b
inductive or (a b : Prop) : Prop := inductive or (a b : Prop) : Prop :=
inl {} : a → or a b, | inl {} : a → or a b
inr {} : b → or a b | inr {} : b → or a b
definition or.intro_left {a : Prop} (b : Prop) (Ha : a) : or a b := definition or.intro_left {a : Prop} (b : Prop) (Ha : a) : or a b :=
or.inl Ha or.inl Ha
@ -76,9 +76,9 @@ mk :: (pr1 : A) (pr2 : B pr1)
-- The parser will generate the terms (pos (bit1 (bit1 (bit0 one)))), zero, and (pos (bit0 (bit1 (bit1 one)))). -- 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). -- This representation can be coerced in whatever we want (e.g., naturals, integers, reals, etc).
inductive pos_num : Type := inductive pos_num : Type :=
one : pos_num, | one : pos_num
bit1 : pos_num → pos_num, | bit1 : pos_num → pos_num
bit0 : pos_num → pos_num | bit0 : pos_num → pos_num
namespace pos_num namespace pos_num
definition succ (a : pos_num) : pos_num := definition succ (a : pos_num) : pos_num :=
@ -86,8 +86,8 @@ namespace pos_num
end pos_num end pos_num
inductive num : Type := inductive num : Type :=
zero : num, | zero : num
pos : pos_num → num | pos : pos_num → num
namespace num namespace num
open pos_num open pos_num
@ -96,20 +96,20 @@ namespace num
end num end num
inductive bool : Type := inductive bool : Type :=
ff : bool, | ff : bool
tt : bool | tt : bool
inductive char : Type := inductive char : Type :=
mk : bool → bool → bool → bool → bool → bool → bool → bool → char mk : bool → bool → bool → bool → bool → bool → bool → bool → char
inductive string : Type := inductive string : Type :=
empty : string, | empty : string
str : char → string → string | str : char → string → string
inductive nat := inductive nat :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
inductive option (A : Type) : Type := inductive option (A : Type) : Type :=
none {} : option A, | none {} : option A
some : A → option A | some : A → option A

View file

@ -241,8 +241,8 @@ theorem exists.elim {A : Type} {p : A → Prop} {B : Prop} (H1 : ∃x, p x) (H2
Exists.rec H2 H1 Exists.rec H2 H1
inductive decidable [class] (p : Prop) : Type := inductive decidable [class] (p : Prop) : Type :=
inl : p → decidable p, | inl : p → decidable p
inr : ¬p → decidable p | inr : ¬p → decidable p
definition decidable_true [instance] : decidable true := definition decidable_true [instance] : decidable true :=
decidable.inl trivial decidable.inl trivial

View file

@ -14,8 +14,8 @@ namespace nat
notation `` := nat notation `` := nat
inductive lt (a : nat) : nat → Prop := inductive lt (a : nat) : nat → Prop :=
base : lt a (succ a), | base : lt a (succ a)
step : Π {b}, lt a b → lt a (succ b) | step : Π {b}, lt a b → lt a (succ b)
notation a < b := lt a b notation a < b := lt a b

View file

@ -38,8 +38,8 @@ namespace prod
-- Lexicographical order based on Ra and Rb -- Lexicographical order based on Ra and Rb
inductive lex : A × B → A × B → Prop := inductive lex : A × B → A × B → Prop :=
left : ∀{a₁ b₁} a₂ b₂, Ra a₁ a₂ → lex (a₁, b₁) (a₂, b₂), | left : ∀{a₁ b₁} a₂ b₂, Ra a₁ a₂ → lex (a₁, b₁) (a₂, b₂)
right : ∀a {b₁ b₂}, Rb b₁ b₂ → lex (a, b₁) (a, b₂) | right : ∀a {b₁ b₂}, Rb b₁ b₂ → lex (a, b₁) (a, b₂)
-- Relational product based on Ra and Rb -- Relational product based on Ra and Rb
inductive rprod : A × B → A × B → Prop := inductive rprod : A × B → A × B → Prop :=

View file

@ -38,5 +38,5 @@ theorem inv_image.irreflexive (f : A → B) (H : irreflexive R) : irreflexive (i
λ (a : A) (H₁ : inv_image R f a a), H (f a) H₁ λ (a : A) (H₁ : inv_image R f a a), H (f a) H₁
inductive tc {A : Type} (R : A → A → Prop) : A → A → Prop := inductive tc {A : Type} (R : A → A → Prop) : A → A → Prop :=
base : ∀a b, R a b → tc R a b, | base : ∀a b, R a b → tc R a b
trans : ∀a b c, tc R a b → tc R b c → tc R a c | trans : ∀a b c, tc R a b → tc R b c → tc R a c

View file

@ -31,8 +31,8 @@ namespace sigma
-- Lexicographical order based on Ra and Rb -- Lexicographical order based on Ra and Rb
inductive lex : sigma B → sigma B → Prop := inductive lex : sigma B → sigma B → Prop :=
left : ∀{a₁ b₁} a₂ b₂, Ra a₁ a₂ → lex ⟨a₁, b₁⟩ ⟨a₂, b₂⟩, | left : ∀{a₁ b₁} a₂ b₂, Ra a₁ a₂ → lex ⟨a₁, b₁⟩ ⟨a₂, b₂⟩
right : ∀a {b₁ b₂}, Rb a b₁ b₂ → lex ⟨a, b₁⟩ ⟨a, b₂⟩ | right : ∀a {b₁ b₂}, Rb a b₁ b₂ → lex ⟨a, b₁⟩ ⟨a, b₂⟩
end end
context context

View file

@ -67,8 +67,8 @@ opaque definition inversion (id : expr) : tactic := builtin
notation a `↦` b:max := rename a b notation a `↦` b:max := rename a b
inductive expr_list : Type := inductive expr_list : Type :=
nil : expr_list, | nil : expr_list
cons : expr → expr_list → expr_list | cons : expr → expr_list → expr_list
-- rewrite_tac is just a marker for the builtin 'rewrite' notation -- rewrite_tac is just a marker for the builtin 'rewrite' notation
-- used to create instances of this tactic. -- used to create instances of this tactic.

View file

@ -234,6 +234,10 @@ struct inductive_cmd_fn {
m_p.add_local(l); m_p.add_local(l);
} }
bool curr_is_intro_prefix() const {
return m_p.curr_is_token(get_bar_tk()) || m_p.curr_is_token(get_comma_tk());
}
/** \brief Parse introduction rules in the scope of m_params. /** \brief Parse introduction rules in the scope of m_params.
Introduction rules with the annotation '{}' are marked for relaxed (aka non-strict) implicit parameter inference. Introduction rules with the annotation '{}' are marked for relaxed (aka non-strict) implicit parameter inference.
@ -241,6 +245,8 @@ struct inductive_cmd_fn {
*/ */
list<intro_rule> parse_intro_rules(name const & ind_name) { list<intro_rule> parse_intro_rules(name const & ind_name) {
buffer<intro_rule> intros; buffer<intro_rule> intros;
if (m_p.curr_is_token(get_bar_tk()))
m_p.next();
while (true) { while (true) {
name intro_name = parse_intro_decl_name(ind_name); name intro_name = parse_intro_decl_name(ind_name);
implicit_infer_kind k = parse_implicit_infer_modifier(m_p); implicit_infer_kind k = parse_implicit_infer_modifier(m_p);
@ -253,7 +259,7 @@ struct inductive_cmd_fn {
expr intro_type = mk_constant(ind_name); expr intro_type = mk_constant(ind_name);
intros.push_back(intro_rule(intro_name, intro_type)); intros.push_back(intro_rule(intro_name, intro_type));
} }
if (!m_p.curr_is_token(get_comma_tk())) if (!curr_is_intro_prefix())
break; break;
m_p.next(); m_p.next();
} }

View file

@ -1,7 +1,7 @@
open eq.ops open eq.ops
inductive Nat : Type := inductive Nat : Type :=
zero : Nat, zero : Nat |
succ : Nat → Nat succ : Nat → Nat
namespace Nat namespace Nat

View file

@ -12,11 +12,11 @@ definition foo : nat → nat → nat
| foo ⌞x⌟ x := x | foo ⌞x⌟ x := x
inductive tree (A : Type) := inductive tree (A : Type) :=
node : tree_list A → tree A, | node : tree_list A → tree A
leaf : A → tree A | leaf : A → tree A
with tree_list := with tree_list :=
nil {} : tree_list A, | nil {} : tree_list A
cons : tree A → tree_list A → tree_list A | cons : tree A → tree_list A → tree_list A
definition is_leaf {A : Type} : tree A → bool definition is_leaf {A : Type} : tree A → bool
with is_leaf_aux : tree_list A → bool with is_leaf_aux : tree_list A → bool

View file

@ -1,6 +1,6 @@
namespace foo namespace foo
open nat open nat
inductive nat : Type := zero, foosucc : nat → nat inductive nat : Type := zero | foosucc : nat → nat
check 0 + nat.zero --error check 0 + nat.zero --error
end foo end foo

View file

@ -1,10 +1,10 @@
inductive list (T : Type) : Type := nil {} : list T, cons : T → list T → list T inductive list (T : Type) : Type := nil {} : list T | cons : T → list T → list T
namespace explicit namespace explicit
inductive ftree.{l₁ l₂} (A : Type.{l₁}) (B : Type.{l₂}) : Type.{max 1 l₁ l₂} := inductive ftree.{l₁ l₂} (A : Type.{l₁}) (B : Type.{l₂}) : Type.{max 1 l₁ l₂} :=
leafa : A → ftree A B, leafa : A → ftree A B |
leafb : B → ftree A B, leafb : B → ftree A B |
node : (A → ftree A B) → (B → ftree A B) → ftree A B node : (A → ftree A B) → (B → ftree A B) → ftree A B
end explicit end explicit
@ -12,7 +12,7 @@ end explicit
namespace implicit namespace implicit
inductive ftree (A : Type) (B : Type) : Type := inductive ftree (A : Type) (B : Type) : Type :=
leafa : ftree A B, leafa : ftree A B |
node : (A → B → ftree A B) → (B → ftree A B) → ftree A B node : (A → B → ftree A B) → (B → ftree A B) → ftree A B
set_option pp.universes true set_option pp.universes true
@ -23,8 +23,8 @@ end implicit
namespace implicit2 namespace implicit2
inductive ftree (A : Type) (B : Type) : Type := inductive ftree (A : Type) (B : Type) : Type :=
leafa : A → ftree A B, leafa : A → ftree A B |
leafb : B → ftree A B, leafb : B → ftree A B |
node : (list A → ftree A B) → (B → ftree A B) → ftree A B node : (list A → ftree A B) → (B → ftree A B) → ftree A B
check ftree check ftree

View file

@ -1,8 +1,8 @@
open nat open nat
inductive vec (A : Type) : nat → Type := inductive vec (A : Type) : nat → Type :=
nil {} : vec A zero, | nil {} : vec A zero
cons : Π {n}, A → vec A n → vec A (succ n) | cons : Π {n}, A → vec A n → vec A (succ n)
namespace vec namespace vec
variables {A B C : Type} variables {A B C : Type}

View file

@ -1,8 +1,8 @@
open nat open nat
inductive vector (A : Type) : nat → Type := inductive vector (A : Type) : nat → Type :=
nil {} : vector A zero, | nil {} : vector A zero
cons : Π {n}, A → vector A n → vector A (succ n) | cons : Π {n}, A → vector A n → vector A (succ n)
infixr `::` := vector.cons infixr `::` := vector.cons

View file

@ -2,10 +2,10 @@ open nat
open eq.ops open eq.ops
inductive even : nat → Type := inductive even : nat → Type :=
even_zero : even zero, | even_zero : even zero
even_succ_of_odd : ∀ {a}, odd a → even (succ a) | even_succ_of_odd : ∀ {a}, odd a → even (succ a)
with odd : nat → Type := with odd : nat → Type :=
odd_succ_of_even : ∀ {a}, even a → odd (succ a) | odd_succ_of_even : ∀ {a}, even a → odd (succ a)
example : even 1 → empty := example : even 1 → empty :=
begin begin

View file

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

View file

@ -5,8 +5,8 @@
open nat open nat
inductive vec (A : Type) : nat → Type := inductive vec (A : Type) : nat → Type :=
nil {} : vec A zero, | nil {} : vec A zero
cons : Π {n}, A → vec A n → vec A (succ n) | cons : Π {n}, A → vec A n → vec A (succ n)
structure S (A : Type) (a : A) (n : nat) (v : vec A n) := structure S (A : Type) (a : A) (n : nat) (v : vec A n) :=
mk :: (fa : A) mk :: (fa : A)

View file

@ -1,8 +1,8 @@
VISIT auto_comp.lean VISIT auto_comp.lean
SYNC 8 SYNC 8
prelude prelude
inductive nat := zero : nat, succ : nat → nat inductive nat := zero : nat | succ : nat → nat
inductive bool := ff, tt inductive bool := ff | tt
namespace nat namespace nat
inductive le : nat → nat → Type.{0} := refl : Π a, le a a inductive le : nat → nat → Type.{0} := refl : Π a, le a a

View file

@ -7,10 +7,10 @@ definition Prop := Type.{0}
inductive eq {A : Type} (a : A) : A → Prop := refl : eq a a inductive eq {A : Type} (a : A) : A → Prop := refl : eq a a
infix `=`:50 := eq infix `=`:50 := eq
definition rfl {A : Type} {a : A} := eq.refl a definition rfl {A : Type} {a : A} := eq.refl a
inductive or (A B : Prop) : Prop := inl {} : A → or A B, inr {} : B → or A B inductive or (A B : Prop) : Prop := inl {} : A → or A B | inr {} : B → or A B
infix ``:35 := or infix ``:35 := or
inductive bool : Type := inductive bool : Type :=
ff : bool, ff : bool|
tt : bool tt : bool
namespace bool namespace bool
protected definition rec_on2 {C : bool → Type} (b : bool) (H₁ : C ff) (H₂ : C tt) : C b := protected definition rec_on2 {C : bool → Type} (b : bool) (H₁ : C ff) (H₂ : C tt) : C b :=

View file

@ -1,7 +1,7 @@
import logic import logic
inductive tree (A : Type) := inductive tree (A : Type) :=
leaf : A → tree A, leaf : A → tree A|
node : tree A → tree A → tree A node : tree A → tree A → tree A
namespace tree namespace tree

View file

@ -9,8 +9,8 @@ import logic.inhabited
-- The parser will generate the terms (pos (bit1 (bit1 (bit0 one)))), zero, and (pos (bit0 (bit1 (bit1 one)))). -- 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). -- This representation can be coerced in whatever we want (e.g., naturals, integers, reals, etc).
inductive pos_num : Type := inductive pos_num : Type :=
one : pos_num, one : pos_num|
bit1 : pos_num → pos_num, bit1 : pos_num → pos_num|
bit0 : pos_num → pos_num bit0 : pos_num → pos_num
theorem pos_num.is_inhabited [instance] : inhabited pos_num := theorem pos_num.is_inhabited [instance] : inhabited pos_num :=

View file

@ -2,8 +2,8 @@ import logic
open nat open nat
inductive vec (A : Type) : nat → Type := inductive vec (A : Type) : nat → Type :=
vnil : vec A zero, vnil : vec A zero |
vone : A → vec A (succ zero), vone : A → vec A (succ zero) |
vtwo : A → A → vec A (succ (succ zero)) vtwo : A → A → vec A (succ (succ zero))
namespace vec namespace vec

View file

@ -2,7 +2,7 @@ import logic
open nat open nat
inductive vector (A : Type) : nat → Type := inductive vector (A : Type) : nat → Type :=
vnil : vector A zero, vnil : vector A zero |
vcons : Π {n : nat}, A → vector A n → vector A (succ n) vcons : Π {n : nat}, A → vector A n → vector A (succ n)
check vector.no_confusion_type check vector.no_confusion_type

View file

@ -1,5 +1,5 @@
import data.num import data.num
inductive list (T : Type) : Type := nil {} : list T, cons : T → list T → list T open list notation h :: t := cons h t notation `[` l:(foldr `,` (h t, cons h t) nil) `]` := l inductive list (T : Type) : Type := nil {} : list T | cons : T → list T → list T open list notation h :: t := cons h t notation `[` l:(foldr `,` (h t, cons h t) nil) `]` := l
infixr `::` := cons infixr `::` := cons
check 1 :: 2 :: nil check 1 :: 2 :: nil
check 1 :: 2 :: 3 :: 4 :: 5 :: nil check 1 :: 2 :: 3 :: 4 :: 5 :: nil

View file

@ -1,5 +1,5 @@
import data.prod data.num import data.prod data.num
inductive list (T : Type) : Type := nil {} : list T, cons : T → list T → list T open list notation h :: t := cons h t notation `[` l:(foldr `,` (h t, cons h t) nil) `]` := l inductive list (T : Type) : Type := nil {} : list T | cons : T → list T → list T open list notation h :: t := cons h t notation `[` l:(foldr `,` (h t, cons h t) nil) `]` := l
open prod num open prod num
constants a b : num constants a b : num
check [a, b, b] check [a, b, b]

View file

@ -1,6 +1,6 @@
import logic data.sigma import logic data.sigma
open sigma open sigma
inductive list (T : Type) : Type := nil {} : list T, cons : T → list T → list T open list notation h :: t := cons h t notation `[` l:(foldr `,` (h t, cons h t) nil) `]` := l inductive list (T : Type) : Type := nil {} : list T | cons : T → list T → list T open list notation h :: t := cons h t notation `[` l:(foldr `,` (h t, cons h t) nil) `]` := l
check ∃ (A : Type₁) (x y : A), x = y check ∃ (A : Type₁) (x y : A), x = y
check ∃ (x : num), x = 0 check ∃ (x : num), x = 0
check Σ (x : num), x = 10 check Σ (x : num), x = 10

View file

@ -1,8 +1,8 @@
import logic import logic
inductive vector (T : Type) : nat → Type := inductive vector (T : Type) : nat → Type :=
nil {} : vector T nat.zero, | nil {} : vector T nat.zero
cons : T → ∀{n}, vector T n → vector T (nat.succ n) | cons : T → ∀{n}, vector T n → vector T (nat.succ n)
#projections or #projections or
#projections and #projections and

View file

@ -1,7 +1,7 @@
namespace nat namespace nat
inductive less (a : nat) : nat → Prop := inductive less (a : nat) : nat → Prop :=
base : less a (succ a), | base : less a (succ a)
step : Π {b}, less a b → less a (succ b) | step : Π {b}, less a b → less a (succ b)
end nat end nat

View file

@ -35,8 +35,8 @@ end algebra
open algebra open algebra
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat namespace nat

View file

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

View file

@ -2,8 +2,8 @@ import logic data.nat
open nat open nat
inductive fin : → Type := inductive fin : → Type :=
zero : Π {n : }, fin (succ n), | zero : Π {n : }, fin (succ n)
succ : Π {n : }, fin n → fin (succ n) | succ : Π {n : }, fin n → fin (succ n)
theorem foo (n m : ) (a : fin n) (b : fin m) (H : n = m) : cast (congr_arg fin H) a = b := theorem foo (n m : ) (a : fin n) (b : fin m) (H : n = m) : cast (congr_arg fin H) a = b :=
have eq : fin n = fin m, from congr_arg fin H, have eq : fin n = fin m, from congr_arg fin H,

View file

@ -2,8 +2,8 @@ prelude
import logic import logic
namespace experiment namespace experiment
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
open eq open eq

View file

@ -13,8 +13,8 @@ end algebra
open algebra open algebra
namespace nat namespace nat
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
constant mul : nat → nat → nat constant mul : nat → nat → nat
constant add : nat → nat → nat constant add : nat → nat → nat

View file

@ -2,8 +2,8 @@ import data.nat
open nat open nat
inductive list (T : Type) : Type := inductive list (T : Type) : Type :=
nil {} : list T, | nil {} : list T
cons : T → list T → list T | cons : T → list T → list T
definition length {T : Type} : list T → nat := list.rec 0 (fun x l m, succ m) definition length {T : Type} : list T → nat := list.rec 0 (fun x l m, succ m)
theorem length_nil {T : Type} : length (@list.nil T) = 0 theorem length_nil {T : Type} : length (@list.nil T) = 0

View file

@ -4,8 +4,8 @@ definition mk_arrow (A : Type) (B : Type) :=
A → A → B A → A → B
inductive confuse (A : Type) := inductive confuse (A : Type) :=
leaf1 : confuse A, | leaf1 : confuse A
leaf2 : num → confuse A, | leaf2 : num → confuse A
node : mk_arrow A (confuse A) → confuse A | node : mk_arrow A (confuse A) → confuse A
check confuse.cases_on check confuse.cases_on

View file

@ -1,12 +1,12 @@
prelude prelude
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat end nat open nat namespace nat end nat open nat
inductive list (A : Type) : Type := inductive list (A : Type) : Type :=
nil {} : list A, | nil {} : list A
cons : A → list A → list A | cons : A → list A → list A
definition nil := @list.nil definition nil := @list.nil
definition cons := @list.cons definition cons := @list.cons
@ -18,8 +18,8 @@ check @nil nat
check cons nat.zero nil check cons nat.zero nil
inductive vector (A : Type) : nat → Type := inductive vector (A : Type) : nat → Type :=
vnil {} : vector A zero, | vnil {} : vector A zero
vcons : forall {n : nat}, A → vector A n → vector A (succ n) | vcons : forall {n : nat}, A → vector A n → vector A (succ n)
namespace vector end vector open vector namespace vector end vector open vector
check vcons zero vnil check vcons zero vnil

View file

@ -1,12 +1,12 @@
prelude prelude
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat end nat open nat namespace nat end nat open nat
inductive list (A : Type) : Type := inductive list (A : Type) : Type :=
nil {} : list A, | nil {} : list A
cons : A → list A → list A | cons : A → list A → list A
namespace list end list open list namespace list end list open list
check nil check nil
check nil.{1} check nil.{1}
@ -16,8 +16,8 @@ check @nil nat
check cons zero nil check cons zero nil
inductive vector (A : Type) : nat → Type := inductive vector (A : Type) : nat → Type :=
vnil {} : vector A zero, | vnil {} : vector A zero
vcons : forall {n : nat}, A → vector A n → vector A (succ n) | vcons : forall {n : nat}, A → vector A n → vector A (succ n)
namespace vector end vector open vector namespace vector end vector open vector
check vcons zero vnil check vcons zero vnil

View file

@ -1,12 +1,12 @@
prelude prelude
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat end nat open nat namespace nat end nat open nat
inductive list (A : Type) : Type := inductive list (A : Type) : Type :=
nil {} : list A, | nil {} : list A
cons : A → list A → list A | cons : A → list A → list A
namespace list end list open list namespace list end list open list
check nil check nil
@ -17,8 +17,8 @@ check @nil nat
check cons zero nil check cons zero nil
inductive vector (A : Type) : nat → Type := inductive vector (A : Type) : nat → Type :=
vnil {} : vector A zero, | vnil {} : vector A zero
vcons : forall {n : nat}, A → vector A n → vector A (succ n) | vcons : forall {n : nat}, A → vector A n → vector A (succ n)
namespace vector end vector open vector namespace vector end vector open vector
check vcons zero vnil check vcons zero vnil

View file

@ -1,15 +1,15 @@
prelude prelude
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
inductive list (A : Type) : Type := inductive list (A : Type) : Type :=
nil {} : list A, | nil {} : list A
cons : A → list A → list A | cons : A → list A → list A
inductive int : Type := inductive int : Type :=
of_nat : nat → int, | of_nat : nat → int
neg : nat → int | neg : nat → int
attribute int.of_nat [coercion] attribute int.of_nat [coercion]

View file

@ -1,15 +1,15 @@
prelude prelude
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
inductive list (A : Type) : Type := inductive list (A : Type) : Type :=
nil {} : list A, | nil {} : list A
cons : A → list A → list A | cons : A → list A → list A
inductive int : Type := inductive int : Type :=
of_nat : nat → int, | of_nat : nat → int
neg : nat → int | neg : nat → int
attribute int.of_nat [coercion] attribute int.of_nat [coercion]

View file

@ -33,8 +33,8 @@ theorem trans {A : Type} {a b c : A} (H1 : a = b) (H2 : b = c) : a = c
:= subst H2 H1 := subst H2 H1
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat end nat open nat namespace nat end nat open nat
print "using strict implicit arguments" print "using strict implicit arguments"

View file

@ -1,9 +1,9 @@
import logic import logic
inductive Three := inductive Three :=
zero : Three, | zero : Three
one : Three, | one : Three
two : Three | two : Three
namespace Three namespace Three

View file

@ -1,5 +1,5 @@
inductive day := inductive day :=
monday, tuesday, wednesday, thursday, friday, saturday, sunday monday | tuesday | wednesday | thursday | friday | saturday | sunday
open day open day

View file

@ -1,10 +1,10 @@
inductive formula := inductive formula :=
eqf : nat → nat → formula, | eqf : nat → nat → formula
andf : formula → formula → formula, | andf : formula → formula → formula
impf : formula → formula → formula, | impf : formula → formula → formula
notf : formula → formula, | notf : formula → formula
orf : formula → formula → formula, | orf : formula → formula → formula
allf : (nat → formula) → formula | allf : (nat → formula) → formula
namespace formula namespace formula
definition implies (a b : Prop) : Prop := a → b definition implies (a b : Prop) : Prop := a → b

View file

@ -1,5 +1,5 @@
inductive day := inductive day :=
monday, tuesday, wednesday, thursday, friday, saturday, sunday monday | tuesday | wednesday | thursday | friday | saturday | sunday
open day open day

View file

@ -1,6 +1,6 @@
inductive formula := inductive formula :=
eqf : nat → nat → formula, | eqf : nat → nat → formula
impf : formula → formula → formula | impf : formula → formula → formula
namespace formula namespace formula
definition denote : formula → Prop definition denote : formula → Prop

View file

@ -1,11 +1,11 @@
open nat open nat
inductive tree (A : Type) := inductive tree (A : Type) :=
leaf : A → tree A, | leaf : A → tree A
node : tree_list A → tree A | node : tree_list A → tree A
with tree_list := with tree_list :=
nil : tree_list A, | nil : tree_list A
cons : tree A → tree_list A → tree_list A | cons : tree A → tree_list A → tree_list A
namespace tree_list namespace tree_list

View file

@ -1,11 +1,11 @@
open nat open nat
inductive tree (A : Type) := inductive tree (A : Type) :=
leaf : A → tree A, | leaf : A → tree A
node : tree_list A → tree A | node : tree_list A → tree A
with tree_list := with tree_list :=
nil : tree_list A, | nil : tree_list A
cons : tree A → tree_list A → tree_list A | cons : tree A → tree_list A → tree_list A
namespace tree namespace tree
open tree_list open tree_list

View file

@ -1,6 +1,6 @@
inductive N := inductive N :=
O : N, | O : N
S : N → N | S : N → N
definition Nat := N definition Nat := N

View file

@ -1,7 +1,7 @@
import logic import logic
inductive day := inductive day :=
monday, tuesday, wednesday, thursday, friday, saturday, sunday monday | tuesday | wednesday | thursday | friday | saturday | sunday
namespace day namespace day

View file

@ -10,8 +10,8 @@ Finite ordinals.
open nat open nat
inductive fin : nat → Type := inductive fin : nat → Type :=
fz : Π n, fin (succ n), | fz : Π n, fin (succ n)
fs : Π {n}, fin n → fin (succ n) | fs : Π {n}, fin n → fin (succ n)
namespace fin namespace fin
definition to_nat : Π {n}, fin n → nat definition to_nat : Π {n}, fin n → nat

View file

@ -10,8 +10,8 @@ Finite ordinals.
open nat open nat
inductive fin : nat → Type := inductive fin : nat → Type :=
fz : Π n, fin (succ n), | fz : Π n, fin (succ n)
fs : Π {n}, fin n → fin (succ n) | fs : Π {n}, fin n → fin (succ n)
namespace fin namespace fin
definition to_nat : Π {n}, fin n → nat definition to_nat : Π {n}, fin n → nat

View file

@ -2,10 +2,10 @@ import data.prod data.unit
open prod open prod
inductive tree (A : Type) : Type := inductive tree (A : Type) : Type :=
node : A → forest A → tree A | node : A → forest A → tree A
with forest : Type := with forest : Type :=
nil : forest A, | nil : forest A
cons : tree A → forest A → forest A | cons : tree A → forest A → forest A
namespace manual namespace manual
definition tree.below.{l₁ l₂} definition tree.below.{l₁ l₂}

View file

@ -2,20 +2,20 @@ import data.prod data.unit
open prod open prod
inductive tree (A : Type) : Type := inductive tree (A : Type) : Type :=
node : A → forest A → tree A | node : A → forest A → tree A
with forest : Type := with forest : Type :=
nil : forest A, | nil : forest A
cons : tree A → forest A → forest A | cons : tree A → forest A → forest A
namespace solution1 namespace solution1
inductive tree_forest (A : Type) := inductive tree_forest (A : Type) :=
of_tree : tree A → tree_forest A, | of_tree : tree A → tree_forest A
of_forest : forest A → tree_forest A | of_forest : forest A → tree_forest A
inductive same_kind {A : Type} : tree_forest A → tree_forest A → Type := inductive same_kind {A : Type} : tree_forest A → tree_forest A → Type :=
is_tree : Π (t₁ t₂ : tree A), same_kind (tree_forest.of_tree t₁) (tree_forest.of_tree t₂), | is_tree : Π (t₁ t₂ : tree A), same_kind (tree_forest.of_tree t₁) (tree_forest.of_tree t₂)
is_forest : Π (f₁ f₂ : forest A), same_kind (tree_forest.of_forest f₁) (tree_forest.of_forest f₂) | is_forest : Π (f₁ f₂ : forest A), same_kind (tree_forest.of_forest f₁) (tree_forest.of_forest f₂)
definition to_tree {A : Type} (tf : tree_forest A) (t : tree A) : same_kind tf (tree_forest.of_tree t) → tree A := definition to_tree {A : Type} (tf : tree_forest A) (t : tree A) : same_kind tf (tree_forest.of_tree t) → tree A :=
tree_forest.cases_on tf tree_forest.cases_on tf
@ -29,8 +29,8 @@ namespace solution2
variables {A B : Type} variables {A B : Type}
inductive same_kind : sum A B → sum A B → Prop := inductive same_kind : sum A B → sum A B → Prop :=
isl : Π (a₁ a₂ : A), same_kind (sum.inl a₁) (sum.inl a₂), | isl : Π (a₁ a₂ : A), same_kind (sum.inl a₁) (sum.inl a₂)
isr : Π (b₁ b₂ : B), same_kind (sum.inr b₁) (sum.inr b₂) | isr : Π (b₁ b₂ : B), same_kind (sum.inr b₁) (sum.inr b₂)
definition to_left (s : sum A B) (a : A) : same_kind s (sum.inl a) → A := definition to_left (s : sum A B) (a : A) : same_kind s (sum.inl a) → A :=
sum.cases_on s sum.cases_on s

View file

@ -2,10 +2,10 @@ import data.nat.basic data.sum data.sigma data.bool
open nat sigma open nat sigma
inductive tree (A : Type) : Type := inductive tree (A : Type) : Type :=
node : A → forest A → tree A | node : A → forest A → tree A
with forest : Type := with forest : Type :=
nil : forest A, | nil : forest A
cons : tree A → forest A → forest A | cons : tree A → forest A → forest A
namespace manual namespace manual
check tree.rec_on check tree.rec_on

View file

@ -1,8 +1,8 @@
import data.unit data.prod import data.unit data.prod
inductive ftree (A : Type) (B : Type) : Type := inductive ftree (A : Type) (B : Type) : Type :=
leafa : ftree A B, | leafa : ftree A B
node : (A → B → ftree A B) → (B → ftree A B) → ftree A B | node : (A → B → ftree A B) → (B → ftree A B) → ftree A B
set_option pp.universes true set_option pp.universes true
check @ftree check @ftree

View file

@ -1,8 +1,8 @@
import data.nat import data.nat
inductive star : Type₁ := inductive star : Type₁ :=
z : star, | z : star
s : (nat → star) → star | s : (nat → star) → star
check @star.rec check @star.rec
check @star.cases_on check @star.cases_on

View file

@ -1,7 +1,7 @@
prelude prelude
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
check nat check nat
check nat.rec.{1} check nat.rec.{1}

View file

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

View file

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

View file

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

View file

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

View file

@ -2,8 +2,8 @@ prelude
definition Prop : Type.{1} := Type.{0} definition Prop : Type.{1} := Type.{0}
inductive or (A B : Prop) : Prop := inductive or (A B : Prop) : Prop :=
intro_left : A → or A B, | intro_left : A → or A B
intro_right : B → or A B | intro_right : B → or A B
check or check or
check or.intro_left check or.intro_left

View file

@ -1,8 +1,8 @@
inductive tree.{u} (A : Type.{u}) : Type.{max u 1} := 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 : Type.{max u 1} := with forest : Type.{max u 1} :=
nil : forest.{u} A, | nil : forest.{u} A
cons : tree.{u} A → forest.{u} A → forest.{u} A | cons : tree.{u} A → forest.{u} A → forest.{u} A
check tree.{1} check tree.{1}
check forest.{1} check forest.{1}

View file

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

View file

@ -1,5 +1,5 @@
inductive day := inductive day :=
monday, tuesday, wednesday, thursday, friday, saturday, sunday monday | tuesday | wednesday | thursday | friday | saturday | sunday
check day.monday check day.monday
open day open day

View file

@ -2,16 +2,16 @@ prelude
definition Prop := Type.{0} definition Prop := Type.{0}
inductive nat := inductive nat :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
inductive list (A : Type) := inductive list (A : Type) :=
nil {} : list A, | nil {} : list A
cons : A → list A → list A | cons : A → list A → list A
inductive list2 (A : Type) : Type := inductive list2 (A : Type) : Type :=
nil2 {} : list2 A, | nil2 {} : list2 A
cons2 : A → list2 A → list2 A | cons2 : A → list2 A → list2 A
inductive and (A B : Prop) : Prop := inductive and (A B : Prop) : Prop :=
and_intro : A → B → and A B and_intro : A → B → and A B

View file

@ -1,23 +1,23 @@
prelude prelude
inductive list (A : Type) : Type := inductive list (A : Type) : Type :=
nil {} : list A, | nil {} : list A
cons : A → list A → list A | cons : A → list A → list A
section section
variable A : Type variable A : Type
inductive list2 : Type := inductive list2 : Type :=
nil2 {} : list2, | nil2 {} : list2
cons2 : A → list2 → list2 | cons2 : A → list2 → list2
end end
constant num : Type.{1} constant num : Type.{1}
namespace Tree namespace Tree
inductive tree (A : Type) : Type := inductive tree (A : Type) : Type :=
node : A → forest A → tree A | node : A → forest A → tree A
with forest : Type := with forest : Type :=
nil : forest A, | nil : forest A
cons : tree A → forest A → forest A | cons : tree A → forest A → forest A
end Tree end Tree
inductive group_struct (A : Type) : Type := inductive group_struct (A : Type) : Type :=

View file

@ -2,8 +2,8 @@ import logic data.nat.basic
open nat open nat
inductive inftree (A : Type) := inductive inftree (A : Type) :=
leaf : A → inftree A, | leaf : A → inftree A
node : (nat → inftree A) → inftree A | node : (nat → inftree A) → inftree A
namespace inftree namespace inftree
inductive dsub {A : Type} : inftree A → inftree A → Prop := inductive dsub {A : Type} : inftree A → inftree A → Prop :=

View file

@ -2,13 +2,13 @@ import logic data.nat.basic
open nat open nat
inductive inftree (A : Type) : Type := inductive inftree (A : Type) : Type :=
leaf : A → inftree A, | leaf : A → inftree A
node : (nat → inftree A) → inftree A → inftree A | node : (nat → inftree A) → inftree A → inftree A
namespace inftree namespace inftree
inductive dsub {A : Type} : inftree A → inftree A → Prop := inductive dsub {A : Type} : inftree A → inftree A → Prop :=
intro₁ : Π (f : nat → inftree A) (a : nat) (t : inftree A), dsub (f a) (node f t), | intro₁ : Π (f : nat → inftree A) (a : nat) (t : inftree A), dsub (f a) (node f t)
intro₂ : Π (f : nat → inftree A) (a : nat) (t : inftree A), dsub t (node f t) | intro₂ : Π (f : nat → inftree A) (a : nat) (t : inftree A), dsub t (node f t)
definition dsub.node.acc {A : Type} (f : nat → inftree A) (hf : ∀a, acc dsub (f a)) definition dsub.node.acc {A : Type} (f : nat → inftree A) (hf : ∀a, acc dsub (f a))
(t : inftree A) (ht : acc dsub t) : acc dsub (node f t) := (t : inftree A) (ht : acc dsub t) : acc dsub (node f t) :=

View file

@ -2,13 +2,13 @@ import logic data.nat.basic
open nat open nat
inductive inftree (A : Type) : Type := inductive inftree (A : Type) : Type :=
leaf : A → inftree A, | leaf : A → inftree A
node : (nat → inftree A) → inftree A → inftree A | node : (nat → inftree A) → inftree A → inftree A
namespace inftree namespace inftree
inductive dsub {A : Type} : inftree A → inftree A → Prop := inductive dsub {A : Type} : inftree A → inftree A → Prop :=
intro₁ : Π (f : nat → inftree A) (a : nat) (t : inftree A), dsub (f a) (node f t), | intro₁ : Π (f : nat → inftree A) (a : nat) (t : inftree A), dsub (f a) (node f t)
intro₂ : Π (f : nat → inftree A) (t : inftree A), dsub t (node f t) | intro₂ : Π (f : nat → inftree A) (t : inftree A), dsub t (node f t)
definition dsub.node.acc {A : Type} (f : nat → inftree A) (hf : ∀a, acc dsub (f a)) definition dsub.node.acc {A : Type} (f : nat → inftree A) (hf : ∀a, acc dsub (f a))
(t : inftree A) (ht : acc dsub t) : acc dsub (node f t) := (t : inftree A) (ht : acc dsub t) : acc dsub (node f t) :=

View file

@ -2,10 +2,10 @@ open nat
open eq.ops open eq.ops
inductive even : nat → Prop := inductive even : nat → Prop :=
even_zero : even zero, | even_zero : even zero
even_succ_of_odd : ∀ {a}, odd a → even (succ a) | even_succ_of_odd : ∀ {a}, odd a → even (succ a)
with odd : nat → Prop := with odd : nat → Prop :=
odd_succ_of_even : ∀ {a}, even a → odd (succ a) | odd_succ_of_even : ∀ {a}, even a → odd (succ a)
example : even 1 → false := example : even 1 → false :=
begin begin

View file

@ -2,8 +2,8 @@ import logic
open tactic open tactic
inductive list (A : Type) : Type := inductive list (A : Type) : Type :=
nil {} : list A, | nil {} : list A
cons : A → list A → list A | cons : A → list A → list A
namespace list end list open list namespace list end list open list
open eq open eq

View file

@ -12,8 +12,8 @@
import data.nat import data.nat
open nat eq.ops open nat eq.ops
inductive list (T : Type) : Type := inductive list (T : Type) : Type :=
nil {} : list T, | nil {} : list T
cons : T → list T → list T | cons : T → list T → list T
namespace list namespace list
theorem list_induction_on {T : Type} {P : list T → Prop} (l : list T) (Hnil : P nil) theorem list_induction_on {T : Type} {P : list T → Prop} (l : list T) (Hnil : P nil)

View file

@ -3,8 +3,8 @@ open decidable
open eq open eq
namespace experiment namespace experiment
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
definition refl := @eq.refl definition refl := @eq.refl
namespace nat namespace nat

View file

@ -2,8 +2,8 @@ import logic
open eq.ops open eq.ops
namespace experiment namespace experiment
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat namespace nat
definition plus (x y : nat) : nat definition plus (x y : nat) : nat

View file

@ -2,8 +2,8 @@ import logic
open eq.ops open eq.ops
namespace experiment namespace experiment
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat namespace nat
definition plus (x y : nat) : nat definition plus (x y : nat) : nat

View file

@ -2,8 +2,8 @@ import logic
open eq.ops open eq.ops
namespace experiment namespace experiment
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat namespace nat
definition add (x y : nat) : nat := nat.rec x (λn r, succ r) y definition add (x y : nat) : nat := nat.rec x (λn r, succ r) y
infixl `+` := add infixl `+` := add

View file

@ -2,8 +2,8 @@ import logic
open eq.ops eq open eq.ops eq
namespace foo namespace foo
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat namespace nat
definition add (x y : nat) : nat := nat.rec x (λn r, succ r) y definition add (x y : nat) : nat := nat.rec x (λn r, succ r) y

View file

@ -2,8 +2,8 @@ import logic
open eq.ops open eq.ops
namespace experiment namespace experiment
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat namespace nat
definition add (x y : nat) : nat := nat.rec x (λn r, succ r) y definition add (x y : nat) : nat := nat.rec x (λn r, succ r) y

View file

@ -1,8 +1,8 @@
import logic import logic
namespace experiment namespace experiment
inductive nat : Type := inductive nat : Type :=
zero : nat, | zero : nat
succ : nat → nat | succ : nat → nat
namespace nat namespace nat
definition add (x y : nat) : nat := nat.rec x (λn r, succ r) y definition add (x y : nat) : nat := nat.rec x (λn r, succ r) y

View file

@ -2,8 +2,8 @@ import logic data.nat.basic
open nat open nat
inductive vector (A : Type) : nat → Type := inductive vector (A : Type) : nat → Type :=
vnil : vector A zero, | vnil : vector A zero
vcons : Π {n : nat}, A → vector A n → vector A (succ n) | vcons : Π {n : nat}, A → vector A n → vector A (succ n)
namespace vector namespace vector
definition no_confusion2 {A : Type} {n : nat} {P : Type} {v₁ v₂ : vector A n} : v₁ = v₂ → vector.no_confusion_type P v₁ v₂ := definition no_confusion2 {A : Type} {n : nat} {P : Type} {v₁ v₂ : vector A n} : v₁ = v₂ → vector.no_confusion_type P v₁ v₂ :=

View file

@ -2,13 +2,13 @@ import data.nat.basic
open nat open nat
inductive fin : nat → Type := inductive fin : nat → Type :=
fz : Π {n : nat}, fin (succ n), | fz : Π {n : nat}, fin (succ n)
fs : Π {n : nat}, fin n → fin (succ n) | fs : Π {n : nat}, fin n → fin (succ n)
namespace fin namespace fin
inductive le : ∀ {n : nat}, fin n → fin n → Prop := inductive le : ∀ {n : nat}, fin n → fin n → Prop :=
lez : ∀ {n : nat} (j : fin (succ n)), le fz j, | lez : ∀ {n : nat} (j : fin (succ n)), le fz j
les : ∀ {n : nat} {i j : fin n}, le i j → le (fs i) (fs j) | les : ∀ {n : nat} {i j : fin n}, le i j → le (fs i) (fs j)
end fin end fin

View file

@ -7,8 +7,8 @@ inductive pone : Type.{0} :=
unit : pone unit : pone
inductive two.{l} : Type.{max 1 l} := inductive two.{l} : Type.{max 1 l} :=
o : two, | o : two
u : two | u : two
inductive wrap.{l} : Type.{max 1 l} := inductive wrap.{l} : Type.{max 1 l} :=
mk : true → wrap mk : true → wrap

View file

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

View file

@ -14,7 +14,7 @@ check point.induction_on
check point.destruct check point.destruct
inductive color := inductive color :=
red, green, blue red | green | blue
structure color_point (A : Type) (B : Type) extends point A B := structure color_point (A : Type) (B : Type) extends point A B :=
mk :: (c : color) mk :: (c : color)

View file

@ -5,7 +5,7 @@ structure point (A : Type) (B : Type) :=
mk :: (x : A) (y : B) mk :: (x : A) (y : B)
inductive color := inductive color :=
red, green, blue red | green | blue
structure color_point (A : Type) (B : Type) extends point A B := structure color_point (A : Type) (B : Type) extends point A B :=
mk :: (c : color) mk :: (c : color)

View file

@ -11,7 +11,7 @@ example (p : point num num) : point.mk (point.x p) (point.y p) = p :=
point.eta p point.eta p
inductive color := inductive color :=
red, green, blue red | green | blue
structure color_point (A : Type) (B : Type) extends point A B := structure color_point (A : Type) (B : Type) extends point A B :=
mk :: (c : color) mk :: (c : color)

View file

@ -9,8 +9,8 @@ check a ◀ b
end end
inductive list (T : Type) : Type := inductive list (T : Type) : Type :=
nil {} : list T, | nil {} : list T
cons : T → list T → list T | cons : T → list T → list T
namespace list namespace list
section section

Some files were not shown because too many files have changed in this diff Show more