2014-01-24 06:31:02 +00:00
|
|
|
variable vec : Nat → Type
|
|
|
|
variable concat {n m : Nat} (v : vec n) (w : vec m) : vec (n + m)
|
|
|
|
infixl 65 ; : concat
|
|
|
|
axiom concat_assoc {n1 n2 n3 : Nat} (v1 : vec n1) (v2 : vec n2) (v3 : vec n3) :
|
2014-02-07 23:03:16 +00:00
|
|
|
(v1 ; v2) ; v3 = cast (to_heq (congr2 vec (symm (Nat::add_assoc n1 n2 n3))))
|
2014-01-24 06:31:02 +00:00
|
|
|
(v1 ; (v2 ; v3))
|
|
|
|
variable empty : vec 0
|
|
|
|
axiom concat_empty {n : Nat} (v : vec n) :
|
2014-02-07 23:03:16 +00:00
|
|
|
v ; empty = cast (to_heq (congr2 vec (symm (Nat::add_zeror n))))
|
2014-01-24 06:31:02 +00:00
|
|
|
v
|
|
|
|
|
|
|
|
rewrite_set simple
|
|
|
|
-- The simplification rules used for Nat and Vectors should "mirror" each other.
|
|
|
|
-- Concatenation is not commutative. So, by adding Nat::add_comm to the
|
|
|
|
-- rewrite set, we prevent the simplifier from reducing the following example
|
|
|
|
add_rewrite concat_assoc concat_empty Nat::add_assoc Nat::add_zeror Nat::add_comm : simple
|
|
|
|
|
2014-02-07 23:03:16 +00:00
|
|
|
universe M >= 1
|
|
|
|
definition TypeM := (Type M)
|
|
|
|
|
2014-01-24 06:31:02 +00:00
|
|
|
variable n : Nat
|
|
|
|
variable v : vec n
|
|
|
|
variable w : vec n
|
|
|
|
variable f {A : TypeM} : A → A
|
|
|
|
|
|
|
|
(*
|
2014-02-04 22:42:28 +00:00
|
|
|
local opts = options({"simplifier", "heq"}, true)
|
2014-01-24 06:31:02 +00:00
|
|
|
local t = parse_lean([[ f ((v ; w) ; empty ; (v ; empty)) ]])
|
|
|
|
print(t)
|
|
|
|
print("===>")
|
2014-02-04 22:42:28 +00:00
|
|
|
local t2, pr = simplify(t, "simple", opts)
|
2014-01-24 06:31:02 +00:00
|
|
|
print(t2)
|
|
|
|
print(pr)
|
|
|
|
get_environment():type_check(pr)
|
|
|
|
*)
|
|
|
|
|
|
|
|
-- Now, if we disable Nat::add_comm, the simplifier works
|
|
|
|
disable_rewrite Nat::add_comm : simple
|
|
|
|
print "After disabling Nat::add_comm"
|
|
|
|
|
|
|
|
(*
|
2014-02-04 22:42:28 +00:00
|
|
|
local opts = options({"simplifier", "heq"}, true)
|
2014-01-24 06:31:02 +00:00
|
|
|
local t = parse_lean([[ f ((v ; w) ; empty ; (v ; empty)) ]])
|
|
|
|
print(t)
|
|
|
|
print("===>")
|
2014-02-04 22:42:28 +00:00
|
|
|
local t2, pr = simplify(t, "simple", opts)
|
2014-01-24 06:31:02 +00:00
|
|
|
print(t2)
|
|
|
|
print(pr)
|
|
|
|
get_environment():type_check(pr)
|
|
|
|
*)
|