From 68aa069172933b875d70a5ef71e9db0ae685a92d Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Sun, 17 Feb 2019 18:34:38 +0100 Subject: Method "quantify" converts product inhabitation into Pure universal statements. Also misc. cleanups. --- Equal.thy | 95 +++++++++++++++++++++++++++++++++++++++++++++++++-------------- 1 file changed, 74 insertions(+), 21 deletions(-) (limited to 'Equal.thy') diff --git a/Equal.thy b/Equal.thy index 99ff268..1d8f6ae 100644 --- a/Equal.thy +++ b/Equal.thy @@ -1,52 +1,105 @@ -(* -Title: Equal.thy -Author: Joshua Chen -Date: 2018 +(******** +Isabelle/HoTT: Equality +Feb 2019 -Equality type -*) +********) theory Equal -imports HoTT_Base +imports Prod HoTT_Methods begin -section \Basic definitions\ +section \Type definitions\ axiomatization - Equal :: "[t, t, t] \ t" and - refl :: "t \ t" and - indEqual :: "[t \ t, t] \ t" ("(1ind\<^sub>=)") + Eq :: "[t, t, t] \ t" and + refl :: "t \ t" and + indEq :: "[[t, t, t] \ t, t \ t, t] \ t" syntax - "_equal" :: "[t, t, t] \ t" ("(3_ =\<^sub>_/ _)" [101, 0, 101] 100) - "_equal_ascii" :: "[t, t, t] \ t" ("(3_ =[_]/ _)" [101, 0, 101] 100) - + "_Eq" :: "[t, t, t] \ t" ("(3_ =[_]/ _)" [101, 0, 101] 100) translations - "a =[A] b" \ "CONST Equal A a b" - "a =\<^sub>A b" \ "CONST Equal A a b" + "a =[A] b" \ "(CONST Eq) A a b" axiomatization where - Equal_form: "\A: U i; a: A; b: A\ \ a =\<^sub>A b : U i" and + Equal_form: "\A: U i; a: A; b: A\ \ a =[A] b: U i" and - Equal_intro: "a : A \ (refl a): a =\<^sub>A a" and + Equal_intro: "a: A \ (refl a): a =[A] a" and Equal_elim: "\ - p: x =\<^sub>A y; + p: x =[A] y; x: A; y: A; \x. x: A \ f x: C x x (refl x); - \x y. \x: A; y: A\ \ C x y: x =\<^sub>A y \ U i \ \ ind\<^sub>= f p : C x y p" and + \x y. \x: A; y: A\ \ C x y: x =[A] y \ U i \ \ indEq C f p : C x y p" and Equal_comp: "\ a: A; \x. x: A \ f x: C x x (refl x); - \x y. \x: A; y: A\ \ C x y: x =\<^sub>A y \ U i \ \ ind\<^sub>= f (refl a) \ f a" + \x y. \x: A; y: A\ \ C x y: x =[A] y \ U i \ \ indEq C f (refl a) \ f a" lemmas Equal_form [form] lemmas Equal_routine [intro] = Equal_form Equal_intro Equal_elim lemmas Equal_comp [comp] +section \Path induction\ + +text \We set up rudimentary automation of path induction:\ + +method path_ind for x y p :: t = + (rule Equal_elim[where ?x=x and ?y=y and ?p=p]; (fact | assumption)?) + + +section \Properties of equality\ + +subsection \Symmetry (path inverse)\ + +definition inv :: "[t, t] \ t" +where "inv A p \ indEq (\x y. ^(y =[A] x)) (\x. refl x) p" + +lemma inv_type: "\A: U i; x: A; y: A; p: x =[A] y\ \ inv A p: y =[A] x" +unfolding inv_def by derive + +lemma inv_comp: "\A: U i; a: A\ \ inv A (refl a) \ refl a" +unfolding inv_def by derive + +declare + inv_type [intro] + inv_comp [comp] + +subsection \Transitivity (path composition)\ + +schematic_goal transitivity: + assumes "A: U i" "x: A" "y: A" "p: x =[A] y" + shows "?p: \z: A. y =[A]z \ x =[A] z" +proof (path_ind x y p, quantify 2) + fix x z show "\p. p: x =[A] z \ p: x =[A] z" . +qed (routine add: assms) + + +syntax "_pathcomp" :: "[t, t] \ t" (infixl "\" 110) +translations "p \ q" \ "CONST pathcomp`p`q" + +lemma pathcomp_type: + assumes "A: U i" "x: A" "y: A" "z: A" + shows "pathcomp: x =\<^sub>A y \ (y =\<^sub>A z) \ (x =\<^sub>A z)" +unfolding pathcomp_def by rule (elim Equal_elim, routine add: assms) + +corollary pathcomp_trans: + assumes "A: U i" and "x: A" "y: A" "z: A" and "p: x =\<^sub>A y" "q: y =\<^sub>A z" + shows "p \ q: x =\<^sub>A z" +by (routine add: assms pathcomp_type) + +lemma pathcomp_comp: + assumes "A: U i" and "a: A" + shows "(refl a) \ (refl a) \ refl a" +unfolding pathcomp_def by (derive lems: assms) + +declare + pathcomp_type [intro] + pathcomp_trans [intro] + pathcomp_comp [comp] + end -- cgit v1.2.3