From 4eca07fed1b54a59718dad04425527f9cc02af9f Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Mon, 18 Jun 2018 18:34:04 +0200 Subject: Dependent projection properties done. --- Sum.thy | 85 ++++++++++++++++++++++++++++++++++++++++++++--------------------- 1 file changed, 58 insertions(+), 27 deletions(-) diff --git a/Sum.thy b/Sum.thy index 0a062cf..08c0de0 100644 --- a/Sum.thy +++ b/Sum.thy @@ -86,7 +86,7 @@ begin "snd_nondep A B \ \<^bold>\p: A \ B. indSum[A, \_. B] (\_. B) (\x y. y) p" end -text "Results about projections:" +text "Typing judgments and computation rules for the dependent and non-dependent projection functions." lemma fst_dep_type: assumes "p : \x:A. B x" @@ -105,7 +105,7 @@ proof \ \The standard reasoner knows to backchain with the produc proof \ \...and with sum elimination here.\ show "A : U" using * .. qed - + qed (rule *) qed (rule assms) @@ -133,57 +133,89 @@ proof - qed +text "In proving results about the second dependent projection function we often use the following two lemmas." + +lemma snd_dep_welltyped: + assumes "p : \x:A. B x" + shows "B (fst[A,B]`p) : U" +proof - + have "\x:A. B x : U" using assms .. + then have *: "B: A \ U" .. + + have "fst[A,B]`p : A" using assms by (rule fst_dep_type) + then show "B (fst[A,B]`p) : U" by (rule *) +qed + +lemma snd_dep_const_type: + assumes "B: A \ U" and "x : A" and "y : B x" + shows "y : B (fst[A,B]`(x,y))" +proof - + have "fst[A,B]`(x,y) \ x" using assms by (rule fst_dep_comp) + then show "y : B (fst[A,B]`(x,y))" using assms by simp +qed + + lemma snd_dep_type: assumes "p : \x:A. B x" shows "snd[A,B]`p : B (fst[A,B]`p)" -oops -\ \Do we need this for the following lemma?\ +proof + have *: "\x:A. B x : U" using assms .. + show "snd[A, B] : \p:(\x:A. B x). B (fst[A,B]`p)" -lemma snd_dep_comp: - assumes "B: A \ U" and "a : A" and "b : B a" - shows "snd[A,B]`(a,b) \ b" -proof - - \ \We use the following two lemmas:\ + proof (unfold snd_dep_def, standard) + show "\p. p : \x:A. B x \ indSum[A,B] (\q. B (fst[A,B]`q)) (\x y. y) p : B (fst[A,B]`p)" - have lemma1: "\p. p : \x:A. B x \ B (fst[A,B]`p) : U" - proof - - fix p assume "p : \x:A. B x" - then have "fst[A,B]`p : A" by (rule fst_dep_type) - then show "B (fst[A,B]`p) : U" by (rule assms(1)) - qed + proof (standard, elim snd_dep_welltyped) + fix x y assume 1: "x : A" and 2: "y : B x" + show "y : B (fst[A,B]`(x,y))" - have lemma2: "\x y. \x : A; y : B x\ \ y : B (fst[A,B]`(x, y))" - proof - - fix x y assume "x : A" and "y : B x" - moreover with assms(1) have "fst[A,B]`(x,y) \ x" by (rule fst_dep_comp) - ultimately show "y : B (fst[A,B]`(x,y))" by simp - qed + proof - + have "B: A \ U" using * .. + then show "y : B (fst[A,B]`(x,y))" using 1 2 by (rule snd_dep_const_type) + qed + + qed + + qed (rule *) + +qed (rule assms) - \ \And now the proof.\ +lemma snd_dep_comp: + assumes "B: A \ U" and "a : A" and "b : B a" + shows "snd[A,B]`(a,b) \ b" +proof - have "snd[A,B]`(a,b) \ indSum[A, B] (\q. B (fst[A,B]`q)) (\x y. y) (a,b)" proof (unfold snd_dep_def, standard) show "(a,b) : \x:A. B x" using assms .. fix p assume *: "p : \x:A. B x" show "indSum[A,B] (\q. B (fst[A,B]`q)) (\x y. y) p : B (fst[A,B]`p)" - by (standard, elim lemma1, elim lemma2) (assumption, rule *) + proof (standard, elim snd_dep_welltyped) + show "\x y. \x : A; y : B x\ \ y : B (fst[A,B]`(x,y))" using assms + by (elim snd_dep_const_type) + qed (rule *) qed also have "indSum[A, B] (\q. B (fst[A,B]`q)) (\x y. y) (a,b) \ b" - by (standard, elim lemma1, elim lemma2) (assumption, (rule assms)+) + proof (standard, elim snd_dep_welltyped) + show "\x y. \x : A; y : B x\ \ y : B (fst[A,B]`(x,y))" using assms + by (elim snd_dep_const_type) + qed (rule assms)+ finally show "snd[A,B]`(a,b) \ b" . qed +text "For non-dependent projection functions:" + lemma fst_nondep_comp: assumes "a : A" and "b : B" shows "fst[A,B]`(a,b) \ a" -proof - +proof (unfold fst_nondep_def) have "A : U" using assms(1) .. - then show "fst[A,B]`(a,b) \ a" unfolding fst_nondep_def by simp + then show "fst[A,B]`(a,b) \ a" unfolding fst_nondep_def qed lemma snd_nondep_comp: "\a : A; b : B\ \ snd[A,B]`(a,b) \ b" @@ -193,7 +225,6 @@ proof - then show "snd[A,B]`(a,b) \ b" unfolding snd_nondep_def by simp qed -lemmas proj_simps [simp] = fst_dep_comp snd_dep_comp fst_nondep_comp snd_nondep_comp end \ No newline at end of file -- cgit v1.2.3