fix(library/unifier): fixes #809
Daniel is correct when he says the interaction between choice
case-splits, delta case-splits, and coercions can be subtle.
I believe the following condition
https://github.com/leanprover/lean/blob/master/src/frontends/lean/elaborator.cpp#L111
reduces counter-intuitive behavior. Example, the coercion should not
influence the resulting type.
BTW, by removing this condition, many files in the library broke when I
tried to compile from scratch
make clean-olean
make
I used the following workaround. Given a delta-delta constraint
f a =?= f b
If the terms are types, and no case-split will be performed, then
the delta-delta constraint is eagerly solved.
In principle, we don't need the condition that the terms are types.
However, many files break if we remove it. The problem is that many files in the standard
library are abusing the higher-order unification procedure. The
elaboration problems are quite tricky to solve.
I use the extra condition "the terms are types" because usually if they
are, "f" is morally injective, and we don't really want to unfold it.
Note that the following two cases do not work
check '{1, 2, 3}
check insert 1 (insert 2 (insert 3 empty))
Well, they work if we the num namespace is open, and they are
interpreted as having type (finset num)
2015-09-01 03:47:50 +00:00
|
|
|
|
import data.finset data.num data.nat data.int algebra.ring
|
|
|
|
|
open finset num nat int algebra
|
|
|
|
|
|
|
|
|
|
check @finset.insert _ _ 1 (@finset.empty ℕ)
|
|
|
|
|
|
2015-10-13 22:39:03 +00:00
|
|
|
|
check '{(1:nat), 2, 3}
|
fix(library/unifier): fixes #809
Daniel is correct when he says the interaction between choice
case-splits, delta case-splits, and coercions can be subtle.
I believe the following condition
https://github.com/leanprover/lean/blob/master/src/frontends/lean/elaborator.cpp#L111
reduces counter-intuitive behavior. Example, the coercion should not
influence the resulting type.
BTW, by removing this condition, many files in the library broke when I
tried to compile from scratch
make clean-olean
make
I used the following workaround. Given a delta-delta constraint
f a =?= f b
If the terms are types, and no case-split will be performed, then
the delta-delta constraint is eagerly solved.
In principle, we don't need the condition that the terms are types.
However, many files break if we remove it. The problem is that many files in the standard
library are abusing the higher-order unification procedure. The
elaboration problems are quite tricky to solve.
I use the extra condition "the terms are types" because usually if they
are, "f" is morally injective, and we don't really want to unfold it.
Note that the following two cases do not work
check '{1, 2, 3}
check insert 1 (insert 2 (insert 3 empty))
Well, they work if we the num namespace is open, and they are
interpreted as having type (finset num)
2015-09-01 03:47:50 +00:00
|
|
|
|
check ('{1, 2, 3} : finset ℕ)
|
|
|
|
|
check ('{1, 2, 3} : finset ℕ) -- finset ℕ
|
|
|
|
|
check ('{1, 2, 3} : finset ℤ) -- finset ℤ
|
|
|
|
|
|
|
|
|
|
example : finset nat :=
|
|
|
|
|
insert 1 (insert 2 (insert 3 empty))
|
|
|
|
|
|
|
|
|
|
check insert 1 (insert 2 (insert 3 empty)) -- finset num
|
|
|
|
|
check (insert 1 (insert 2 (insert 3 empty)) : finset nat)
|
|
|
|
|
|
|
|
|
|
check (insert (1:nat) (insert (2:nat) (insert (3:nat) empty)))
|
|
|
|
|
|
|
|
|
|
definition foo_nat (x : finset ℕ) : finset ℕ := x
|
|
|
|
|
definition foo_int (x : finset ℤ) : finset ℤ := x
|
|
|
|
|
|
|
|
|
|
check foo_nat '{1, 2, 3} -- finset ℕ
|
|
|
|
|
check foo_int '{1, 2, 3} -- finset ℤ
|