(* Title: ex/Methods.thy Author: Joshua Chen Date: 2018 Basic HoTT method usage examples. *) theory Methods imports "../HoTT" begin 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 (routine add: assms) \ \Correctly determines the type of the pair.\ schematic_goal "\A: U(i); B: U(i); a : A; b : B\ \ : ?A" by routine \ \Finds pair (too easy).\ schematic_goal "\A: U(i); B: U(i); a : A; b : B\ \ ?x : A \ B" apply (rule intros) apply assumption+ done \ \Function application. We still often have to explicitly specify types.\ lemma "\A: U i; a: A\ \ (\<^bold>\x. )`a \ " proof compute show "\x. x: A \ : A \ \" by routine qed text \ The proof below takes a little more work than one might expect; it would be nice to have a one-line method or proof. \ lemma "\A: U i; B: A \ U i; a: A; b: B a\ \ (\<^bold>\x y. )`a`b \ " proof (compute, routine) show "\A: U i; B: A \ U i; a: A; b: B a\ \ (\<^bold>\y. )`b \ " proof compute show "\b. \A: U i; B: A \ U i; a: A; b: B a\ \ : \x:A. B x" by routine qed qed end