(* Title: tests/Test.thy Author: Joshua Chen Date: 2018 This is an old test suite from early implementations. It is not always guaranteed to be up to date or to reflect most recent versions of the theory. *) theory Test imports "../HoTT" begin text \ A bunch of theorems and other statements for sanity-checking, as well as things that should be automatically simplified. Things that *should* be automated: \<^item> Checking that @{term A} is a well-formed type, when writing things like @{prop "x: A"} and @{prop "A: U i"}. \<^item> Checking that the argument to a (dependent/non-dependent) function matches the type? Also the arguments to a pair? \ declare[[unify_trace_simp, unify_trace_types, simp_trace, simp_trace_depth_limit=5]] \ \Turn on trace for unification and the Simplifier, for debugging.\ section \\-type\ subsection \Typing functions\ text \Declaring @{thm Prod_intro} with the @{attribute intro} attribute enables @{method rule} to prove the following.\ proposition assumes "A : U(i)" shows "\<^bold>\x. x: A \ A" by (routine add: assms) proposition assumes "A : U(i)" and "A \ B" shows "\<^bold>\x. x : B \ A" proof - have "A \ A \ B \ A" using assms by simp moreover have "\<^bold>\x. x : A \ A" by (routine add: assms) ultimately show "\<^bold>\x. x : B \ A" by simp qed proposition assumes "A : U(i)" and "B : U(i)" shows "\<^bold>\x y. x : A \ B \ A" by (routine add: assms) subsection \Function application\ proposition assumes "a : A" shows "(\<^bold>\x. x)`a \ a" by (derive lems: assms) lemma assumes "a : A" and "\x. x: A \ B(x): U(i)" shows "(\<^bold>\x y. y)`a \ \<^bold>\y. y" by (derive lems: assms) lemma "\A: U(i); B: U(i); a : A; b : B\ \ (\<^bold>\x y. y)`a`b \ b" by derive lemma "\A: U(i); a : A\ \ (\<^bold>\x y. f x y)`a \ \<^bold>\y. f a y" proof derive oops \ \Missing some premises.\ 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" proof derive oops subsection \Currying functions\ proposition curried_function_formation: assumes "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 (routine add: assms) proposition higher_order_currying_formation: assumes "A: U(i)" and "B: A \ U(i)" and "\x y. y: B(x) \ C(x)(y): U(i)" and "\x y z. z : C(x)(y) \ D(x)(y)(z): U(i)" shows "\x:A. \y:B(x). \z:C(x)(y). D(x)(y)(z): U(i)" by (routine add: assms) lemma curried_type_judgment: assumes "A: U(i)" and "B: A \ U(i)" and "\x y. \x : A; y : B(x)\ \ f x y : C x y" shows "\<^bold>\x y. f x y : \x:A. \y:B(x). C x y" by (routine add: assms) text \ Polymorphic identity function is now trivial due to lambda expression polymorphism. It was more involved in previous monomorphic incarnations. \ lemma "x: A \ id`x \ x" by derive section \Natural numbers\ text \Automatic proof methods recognize natural numbers.\ proposition "succ(succ(succ 0)): \" by routine end