(* Title: HoTT/ex/HoTT book/Ch1.thy Author: Josh Chen A formalization of some content of Chapter 1 of the Homotopy Type Theory book. *) theory Ch1 imports "../../HoTT" begin chapter \HoTT Book, Chapter 1\ section \1.6 Dependent pair types (\-types)\ text "Propositional uniqueness principle:" schematic_goal assumes "(\x:A. B(x)): U(i)" and "p: \x:A. B(x)" shows "?a: p =[\x:A. B(x)] " text "Proof by induction on \p: \x:A. B(x)\:" proof (rule Sum_elim[where ?p=p]) text "We just need to prove the base case; the rest will be taken care of automatically." fix x y assume asm: "x: A" "y: B(x)" show "refl(): =[\x:A. B(x)] , snd >" proof (subst (0 1) comp) text " The computation rules for \fst\ and \snd\ require that \x\ and \y\ have appropriate types. The automatic proof methods have trouble picking the appropriate types, so we state them explicitly, " show "x: A" and "y: B(x)" by (fact asm)+ text "...twice, once each for the substitutions of \fst\ and \snd\." show "x: A" and "y: B(x)" by (fact asm)+ qed (derive lems: assms asm) qed (derive lems: assms) section \Exercises\ text "Exercise 1.13" abbreviation "not" ("\'(_')") where "\(A) \ A \ \" text "This proof requires the use of universe cumulativity." proposition assumes "A: U(i)" shows "\<^bold>\f. f`(inr(\<^bold>\a. f`inl(a))): \(\(A + \(A)))" by (derive lems: assms U_cumulative[where ?A=\ and ?i=O and ?j=i]) end