From 0c786e631f32c1f92c508d3fe8aad6f433bb5ce0 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Fri, 17 Aug 2018 14:20:39 +0200 Subject: Partway through updating proofs after the change of the function composition definition --- ProdProps.thy | 57 +++++++++++---------------------------------------------- 1 file changed, 11 insertions(+), 46 deletions(-) (limited to 'ProdProps.thy') diff --git a/ProdProps.thy b/ProdProps.thy index e48b3e6..2145bf9 100644 --- a/ProdProps.thy +++ b/ProdProps.thy @@ -15,22 +15,18 @@ begin section \Composition\ text " - The proof of associativity is surprisingly intricate; it involves telling Isabelle to use the correct rule for \-type judgmental equality, and the correct substitutions in the subgoals thereafter, among other things. + The proof of associativity needs some guidance; it involves telling Isabelle to use the correct rule for \-type definitional equality, and the correct substitutions in the subgoals thereafter. " lemma compose_assoc: - assumes - "A \ B: U(i)" "B \ C: U(i)" and - "\x:C. D(x): U(i)" and - "f: A \ B" "g: B \ C" "h: \x:C. D(x)" + assumes "A: U(i)" and "f: A \ B" "g: B \ C" "h: \x:C. D(x)" shows "(h \ g) \ f \ h \ (g \ f)" proof (subst (0 1 2 3) compose_def) - text "Compute the different bracketings and show they are equivalent:" - show "\<^bold>\x. (\<^bold>\y. h`(g`y))`(f`x) \ \<^bold>\x. h`((\<^bold>\y. g`(f`y))`x)" proof (subst Prod_eq) - show "\<^bold>\x. h`((\<^bold>\y. g`(f`y))`x) \ \<^bold>\x. h`((\<^bold>\y. g`(f`y))`x)" by simp + \ \Todo: set the Simplifier (or other simplification methods) up to use \Prod_eq\!\ + show "\x. x: A \ (\<^bold>\y. h`(g`y))`(f`x) \ h`((\<^bold>\y. g`(f`y))`x)" proof compute show "\x. x: A \ h`(g`(f`x)) \ h`((\<^bold>\y. g`(f`y))`x)" @@ -39,49 +35,18 @@ proof (subst (0 1 2 3) compose_def) qed show "\x. x: B \ h`(g`x): D(g`x)" by (simple lems: assms) qed (simple lems: assms) - qed (wellformed lems: assms) - - text " - Finish off any remaining subgoals that Isabelle can't prove automatically. - These typically require the application of specific substitutions or computation rules. - " - show "g \ f: A \ C" by (subst compose_def) (derive lems: assms) - - show "h \ g: \x:B. D(g`x)" - proof (subst compose_def) - show "\<^bold>\x. h`(g`x): \x:B. D(g`x)" - proof - show "\x. x: B \ h`(g`x): D(g`x)" - proof - show "\x. x: B \ g`x: C" by (simple lems: assms) - qed (simple lems: assms) - qed (wellformed lems: assms) - qed fact+ - - show "\x:B. D (g`x): U(i)" - proof - show "\x. x: B \ D(g`x): U(i)" - proof - - have *: "D: C \ U(i)" by (wellformed lems: assms) - - fix x assume asm: "x: B" - have "g`x: C" by (simple lems: assms asm) - then show "D(g`x): U(i)" by (rule *) - qed - qed (wellformed lems: assms) - - have "A: U(i)" by (wellformed lems: assms) - moreover have "C: U(i)" by (wellformed lems: assms) - ultimately show "A \ C: U(i)" .. -qed (simple lems: assms) - + qed fact +qed -text "The following rule is correct, but we cannot prove it using just the rules of the theory." lemma compose_comp': assumes "A: U(i)" and "\x. x: A \ b(x): B" and "\x. x: B \ c(x): C(x)" shows "(\<^bold>\x. c(x)) \ (\<^bold>\x. b(x)) \ \<^bold>\x. c(b(x))" -oops +proof (subst compose_def, subst Prod_eq) + show "\x. x: A \ (\<^bold>\x. c(x))`((\<^bold>\x. b(x))`x) \ \<^bold>\x. c (b x)" + proof compute + + text "However we can derive a variant with more explicit premises:" -- cgit v1.2.3