(* Title: HoTT/ProdProps.thy Author: Josh Chen Date: Aug 2018 Properties of the dependent product. *) theory ProdProps imports HoTT_Methods Prod begin section \Composition\ text " The proof of associativity is surprisingly intricate; it involves telling Isabelle to use the correct rule for \-type judgmental equality, and the correct substitutions in the subgoals thereafter, among other things. " lemma compose_assoc: assumes "A \ B: U(i)" "B \ C: U(i)" and "\x:C. D(x): U(i)" and "f: A \ B" "g: B \ C" "h: \x:C. D(x)" shows "(h \ g) \ f \ h \ (g \ f)" proof (subst (0 1 2 3) compose_def) text "Compute the different bracketings and show they are equivalent:" show "\<^bold>\x. (\<^bold>\y. h`(g`y))`(f`x) \ \<^bold>\x. h`((\<^bold>\y. g`(f`y))`x)" proof (subst Prod_eq) show "\<^bold>\x. h`((\<^bold>\y. g`(f`y))`x) \ \<^bold>\x. h`((\<^bold>\y. g`(f`y))`x)" by simp show "\x. x: A \ (\<^bold>\y. h`(g`y))`(f`x) \ h`((\<^bold>\y. g`(f`y))`x)" proof compute show "\x. x: A \ h`(g`(f`x)) \ h`((\<^bold>\y. g`(f`y))`x)" proof compute show "\x. x: A \ g`(f`x): C" by (simple lems: assms) qed show "\x. x: B \ h`(g`x): D(g`x)" by (simple lems: assms) qed (simple lems: assms) qed (wellformed lems: assms) text " Finish off any remaining subgoals that Isabelle can't prove automatically. These typically require the application of specific substitutions or computation rules. " show "g \ f: A \ C" by (subst compose_def) (derive lems: assms) show "h \ g: \x:B. D(g`x)" proof (subst compose_def) show "\<^bold>\x. h`(g`x): \x:B. D(g`x)" proof show "\x. x: B \ h`(g`x): D(g`x)" proof show "\x. x: B \ g`x: C" by (simple lems: assms) qed (simple lems: assms) qed (wellformed lems: assms) qed fact+ show "\x:B. D (g`x): U(i)" proof show "\x. x: B \ D(g`x): U(i)" proof - have *: "D: C \ U(i)" by (wellformed lems: assms) fix x assume asm: "x: B" have "g`x: C" by (simple lems: assms asm) then show "D(g`x): U(i)" by (rule *) qed qed (wellformed lems: assms) have "A: U(i)" by (wellformed lems: assms) moreover have "C: U(i)" by (wellformed lems: assms) ultimately show "A \ C: U(i)" .. qed (simple lems: assms) text "The following rule is correct, but we cannot prove it using just the rules of the theory." lemma compose_comp': assumes "A: U(i)" and "\x. x: A \ b(x): B" and "\x. x: B \ c(x): C(x)" shows "(\<^bold>\x. c(x)) \ (\<^bold>\x. b(x)) \ \<^bold>\x. c(b(x))" oops text "However we can derive a variant with more explicit premises:" lemma compose_comp: assumes "A: U(i)" "B: U(i)" "C: B \ U(i)" and "\x. x: A \ b(x): B" and "\x. x: B \ c(x): C(x)" shows "(\<^bold>\x. c(x)) \ (\<^bold>\x. b(x)) \ \<^bold>\x. c(b(x))" proof (subst compose_def) show "\<^bold>\x. (\<^bold>\x. c(x))`((\<^bold>\x. b(x))`x) \ \<^bold>\x. c(b(x))" proof show "\a. a: A \ (\<^bold>\x. c(x))`((\<^bold>\x. b(x))`a) \ c(b(a))" proof compute show "\a. a: A \ c((\<^bold>\x. b(x))`a) \ c(b(a))" by compute (simple lems: assms) qed (simple lems: assms) qed fact qed (simple lems: assms) lemmas compose_comps [comp] = compose_def compose_comp end