2015-11-29 05:33:09 +00:00
|
|
|
|
-- Test [light] annotation
|
|
|
|
|
-- Remark: it will take some additional work to get ⁻¹ to rewrite well
|
|
|
|
|
-- when there is a proof obligation.
|
2015-12-05 01:57:03 +00:00
|
|
|
|
import algebra.ring algebra.field data.set data.finset
|
|
|
|
|
open algebra
|
2015-12-18 06:42:31 +00:00
|
|
|
|
attribute neg [light 3]
|
|
|
|
|
attribute inv [light 3]
|
2015-11-29 05:33:09 +00:00
|
|
|
|
|
|
|
|
|
attribute add.right_inv [simp]
|
|
|
|
|
attribute add_neg_cancel_left [simp]
|
|
|
|
|
|
|
|
|
|
attribute mul.right_inv [simp]
|
|
|
|
|
attribute mul_inv_cancel_left [simp]
|
|
|
|
|
|
2015-12-05 01:57:03 +00:00
|
|
|
|
open simplifier.unit simplifier.ac
|
|
|
|
|
|
2015-11-29 05:33:09 +00:00
|
|
|
|
namespace ag
|
|
|
|
|
universe l
|
|
|
|
|
constants (A : Type.{l}) (s1 : add_comm_group A) (s2 : has_one A)
|
|
|
|
|
attribute s1 [instance]
|
|
|
|
|
attribute s2 [instance]
|
|
|
|
|
constants (x y z w v : A)
|
|
|
|
|
|
|
|
|
|
#simplify eq env 0 x + y + - x + -y + z + -z
|
|
|
|
|
#simplify eq env 0 -100 + -v + -v + x + -v + y + - x + -y + z + -z + v + v + v + 100
|
|
|
|
|
end ag
|
|
|
|
|
|
|
|
|
|
namespace mg
|
|
|
|
|
universe l
|
|
|
|
|
constants (A : Type.{l}) (s1 : comm_group A) (s2 : has_add A)
|
|
|
|
|
attribute s1 [instance]
|
|
|
|
|
attribute s2 [instance]
|
|
|
|
|
constants (x y z w v : A)
|
|
|
|
|
|
|
|
|
|
#simplify eq env 0 x⁻¹ * y⁻¹ * z⁻¹ * 100⁻¹ * x * y * z * 100
|
|
|
|
|
|
|
|
|
|
end mg
|
|
|
|
|
|
|
|
|
|
namespace s
|
|
|
|
|
open set
|
|
|
|
|
universe l
|
|
|
|
|
constants (A : Type.{l}) (x y z v w : set A)
|
2015-12-18 06:42:31 +00:00
|
|
|
|
attribute complement [light 2]
|
2015-11-29 05:33:09 +00:00
|
|
|
|
|
|
|
|
|
-- TODO(dhs, leo): Where do we put this group of simp rules?
|
|
|
|
|
attribute union_comp_self [simp]
|
|
|
|
|
lemma union_comp_self_left [simp] {X : Type} (s t : set X) : s ∪ (-s ∪ t)= univ := sorry
|
|
|
|
|
|
2016-01-02 20:51:20 +00:00
|
|
|
|
attribute union_comm [simp]
|
|
|
|
|
attribute union_assoc [simp]
|
|
|
|
|
attribute union_left_comm [simp]
|
2015-11-29 05:33:09 +00:00
|
|
|
|
|
|
|
|
|
#simplify eq env 0 x ∪ y ∪ z ∪ -x
|
|
|
|
|
|
|
|
|
|
attribute inter_comp_self [simp]
|
|
|
|
|
lemma inter_comp_self_left [simp] {X : Type} (s t : set X) : s ∩ (-s ∩ t)= empty := sorry
|
|
|
|
|
|
2016-01-02 20:51:20 +00:00
|
|
|
|
attribute inter_comm [simp]
|
|
|
|
|
attribute inter_assoc [simp]
|
|
|
|
|
attribute inter_left_comm [simp]
|
2015-11-29 05:33:09 +00:00
|
|
|
|
|
|
|
|
|
#simplify eq env 0 x ∩ y ∩ z ∩ -x
|
|
|
|
|
|
|
|
|
|
end s
|