# HG changeset patch # User wenzelm # Date 1467754789 -7200 # Node ID 249fa34faba26e85a39f0fc24fa678d6f45b0cc1 # Parent d1742d1b7f0f643ea0b54fa533e34abe99046a1e misc tuning and modernization; diff -r d1742d1b7f0f -r 249fa34faba2 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Tue Jul 05 22:47:48 2016 +0200 +++ b/src/HOL/Fun.thy Tue Jul 05 23:39:49 2016 +0200 @@ -21,6 +21,7 @@ lemma b_uniq_choice: "\x\S. \!y. Q x y \ \f. \x\S. Q x (f x)" by (force intro: theI') + subsection \The Identity Function \id\\ definition id :: "'a \ 'a" @@ -283,9 +284,8 @@ shows "inj f" \ \Courtesy of Stephan Merz\ proof (rule inj_onI) - fix x y - assume f_eq: "f x = f y" - show "x = y" by (rule linorder_cases) (auto dest: hyp simp: f_eq) + show "x = y" if "f x = f y" for x y + by (rule linorder_cases) (auto dest: hyp simp: that) qed lemma surj_def: "surj f \ (\y. \x. y = f x)" @@ -427,9 +427,8 @@ by (auto simp add: bij_betw_def) lemma bij_betw_combine: - assumes "bij_betw f A B" "bij_betw f C D" "B \ D = {}" - shows "bij_betw f (A \ C) (B \ D)" - using assms unfolding bij_betw_def inj_on_Un image_Un by auto + "bij_betw f A B \ bij_betw f C D \ B \ D = {} \ bij_betw f (A \ C) (B \ D)" + unfolding bij_betw_def inj_on_Un image_Un by auto lemma bij_betw_subset: "bij_betw f A A' \ B \ A \ f ` B = B' \ bij_betw f B B'" by (auto simp add: bij_betw_def inj_on_def) @@ -531,14 +530,15 @@ and img2: "f' ` A' \ A" shows "bij_betw f A A'" using assms -proof (unfold bij_betw_def inj_on_def, safe) + unfolding bij_betw_def inj_on_def +proof safe fix a b assume *: "a \ A" "b \ A" and **: "f a = f b" have "a = f' (f a) \ b = f'(f b)" using * left by simp with ** show "a = b" by simp next fix a' assume *: "a' \ A'" - hence "f' a' \ A" using img2 by blast + then have "f' a' \ A" using img2 by blast moreover have "a' = f (f' a')" using * right by simp ultimately show "a' \ f ` A" by blast @@ -823,7 +823,7 @@ subsubsection \Proof tools\ -text \Simplify terms of the form f(...,x:=y,...,x:=z,...) to f(...,x:=z,...)\ +text \Simplify terms of the form \f(\,x:=y,\,x:=z,\)\ to \f(\,x:=z,\)\\ simproc_setup fun_upd2 ("f(v := w, x := y)") = \fn _ => let @@ -843,14 +843,14 @@ let val t = Thm.term_of ct in - case find_double t of + (case find_double t of (T, NONE) => NONE | (T, SOME rhs) => SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs)) (fn _ => resolve_tac ctxt [eq_reflection] 1 THEN resolve_tac ctxt @{thms ext} 1 THEN - simp_tac (put_simpset ss ctxt) 1)) + simp_tac (put_simpset ss ctxt) 1))) end in proc end \ diff -r d1742d1b7f0f -r 249fa34faba2 src/HOL/Inductive.thy --- a/src/HOL/Inductive.thy Tue Jul 05 22:47:48 2016 +0200 +++ b/src/HOL/Inductive.thy Tue Jul 05 23:39:49 2016 +0200 @@ -19,48 +19,45 @@ context complete_lattice begin -definition - lfp :: "('a \ 'a) \ 'a" where - "lfp f = Inf {u. f u \ u}" \\least fixed point\ +definition lfp :: "('a \ 'a) \ 'a" \ \least fixed point\ + where "lfp f = Inf {u. f u \ u}" -definition - gfp :: "('a \ 'a) \ 'a" where - "gfp f = Sup {u. u \ f u}" \\greatest fixed point\ +definition gfp :: "('a \ 'a) \ 'a" \ \greatest fixed point\ + where "gfp f = Sup {u. u \ f u}" -subsection\Proof of Knaster-Tarski Theorem using @{term lfp}\ +subsection \Proof of Knaster-Tarski Theorem using @{term lfp}\ -text\@{term "lfp f"} is the least upper bound of - the set @{term "{u. f(u) \ u}"}\ +text \@{term "lfp f"} is the least upper bound of the set @{term "{u. f u \ u}"}\ -lemma lfp_lowerbound: "f A \ A ==> lfp f \ A" +lemma lfp_lowerbound: "f A \ A \ lfp f \ A" by (auto simp add: lfp_def intro: Inf_lower) -lemma lfp_greatest: "(!!u. f u \ u ==> A \ u) ==> A \ lfp f" +lemma lfp_greatest: "(\u. f u \ u \ A \ u) \ A \ lfp f" by (auto simp add: lfp_def intro: Inf_greatest) end -lemma lfp_lemma2: "mono f ==> f (lfp f) \ lfp f" +lemma lfp_lemma2: "mono f \ f (lfp f) \ lfp f" by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound) -lemma lfp_lemma3: "mono f ==> lfp f \ f (lfp f)" +lemma lfp_lemma3: "mono f \ lfp f \ f (lfp f)" by (iprover intro: lfp_lemma2 monoD lfp_lowerbound) -lemma lfp_unfold: "mono f ==> lfp f = f (lfp f)" +lemma lfp_unfold: "mono f \ lfp f = f (lfp f)" by (iprover intro: order_antisym lfp_lemma2 lfp_lemma3) lemma lfp_const: "lfp (\x. t) = t" - by (rule lfp_unfold) (simp add:mono_def) + by (rule lfp_unfold) (simp add: mono_def) subsection \General induction rules for least fixed points\ -lemma lfp_ordinal_induct[case_names mono step union]: +lemma lfp_ordinal_induct [case_names mono step union]: fixes f :: "'a::complete_lattice \ 'a" assumes mono: "mono f" - and P_f: "\S. P S \ S \ lfp f \ P (f S)" - and P_Union: "\M. \S\M. P S \ P (Sup M)" + and P_f: "\S. P S \ S \ lfp f \ P (f S)" + and P_Union: "\M. \S\M. P S \ P (Sup M)" shows "P (lfp f)" proof - let ?M = "{S. S \ lfp f \ P S}" @@ -68,96 +65,94 @@ also have "Sup ?M = lfp f" proof (rule antisym) show "Sup ?M \ lfp f" by (blast intro: Sup_least) - hence "f (Sup ?M) \ f (lfp f)" by (rule mono [THEN monoD]) - hence "f (Sup ?M) \ lfp f" using mono [THEN lfp_unfold] by simp - hence "f (Sup ?M) \ ?M" using P_Union by simp (intro P_f Sup_least, auto) - hence "f (Sup ?M) \ Sup ?M" by (rule Sup_upper) - thus "lfp f \ Sup ?M" by (rule lfp_lowerbound) + then have "f (Sup ?M) \ f (lfp f)" + by (rule mono [THEN monoD]) + then have "f (Sup ?M) \ lfp f" + using mono [THEN lfp_unfold] by simp + then have "f (Sup ?M) \ ?M" + using P_Union by simp (intro P_f Sup_least, auto) + then have "f (Sup ?M) \ Sup ?M" + by (rule Sup_upper) + then show "lfp f \ Sup ?M" + by (rule lfp_lowerbound) qed finally show ?thesis . -qed +qed theorem lfp_induct: - assumes mono: "mono f" and ind: "f (inf (lfp f) P) \ P" + assumes mono: "mono f" + and ind: "f (inf (lfp f) P) \ P" shows "lfp f \ P" proof (induction rule: lfp_ordinal_induct) - case (step S) then show ?case + case (step S) + then show ?case by (intro order_trans[OF _ ind] monoD[OF mono]) auto qed (auto intro: mono Sup_least) lemma lfp_induct_set: - assumes lfp: "a: lfp(f)" - and mono: "mono(f)" - and indhyp: "!!x. [| x: f(lfp(f) Int {x. P(x)}) |] ==> P(x)" - shows "P(a)" - by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp]) - (auto simp: intro: indhyp) + assumes lfp: "a \ lfp f" + and mono: "mono f" + and hyp: "\x. x \ f (lfp f \ {x. P x}) \ P x" + shows "P a" + by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp]) (auto intro: hyp) -lemma lfp_ordinal_induct_set: +lemma lfp_ordinal_induct_set: assumes mono: "mono f" - and P_f: "!!S. P S ==> P(f S)" - and P_Union: "!!M. !S:M. P S ==> P (\M)" - shows "P(lfp f)" + and P_f: "\S. P S \ P (f S)" + and P_Union: "\M. \S\M. P S \ P (\M)" + shows "P (lfp f)" using assms by (rule lfp_ordinal_induct) -text\Definition forms of \lfp_unfold\ and \lfp_induct\, - to control unfolding\ +text \Definition forms of \lfp_unfold\ and \lfp_induct\, to control unfolding.\ -lemma def_lfp_unfold: "[| h==lfp(f); mono(f) |] ==> h = f(h)" +lemma def_lfp_unfold: "h \ lfp f \ mono f \ h = f h" by (auto intro!: lfp_unfold) -lemma def_lfp_induct: - "[| A == lfp(f); mono(f); - f (inf A P) \ P - |] ==> A \ P" +lemma def_lfp_induct: "A \ lfp f \ mono f \ f (inf A P) \ P \ A \ P" by (blast intro: lfp_induct) -lemma def_lfp_induct_set: - "[| A == lfp(f); mono(f); a:A; - !!x. [| x: f(A Int {x. P(x)}) |] ==> P(x) - |] ==> P(a)" +lemma def_lfp_induct_set: + "A \ lfp f \ mono f \ a \ A \ (\x. x \ f (A \ {x. P x}) \ P x) \ P a" by (blast intro: lfp_induct_set) -(*Monotonicity of lfp!*) -lemma lfp_mono: "(!!Z. f Z \ g Z) ==> lfp f \ lfp g" - by (rule lfp_lowerbound [THEN lfp_greatest], blast intro: order_trans) +text \Monotonicity of \lfp\!\ +lemma lfp_mono: "(\Z. f Z \ g Z) \ lfp f \ lfp g" + by (rule lfp_lowerbound [THEN lfp_greatest]) (blast intro: order_trans) -subsection \Proof of Knaster-Tarski Theorem using @{term gfp}\ +subsection \Proof of Knaster-Tarski Theorem using \gfp\\ -text\@{term "gfp f"} is the greatest lower bound of - the set @{term "{u. u \ f(u)}"}\ +text \@{term "gfp f"} is the greatest lower bound of the set @{term "{u. u \ f u}"}\ -lemma gfp_upperbound: "X \ f X ==> X \ gfp f" +lemma gfp_upperbound: "X \ f X \ X \ gfp f" by (auto simp add: gfp_def intro: Sup_upper) -lemma gfp_least: "(!!u. u \ f u ==> u \ X) ==> gfp f \ X" +lemma gfp_least: "(\u. u \ f u \ u \ X) \ gfp f \ X" by (auto simp add: gfp_def intro: Sup_least) -lemma gfp_lemma2: "mono f ==> gfp f \ f (gfp f)" +lemma gfp_lemma2: "mono f \ gfp f \ f (gfp f)" by (iprover intro: gfp_least order_trans monoD gfp_upperbound) -lemma gfp_lemma3: "mono f ==> f (gfp f) \ gfp f" +lemma gfp_lemma3: "mono f \ f (gfp f) \ gfp f" by (iprover intro: gfp_lemma2 monoD gfp_upperbound) -lemma gfp_unfold: "mono f ==> gfp f = f (gfp f)" +lemma gfp_unfold: "mono f \ gfp f = f (gfp f)" by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3) subsection \Coinduction rules for greatest fixed points\ -text\weak version\ -lemma weak_coinduct: "[| a: X; X \ f(X) |] ==> a : gfp(f)" +text \Weak version.\ +lemma weak_coinduct: "a \ X \ X \ f X \ a \ gfp f" by (rule gfp_upperbound [THEN subsetD]) auto -lemma weak_coinduct_image: "!!X. [| a : X; g`X \ f (g`X) |] ==> g a : gfp f" +lemma weak_coinduct_image: "a \ X \ g`X \ f (g`X) \ g a \ gfp f" apply (erule gfp_upperbound [THEN subsetD]) apply (erule imageI) done -lemma coinduct_lemma: - "[| X \ f (sup X (gfp f)); mono f |] ==> sup X (gfp f) \ f (sup X (gfp f))" +lemma coinduct_lemma: "X \ f (sup X (gfp f)) \ mono f \ sup X (gfp f) \ f (sup X (gfp f))" apply (frule gfp_lemma2) apply (drule mono_sup) apply (rule le_supI) @@ -169,112 +164,113 @@ apply assumption done -text\strong version, thanks to Coen and Frost\ -lemma coinduct_set: "[| mono(f); a: X; X \ f(X Un gfp(f)) |] ==> a : gfp(f)" +text \Strong version, thanks to Coen and Frost.\ +lemma coinduct_set: "mono f \ a \ X \ X \ f (X \ gfp f) \ a \ gfp f" by (rule weak_coinduct[rotated], rule coinduct_lemma) blast+ -lemma gfp_fun_UnI2: "[| mono(f); a: gfp(f) |] ==> a: f(X Un gfp(f))" +lemma gfp_fun_UnI2: "mono f \ a \ gfp f \ a \ f (X \ gfp f)" by (blast dest: gfp_lemma2 mono_Un) lemma gfp_ordinal_induct[case_names mono step union]: fixes f :: "'a::complete_lattice \ 'a" assumes mono: "mono f" - and P_f: "\S. P S \ gfp f \ S \ P (f S)" - and P_Union: "\M. \S\M. P S \ P (Inf M)" + and P_f: "\S. P S \ gfp f \ S \ P (f S)" + and P_Union: "\M. \S\M. P S \ P (Inf M)" shows "P (gfp f)" proof - let ?M = "{S. gfp f \ S \ P S}" have "P (Inf ?M)" using P_Union by simp also have "Inf ?M = gfp f" proof (rule antisym) - show "gfp f \ Inf ?M" by (blast intro: Inf_greatest) - hence "f (gfp f) \ f (Inf ?M)" by (rule mono [THEN monoD]) - hence "gfp f \ f (Inf ?M)" using mono [THEN gfp_unfold] by simp - hence "f (Inf ?M) \ ?M" using P_Union by simp (intro P_f Inf_greatest, auto) - hence "Inf ?M \ f (Inf ?M)" by (rule Inf_lower) - thus "Inf ?M \ gfp f" by (rule gfp_upperbound) + show "gfp f \ Inf ?M" + by (blast intro: Inf_greatest) + then have "f (gfp f) \ f (Inf ?M)" + by (rule mono [THEN monoD]) + then have "gfp f \ f (Inf ?M)" + using mono [THEN gfp_unfold] by simp + then have "f (Inf ?M) \ ?M" + using P_Union by simp (intro P_f Inf_greatest, auto) + then have "Inf ?M \ f (Inf ?M)" + by (rule Inf_lower) + then show "Inf ?M \ gfp f" + by (rule gfp_upperbound) qed finally show ?thesis . -qed +qed -lemma coinduct: assumes mono: "mono f" and ind: "X \ f (sup X (gfp f))" shows "X \ gfp f" +lemma coinduct: + assumes mono: "mono f" + and ind: "X \ f (sup X (gfp f))" + shows "X \ gfp f" proof (induction rule: gfp_ordinal_induct) case (step S) then show ?case by (intro order_trans[OF ind _] monoD[OF mono]) auto qed (auto intro: mono Inf_greatest) + subsection \Even Stronger Coinduction Rule, by Martin Coen\ -text\Weakens the condition @{term "X \ f(X)"} to one expressed using both +text \Weakens the condition @{term "X \ f X"} to one expressed using both @{term lfp} and @{term gfp}\ - -lemma coinduct3_mono_lemma: "mono(f) ==> mono(%x. f(x) Un X Un B)" -by (iprover intro: subset_refl monoI Un_mono monoD) +lemma coinduct3_mono_lemma: "mono f \ mono (\x. f x \ X \ B)" + by (iprover intro: subset_refl monoI Un_mono monoD) lemma coinduct3_lemma: - "[| X \ f(lfp(%x. f(x) Un X Un gfp(f))); mono(f) |] - ==> lfp(%x. f(x) Un X Un gfp(f)) \ f(lfp(%x. f(x) Un X Un gfp(f)))" -apply (rule subset_trans) -apply (erule coinduct3_mono_lemma [THEN lfp_lemma3]) -apply (rule Un_least [THEN Un_least]) -apply (rule subset_refl, assumption) -apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption) -apply (rule monoD, assumption) -apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto) -done + "X \ f (lfp (\x. f x \ X \ gfp f)) \ mono f \ + lfp (\x. f x \ X \ gfp f) \ f (lfp (\x. f x \ X \ gfp f))" + apply (rule subset_trans) + apply (erule coinduct3_mono_lemma [THEN lfp_lemma3]) + apply (rule Un_least [THEN Un_least]) + apply (rule subset_refl, assumption) + apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption) + apply (rule monoD, assumption) + apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto) + done -lemma coinduct3: - "[| mono(f); a:X; X \ f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)" -apply (rule coinduct3_lemma [THEN [2] weak_coinduct]) -apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst]) -apply (simp_all) -done +lemma coinduct3: "mono f \ a \ X \ X \ f (lfp (\x. f x \ X \ gfp f)) \ a \ gfp f" + apply (rule coinduct3_lemma [THEN [2] weak_coinduct]) + apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst]) + apply simp_all + done -text\Definition forms of \gfp_unfold\ and \coinduct\, - to control unfolding\ +text \Definition forms of \gfp_unfold\ and \coinduct\, to control unfolding.\ -lemma def_gfp_unfold: "[| A==gfp(f); mono(f) |] ==> A = f(A)" +lemma def_gfp_unfold: "A \ gfp f \ mono f \ A = f A" by (auto intro!: gfp_unfold) -lemma def_coinduct: - "[| A==gfp(f); mono(f); X \ f(sup X A) |] ==> X \ A" +lemma def_coinduct: "A \ gfp f \ mono f \ X \ f (sup X A) \ X \ A" by (iprover intro!: coinduct) -lemma def_coinduct_set: - "[| A==gfp(f); mono(f); a:X; X \ f(X Un A) |] ==> a: A" +lemma def_coinduct_set: "A \ gfp f \ mono f \ a \ X \ X \ f (X \ A) \ a \ A" by (auto intro!: coinduct_set) -(*The version used in the induction/coinduction package*) lemma def_Collect_coinduct: - "[| A == gfp(%w. Collect(P(w))); mono(%w. Collect(P(w))); - a: X; !!z. z: X ==> P (X Un A) z |] ==> - a : A" + "A \ gfp (\w. Collect (P w)) \ mono (\w. Collect (P w)) \ a \ X \ + (\z. z \ X \ P (X \ A) z) \ a \ A" by (erule def_coinduct_set) auto -lemma def_coinduct3: - "[| A==gfp(f); mono(f); a:X; X \ f(lfp(%x. f(x) Un X Un A)) |] ==> a: A" +lemma def_coinduct3: "A \ gfp f \ mono f \ a \ X \ X \ f (lfp (\x. f x \ X \ A)) \ a \ A" by (auto intro!: coinduct3) -text\Monotonicity of @{term gfp}!\ -lemma gfp_mono: "(!!Z. f Z \ g Z) ==> gfp f \ gfp g" - by (rule gfp_upperbound [THEN gfp_least], blast intro: order_trans) +text \Monotonicity of @{term gfp}!\ +lemma gfp_mono: "(\Z. f Z \ g Z) \ gfp f \ gfp g" + by (rule gfp_upperbound [THEN gfp_least]) (blast intro: order_trans) + subsection \Rules for fixed point calculus\ - lemma lfp_rolling: assumes "mono g" "mono f" shows "g (lfp (\x. f (g x))) = lfp (\x. g (f x))" proof (rule antisym) have *: "mono (\x. f (g x))" using assms by (auto simp: mono_def) - show "lfp (\x. g (f x)) \ g (lfp (\x. f (g x)))" by (rule lfp_lowerbound) (simp add: lfp_unfold[OF *, symmetric]) - show "g (lfp (\x. f (g x))) \ lfp (\x. g (f x))" proof (rule lfp_greatest) - fix u assume "g (f u) \ u" + fix u + assume "g (f u) \ u" moreover then have "g (lfp (\x. f (g x))) \ g (f u)" by (intro assms[THEN monoD] lfp_lowerbound) ultimately show "g (lfp (\x. f (g x))) \ u" @@ -309,7 +305,6 @@ using assms by (auto simp: mono_def) show "g (gfp (\x. f (g x))) \ gfp (\x. g (f x))" by (rule gfp_upperbound) (simp add: gfp_unfold[OF *, symmetric]) - show "gfp (\x. g (f x)) \ g (gfp (\x. f (g x)))" proof (rule gfp_least) fix u assume "u \ g (f u)" @@ -339,6 +334,7 @@ qed qed + subsection \Inductive predicates and sets\ text \Package setup.\ diff -r d1742d1b7f0f -r 249fa34faba2 src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Tue Jul 05 22:47:48 2016 +0200 +++ b/src/HOL/Product_Type.thy Tue Jul 05 23:39:49 2016 +0200 @@ -37,12 +37,11 @@ declare case_split [cases type: bool] \ "prefer plain propositional version" -lemma - shows [code]: "HOL.equal False P \ \ P" - and [code]: "HOL.equal True P \ P" - and [code]: "HOL.equal P False \ \ P" - and [code]: "HOL.equal P True \ P" - and [code nbe]: "HOL.equal P P \ True" +lemma [code]: "HOL.equal False P \ \ P" + and [code]: "HOL.equal True P \ P" + and [code]: "HOL.equal P False \ \ P" + and [code]: "HOL.equal P True \ P" + and [code nbe]: "HOL.equal P P \ True" by (simp_all add: equal) lemma If_case_cert: @@ -101,23 +100,23 @@ setup \Sign.parent_path\ -lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()" +lemma unit_all_eq1: "(\x::unit. PROP P x) \ PROP P ()" by simp -lemma unit_all_eq2: "(!!x::unit. PROP P) == PROP P" +lemma unit_all_eq2: "(\x::unit. PROP P) \ PROP P" by (rule triv_forall_equality) text \ This rewrite counters the effect of simproc \unit_eq\ on @{term - [source] "%u::unit. f u"}, replacing it by @{term [source] - f} rather than by @{term [source] "%u. f ()"}. + [source] "\u::unit. f u"}, replacing it by @{term [source] + f} rather than by @{term [source] "\u. f ()"}. \ -lemma unit_abs_eta_conv [simp]: "(%u::unit. f ()) = f" +lemma unit_abs_eta_conv [simp]: "(\u::unit. f ()) = f" by (rule ext) simp -lemma UNIV_unit: - "UNIV = {()}" by auto +lemma UNIV_unit: "UNIV = {()}" + by auto instantiation unit :: default begin @@ -128,52 +127,41 @@ end -instantiation unit :: "{complete_boolean_algebra, complete_linorder, wellorder}" +instantiation unit :: "{complete_boolean_algebra,complete_linorder,wellorder}" begin definition less_eq_unit :: "unit \ unit \ bool" -where - "(_::unit) \ _ \ True" + where "(_::unit) \ _ \ True" -lemma less_eq_unit [iff]: - "(u::unit) \ v" +lemma less_eq_unit [iff]: "u \ v" for u v :: unit by (simp add: less_eq_unit_def) definition less_unit :: "unit \ unit \ bool" -where - "(_::unit) < _ \ False" + where "(_::unit) < _ \ False" -lemma less_unit [iff]: - "\ (u::unit) < v" +lemma less_unit [iff]: "\ u < v" for u v :: unit by (simp_all add: less_eq_unit_def less_unit_def) definition bot_unit :: unit -where - [code_unfold]: "\ = ()" + where [code_unfold]: "\ = ()" definition top_unit :: unit -where - [code_unfold]: "\ = ()" + where [code_unfold]: "\ = ()" definition inf_unit :: "unit \ unit \ unit" -where - [simp]: "_ \ _ = ()" + where [simp]: "_ \ _ = ()" definition sup_unit :: "unit \ unit \ unit" -where - [simp]: "_ \ _ = ()" + where [simp]: "_ \ _ = ()" definition Inf_unit :: "unit set \ unit" -where - [simp]: "\_ = ()" + where [simp]: "\_ = ()" definition Sup_unit :: "unit set \ unit" -where - [simp]: "\_ = ()" + where [simp]: "\_ = ()" definition uminus_unit :: "unit \ unit" -where - [simp]: "- _ = ()" + where [simp]: "- _ = ()" declare less_eq_unit_def [abs_def, code_unfold] less_unit_def [abs_def, code_unfold] @@ -188,8 +176,8 @@ end -lemma [code]: - "HOL.equal (u::unit) v \ True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+ +lemma [code]: "HOL.equal u v \ True" for u v :: unit + unfolding equal unit_eq [of u] unit_eq [of v] by rule+ code_printing type_constructor unit \ @@ -221,8 +209,8 @@ subsubsection \Type definition\ -definition Pair_Rep :: "'a \ 'b \ 'a \ 'b \ bool" where - "Pair_Rep a b = (\x y. x = a \ y = b)" +definition Pair_Rep :: "'a \ 'b \ 'a \ 'b \ bool" + where "Pair_Rep a b = (\x y. x = a \ y = b)" definition "prod = {f. \a b. f = Pair_Rep (a::'a) (b::'b)}" @@ -232,8 +220,8 @@ type_notation (ASCII) prod (infixr "*" 20) -definition Pair :: "'a \ 'b \ 'a \ 'b" where - "Pair a b = Abs_prod (Pair_Rep a b)" +definition Pair :: "'a \ 'b \ 'a \ 'b" + where "Pair a b = Abs_prod (Pair_Rep a b)" lemma prod_cases: "(\a b. P (Pair a b)) \ P p" by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def) @@ -258,7 +246,7 @@ setup \Sign.mandatory_path "old"\ old_rep_datatype Pair -by (erule prod_cases) (rule prod.inject) + by (erule prod_cases) (rule prod.inject) setup \Sign.parent_path\ @@ -297,16 +285,14 @@ \ nonterminal tuple_args and patterns - syntax - "_tuple" :: "'a => tuple_args => 'a * 'b" ("(1'(_,/ _'))") - "_tuple_arg" :: "'a => tuple_args" ("_") - "_tuple_args" :: "'a => tuple_args => tuple_args" ("_,/ _") - "_pattern" :: "[pttrn, patterns] => pttrn" ("'(_,/ _')") - "" :: "pttrn => patterns" ("_") - "_patterns" :: "[pttrn, patterns] => patterns" ("_,/ _") + "_tuple" :: "'a \ tuple_args \ 'a \ 'b" ("(1'(_,/ _'))") + "_tuple_arg" :: "'a \ tuple_args" ("_") + "_tuple_args" :: "'a \ tuple_args \ tuple_args" ("_,/ _") + "_pattern" :: "pttrn \ patterns \ pttrn" ("'(_,/ _')") + "" :: "pttrn \ patterns" ("_") + "_patterns" :: "pttrn \ patterns \ patterns" ("_,/ _") "_unit" :: pttrn ("'(')") - translations "(x, y)" \ "CONST Pair x y" "_pattern x y" \ "CONST Pair x y" @@ -356,9 +342,9 @@ in [(@{const_syntax case_prod}, K case_prod_guess_names_tr')] end \ -text \reconstruct pattern from (nested) @{const case_prod}s, +text \Reconstruct pattern from (nested) @{const case_prod}s, avoiding eta-contraction of body; required for enclosing "let", - if "let" does not avoid eta-contraction, which has been observed to occur\ + if "let" does not avoid eta-contraction, which has been observed to occur.\ print_translation \ let @@ -421,19 +407,16 @@ subsubsection \Fundamental operations and properties\ -lemma Pair_inject: - assumes "(a, b) = (a', b')" - and "a = a' \ b = b' \ R" - shows R - using assms by simp +lemma Pair_inject: "(a, b) = (a', b') \ (a = a' \ b = b' \ R) \ R" + by simp -lemma surj_pair [simp]: "EX x y. p = (x, y)" +lemma surj_pair [simp]: "\x y. p = (x, y)" by (cases p) simp -lemma fst_eqD: "fst (x, y) = a ==> x = a" +lemma fst_eqD: "fst (x, y) = a \ x = a" by simp -lemma snd_eqD: "snd (x, y) = a ==> y = a" +lemma snd_eqD: "snd (x, y) = a \ y = a" by simp lemma case_prod_unfold [nitpick_unfold]: "case_prod = (\c p. c (fst p) (snd p))" @@ -469,25 +452,25 @@ lemma The_case_prod: "The (case_prod P) = (THE xy. P (fst xy) (snd xy))" by (simp add: case_prod_unfold) -lemma cond_case_prod_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g" +lemma cond_case_prod_eta: "(\x y. f x y = g (x, y)) \ (\(x, y). f x y) = g" by (simp add: case_prod_eta) -lemma split_paired_all [no_atp]: "(!!x. PROP P x) == (!!a b. PROP P (a, b))" +lemma split_paired_all [no_atp]: "(\x. PROP P x) \ (\a b. PROP P (a, b))" proof fix a b - assume "!!x. PROP P x" + assume "\x. PROP P x" then show "PROP P (a, b)" . next fix x - assume "!!a b. PROP P (a, b)" + assume "\a b. PROP P (a, b)" from \PROP P (fst x, snd x)\ show "PROP P x" by simp qed text \ The rule @{thm [source] split_paired_all} does not work with the Simplifier because it also affects premises in congrence rules, - where this can lead to premises of the form \!!a b. ... = - ?P(a, b)\ which cannot be solved by reflexivity. + where this can lead to premises of the form \\a b. \ = ?P(a, b)\ + which cannot be solved by reflexivity. \ lemmas split_tupled_all = split_paired_all unit_all_eq2 @@ -520,11 +503,11 @@ setup \map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\ -lemma split_paired_All [simp, no_atp]: "(ALL x. P x) = (ALL a b. P (a, b))" +lemma split_paired_All [simp, no_atp]: "(\x. P x) \ (\a b. P (a, b))" \ \\[iff]\ is not a good idea because it makes \blast\ loop\ by fast -lemma split_paired_Ex [simp, no_atp]: "(EX x. P x) = (EX a b. P (a, b))" +lemma split_paired_Ex [simp, no_atp]: "(\x. P x) \ (\a b. P (a, b))" by fast lemma split_paired_The [no_atp]: "(THE x. P x) = (THE (a, b). P (a, b))" @@ -590,9 +573,9 @@ by (auto simp: fun_eq_iff) text \ - \medskip @{const case_prod} used as a logical connective or set former. + \<^medskip> @{const case_prod} used as a logical connective or set former. - \medskip These rules are for use with \blast\; could instead + \<^medskip> These rules are for use with \blast\; could instead call \simp\ using @{thm [source] prod.split} as rewrite.\ lemma case_prodI2: @@ -621,12 +604,10 @@ using q by (simp add: case_prod_unfold) qed -lemma case_prodD': - "(case (a, b) of (c, d) \ R c d) c \ R a b c" +lemma case_prodD': "(case (a, b) of (c, d) \ R c d) c \ R a b c" by simp -lemma mem_case_prodI: - "z \ c a b \ z \ (case (a, b) of (d, e) \ c d e)" +lemma mem_case_prodI: "z \ c a b \ z \ (case (a, b) of (d, e) \ c d e)" by simp lemma mem_case_prodI2 [intro!]: @@ -661,13 +642,13 @@ to quite time-consuming computations (in particular for nested tuples) *) setup \map_theory_claset (fn ctxt => ctxt addSbefore ("split_conv_tac", split_conv_tac))\ -lemma split_eta_SetCompr [simp, no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P" +lemma split_eta_SetCompr [simp, no_atp]: "(\u. \x y. u = (x, y) \ P (x, y)) = P" by (rule ext) fast -lemma split_eta_SetCompr2 [simp, no_atp]: "(%u. EX x y. u = (x, y) & P x y) = case_prod P" +lemma split_eta_SetCompr2 [simp, no_atp]: "(\u. \x y. u = (x, y) \ P x y) = case_prod P" by (rule ext) fast -lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & case_prod Q ab)" +lemma split_part [simp]: "(\(a,b). P \ Q a b) = (\ab. P \ case_prod Q ab)" \ \Allows simplifications of nested splits in case of independent predicates.\ by (rule ext) blast @@ -676,16 +657,15 @@ b) can lead to nontermination in the presence of split_def *) lemma split_comp_eq: - fixes f :: "'a => 'b => 'c" and g :: "'d => 'a" - shows "(%u. f (g (fst u)) (snd u)) = (case_prod (%x. f (g x)))" + fixes f :: "'a \ 'b \ 'c" + and g :: "'d \ 'a" + shows "(\u. f (g (fst u)) (snd u)) = case_prod (\x. f (g x))" by (rule ext) auto -lemma pair_imageI [intro]: "(a, b) : A ==> f a b : (%(a, b). f a b) ` A" - apply (rule_tac x = "(a, b)" in image_eqI) - apply auto - done +lemma pair_imageI [intro]: "(a, b) \ A \ f a b \ (\(a, b). f a b) ` A" + by (rule image_eqI [where x = "(a, b)"]) auto -lemma The_split_eq [simp]: "(THE (x',y'). x = x' & y = y') = (x, y)" +lemma The_split_eq [simp]: "(THE (x',y'). x = x' \ y = y') = (x, y)" by blast (* @@ -701,8 +681,7 @@ qed "The_split_eq"; *) -lemma case_prod_beta: - "case_prod f p = f (fst p) (snd p)" +lemma case_prod_beta: "case_prod f p = f (fst p) (snd p)" by (fact prod.case_eq_if) lemma prod_cases3 [cases type]: @@ -710,7 +689,7 @@ by (cases y, case_tac b) blast lemma prod_induct3 [case_names fields, induct type]: - "(!!a b c. P (a, b, c)) ==> P x" + "(\a b c. P (a, b, c)) \ P x" by (cases x) blast lemma prod_cases4 [cases type]: @@ -718,7 +697,7 @@ by (cases y, case_tac c) blast lemma prod_induct4 [case_names fields, induct type]: - "(!!a b c d. P (a, b, c, d)) ==> P x" + "(\a b c d. P (a, b, c, d)) \ P x" by (cases x) blast lemma prod_cases5 [cases type]: @@ -726,7 +705,7 @@ by (cases y, case_tac d) blast lemma prod_induct5 [case_names fields, induct type]: - "(!!a b c d e. P (a, b, c, d, e)) ==> P x" + "(\a b c d e. P (a, b, c, d, e)) \ P x" by (cases x) blast lemma prod_cases6 [cases type]: @@ -734,7 +713,7 @@ by (cases y, case_tac e) blast lemma prod_induct6 [case_names fields, induct type]: - "(!!a b c d e f. P (a, b, c, d, e, f)) ==> P x" + "(\a b c d e f. P (a, b, c, d, e, f)) \ P x" by (cases x) blast lemma prod_cases7 [cases type]: @@ -742,11 +721,11 @@ by (cases y, case_tac f) blast lemma prod_induct7 [case_names fields, induct type]: - "(!!a b c d e f g. P (a, b, c, d, e, f, g)) ==> P x" + "(\a b c d e f g. P (a, b, c, d, e, f, g)) \ P x" by (cases x) blast -definition internal_case_prod :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where - "internal_case_prod == case_prod" +definition internal_case_prod :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" + where "internal_case_prod \ case_prod" lemma internal_case_prod_conv: "internal_case_prod c (a, b) = c a b" by (simp only: internal_case_prod_def case_prod_conv) @@ -758,8 +737,8 @@ subsubsection \Derived operations\ -definition curry :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where - "curry = (\c x y. c (x, y))" +definition curry :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" + where "curry = (\c x y. c (x, y))" lemma curry_conv [simp, code]: "curry f a b = f (a, b)" by (simp add: curry_def) @@ -780,16 +759,14 @@ by (simp add: curry_def case_prod_unfold) lemma curry_K: "curry (\x. c) = (\x y. c)" -by(simp add: fun_eq_iff) + by (simp add: fun_eq_iff) -text \ - The composition-uncurry combinator. -\ +text \The composition-uncurry combinator.\ notation fcomp (infixl "\>" 60) -definition scomp :: "('a \ 'b \ 'c) \ ('b \ 'c \ 'd) \ 'a \ 'd" (infixl "\\" 60) where - "f \\ g = (\x. case_prod g (f x))" +definition scomp :: "('a \ 'b \ 'c) \ ('b \ 'c \ 'd) \ 'a \ 'd" (infixl "\\" 60) + where "f \\ g = (\x. case_prod g (f x))" lemma scomp_unfold: "scomp = (\f g x. g (fst (f x)) (snd (f x)))" by (simp add: fun_eq_iff scomp_def case_prod_unfold) @@ -819,46 +796,37 @@ no_notation scomp (infixl "\\" 60) text \ - @{term map_prod} --- action of the product functor upon - functions. + @{term map_prod} --- action of the product functor upon functions. \ -definition map_prod :: "('a \ 'c) \ ('b \ 'd) \ 'a \ 'b \ 'c \ 'd" where - "map_prod f g = (\(x, y). (f x, g y))" +definition map_prod :: "('a \ 'c) \ ('b \ 'd) \ 'a \ 'b \ 'c \ 'd" + where "map_prod f g = (\(x, y). (f x, g y))" -lemma map_prod_simp [simp, code]: - "map_prod f g (a, b) = (f a, g b)" +lemma map_prod_simp [simp, code]: "map_prod f g (a, b) = (f a, g b)" by (simp add: map_prod_def) functor map_prod: map_prod by (auto simp add: split_paired_all) -lemma fst_map_prod [simp]: - "fst (map_prod f g x) = f (fst x)" +lemma fst_map_prod [simp]: "fst (map_prod f g x) = f (fst x)" by (cases x) simp_all -lemma snd_map_prod [simp]: - "snd (map_prod f g x) = g (snd x)" +lemma snd_map_prod [simp]: "snd (map_prod f g x) = g (snd x)" by (cases x) simp_all -lemma fst_comp_map_prod [simp]: - "fst \ map_prod f g = f \ fst" +lemma fst_comp_map_prod [simp]: "fst \ map_prod f g = f \ fst" by (rule ext) simp_all -lemma snd_comp_map_prod [simp]: - "snd \ map_prod f g = g \ snd" +lemma snd_comp_map_prod [simp]: "snd \ map_prod f g = g \ snd" by (rule ext) simp_all -lemma map_prod_compose: - "map_prod (f1 o f2) (g1 o g2) = (map_prod f1 g1 o map_prod f2 g2)" +lemma map_prod_compose: "map_prod (f1 \ f2) (g1 \ g2) = (map_prod f1 g1 \ map_prod f2 g2)" by (rule ext) (simp add: map_prod.compositionality comp_def) -lemma map_prod_ident [simp]: - "map_prod (%x. x) (%y. y) = (%z. z)" +lemma map_prod_ident [simp]: "map_prod (\x. x) (\y. y) = (\z. z)" by (rule ext) (simp add: map_prod.identity) -lemma map_prod_imageI [intro]: - "(a, b) \ R \ (f a, g b) \ map_prod f g ` R" +lemma map_prod_imageI [intro]: "(a, b) \ R \ (f a, g b) \ map_prod f g ` R" by (rule image_eqI) simp_all lemma prod_fun_imageE [elim!]: @@ -871,86 +839,67 @@ apply simp_all done -definition apfst :: "('a \ 'c) \ 'a \ 'b \ 'c \ 'b" where - "apfst f = map_prod f id" +definition apfst :: "('a \ 'c) \ 'a \ 'b \ 'c \ 'b" + where "apfst f = map_prod f id" -definition apsnd :: "('b \ 'c) \ 'a \ 'b \ 'a \ 'c" where - "apsnd f = map_prod id f" +definition apsnd :: "('b \ 'c) \ 'a \ 'b \ 'a \ 'c" + where "apsnd f = map_prod id f" -lemma apfst_conv [simp, code]: - "apfst f (x, y) = (f x, y)" +lemma apfst_conv [simp, code]: "apfst f (x, y) = (f x, y)" by (simp add: apfst_def) -lemma apsnd_conv [simp, code]: - "apsnd f (x, y) = (x, f y)" +lemma apsnd_conv [simp, code]: "apsnd f (x, y) = (x, f y)" by (simp add: apsnd_def) -lemma fst_apfst [simp]: - "fst (apfst f x) = f (fst x)" +lemma fst_apfst [simp]: "fst (apfst f x) = f (fst x)" by (cases x) simp -lemma fst_comp_apfst [simp]: - "fst \ apfst f = f \ fst" +lemma fst_comp_apfst [simp]: "fst \ apfst f = f \ fst" by (simp add: fun_eq_iff) -lemma fst_apsnd [simp]: - "fst (apsnd f x) = fst x" +lemma fst_apsnd [simp]: "fst (apsnd f x) = fst x" by (cases x) simp -lemma fst_comp_apsnd [simp]: - "fst \ apsnd f = fst" +lemma fst_comp_apsnd [simp]: "fst \ apsnd f = fst" by (simp add: fun_eq_iff) -lemma snd_apfst [simp]: - "snd (apfst f x) = snd x" +lemma snd_apfst [simp]: "snd (apfst f x) = snd x" by (cases x) simp -lemma snd_comp_apfst [simp]: - "snd \ apfst f = snd" +lemma snd_comp_apfst [simp]: "snd \ apfst f = snd" by (simp add: fun_eq_iff) -lemma snd_apsnd [simp]: - "snd (apsnd f x) = f (snd x)" +lemma snd_apsnd [simp]: "snd (apsnd f x) = f (snd x)" by (cases x) simp -lemma snd_comp_apsnd [simp]: - "snd \ apsnd f = f \ snd" +lemma snd_comp_apsnd [simp]: "snd \ apsnd f = f \ snd" by (simp add: fun_eq_iff) -lemma apfst_compose: - "apfst f (apfst g x) = apfst (f \ g) x" +lemma apfst_compose: "apfst f (apfst g x) = apfst (f \ g) x" by (cases x) simp -lemma apsnd_compose: - "apsnd f (apsnd g x) = apsnd (f \ g) x" +lemma apsnd_compose: "apsnd f (apsnd g x) = apsnd (f \ g) x" by (cases x) simp -lemma apfst_apsnd [simp]: - "apfst f (apsnd g x) = (f (fst x), g (snd x))" +lemma apfst_apsnd [simp]: "apfst f (apsnd g x) = (f (fst x), g (snd x))" by (cases x) simp -lemma apsnd_apfst [simp]: - "apsnd f (apfst g x) = (g (fst x), f (snd x))" +lemma apsnd_apfst [simp]: "apsnd f (apfst g x) = (g (fst x), f (snd x))" by (cases x) simp -lemma apfst_id [simp] : - "apfst id = id" +lemma apfst_id [simp]: "apfst id = id" by (simp add: fun_eq_iff) -lemma apsnd_id [simp] : - "apsnd id = id" +lemma apsnd_id [simp]: "apsnd id = id" by (simp add: fun_eq_iff) -lemma apfst_eq_conv [simp]: - "apfst f x = apfst g x \ f (fst x) = g (fst x)" +lemma apfst_eq_conv [simp]: "apfst f x = apfst g x \ f (fst x) = g (fst x)" by (cases x) simp -lemma apsnd_eq_conv [simp]: - "apsnd f x = apsnd g x \ f (snd x) = g (snd x)" +lemma apsnd_eq_conv [simp]: "apsnd f x = apsnd g x \ f (snd x) = g (snd x)" by (cases x) simp -lemma apsnd_apfst_commute: - "apsnd f (apfst g p) = apfst g (apsnd f p)" +lemma apsnd_apfst_commute: "apsnd f (apfst g p) = apfst g (apsnd f p)" by simp context @@ -959,104 +908,83 @@ local_setup \Local_Theory.map_background_naming (Name_Space.mandatory_path "prod")\ definition swap :: "'a \ 'b \ 'b \ 'a" -where - "swap p = (snd p, fst p)" + where "swap p = (snd p, fst p)" end -lemma swap_simp [simp]: - "prod.swap (x, y) = (y, x)" +lemma swap_simp [simp]: "prod.swap (x, y) = (y, x)" by (simp add: prod.swap_def) -lemma swap_swap [simp]: - "prod.swap (prod.swap p) = p" +lemma swap_swap [simp]: "prod.swap (prod.swap p) = p" by (cases p) simp -lemma swap_comp_swap [simp]: - "prod.swap \ prod.swap = id" +lemma swap_comp_swap [simp]: "prod.swap \ prod.swap = id" by (simp add: fun_eq_iff) -lemma pair_in_swap_image [simp]: - "(y, x) \ prod.swap ` A \ (x, y) \ A" +lemma pair_in_swap_image [simp]: "(y, x) \ prod.swap ` A \ (x, y) \ A" by (auto intro!: image_eqI) -lemma inj_swap [simp]: - "inj_on prod.swap A" +lemma inj_swap [simp]: "inj_on prod.swap A" by (rule inj_onI) auto -lemma swap_inj_on: - "inj_on (\(i, j). (j, i)) A" +lemma swap_inj_on: "inj_on (\(i, j). (j, i)) A" by (rule inj_onI) auto -lemma surj_swap [simp]: - "surj prod.swap" +lemma surj_swap [simp]: "surj prod.swap" by (rule surjI [of _ prod.swap]) simp -lemma bij_swap [simp]: - "bij prod.swap" +lemma bij_swap [simp]: "bij prod.swap" by (simp add: bij_def) -lemma case_swap [simp]: - "(case prod.swap p of (y, x) \ f x y) = (case p of (x, y) \ f x y)" +lemma case_swap [simp]: "(case prod.swap p of (y, x) \ f x y) = (case p of (x, y) \ f x y)" by (cases p) simp lemma fst_swap [simp]: "fst (prod.swap x) = snd x" -by(cases x) simp + by (cases x) simp lemma snd_swap [simp]: "snd (prod.swap x) = fst x" -by(cases x) simp + by (cases x) simp -text \ - Disjoint union of a family of sets -- Sigma. -\ +text \Disjoint union of a family of sets -- Sigma.\ -definition Sigma :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" where - Sigma_def: "Sigma A B == UN x:A. UN y:B x. {Pair x y}" +definition Sigma :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" + where "Sigma A B \ \x\A. \y\B x. {Pair x y}" -abbreviation - Times :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 80) where - "A \ B == Sigma A (%_. B)" +abbreviation Times :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 80) + where "A \ B \ Sigma A (\_. B)" hide_const (open) Times syntax - "_Sigma" :: "[pttrn, 'a set, 'b set] => ('a * 'b) set" ("(3SIGMA _:_./ _)" [0, 0, 10] 10) + "_Sigma" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3SIGMA _:_./ _)" [0, 0, 10] 10) translations - "SIGMA x:A. B" == "CONST Sigma A (%x. B)" - -lemma SigmaI [intro!]: "[| a:A; b:B(a) |] ==> (a,b) : Sigma A B" - by (unfold Sigma_def) blast + "SIGMA x:A. B" \ "CONST Sigma A (\x. B)" -lemma SigmaE [elim!]: - "[| c: Sigma A B; - !!x y.[| x:A; y:B(x); c=(x,y) |] ==> P - |] ==> P" +lemma SigmaI [intro!]: "a \ A \ b \ B a \ (a, b) \ Sigma A B" + unfolding Sigma_def by blast + +lemma SigmaE [elim!]: "c \ Sigma A B \ (\x y. x \ A \ y \ B x \ c = (x, y) \ P) \ P" \ \The general elimination rule.\ - by (unfold Sigma_def) blast + unfolding Sigma_def by blast text \ - Elimination of @{term "(a, b) : A \ B"} -- introduces no + Elimination of @{term "(a, b) \ A \ B"} -- introduces no eigenvariables. \ -lemma SigmaD1: "(a, b) : Sigma A B ==> a : A" +lemma SigmaD1: "(a, b) \ Sigma A B \ a \ A" by blast -lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a" +lemma SigmaD2: "(a, b) \ Sigma A B \ b \ B a" by blast -lemma SigmaE2: - "[| (a, b) : Sigma A B; - [| a:A; b:B(a) |] ==> P - |] ==> P" +lemma SigmaE2: "(a, b) \ Sigma A B \ (a \ A \ b \ B a \ P) \ P" by blast -lemma Sigma_cong: - "\A = B; !!x. x \ B \ C x = D x\ - \ (SIGMA x: A. C x) = (SIGMA x: B. D x)" +lemma Sigma_cong: "A = B \ (\x. x \ B \ C x = D x) \ (SIGMA x:A. C x) = (SIGMA x:B. D x)" by auto -lemma Sigma_mono: "[| A <= C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D" +lemma Sigma_mono: "A \ C \ (\x. x \ A \ B x \ D x) \ Sigma A B \ Sigma C D" by blast lemma Sigma_empty1 [simp]: "Sigma {} B = {}" @@ -1074,7 +1002,7 @@ lemma Compl_Times_UNIV2 [simp]: "- (A \ UNIV) = (-A) \ UNIV" by auto -lemma mem_Sigma_iff [iff]: "((a,b): Sigma A B) = (a:A & b:B(a))" +lemma mem_Sigma_iff [iff]: "(a, b) \ Sigma A B \ a \ A \ b \ B a" by blast lemma mem_Times_iff: "x \ A \ B \ fst x \ A \ snd x \ B" @@ -1083,26 +1011,22 @@ lemma Sigma_empty_iff: "(SIGMA i:I. X i) = {} \ (\i\I. X i = {})" by auto -lemma Times_subset_cancel2: "x:C ==> (A \ C <= B \ C) = (A <= B)" - by blast - -lemma Times_eq_cancel2: "x:C ==> (A \ C = B \ C) = (A = B)" - by (blast elim: equalityE) - -lemma Collect_case_prod_Sigma: - "{(x, y). P x \ Q x y} = (SIGMA x:Collect P. Collect (Q x))" +lemma Times_subset_cancel2: "x \ C \ A \ C \ B \ C \ A \ B" by blast -lemma Collect_case_prod [simp]: - "{(a, b). P a \ Q b} = Collect P \ Collect Q " +lemma Times_eq_cancel2: "x \ C \ A \ C = B \ C \ A = B" + by (blast elim: equalityE) + +lemma Collect_case_prod_Sigma: "{(x, y). P x \ Q x y} = (SIGMA x:Collect P. Collect (Q x))" + by blast + +lemma Collect_case_prod [simp]: "{(a, b). P a \ Q b} = Collect P \ Collect Q " by (fact Collect_case_prod_Sigma) -lemma Collect_case_prodD: - "x \ Collect (case_prod A) \ A (fst x) (snd x)" +lemma Collect_case_prodD: "x \ Collect (case_prod A) \ A (fst x) (snd x)" by auto -lemma Collect_case_prod_mono: - "A \ B \ Collect (case_prod A) \ Collect (case_prod B)" +lemma Collect_case_prod_mono: "A \ B \ Collect (case_prod A) \ Collect (case_prod B)" by auto (auto elim!: le_funE) lemma Collect_split_mono_strong: @@ -1110,45 +1034,35 @@ \ A \ Collect (case_prod P) \ A \ Collect (case_prod Q)" by fastforce -lemma UN_Times_distrib: - "(\(a, b)\A \ B. E a \ F b) = UNION A E \ UNION B F" +lemma UN_Times_distrib: "(\(a, b)\A \ B. E a \ F b) = UNION A E \ UNION B F" \ \Suggested by Pierre Chartier\ by blast -lemma split_paired_Ball_Sigma [simp, no_atp]: - "(\z\Sigma A B. P z) \ (\x\A. \y\B x. P (x, y))" +lemma split_paired_Ball_Sigma [simp, no_atp]: "(\z\Sigma A B. P z) \ (\x\A. \y\B x. P (x, y))" by blast -lemma split_paired_Bex_Sigma [simp, no_atp]: - "(\z\Sigma A B. P z) \ (\x\A. \y\B x. P (x, y))" +lemma split_paired_Bex_Sigma [simp, no_atp]: "(\z\Sigma A B. P z) \ (\x\A. \y\B x. P (x, y))" by blast -lemma Sigma_Un_distrib1: - "Sigma (I \ J) C = Sigma I C \ Sigma J C" +lemma Sigma_Un_distrib1: "Sigma (I \ J) C = Sigma I C \ Sigma J C" by blast -lemma Sigma_Un_distrib2: - "(SIGMA i:I. A i \ B i) = Sigma I A \ Sigma I B" +lemma Sigma_Un_distrib2: "(SIGMA i:I. A i \ B i) = Sigma I A \ Sigma I B" by blast -lemma Sigma_Int_distrib1: - "Sigma (I \ J) C = Sigma I C \ Sigma J C" +lemma Sigma_Int_distrib1: "Sigma (I \ J) C = Sigma I C \ Sigma J C" by blast -lemma Sigma_Int_distrib2: - "(SIGMA i:I. A i \ B i) = Sigma I A \ Sigma I B" +lemma Sigma_Int_distrib2: "(SIGMA i:I. A i \ B i) = Sigma I A \ Sigma I B" by blast -lemma Sigma_Diff_distrib1: - "Sigma (I - J) C = Sigma I C - Sigma J C" +lemma Sigma_Diff_distrib1: "Sigma (I - J) C = Sigma I C - Sigma J C" by blast -lemma Sigma_Diff_distrib2: - "(SIGMA i:I. A i - B i) = Sigma I A - Sigma I B" +lemma Sigma_Diff_distrib2: "(SIGMA i:I. A i - B i) = Sigma I A - Sigma I B" by blast -lemma Sigma_Union: - "Sigma (\X) B = (\A\X. Sigma A B)" +lemma Sigma_Union: "Sigma (\X) B = (\A\X. Sigma A B)" by blast lemma Pair_vimage_Sigma: "Pair x -` Sigma A f = (if x \ A then f x else {})" @@ -1159,79 +1073,62 @@ matching, especially when the rules are re-oriented. \ -lemma Times_Un_distrib1: - "(A \ B) \ C = A \ C \ B \ C " +lemma Times_Un_distrib1: "(A \ B) \ C = A \ C \ B \ C " by (fact Sigma_Un_distrib1) -lemma Times_Int_distrib1: - "(A \ B) \ C = A \ C \ B \ C " +lemma Times_Int_distrib1: "(A \ B) \ C = A \ C \ B \ C " by (fact Sigma_Int_distrib1) -lemma Times_Diff_distrib1: - "(A - B) \ C = A \ C - B \ C " +lemma Times_Diff_distrib1: "(A - B) \ C = A \ C - B \ C " by (fact Sigma_Diff_distrib1) -lemma Times_empty [simp]: - "A \ B = {} \ A = {} \ B = {}" +lemma Times_empty [simp]: "A \ B = {} \ A = {} \ B = {}" by auto -lemma times_eq_iff: - "A \ B = C \ D \ A = C \ B = D \ (A = {} \ B = {}) \ (C = {} \ D = {})" +lemma times_eq_iff: "A \ B = C \ D \ A = C \ B = D \ (A = {} \ B = {}) \ (C = {} \ D = {})" by auto -lemma fst_image_times [simp]: - "fst ` (A \ B) = (if B = {} then {} else A)" +lemma fst_image_times [simp]: "fst ` (A \ B) = (if B = {} then {} else A)" by force -lemma snd_image_times [simp]: - "snd ` (A \ B) = (if A = {} then {} else B)" +lemma snd_image_times [simp]: "snd ` (A \ B) = (if A = {} then {} else B)" by force -lemma fst_image_Sigma: - "fst ` (Sigma A B) = {x \ A. B(x) \ {}}" +lemma fst_image_Sigma: "fst ` (Sigma A B) = {x \ A. B(x) \ {}}" by force -lemma snd_image_Sigma: - "snd ` (Sigma A B) = (\ x \ A. B x)" +lemma snd_image_Sigma: "snd ` (Sigma A B) = (\ x \ A. B x)" by force -lemma vimage_fst: - "fst -` A = A \ UNIV" +lemma vimage_fst: "fst -` A = A \ UNIV" by auto -lemma vimage_snd: - "snd -` A = UNIV \ A" +lemma vimage_snd: "snd -` A = UNIV \ A" by auto -lemma insert_times_insert[simp]: - "insert a A \ insert b B = - insert (a,b) (A \ insert b B \ insert a A \ B)" +lemma insert_times_insert [simp]: + "insert a A \ insert b B = insert (a,b) (A \ insert b B \ insert a A \ B)" by blast -lemma vimage_Times: - "f -` (A \ B) = (fst \ f) -` A \ (snd \ f) -` B" +lemma vimage_Times: "f -` (A \ B) = (fst \ f) -` A \ (snd \ f) -` B" proof (rule set_eqI) - fix x - show "x \ f -` (A \ B) \ x \ (fst \ f) -` A \ (snd \ f) -` B" + show "x \ f -` (A \ B) \ x \ (fst \ f) -` A \ (snd \ f) -` B" for x by (cases "f x") (auto split: prod.split) qed -lemma times_Int_times: - "A \ B \ C \ D = (A \ C) \ (B \ D)" +lemma times_Int_times: "A \ B \ C \ D = (A \ C) \ (B \ D)" by auto -lemma product_swap: - "prod.swap ` (A \ B) = B \ A" +lemma product_swap: "prod.swap ` (A \ B) = B \ A" by (auto simp add: set_eq_iff) -lemma swap_product: - "(\(i, j). (j, i)) ` (A \ B) = B \ A" +lemma swap_product: "(\(i, j). (j, i)) ` (A \ B) = B \ A" by (auto simp add: set_eq_iff) -lemma image_split_eq_Sigma: - "(\x. (f x, g x)) ` A = Sigma (f ` A) (\x. g ` (f -` {x} \ A))" +lemma image_split_eq_Sigma: "(\x. (f x, g x)) ` A = Sigma (f ` A) (\x. g ` (f -` {x} \ A))" proof (safe intro!: imageI) - fix a b assume *: "a \ A" "b \ A" and eq: "f a = f b" + fix a b + assume *: "a \ A" "b \ A" and eq: "f a = f b" show "(f b, g a) \ (\x. (f x, g x)) ` A" using * eq[symmetric] by auto qed simp_all @@ -1240,16 +1137,16 @@ by force lemma inj_on_apfst [simp]: "inj_on (apfst f) (A \ UNIV) \ inj_on f A" -by(auto simp add: inj_on_def) + by (auto simp add: inj_on_def) lemma inj_apfst [simp]: "inj (apfst f) \ inj f" -using inj_on_apfst[of f UNIV] by simp + using inj_on_apfst[of f UNIV] by simp lemma inj_on_apsnd [simp]: "inj_on (apsnd f) (UNIV \ A) \ inj_on f A" -by(auto simp add: inj_on_def) + by (auto simp add: inj_on_def) lemma inj_apsnd [simp]: "inj (apsnd f) \ inj f" -using inj_on_apsnd[of f UNIV] by simp + using inj_on_apsnd[of f UNIV] by simp context begin @@ -1257,9 +1154,8 @@ qualified definition product :: "'a set \ 'b set \ ('a \ 'b) set" where [code_abbrev]: "product A B = A \ B" -lemma member_product: - "x \ Product_Type.product A B \ x \ A \ B" - by (simp add: Product_Type.product_def) +lemma member_product: "x \ Product_Type.product A B \ x \ A \ B" + by (simp add: product_def) end @@ -1269,59 +1165,80 @@ assumes "inj_on f A" and "inj_on g B" shows "inj_on (map_prod f g) (A \ B)" proof (rule inj_onI) - fix x :: "'a \ 'c" and y :: "'a \ 'c" - assume "x \ A \ B" hence "fst x \ A" and "snd x \ B" by auto - assume "y \ A \ B" hence "fst y \ A" and "snd y \ B" by auto + fix x :: "'a \ 'c" + fix y :: "'a \ 'c" + assume "x \ A \ B" + then have "fst x \ A" and "snd x \ B" by auto + assume "y \ A \ B" + then have "fst y \ A" and "snd y \ B" by auto assume "map_prod f g x = map_prod f g y" - hence "fst (map_prod f g x) = fst (map_prod f g y)" by (auto) - hence "f (fst x) = f (fst y)" by (cases x,cases y,auto) - with \inj_on f A\ and \fst x \ A\ and \fst y \ A\ - have "fst x = fst y" by (auto dest:dest:inj_onD) + then have "fst (map_prod f g x) = fst (map_prod f g y)" by auto + then have "f (fst x) = f (fst y)" by (cases x, cases y) auto + with \inj_on f A\ and \fst x \ A\ and \fst y \ A\ have "fst x = fst y" + by (auto dest: inj_onD) moreover from \map_prod f g x = map_prod f g y\ - have "snd (map_prod f g x) = snd (map_prod f g y)" by (auto) - hence "g (snd x) = g (snd y)" by (cases x,cases y,auto) - with \inj_on g B\ and \snd x \ B\ and \snd y \ B\ - have "snd x = snd y" by (auto dest:dest:inj_onD) - ultimately show "x = y" by(rule prod_eqI) + have "snd (map_prod f g x) = snd (map_prod f g y)" by auto + then have "g (snd x) = g (snd y)" by (cases x, cases y) auto + with \inj_on g B\ and \snd x \ B\ and \snd y \ B\ have "snd x = snd y" + by (auto dest: inj_onD) + ultimately show "x = y" by (rule prod_eqI) qed lemma map_prod_surj: - fixes f :: "'a \ 'b" and g :: "'c \ 'd" + fixes f :: "'a \ 'b" + and g :: "'c \ 'd" assumes "surj f" and "surj g" shows "surj (map_prod f g)" -unfolding surj_def + unfolding surj_def proof fix y :: "'b \ 'd" - from \surj f\ obtain a where "fst y = f a" by (auto elim:surjE) + from \surj f\ obtain a where "fst y = f a" + by (auto elim: surjE) moreover - from \surj g\ obtain b where "snd y = g b" by (auto elim:surjE) - ultimately have "(fst y, snd y) = map_prod f g (a,b)" by auto - thus "\x. y = map_prod f g x" by auto + from \surj g\ obtain b where "snd y = g b" + by (auto elim: surjE) + ultimately have "(fst y, snd y) = map_prod f g (a,b)" + by auto + then show "\x. y = map_prod f g x" + by auto qed lemma map_prod_surj_on: assumes "f ` A = A'" and "g ` B = B'" shows "map_prod f g ` (A \ B) = A' \ B'" -unfolding image_def -proof(rule set_eqI,rule iffI) + unfolding image_def +proof (rule set_eqI, rule iffI) fix x :: "'a \ 'c" assume "x \ {y::'a \ 'c. \x::'b \ 'd\A \ B. y = map_prod f g x}" - then obtain y where "y \ A \ B" and "x = map_prod f g y" by blast - from \image f A = A'\ and \y \ A \ B\ have "f (fst y) \ A'" by auto - moreover from \image g B = B'\ and \y \ A \ B\ have "g (snd y) \ B'" by auto - ultimately have "(f (fst y), g (snd y)) \ (A' \ B')" by auto - with \x = map_prod f g y\ show "x \ A' \ B'" by (cases y, auto) + then obtain y where "y \ A \ B" and "x = map_prod f g y" + by blast + from \image f A = A'\ and \y \ A \ B\ have "f (fst y) \ A'" + by auto + moreover from \image g B = B'\ and \y \ A \ B\ have "g (snd y) \ B'" + by auto + ultimately have "(f (fst y), g (snd y)) \ (A' \ B')" + by auto + with \x = map_prod f g y\ show "x \ A' \ B'" + by (cases y) auto next fix x :: "'a \ 'c" - assume "x \ A' \ B'" hence "fst x \ A'" and "snd x \ B'" by auto - from \image f A = A'\ and \fst x \ A'\ have "fst x \ image f A" by auto - then obtain a where "a \ A" and "fst x = f a" by (rule imageE) - moreover from \image g B = B'\ and \snd x \ B'\ - obtain b where "b \ B" and "snd x = g b" by auto - ultimately have "(fst x, snd x) = map_prod f g (a,b)" by auto - moreover from \a \ A\ and \b \ B\ have "(a , b) \ A \ B" by auto - ultimately have "\y \ A \ B. x = map_prod f g y" by auto - thus "x \ {x. \y \ A \ B. x = map_prod f g y}" by auto + assume "x \ A' \ B'" + then have "fst x \ A'" and "snd x \ B'" + by auto + from \image f A = A'\ and \fst x \ A'\ have "fst x \ image f A" + by auto + then obtain a where "a \ A" and "fst x = f a" + by (rule imageE) + moreover from \image g B = B'\ and \snd x \ B'\ obtain b where "b \ B" and "snd x = g b" + by auto + ultimately have "(fst x, snd x) = map_prod f g (a, b)" + by auto + moreover from \a \ A\ and \b \ B\ have "(a , b) \ A \ B" + by auto + ultimately have "\y \ A \ B. x = map_prod f g y" + by auto + then show "x \ {x. \y \ A \ B. x = map_prod f g y}" + by auto qed diff -r d1742d1b7f0f -r 249fa34faba2 src/HOL/Set.thy --- a/src/HOL/Set.thy Tue Jul 05 22:47:48 2016 +0200 +++ b/src/HOL/Set.thy Tue Jul 05 23:39:49 2016 +0200 @@ -466,7 +466,7 @@ \ \Classical elimination rule.\ by (auto simp add: less_eq_set_def le_fun_def) -lemma subset_eq: "A \ B = (\x\A. x \ B)" +lemma subset_eq: "A \ B \ (\x\A. x \ B)" by blast lemma contra_subsetD: "A \ B \ c \ B \ c \ A" diff -r d1742d1b7f0f -r 249fa34faba2 src/HOL/Sum_Type.thy --- a/src/HOL/Sum_Type.thy Tue Jul 05 22:47:48 2016 +0200 +++ b/src/HOL/Sum_Type.thy Tue Jul 05 23:39:49 2016 +0200 @@ -11,15 +11,15 @@ subsection \Construction of the sum type and its basic abstract operations\ -definition Inl_Rep :: "'a \ 'a \ 'b \ bool \ bool" where - "Inl_Rep a x y p \ x = a \ p" +definition Inl_Rep :: "'a \ 'a \ 'b \ bool \ bool" + where "Inl_Rep a x y p \ x = a \ p" -definition Inr_Rep :: "'b \ 'a \ 'b \ bool \ bool" where - "Inr_Rep b x y p \ y = b \ \ p" +definition Inr_Rep :: "'b \ 'a \ 'b \ bool \ bool" + where "Inr_Rep b x y p \ y = b \ \ p" definition "sum = {f. (\a. f = Inl_Rep (a::'a)) \ (\b. f = Inr_Rep (b::'b))}" -typedef ('a, 'b) sum (infixr "+" 10) = "sum :: ('a => 'b => bool => bool) set" +typedef ('a, 'b) sum (infixr "+" 10) = "sum :: ('a \ 'b \ bool \ bool) set" unfolding sum_def by auto lemma Inl_RepI: "Inl_Rep a \ sum" @@ -46,48 +46,51 @@ lemma Inl_Rep_not_Inr_Rep: "Inl_Rep a \ Inr_Rep b" by (auto simp add: Inl_Rep_def Inr_Rep_def fun_eq_iff) -definition Inl :: "'a \ 'a + 'b" where - "Inl = Abs_sum \ Inl_Rep" +definition Inl :: "'a \ 'a + 'b" + where "Inl = Abs_sum \ Inl_Rep" -definition Inr :: "'b \ 'a + 'b" where - "Inr = Abs_sum \ Inr_Rep" +definition Inr :: "'b \ 'a + 'b" + where "Inr = Abs_sum \ Inr_Rep" lemma inj_Inl [simp]: "inj_on Inl A" -by (auto simp add: Inl_def intro!: comp_inj_on Inl_Rep_inject inj_on_Abs_sum Inl_RepI) + by (auto simp add: Inl_def intro!: comp_inj_on Inl_Rep_inject inj_on_Abs_sum Inl_RepI) lemma Inl_inject: "Inl x = Inl y \ x = y" -using inj_Inl by (rule injD) + using inj_Inl by (rule injD) lemma inj_Inr [simp]: "inj_on Inr A" -by (auto simp add: Inr_def intro!: comp_inj_on Inr_Rep_inject inj_on_Abs_sum Inr_RepI) + by (auto simp add: Inr_def intro!: comp_inj_on Inr_Rep_inject inj_on_Abs_sum Inr_RepI) lemma Inr_inject: "Inr x = Inr y \ x = y" -using inj_Inr by (rule injD) + using inj_Inr by (rule injD) lemma Inl_not_Inr: "Inl a \ Inr b" proof - - from Inl_RepI [of a] Inr_RepI [of b] have "{Inl_Rep a, Inr_Rep b} \ sum" by auto + have "{Inl_Rep a, Inr_Rep b} \ sum" + using Inl_RepI [of a] Inr_RepI [of b] by auto with inj_on_Abs_sum have "inj_on Abs_sum {Inl_Rep a, Inr_Rep b}" . - with Inl_Rep_not_Inr_Rep [of a b] inj_on_contraD have "Abs_sum (Inl_Rep a) \ Abs_sum (Inr_Rep b)" by auto - then show ?thesis by (simp add: Inl_def Inr_def) + with Inl_Rep_not_Inr_Rep [of a b] inj_on_contraD have "Abs_sum (Inl_Rep a) \ Abs_sum (Inr_Rep b)" + by auto + then show ?thesis + by (simp add: Inl_def Inr_def) qed -lemma Inr_not_Inl: "Inr b \ Inl a" +lemma Inr_not_Inl: "Inr b \ Inl a" using Inl_not_Inr by (rule not_sym) -lemma sumE: +lemma sumE: assumes "\x::'a. s = Inl x \ P" and "\y::'b. s = Inr y \ P" shows P proof (rule Abs_sum_cases [of s]) - fix f + fix f assume "s = Abs_sum f" and "f \ sum" with assms show P by (auto simp add: sum_def Inl_def Inr_def) qed free_constructors case_sum for - isl: Inl projl - | Inr projr + isl: Inl projl +| Inr projr by (erule sumE, assumption) (auto dest: Inl_inject Inr_inject simp add: Inl_not_Inr) text \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ @@ -119,24 +122,21 @@ setup \Sign.parent_path\ -primrec map_sum :: "('a \ 'c) \ ('b \ 'd) \ 'a + 'b \ 'c + 'd" where +primrec map_sum :: "('a \ 'c) \ ('b \ 'd) \ 'a + 'b \ 'c + 'd" +where "map_sum f1 f2 (Inl a) = Inl (f1 a)" | "map_sum f1 f2 (Inr a) = Inr (f2 a)" -functor map_sum: map_sum proof - - fix f g h i - show "map_sum f g \ map_sum h i = map_sum (f \ h) (g \ i)" +functor map_sum: map_sum +proof - + show "map_sum f g \ map_sum h i = map_sum (f \ h) (g \ i)" for f g h i proof - fix s - show "(map_sum f g \ map_sum h i) s = map_sum (f \ h) (g \ i) s" + show "(map_sum f g \ map_sum h i) s = map_sum (f \ h) (g \ i) s" for s by (cases s) simp_all qed -next - fix s show "map_sum id id = id" proof - fix s - show "map_sum id id s = id s" + show "map_sum id id s = id s" for s by (cases s) simp_all qed qed @@ -145,7 +145,8 @@ by (auto intro: sum.induct) lemma split_sum_ex: "(\x. P x) \ (\x. P (Inl x)) \ (\x. P (Inr x))" -using split_sum_all[of "\x. \P x"] by blast + using split_sum_all[of "\x. \P x"] by blast + subsection \Projections\ @@ -161,29 +162,32 @@ lemma case_sum_inject: assumes a: "case_sum f1 f2 = case_sum g1 g2" - assumes r: "f1 = g1 \ f2 = g2 \ P" + and r: "f1 = g1 \ f2 = g2 \ P" shows P proof (rule r) - show "f1 = g1" proof + show "f1 = g1" + proof fix x :: 'a from a have "case_sum f1 f2 (Inl x) = case_sum g1 g2 (Inl x)" by simp then show "f1 x = g1 x" by simp qed - show "f2 = g2" proof + show "f2 = g2" + proof fix y :: 'b from a have "case_sum f1 f2 (Inr y) = case_sum g1 g2 (Inr y)" by simp then show "f2 y = g2 y" by simp qed qed -primrec Suml :: "('a \ 'c) \ 'a + 'b \ 'c" where - "Suml f (Inl x) = f x" +primrec Suml :: "('a \ 'c) \ 'a + 'b \ 'c" + where "Suml f (Inl x) = f x" -primrec Sumr :: "('b \ 'c) \ 'a + 'b \ 'c" where - "Sumr f (Inr x) = f x" +primrec Sumr :: "('b \ 'c) \ 'a + 'b \ 'c" + where "Sumr f (Inr x) = f x" lemma Suml_inject: - assumes "Suml f = Suml g" shows "f = g" + assumes "Suml f = Suml g" + shows "f = g" proof fix x :: 'a let ?s = "Inl x :: 'a + 'b" @@ -192,7 +196,8 @@ qed lemma Sumr_inject: - assumes "Sumr f = Sumr g" shows "f = g" + assumes "Sumr f = Sumr g" + shows "f = g" proof fix x :: 'b let ?s = "Inr x :: 'a + 'b" @@ -203,25 +208,25 @@ subsection \The Disjoint Sum of Sets\ -definition Plus :: "'a set \ 'b set \ ('a + 'b) set" (infixr "<+>" 65) where - "A <+> B = Inl ` A \ Inr ` B" +definition Plus :: "'a set \ 'b set \ ('a + 'b) set" (infixr "<+>" 65) + where "A <+> B = Inl ` A \ Inr ` B" -hide_const (open) Plus \"Valuable identifier" +hide_const (open) Plus \ "Valuable identifier" lemma InlI [intro!]: "a \ A \ Inl a \ A <+> B" -by (simp add: Plus_def) + by (simp add: Plus_def) lemma InrI [intro!]: "b \ B \ Inr b \ A <+> B" -by (simp add: Plus_def) + by (simp add: Plus_def) text \Exhaustion rule for sums, a degenerate form of induction\ -lemma PlusE [elim!]: +lemma PlusE [elim!]: "u \ A <+> B \ (\x. x \ A \ u = Inl x \ P) \ (\y. y \ B \ u = Inr y \ P) \ P" -by (auto simp add: Plus_def) + by (auto simp add: Plus_def) lemma Plus_eq_empty_conv [simp]: "A <+> B = {} \ A = {} \ B = {}" -by auto + by auto lemma UNIV_Plus_UNIV [simp]: "UNIV <+> UNIV = UNIV" proof (rule set_eqI) @@ -229,14 +234,11 @@ show "u \ UNIV <+> UNIV \ u \ UNIV" by (cases u) auto qed -lemma UNIV_sum: - "UNIV = Inl ` UNIV \ Inr ` UNIV" +lemma UNIV_sum: "UNIV = Inl ` UNIV \ Inr ` UNIV" proof - - { fix x :: "'a + 'b" - assume "x \ range Inr" - then have "x \ range Inl" - by (cases x) simp_all - } then show ?thesis by auto + have "x \ range Inl" if "x \ range Inr" for x :: "'a + 'b" + using that by (cases x) simp_all + then show ?thesis by auto qed hide_const (open) Suml Sumr sum