From 10757b7f628655f962b2dd1c7849c75098320ed1 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Fri, 17 Aug 2018 10:48:42 +0200 Subject: Testing with the new theory --- HoTT_Test.thy | 54 +++++++++++++++++++++++++----------------------------- 1 file changed, 25 insertions(+), 29 deletions(-) diff --git a/HoTT_Test.thy b/HoTT_Test.thy index bce2f00..67284cf 100644 --- a/HoTT_Test.thy +++ b/HoTT_Test.thy @@ -2,8 +2,8 @@ Author: Josh Chen Date: Aug 2018 -This is an old "test suite" from early implementations of the theory, updated for the most recent version. -It is not always guaranteed to be up to date. +This is an old "test suite" from early implementations of the theory. +It is not always guaranteed to be up to date, or reflect most recent versions of the theory. *) theory HoTT_Test @@ -50,8 +50,6 @@ by (simple lems: assms) subsection \Function application\ -text "silly test" - proposition assumes "a : A" shows "(\<^bold>\x. x)`a \ a" by (simple lems: assms) text "Currying:" @@ -63,35 +61,30 @@ proof show "\x. x : A \ \<^bold>\y. y : B(x) \ B(x)" by (simple lems: assms) qed fact -lemma "\a : A; b : B\ \ (\<^bold>\x:A. \<^bold>\y:B(x). y)`a`b \ b" by simp +lemma "\A: U(i); B: U(i); a : A; b : B\ \ (\<^bold>\x y. y)`a`b \ b" by (compute, simple) -lemma "a : A \ (\<^bold>\x:A. \<^bold>\y:B(x). f x y)`a \ \<^bold>\y:B(a). f a y" by simp +lemma "\A: U(i); a : A \ \ (\<^bold>\x y. f x y)`a \ \<^bold>\y. f a y" +proof compute + show "\x. \A: U(i); x: A\ \ \<^bold>\y. f x y: \y:B(x). C x y" + proof + oops -lemma "\a : A; b : B(a); c : C(a)(b)\ \ (\<^bold>\x:A. \<^bold>\y:B(x). \<^bold>\z:C(x)(y). f x y z)`a`b`c \ f a b c" by simp +lemma "\a : A; b : B(a); c : C(a)(b)\ \ (\<^bold>\x. \<^bold>\y. \<^bold>\z. f x y z)`a`b`c \ f a b c" + oops subsection \Currying functions\ proposition curried_function_formation: - fixes - A::Term and - B::"Term \ Term" and - C::"Term \ Term \ Term" + fixes A B C assumes - "A : U" and - "B: A \ U" and - "\x::Term. C(x): B(x) \ U" - shows "\x:A. \y:B(x). C x y : U" -proof - fix x::Term - assume *: "x : A" - show "\y:B(x). C x y : U" - proof - show "B(x) : U" using * by (rule assms) - qed (rule assms) -qed (rule assms) + "A : U(i)" and + "B: A \ U(i)" and + "\x. C(x): B(x) \ U(i)" + shows "\x:A. \y:B(x). C x y : U(i)" + by (simple lems: assms) + -(**** GOOD CANDIDATE FOR AUTOMATION - EISBACH! ****) proposition higher_order_currying_formation: fixes A::Term and @@ -138,8 +131,10 @@ proof qed qed -text "Note that the propositions and proofs above often say nothing about the well-formedness of the types, or the well-typedness of the lambdas involved; one has to be very explicit and prove such things separately! -This is the result of the choices made regarding the premises of the type rules." +text " + Note that the propositions and proofs above often say nothing about the well-formedness of the types, or the well-typedness of the lambdas involved; one has to be very explicit and prove such things separately! + This is the result of the choices made regarding the premises of the type rules. +" section \\ type\ @@ -189,13 +184,14 @@ have "p : \x:A. B(x) \ indSum(C)`f`p : C(p)" using * ** by end + section \Universes and polymorphism\ -text "Polymorphic identity function." +text "Polymorphic identity function. Uses universe types." + -consts Ui::Term -definition Id where "Id \ \<^bold>\A:Ui. \<^bold>\x:A. x" +definition Id :: "Ord \ Term" where "Id(i) \ \<^bold>\A x. x" (* -- cgit v1.2.3