(* Title: ex/Book/Ch1.thy Author: Josh Chen Date: 2018 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)\ paragraph \Propositional uniqueness principle.\ schematic_goal assumes "A: U i" and "B: A \ U i" and "p: \x:A. B x" shows "?a: p =[\x:A. B x] " text \Proof by induction on @{term "p: \x:A. B x"}:\ proof (rule Sum_elim[where ?p=p]) text \We prove the base case.\ fix x y assume asm: "x: A" "y: B x" show "refl : =[\x:A. B x] , snd >" proof compute show "x: A" and "y: B x" by (fact asm)+ \ \Hint the correct types.\ text \And now @{method derive} takes care of the rest. \ qed (derive lems: assms asm) qed (derive lems: assms) section \Exercises\ paragraph \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) end