From ca8e7a2681c133abdd082cfa29d6994fa73f2d47 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Wed, 15 Aug 2018 17:17:34 +0200 Subject: Rename to distinguish function and path composition; function composition proofs, which have issues... --- EqualProps.thy | 40 ++++++++++++++++++++-------------------- 1 file changed, 20 insertions(+), 20 deletions(-) (limited to 'EqualProps.thy') diff --git a/EqualProps.thy b/EqualProps.thy index 2a13ed2..f3b355a 100644 --- a/EqualProps.thy +++ b/EqualProps.thy @@ -48,17 +48,17 @@ text " Raw composition function, of type \\x:A. \y:A. x =\<^sub>A y \ (\z:A. y =\<^sub>A z \ x =\<^sub>A z)\ polymorphic over the type \A\. " -axiomatization reqcompose :: Term where - reqcompose_def: "reqcompose \ \<^bold>\x y p. ind\<^sub>= (\_. \<^bold>\z q. ind\<^sub>= (\x. refl(x)) q) p" +axiomatization rpathcomp :: Term where + rpathcomp_def: "rpathcomp \ \<^bold>\x y p. ind\<^sub>= (\_. \<^bold>\z q. ind\<^sub>= (\x. refl(x)) q) p" text " More complicated proofs---the nested path inductions require more explicit step-by-step rule applications: " -lemma reqcompose_type: +lemma rpathcomp_type: assumes "A: U(i)" - shows "reqcompose: \x:A. \y:A. x =\<^sub>A y \ (\z:A. y =\<^sub>A z \ x =\<^sub>A z)" -unfolding reqcompose_def + shows "rpathcomp: \x:A. \y:A. x =\<^sub>A y \ (\z:A. y =\<^sub>A z \ x =\<^sub>A z)" +unfolding rpathcomp_def proof fix x assume 1: "x: A" show "\<^bold>\y p. ind\<^sub>= (\_. \<^bold>\z q. ind\<^sub>= refl q) p: \y:A. x =\<^sub>A y \ (\z:A. y =\<^sub>A z \ x =\<^sub>A z)" @@ -85,17 +85,17 @@ qed fact corollary assumes "A: U(i)" "x: A" "y: A" "z: A" "p: x =\<^sub>A y" "q: y =\<^sub>A z" - shows "reqcompose`x`y`p`z`q: x =\<^sub>A z" - by (simple lems: assms reqcompose_type) + shows "rpathcomp`x`y`p`z`q: x =\<^sub>A z" + by (simple lems: assms rpathcomp_type) text " The following proof is very long, chiefly because for every application of \`\ we have to show the wellformedness of the type family appearing in the equality computation rule. " -lemma reqcompose_comp: +lemma rpathcomp_comp: assumes "A: U(i)" and "a: A" - shows "reqcompose`a`a`refl(a)`a`refl(a) \ refl(a)" -unfolding reqcompose_def + shows "rpathcomp`a`a`refl(a)`a`refl(a) \ refl(a)" +unfolding rpathcomp_def proof compute { fix x assume 1: "x: A" show "\<^bold>\y p. ind\<^sub>= (\_. \<^bold>\z q. ind\<^sub>= refl q) p: \y:A. x =\<^sub>A y \ (\z:A. y =\<^sub>A z \ x =\<^sub>A z)" @@ -197,28 +197,28 @@ qed fact text "The raw object lambda term is cumbersome to use, so we define a simpler constant instead." -axiomatization eqcompose :: "[Term, Term] \ Term" (infixl "\" 60) where - eqcompose_def: "\ +axiomatization pathcomp :: "[Term, Term] \ Term" (infixl "\" 60) where + pathcomp_def: "\ A: U(i); x: A; y: A; z: A; p: x =\<^sub>A y; q: y =\<^sub>A z - \ \ p \ q \ reqcompose`x`y`p`z`q" + \ \ p \ q \ rpathcomp`x`y`p`z`q" -lemma eqcompose_type: +lemma pathcomp_type: assumes "A: U(i)" "x: A" "y: A" "z: A" "p: x =\<^sub>A y" "q: y =\<^sub>A z" shows "p \ q: x =\<^sub>A z" -proof (subst eqcompose_def) +proof (subst pathcomp_def) show "A: U(i)" "x: A" "y: A" "z: A" "p: x =\<^sub>A y" "q: y =\<^sub>A z" by fact+ -qed (simple lems: assms reqcompose_type) +qed (simple lems: assms rpathcomp_type) -lemma eqcompose_comp: +lemma pathcomp_comp: assumes "A : U(i)" and "a : A" shows "refl(a) \ refl(a) \ refl(a)" -by (subst eqcompose_def) (simple lems: assms reqcompose_comp) +by (subst pathcomp_def) (simple lems: assms rpathcomp_comp) -lemmas EqualProps_rules [intro] = inv_type eqcompose_type -lemmas EqualProps_comps [comp] = inv_comp eqcompose_comp +lemmas EqualProps_rules [intro] = inv_type pathcomp_type +lemmas EqualProps_comps [comp] = inv_comp pathcomp_comp end \ No newline at end of file -- cgit v1.2.3