From 6857e783fa5cb91f058be322a18fb9ea583f2aad Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Tue, 18 Sep 2018 11:38:54 +0200 Subject: Overhaul of the theory presentations. New methods in HoTT_Methods.thy for handling universes. Commit for release 0.1.0! --- ex/Methods.thy | 73 ++++++++++++++++++---------------------------------------- 1 file changed, 23 insertions(+), 50 deletions(-) (limited to 'ex/Methods.thy') diff --git a/ex/Methods.thy b/ex/Methods.thy index c78af14..09975b0 100644 --- a/ex/Methods.thy +++ b/ex/Methods.thy @@ -1,76 +1,49 @@ -(* Title: HoTT/ex/Methods.thy - Author: Josh Chen +(* +Title: ex/Methods.thy +Author: Joshua Chen +Date: 2018 -HoTT method usage examples +Basic HoTT method usage examples. *) theory Methods - imports "../HoTT" -begin +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 (routine 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:" + 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\ +\ \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 Sum_intro) +apply (rule intros) 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 \ " +\ \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 (routine lems: assms) - +qed -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 (routine lems: assms) +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. +\ - show "(\<^bold>\b. )`b \ " +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. b: B(a) \ : \x:A. B(x)" by (routine lems: assms) - qed fact -qed fact + show "\b. \A: U i; B: A \ U i; a: A; b: B a\ \ : \x:A. B x" by routine + qed +qed end -- cgit v1.2.3