From 7823d59b2d9436f1bf0042fff62ee2bcc4c29ec0 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Tue, 3 Jul 2018 18:57:57 +0200 Subject: Refactor HoTT_Methods.thy, proved more stuff with new methods. --- ex/Methods.thy | 41 +++++++++++++++++++++++++++++++++++++++++ 1 file changed, 41 insertions(+) create mode 100644 ex/Methods.thy (limited to 'ex/Methods.thy') diff --git a/ex/Methods.thy b/ex/Methods.thy new file mode 100644 index 0000000..d174dde --- /dev/null +++ b/ex/Methods.thy @@ -0,0 +1,41 @@ +(* Title: HoTT/ex/Methods.thy + Author: Josh Chen + Date: Jul 2018 + +HoTT method usage examples +*) + +theory Methods + imports "../HoTT" +begin + +lemma + assumes "A : U" "B: A \ U" "\x. x : A \ C x: B x \ U" + shows "\x:A. \y:B x. \z:C x y. \w:A. x =\<^sub>A w : U" +by (simple lems: assms) + + +lemma + assumes "f : \x:A. \y: B x. \z: C x y. D x y z" + shows + "A : U" and + "B: A \ U" and + "\x. x : A \ C x: B x \ U" and + "\x y. \x : A; y : B x\ \ D x y: C x y \ U" +proof - + show "A : U" by (wellformed jdgmt: assms) + show "B: A \ U" by (wellformed jdgmt: assms) + show "\x. x : A \ C x: B x \ U" by (wellformed jdgmt: assms) + show "\x y. \x : A; y : B x\ \ D x y: C x y \ U" by (wellformed jdgmt: assms) +qed + + +text "Typechecking:" + +\ \Correctly determines the type of the pair\ +schematic_goal "\a : A; b : B\ \ (a, b) : ?A" by simple + +\ \Finds element\ +schematic_goal "\a : A; b : B\ \ ?x : A \ B" by simple + +end \ No newline at end of file -- cgit v1.2.3