fix(tests/lean): adjust tests to reflect changes in the pretty printer

This commit is contained in:
Leonardo de Moura 2015-09-30 16:52:56 -07:00
parent ede23a3267
commit 1c4dea9941
21 changed files with 41 additions and 41 deletions

View file

@ -19,7 +19,7 @@ namespace bijection
(by rewrite [compose.assoc, -{finv f ∘ _}compose.assoc, linv f, compose.left_id, linv g]) (by rewrite [compose.assoc, -{finv f ∘ _}compose.assoc, linv f, compose.left_id, linv g])
(by rewrite [-compose.assoc, {_ ∘ finv g}compose.assoc, rinv g, compose.right_id, rinv f]) (by rewrite [-compose.assoc, {_ ∘ finv g}compose.assoc, rinv g, compose.right_id, rinv f])
infixr `∘b`:100 := compose infixr ` ∘b `:100 := compose
lemma compose.assoc (f g h : bijection A) : (f ∘b g) ∘b h = f ∘b (g ∘b h) := rfl lemma compose.assoc (f g h : bijection A) : (f ∘b g) ∘b h = f ∘b (g ∘b h) := rfl

View file

@ -12,8 +12,8 @@ end
section section
parameter {A : Type} parameter {A : Type}
definition relation' : A → A → Type := λa b, a = b definition relation' : A → A → Type := λa b, a = b
local infix `~1`:50 := relation' local infix ` ~1 `:50 := relation'
local infix [parsing-only] `~2`:50 := relation' local infix [parsing-only] ` ~2 `:50 := relation'
variable {a : A} variable {a : A}
check relation' a a check relation' a a
check a ~1 a check a ~1 a
@ -23,7 +23,7 @@ end
section section
parameter {A : Type} parameter {A : Type}
definition relation'' : A → A → Type := λa b, a = b definition relation'' : A → A → Type := λa b, a = b
local infix [parsing-only] `~2`:50 := relation'' local infix [parsing-only] ` ~2 `:50 := relation''
variable {a : A} variable {a : A}
check relation'' a a check relation'' a a
check a ~2 a check a ~2 a

View file

@ -7,14 +7,14 @@ u₂ : B u₁,
v₁ : A, v₁ : A,
v₂ : B v₁, v₂ : B v₁,
p : ⟨u₁, u₂⟩.1 = ⟨v₁, v₂⟩.1, p : ⟨u₁, u₂⟩.1 = ⟨v₁, v₂⟩.1,
q : u₂ =[ p ] v₂ q : u₂ =[p] v₂
⊢ ⟨(sigma_eq p q)..1, (sigma_eq p q)..2⟩ = ⟨p, q⟩ ⊢ ⟨(sigma_eq p q)..1, (sigma_eq p q)..2⟩ = ⟨p, q⟩
690.hlean:12:0: error: don't know how to synthesize placeholder 690.hlean:12:0: error: don't know how to synthesize placeholder
A : Type, A : Type,
B : A → Type, B : A → Type,
u v : Σ (a : A), B a, u v : Σ (a : A), B a,
p : u.1 = v.1, p : u.1 = v.1,
q : u.2 =[ p ] v.2 q : u.2 =[p] v.2
⊢ ⟨(sigma_eq p q)..1, (sigma_eq p q)..2⟩ = ⟨p, q⟩ ⊢ ⟨(sigma_eq p q)..1, (sigma_eq p q)..2⟩ = ⟨p, q⟩
690.hlean:12:0: error: failed to add declaration 'dpair_sigma_eq' to environment, value has metavariables 690.hlean:12:0: error: failed to add declaration 'dpair_sigma_eq' to environment, value has metavariables
remark: set 'formatter.hide_full_terms' to false to see the complete term remark: set 'formatter.hide_full_terms' to false to see the complete term

View file

@ -1,6 +1,6 @@
K_bug.lean:14:24: error: type mismatch at term K_bug.lean:14:24: error: type mismatch at term
pred_succ n ⁻¹ pred_succ n⁻¹
has type has type
pred (succ n ⁻¹) = n ⁻¹ pred (succ n⁻¹) = n⁻¹
but is expected to have type but is expected to have type
n = pred (succ n) n = pred (succ n)

View file

@ -1,5 +1,5 @@
{x : ∈ S| x > 0} : set {x : ∈ S | x > 0} : set
{x : ∈ s| x > 0} : finset {x : ∈ s | x > 0} : finset
@set.sep.{1} nat (λ (x : nat), nat.gt x (nat.of_num 0)) S : set.{1} nat @set.sep.{1} nat (λ (x : nat), nat.gt x (nat.of_num 0)) S : set.{1} nat
@finset.sep.{1} nat (λ (a b : nat), nat.has_decidable_eq a b) (λ (x : nat), nat.gt x (nat.of_num 0)) @finset.sep.{1} nat (λ (a b : nat), nat.has_decidable_eq a b) (λ (x : nat), nat.gt x (nat.of_num 0))
(λ (a : nat), nat.decidable_ge a (nat.succ (nat.of_num 0))) (λ (a : nat), nat.decidable_ge a (nat.succ (nat.of_num 0)))

View file

@ -1,6 +1,6 @@
prelude definition bool : Type.{1} := Type.{0} prelude definition bool : Type.{1} := Type.{0}
definition and (p q : bool) : bool := ∀ c : bool, (p → q → c) → c definition and (p q : bool) : bool := ∀ c : bool, (p → q → c) → c
infixl ``:25 := and infixl ``:25 := and
constant a : bool constant a : bool

View file

@ -1,12 +1,12 @@
prelude constant A : Type.{1} prelude constant A : Type.{1}
definition bool : Type.{1} := Type.{0} definition bool : Type.{1} := Type.{0}
constant eq : A → A → bool constant eq : A → A → bool
infixl `=`:50 := eq infixl ` = `:50 := eq
axiom subst (P : A → bool) (a b : A) (H1 : a = b) (H2 : P a) : P b axiom subst (P : A → bool) (a b : A) (H1 : a = b) (H2 : P a) : P b
axiom eq_trans (a b c : A) (H1 : a = b) (H2 : b = c) : a = c axiom eq_trans (a b c : A) (H1 : a = b) (H2 : b = c) : a = c
axiom eq_refl (a : A) : a = a axiom eq_refl (a : A) : a = a
constant le : A → A → bool constant le : A → A → bool
infixl ``:50 := le infixl ``:50 := le
axiom le_trans (a b c : A) (H1 : a ≤ b) (H2 : b ≤ c) : a ≤ c axiom le_trans (a b c : A) (H1 : a ≤ b) (H2 : b ≤ c) : a ≤ c
axiom le_refl (a : A) : a ≤ a axiom le_refl (a : A) : a ≤ a
axiom eq_le_trans (a b c : A) (H1 : a = b) (H2 : b ≤ c) : a ≤ c axiom eq_le_trans (a b c : A) (H1 : a = b) (H2 : b ≤ c) : a ≤ c
@ -29,7 +29,7 @@ check calc a = b : H1
... = e : H4 ... = e : H4
constant lt : A → A → bool constant lt : A → A → bool
infixl `<`:50 := lt infixl ` < `:50 := lt
axiom lt_trans (a b c : A) (H1 : a < b) (H2 : b < c) : a < c axiom lt_trans (a b c : A) (H1 : a < b) (H2 : b < c) : a < c
axiom le_lt_trans (a b c : A) (H1 : a ≤ b) (H2 : b < c) : a < c axiom le_lt_trans (a b c : A) (H1 : a ≤ b) (H2 : b < c) : a < c
axiom lt_le_trans (a b c : A) (H1 : a < b) (H2 : b ≤ c) : a < c axiom lt_le_trans (a b c : A) (H1 : a < b) (H2 : b ≤ c) : a < c
@ -41,7 +41,7 @@ check calc b ≤ c : H2
... < d : H5 ... < d : H5
constant le2 : A → A → bool constant le2 : A → A → bool
infixl ``:50 := le2 infixl ``:50 := le2
constant le2_trans (a b c : A) (H1 : le2 a b) (H2 : le2 b c) : le2 a c constant le2_trans (a b c : A) (H1 : le2 a b) (H2 : le2 b c) : le2 a c
attribute le2_trans [trans] attribute le2_trans [trans]
print raw calc b ≤ c : H2 print raw calc b ≤ c : H2

View file

@ -39,7 +39,7 @@ namespace PropF
definition valuation := PropVar → bool definition valuation := PropVar → bool
reserve infix ``:26 reserve infix ``:26
/- Provability -/ /- Provability -/

View file

@ -1,17 +1,17 @@
prelude prelude
definition Prop := Type.{0} inductive true : Prop := intro : true inductive false : Prop constant num : Type definition Prop := Type.{0} inductive true : Prop := intro : true inductive false : Prop constant num : Type
inductive prod (A B : Type) := mk : A → B → prod A B infixl `×`:30 := prod inductive prod (A B : Type) := mk : A → B → prod A B infixl ` × `:30 := prod
variables a b c : num variables a b c : num
section section
local notation `(` t:(foldr `,` (e r, prod.mk e r)) `)` := t local notation `(` t:(foldr `, ` (e r, prod.mk e r)) `)` := t
check (a, false, b, true, c) check (a, false, b, true, c)
set_option pp.notation false set_option pp.notation false
check (a, false, b, true, c) check (a, false, b, true, c)
end end
section section
local notation `(` t:(foldr `,` (e r, prod.mk r e)) `)` := t local notation `(` t:(foldr `, ` (e r, prod.mk r e)) `)` := t
set_option pp.notation true set_option pp.notation true
check (a, false, b, true, c) check (a, false, b, true, c)
set_option pp.notation false set_option pp.notation false
@ -19,7 +19,7 @@ section
end end
section section
local notation `(` t:(foldl `,` (e r, prod.mk r e)) `)` := t local notation `(` t:(foldl `, ` (e r, prod.mk r e)) `)` := t
set_option pp.notation true set_option pp.notation true
check (a, false, b, true, c) check (a, false, b, true, c)
set_option pp.notation false set_option pp.notation false
@ -27,7 +27,7 @@ section
end end
section section
local notation `(` t:(foldl `,` (e r, prod.mk e r)) `)` := t local notation `(` t:(foldl `, ` (e r, prod.mk e r)) `)` := t
set_option pp.notation true set_option pp.notation true
check (a, false, b, true, c) check (a, false, b, true, c)
set_option pp.notation false set_option pp.notation false

View file

@ -3,6 +3,6 @@ have e2 : a = c, from e1 ⬝ H2,
have e3 : c = a, from e2⁻¹, have e3 : c = a, from e2⁻¹,
assert e4 : b = a, from e1⁻¹, assert e4 : b = a, from e1⁻¹,
have e5 : b = c, from e4 ⬝ e2, have e5 : b = c, from e4 ⬝ e2,
have e6 : a = a, from H1 ⬝ H2 ⬝ H2 ⁻¹ ⬝ H1 ⁻¹ ⬝ H1 ⬝ H2 ⬝ H2 ⁻¹ ⬝ H1 ⁻¹, have e6 : a = a, from H1 ⬝ H2 ⬝ H2⁻¹ ⬝ H1⁻¹ ⬝ H1 ⬝ H2 ⬝ H2⁻¹ ⬝ H1⁻¹,
e3 ⬝ e2 : e3 ⬝ e2 :
c = c c = c

View file

@ -7,8 +7,8 @@ constant lst.nil {A : Type} : lst A
constant vec.cons {A : Type} : A → vec A → vec A constant vec.cons {A : Type} : A → vec A → vec A
constant lst.cons {A : Type} : A → lst A → lst A constant lst.cons {A : Type} : A → lst A → lst A
notation `[` l:(foldr `,` (h t, vec.cons h t) vec.nil `]`) := l notation `[` l:(foldr `, ` (h t, vec.cons h t) vec.nil `]`) := l
notation `[` l:(foldr `,` (h t, lst.cons h t) lst.nil `]`) := l notation `[` l:(foldr `, ` (h t, lst.cons h t) lst.nil `]`) := l
constant A : Type.{1} constant A : Type.{1}
variables a b c : A variables a b c : A

View file

@ -1,2 +1,2 @@
1 :: 2 :: nil : list num 1::2::nil : list num
1 :: 2 :: 3 :: 4 :: 5 :: nil : list num 1::2::3::4::5::nil : list num

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,4 +1,4 @@
[ a, b, b ] : list num [a, b, b] : list num
(a, true, a = b, b) : num × Prop × Prop × num (a, true, a = b, b) : num × Prop × Prop × num
(a, b) : num × num (a, b) : num × num
[ 1, 2 + 2, 3 ] : list num [1, 2 + 2, 3] : list num

View file

@ -1,2 +1,2 @@
g 0 :+1 :+1 (1 :+1 + 2 :+1) :+1 : num g 0:+1:+1 (1:+1 + 2:+1):+1 : num
g (f (f 0)) (f (add (f 1) (f 2))) : num g (f (f 0)) (f (add (f 1) (f 2))) : num

View file

@ -1,3 +1,3 @@
10 +++ : num 10+++ : num
g 10 : num g 10 : num
Type : Type Type : Type

View file

@ -8,7 +8,7 @@ constant q : B
constant x : N constant x : N
constant y : N constant y : N
constant z : N constant z : N
infixr ``:25 := and infixr ``:25 := and
notation `if` c `then` t:45 `else` e:45 := ite c t e notation `if` c `then` t:45 `else` e:45 := ite c t e
check if p ∧ q then f x else y check if p ∧ q then f x else y
check if p ∧ q then q else y check if p ∧ q then q else y
@ -16,7 +16,7 @@ constant list : Type.{1}
constant nil : list constant nil : list
constant cons : N → list → list constant cons : N → list → list
-- Non empty lists -- Non empty lists
notation `[` l:(foldr `,` (h t, cons h t) nil) `]` := l notation `[` l:(foldr `, ` (h t, cons h t) nil) `]` := l
check [x, y, z, x, y, y] check [x, y, z, x, y, y]
check [x] check [x]
notation `[` `]` := nil notation `[` `]` := nil

View file

@ -7,6 +7,6 @@ has type
B B
but is expected to have type but is expected to have type
N N
[ x, y, z, x, y, y ] : list [x, y, z, x, y, y] : list
[ x ] : list [x] : list
[ ] : list [] : list

View file

@ -1,8 +1,8 @@
prelude constant A : Type.{1} prelude constant A : Type.{1}
definition bool : Type.{1} := Type.{0} definition bool : Type.{1} := Type.{0}
constant Exists (P : A → bool) : bool constant Exists (P : A → bool) : bool
notation `exists` binders `,` b:(scoped b, Exists b) := b notation `exists` binders `, ` b:(scoped b, Exists b) := b
notation `∃` binders `,` b:(scoped b, Exists b) := b notation `∃` binders `, ` b:(scoped b, Exists b) := b
constant p : A → bool constant p : A → bool
constant q : A → A → bool constant q : A → A → bool
check exists x : A, p x check exists x : A, p x

View file

@ -32,7 +32,7 @@ end
namespace foo namespace foo
constant f : A → A → A constant f : A → A → A
infix `*`:75 := f infix ` * `:75 := f
end foo end foo
section section

View file

@ -1,3 +1,3 @@
a + b * a : N a+b*a : N
t9.lean:16:7: error: invalid expression t9.lean:16:7: error: invalid expression
a + b * a : N a+b*a : N