From 593faab277de53cbe2cb0c2feca5de307d9334ac Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Sat, 9 Jun 2018 00:11:39 +0200 Subject: Reorganize code --- Equal.thy | 81 +++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 81 insertions(+) create mode 100644 Equal.thy (limited to 'Equal.thy') diff --git a/Equal.thy b/Equal.thy new file mode 100644 index 0000000..b9f676f --- /dev/null +++ b/Equal.thy @@ -0,0 +1,81 @@ +theory Equal + imports HoTT_Base Prod + +begin + +subsection \Equality type\ + + axiomatization + Equal :: "[Term, Term, Term] \ Term" + + syntax + "_EQUAL" :: "[Term, Term, Term] \ Term" ("(3_ =\<^sub>_/ _)" [101, 101] 100) + "_EQUAL_ASCII" :: "[Term, Term, Term] \ Term" ("(3_ =[_]/ _)" [101, 0, 101] 100) + translations + "a =[A] b" \ "CONST Equal A a b" + "a =\<^sub>A b" \ "CONST Equal A a b" + + axiomatization + refl :: "Term \ Term" ("(refl'(_'))") and + indEqual :: "[Term, [Term, Term, Term] \ Term] \ Term" ("(indEqual[_])") + where + Equal_form: "\A a b::Term. \A : U; a : A; b : A\ \ a =\<^sub>A b : U" + (* Should I write a permuted version \\A : U; b : A; a : A\ \ \\? *) + and + Equal_intro [intro]: "\A x::Term. x : A \ refl(x) : x =\<^sub>A x" + and + Equal_elim [elim]: + "\(A::Term) (C::[Term, Term, Term] \ Term) (f::Term) (a::Term) (b::Term) (p::Term). + \ \x y::Term. \x : A; y : A\ \ C(x)(y): x =\<^sub>A y \ U; + f : \x:A. C(x)(x)(refl(x)); + a : A; + b : A; + p : a =\<^sub>A b \ + \ indEqual[A](C)`f`a`b`p : C(a)(b)(p)" + and + Equal_comp [simp]: + "\(A::Term) (C::[Term, Term, Term] \ Term) (f::Term) (a::Term). indEqual[A](C)`f`a`a`refl(a) \ f`a" + + lemmas Equal_formation [intro] = Equal_form Equal_form[rotated 1] Equal_form[rotated 2] + + subsubsection \Properties of equality\ + + text "Symmetry/Path inverse" + + definition inv :: "[Term, Term, Term] \ Term" ("(1inv[_,/ _,/ _])") + where "inv[A,x,y] \ indEqual[A](\x y _. y =\<^sub>A x)`(\<^bold>\x:A. refl(x))`x`y" + + lemma inv_comp: "\A a::Term. a : A \ inv[A,a,a]`refl(a) \ refl(a)" unfolding inv_def by simp + + text "Transitivity/Path composition" + + \ \"Raw" composition function\ + definition compose' :: "Term \ Term" ("(1compose''[_])") + where "compose'[A] \ indEqual[A](\x y _. \z:A. \q: y =\<^sub>A z. x =\<^sub>A z)`(indEqual[A](\x z _. x =\<^sub>A z)`(\<^bold>\x:A. refl(x)))" + + \ \"Natural" composition function\ + abbreviation compose :: "[Term, Term, Term, Term] \ Term" ("(1compose[_,/ _,/ _,/ _])") + where "compose[A,x,y,z] \ \<^bold>\p:x =\<^sub>A y. \<^bold>\q:y =\<^sub>A z. compose'[A]`x`y`p`z`q" + + (**** GOOD CANDIDATE FOR AUTOMATION ****) + lemma compose_comp: + assumes "a : A" + shows "compose[A,a,a,a]`refl(a)`refl(a) \ refl(a)" using assms Equal_intro[OF assms] unfolding compose'_def by simp + + text "The above proof is a good candidate for proof automation; in particular we would like the system to be able to automatically find the conditions of the \using\ clause in the proof. + This would likely involve something like: + 1. Recognizing that there is a function application that can be simplified. + 2. Noting that the obstruction to applying \Prod_comp\ is the requirement that \refl(a) : a =\<^sub>A a\. + 3. Obtaining such a condition, using the known fact \a : A\ and the introduction rule \Equal_intro\." + + lemmas Equal_simps [simp] = inv_comp compose_comp + + subsubsection \Pretty printing\ + + abbreviation inv_pretty :: "[Term, Term, Term, Term] \ Term" ("(1_\<^sup>-\<^sup>1[_, _, _])" 500) + where "p\<^sup>-\<^sup>1[A,x,y] \ inv[A,x,y]`p" + + abbreviation compose_pretty :: "[Term, Term, Term, Term, Term, Term] \ Term" ("(1_ \[_, _, _, _]/ _)") + where "p \[A,x,y,z] q \ compose[A,x,y,z]`p`q" + +end \ No newline at end of file -- cgit v1.2.3