(* Title: HoTT/ex/Methods.thy Author: Josh Chen Date: Aug 2018 HoTT method usage examples *) theory Methods imports "../HoTT" begin text "Wellformedness results, metatheorems written into the object theory using the wellformedness rules." lemma assumes "A : U(i)" "B: A \ U(i)" "\x. x : A \ C x: B x \ U(i)" shows "\x:A. \y:B x. \z:C x y. \w:A. x =\<^sub>A w : U(i)" by (simple lems: assms) lemma assumes "\x:A. \y: B x. \z: C x y. D x y z: U(i)" shows "A : U(i)" and "B: A \ U(i)" and "\x. x : A \ C x: B x \ U(i)" and "\x y. \x : A; y : B x\ \ D x y: C x y \ U(i)" proof - show "A : U(i)" and "B: A \ U(i)" and "\x. x : A \ C x: B x \ U(i)" and "\x y. \x : A; y : B x\ \ D x y: C x y \ U(i)" by (derive lems: assms) qed text "Typechecking and constructing inhabitants:" \ \Correctly determines the type of the pair\ schematic_goal "\A: U(i); B: U(i); a : A; b : B\ \ : ?A" by simple \ \Finds pair (too easy).\ schematic_goal "\A: U(i); B: U(i); a : A; b : B\ \ ?x : A \ B" apply (rule Sum_intro) apply assumption+ done text " Function application. The proof methods are not yet automated as well as I would like; we still often have to explicitly specify types. " lemma assumes "A: U(i)" "a: A" shows "(\<^bold>\x. )`a \ " proof compute show "\x. x: A \ : A \ \" by simple qed (simple lems: assms) lemma assumes "A: U(i)" "B: A \ U(i)" "a: A" "b: B(a)" shows "(\<^bold>\x y. )`a`b \ " proof compute show "\x. x: A \ \<^bold>\y. : \y:B(x). \x:A. B(x)" by (simple lems: assms) show "(\<^bold>\b. )`b \ " proof compute show "\b. b: B(a) \ : \x:A. B(x)" by (simple lems: assms) qed fact qed fact end