# HG changeset patch # User haftmann # Date 1329808542 -3600 # Node ID ae926869a31126a912779c5e34818ad57106d17b # Parent 2848e36e0348026c900cc6c3f4cd581627e36d72 reverting changesets from 5d33a3269029 on: change of order of declaration of classical rules makes serious problems diff -r 2848e36e0348 -r ae926869a311 src/HOL/Big_Operators.thy --- a/src/HOL/Big_Operators.thy Mon Feb 20 21:04:00 2012 +0100 +++ b/src/HOL/Big_Operators.thy Tue Feb 21 08:15:42 2012 +0100 @@ -786,13 +786,15 @@ by (simp only: card_def setsum_def) lemma card_UN_disjoint: - assumes "finite I" and "\i\I. finite (A i)" - and "\i\I. \j\I. i \ j \ A i \ A j = {}" - shows "card (UNION I A) = (\i\I. card(A i))" -proof - - have "(\i\I. card (A i)) = (\i\I. \x\A i. 1)" by simp - with assms show ?thesis by (simp add: card_eq_setsum setsum_UN_disjoint del: setsum_constant) -qed + "finite I ==> (ALL i:I. finite (A i)) ==> + (ALL i:I. ALL j:I. i \ j --> A i Int A j = {}) + ==> card (UNION I A) = (\i\I. card(A i))" +apply (simp add: card_eq_setsum del: setsum_constant) +apply (subgoal_tac + "setsum (%i. card (A i)) I = setsum (%i. (setsum (%x. 1) (A i))) I") +apply (simp add: setsum_UN_disjoint del: setsum_constant) +apply simp +done lemma card_Union_disjoint: "finite C ==> (ALL A:C. finite A) ==> diff -r 2848e36e0348 -r ae926869a311 src/HOL/Complete_Lattices.thy --- a/src/HOL/Complete_Lattices.thy Mon Feb 20 21:04:00 2012 +0100 +++ b/src/HOL/Complete_Lattices.thy Tue Feb 21 08:15:42 2012 +0100 @@ -536,7 +536,7 @@ end -subsection {* @{typ bool} as complete lattice *} +subsection {* @{typ bool} and @{typ "_ \ _"} as complete lattice *} instantiation bool :: complete_lattice begin @@ -573,9 +573,6 @@ instance bool :: complete_boolean_algebra proof qed (auto intro: bool_induct) - -subsection {* @{typ "_ \ _"} as complete lattice *} - instantiation "fun" :: (type, complete_lattice) complete_lattice begin @@ -611,54 +608,6 @@ instance "fun" :: (type, complete_boolean_algebra) complete_boolean_algebra .. - -subsection {* Unary and binary predicates as complete lattice *} - -lemma INF1_iff: "(\x\A. B x) b = (\x\A. B x b)" - by (simp add: INF_apply) - -lemma INF2_iff: "(\x\A. B x) b c = (\x\A. B x b c)" - by (simp add: INF_apply) - -lemma INF1_I [intro!]: "(\x. x \ A \ B x b) \ (\x\A. B x) b" - by (auto simp add: INF_apply) - -lemma INF2_I [intro!]: "(\x. x \ A \ B x b c) \ (\x\A. B x) b c" - by (auto simp add: INF_apply) - -lemma INF1_D [elim]: "(\x\A. B x) b \ a \ A \ B a b" - by (auto simp add: INF_apply) - -lemma INF2_D [elim]: "(\x\A. B x) b c \ a \ A \ B a b c" - by (auto simp add: INF_apply) - -lemma INF1_E [elim]: "(\x\A. B x) b \ (B a b \ R) \ (a \ A \ R) \ R" - by (auto simp add: INF_apply) - -lemma INF2_E [elim]: "(\x\A. B x) b c \ (B a b c \ R) \ (a \ A \ R) \ R" - by (auto simp add: INF_apply) - -lemma SUP1_iff: "(\x\A. B x) b = (\x\A. B x b)" - by (simp add: SUP_apply) - -lemma SUP2_iff: "(\x\A. B x) b c = (\x\A. B x b c)" - by (simp add: SUP_apply) - -lemma SUP1_I [intro]: "a \ A \ B a b \ (\x\A. B x) b" - by (auto simp add: SUP_apply) - -lemma SUP2_I [intro]: "a \ A \ B a b c \ (\x\A. B x) b c" - by (auto simp add: SUP_apply) - -lemma SUP1_E [elim!]: "(\x\A. B x) b \ (\x. x \ A \ B x b \ R) \ R" - by (auto simp add: SUP_apply) - -lemma SUP2_E [elim!]: "(\x\A. B x) b c \ (\x. x \ A \ B x b c \ R) \ R" - by (auto simp add: SUP_apply) - - -subsection {* @{typ "_ set"} as complete lattice *} - instantiation "set" :: (type) complete_lattice begin @@ -678,7 +627,7 @@ qed (auto simp add: INF_def SUP_def Inf_set_def Sup_set_def image_def) -subsubsection {* Inter *} +subsection {* Inter *} abbreviation Inter :: "'a set set \ 'a set" where "Inter S \ \S" @@ -750,7 +699,7 @@ by (fact Inf_superset_mono) -subsubsection {* Intersections of families *} +subsection {* Intersections of families *} abbreviation INTER :: "'a set \ ('a \ 'b set) \ 'b set" where "INTER \ INFI" @@ -861,7 +810,7 @@ by blast -subsubsection {* Union *} +subsection {* Union *} abbreviation Union :: "'a set set \ 'a set" where "Union S \ \S" @@ -930,7 +879,7 @@ by (fact Sup_subset_mono) -subsubsection {* Unions of families *} +subsection {* Unions of families *} abbreviation UNION :: "'a set \ ('a \ 'b set) \ 'b set" where "UNION \ SUPR" @@ -1093,7 +1042,7 @@ by blast -subsubsection {* Distributive laws *} +subsection {* Distributive laws *} lemma Int_Union: "A \ \B = (\C\B. A \ C)" by (fact inf_Sup) @@ -1135,7 +1084,7 @@ by (fact Sup_inf_eq_bot_iff) -subsubsection {* Complement *} +subsection {* Complement *} lemma Compl_INT [simp]: "- (\x\A. B x) = (\x\A. -B x)" by (fact uminus_INF) @@ -1144,7 +1093,7 @@ by (fact uminus_SUP) -subsubsection {* Miniscoping and maxiscoping *} +subsection {* Miniscoping and maxiscoping *} text {* \medskip Miniscoping: pushing in quantifiers and big Unions and Intersections. *} diff -r 2848e36e0348 -r ae926869a311 src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Mon Feb 20 21:04:00 2012 +0100 +++ b/src/HOL/Lattices.thy Tue Feb 21 08:15:42 2012 +0100 @@ -701,63 +701,6 @@ instance "fun" :: (type, boolean_algebra) boolean_algebra proof qed (rule ext, simp_all add: inf_apply sup_apply bot_apply top_apply uminus_apply minus_apply inf_compl_bot sup_compl_top diff_eq)+ - -subsection {* Unary and binary predicates as lattice *} - -lemma inf1I [intro!]: "A x \ B x \ (A \ B) x" - by (simp add: inf_fun_def) - -lemma inf2I [intro!]: "A x y \ B x y \ (A \ B) x y" - by (simp add: inf_fun_def) - -lemma inf1E [elim!]: "(A \ B) x \ (A x \ B x \ P) \ P" - by (simp add: inf_fun_def) - -lemma inf2E [elim!]: "(A \ B) x y \ (A x y \ B x y \ P) \ P" - by (simp add: inf_fun_def) - -lemma inf1D1: "(A \ B) x \ A x" - by (simp add: inf_fun_def) - -lemma inf2D1: "(A \ B) x y \ A x y" - by (simp add: inf_fun_def) - -lemma inf1D2: "(A \ B) x \ B x" - by (simp add: inf_fun_def) - -lemma inf2D2: "(A \ B) x y \ B x y" - by (simp add: inf_fun_def) - -lemma sup1I1 [intro?]: "A x \ (A \ B) x" - by (simp add: sup_fun_def) - -lemma sup2I1 [intro?]: "A x y \ (A \ B) x y" - by (simp add: sup_fun_def) - -lemma sup1I2 [intro?]: "B x \ (A \ B) x" - by (simp add: sup_fun_def) - -lemma sup2I2 [intro?]: "B x y \ (A \ B) x y" - by (simp add: sup_fun_def) - -lemma sup1E [elim!]: "(A \ B) x \ (A x \ P) \ (B x \ P) \ P" - by (simp add: sup_fun_def) iprover - -lemma sup2E [elim!]: "(A \ B) x y \ (A x y \ P) \ (B x y \ P) \ P" - by (simp add: sup_fun_def) iprover - -text {* - \medskip Classical introduction rule: no commitment to @{text A} vs - @{text B}. -*} - -lemma sup1CI [intro!]: "(\ B x \ A x) \ (A \ B) x" - by (auto simp add: sup_fun_def) - -lemma sup2CI [intro!]: "(\ B x y \ A x y) \ (A \ B) x y" - by (auto simp add: sup_fun_def) - - no_notation less_eq (infix "\" 50) and less (infix "\" 50) and @@ -767,4 +710,3 @@ bot ("\") end - diff -r 2848e36e0348 -r ae926869a311 src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Mon Feb 20 21:04:00 2012 +0100 +++ b/src/HOL/Orderings.thy Tue Feb 21 08:15:42 2012 +0100 @@ -1111,6 +1111,10 @@ end +no_notation + bot ("\") and + top ("\") + subsection {* Dense orders *} @@ -1235,10 +1239,10 @@ [simp]: "(P\bool) < Q \ \ P \ Q" definition - [simp]: "\ \ False" + [simp]: "bot \ False" definition - [simp]: "\ \ True" + [simp]: "top \ True" instance proof qed auto @@ -1257,10 +1261,10 @@ lemma le_boolD: "P \ Q \ P \ Q" by simp -lemma bot_boolE: "\ \ P" +lemma bot_boolE: "bot \ P" by simp -lemma top_boolI: \ +lemma top_boolI: top by simp lemma [code]: @@ -1297,10 +1301,10 @@ begin definition - "\ = (\x. \)" + "bot = (\x. bot)" lemma bot_apply: - "\ x = \" + "bot x = bot" by (simp add: bot_fun_def) instance proof @@ -1312,11 +1316,11 @@ begin definition - [no_atp]: "\ = (\x. \)" + [no_atp]: "top = (\x. top)" declare top_fun_def_raw [no_atp] lemma top_apply: - "\ x = \" + "top x = top" by (simp add: top_fun_def) instance proof @@ -1334,61 +1338,6 @@ unfolding le_fun_def by simp -subsection {* Order on unary and binary predicates *} - -lemma predicate1I: - assumes PQ: "\x. P x \ Q x" - shows "P \ Q" - apply (rule le_funI) - apply (rule le_boolI) - apply (rule PQ) - apply assumption - done - -lemma predicate1D [Pure.dest?, dest?]: - "P \ Q \ P x \ Q x" - apply (erule le_funE) - apply (erule le_boolE) - apply assumption+ - done - -lemma rev_predicate1D: - "P x \ P \ Q \ Q x" - by (rule predicate1D) - -lemma predicate2I [Pure.intro!, intro!]: - assumes PQ: "\x y. P x y \ Q x y" - shows "P \ Q" - apply (rule le_funI)+ - apply (rule le_boolI) - apply (rule PQ) - apply assumption - done - -lemma predicate2D [Pure.dest, dest]: - "P \ Q \ P x y \ Q x y" - apply (erule le_funE)+ - apply (erule le_boolE) - apply assumption+ - done - -lemma rev_predicate2D: - "P x y \ P \ Q \ Q x y" - by (rule predicate2D) - -lemma bot1E [no_atp, elim!]: "\ x \ P" - by (simp add: bot_fun_def) - -lemma bot2E [elim!]: "\ x y \ P" - by (simp add: bot_fun_def) - -lemma top1I [intro!]: "\ x" - by (simp add: top_fun_def) - -lemma top2I [intro!]: "\ x y" - by (simp add: top_fun_def) - - subsection {* Name duplicates *} lemmas order_eq_refl = preorder_class.eq_refl @@ -1426,8 +1375,4 @@ lemmas linorder_antisym_conv2 = linorder_class.antisym_conv2 lemmas linorder_antisym_conv3 = linorder_class.antisym_conv3 -no_notation - top ("\") and - bot ("\") - end diff -r 2848e36e0348 -r ae926869a311 src/HOL/Predicate.thy --- a/src/HOL/Predicate.thy Mon Feb 20 21:04:00 2012 +0100 +++ b/src/HOL/Predicate.thy Tue Feb 21 08:15:42 2012 +0100 @@ -25,6 +25,52 @@ subsection {* Predicates as (complete) lattices *} +text {* + Handy introduction and elimination rules for @{text "\"} + on unary and binary predicates +*} + +lemma predicate1I: + assumes PQ: "\x. P x \ Q x" + shows "P \ Q" + apply (rule le_funI) + apply (rule le_boolI) + apply (rule PQ) + apply assumption + done + +lemma predicate1D [Pure.dest?, dest?]: + "P \ Q \ P x \ Q x" + apply (erule le_funE) + apply (erule le_boolE) + apply assumption+ + done + +lemma rev_predicate1D: + "P x \ P \ Q \ Q x" + by (rule predicate1D) + +lemma predicate2I [Pure.intro!, intro!]: + assumes PQ: "\x y. P x y \ Q x y" + shows "P \ Q" + apply (rule le_funI)+ + apply (rule le_boolI) + apply (rule PQ) + apply assumption + done + +lemma predicate2D [Pure.dest, dest]: + "P \ Q \ P x y \ Q x y" + apply (erule le_funE)+ + apply (erule le_boolE) + apply assumption+ + done + +lemma rev_predicate2D: + "P x y \ P \ Q \ Q x y" + by (rule predicate2D) + + subsubsection {* Equality *} lemma pred_equals_eq [pred_set_conv]: "((\x. x \ R) = (\x. x \ S)) \ (R = S)" @@ -45,15 +91,51 @@ subsubsection {* Top and bottom elements *} +lemma bot1E [no_atp, elim!]: "\ x \ P" + by (simp add: bot_fun_def) + +lemma bot2E [elim!]: "\ x y \ P" + by (simp add: bot_fun_def) + lemma bot_empty_eq: "\ = (\x. x \ {})" by (auto simp add: fun_eq_iff) lemma bot_empty_eq2: "\ = (\x y. (x, y) \ {})" by (auto simp add: fun_eq_iff) +lemma top1I [intro!]: "\ x" + by (simp add: top_fun_def) + +lemma top2I [intro!]: "\ x y" + by (simp add: top_fun_def) + subsubsection {* Binary intersection *} +lemma inf1I [intro!]: "A x \ B x \ (A \ B) x" + by (simp add: inf_fun_def) + +lemma inf2I [intro!]: "A x y \ B x y \ (A \ B) x y" + by (simp add: inf_fun_def) + +lemma inf1E [elim!]: "(A \ B) x \ (A x \ B x \ P) \ P" + by (simp add: inf_fun_def) + +lemma inf2E [elim!]: "(A \ B) x y \ (A x y \ B x y \ P) \ P" + by (simp add: inf_fun_def) + +lemma inf1D1: "(A \ B) x \ A x" + by (simp add: inf_fun_def) + +lemma inf2D1: "(A \ B) x y \ A x y" + by (simp add: inf_fun_def) + +lemma inf1D2: "(A \ B) x \ B x" + by (simp add: inf_fun_def) + +lemma inf2D2: "(A \ B) x y \ B x y" + by (simp add: inf_fun_def) + lemma inf_Int_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) = (\x. x \ R \ S)" by (simp add: inf_fun_def) @@ -63,6 +145,35 @@ subsubsection {* Binary union *} +lemma sup1I1 [intro?]: "A x \ (A \ B) x" + by (simp add: sup_fun_def) + +lemma sup2I1 [intro?]: "A x y \ (A \ B) x y" + by (simp add: sup_fun_def) + +lemma sup1I2 [intro?]: "B x \ (A \ B) x" + by (simp add: sup_fun_def) + +lemma sup2I2 [intro?]: "B x y \ (A \ B) x y" + by (simp add: sup_fun_def) + +lemma sup1E [elim!]: "(A \ B) x \ (A x \ P) \ (B x \ P) \ P" + by (simp add: sup_fun_def) iprover + +lemma sup2E [elim!]: "(A \ B) x y \ (A x y \ P) \ (B x y \ P) \ P" + by (simp add: sup_fun_def) iprover + +text {* + \medskip Classical introduction rule: no commitment to @{text A} vs + @{text B}. +*} + +lemma sup1CI [intro!]: "(\ B x \ A x) \ (A \ B) x" + by (auto simp add: sup_fun_def) + +lemma sup2CI [intro!]: "(\ B x y \ A x y) \ (A \ B) x y" + by (auto simp add: sup_fun_def) + lemma sup_Un_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) = (\x. x \ R \ S)" by (simp add: sup_fun_def) @@ -72,15 +183,57 @@ subsubsection {* Intersections of families *} -lemma INF_INT_eq: "(\i. (\x. x \ r i)) = (\x. x \ (\i. r i))" +lemma INF1_iff: "(\x\A. B x) b = (\x\A. B x b)" + by (simp add: INF_apply) + +lemma INF2_iff: "(\x\A. B x) b c = (\x\A. B x b c)" + by (simp add: INF_apply) + +lemma INF1_I [intro!]: "(\x. x \ A \ B x b) \ (\x\A. B x) b" + by (auto simp add: INF_apply) + +lemma INF2_I [intro!]: "(\x. x \ A \ B x b c) \ (\x\A. B x) b c" + by (auto simp add: INF_apply) + +lemma INF1_D [elim]: "(\x\A. B x) b \ a \ A \ B a b" + by (auto simp add: INF_apply) + +lemma INF2_D [elim]: "(\x\A. B x) b c \ a \ A \ B a b c" + by (auto simp add: INF_apply) + +lemma INF1_E [elim]: "(\x\A. B x) b \ (B a b \ R) \ (a \ A \ R) \ R" + by (auto simp add: INF_apply) + +lemma INF2_E [elim]: "(\x\A. B x) b c \ (B a b c \ R) \ (a \ A \ R) \ R" + by (auto simp add: INF_apply) + +lemma INF_INT_eq: "(\i. (\x. x \ r i)) = (\x. x \ (\i. r i))" by (simp add: INF_apply fun_eq_iff) -lemma INF_INT_eq2: "(\i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i. r i))" +lemma INF_INT_eq2: "(\i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i. r i))" by (simp add: INF_apply fun_eq_iff) subsubsection {* Unions of families *} +lemma SUP1_iff: "(\x\A. B x) b = (\x\A. B x b)" + by (simp add: SUP_apply) + +lemma SUP2_iff: "(\x\A. B x) b c = (\x\A. B x b c)" + by (simp add: SUP_apply) + +lemma SUP1_I [intro]: "a \ A \ B a b \ (\x\A. B x) b" + by (auto simp add: SUP_apply) + +lemma SUP2_I [intro]: "a \ A \ B a b c \ (\x\A. B x) b c" + by (auto simp add: SUP_apply) + +lemma SUP1_E [elim!]: "(\x\A. B x) b \ (\x. x \ A \ B x b \ R) \ R" + by (auto simp add: SUP_apply) + +lemma SUP2_E [elim!]: "(\x\A. B x) b c \ (\x. x \ A \ B x b c \ R) \ R" + by (auto simp add: SUP_apply) + lemma SUP_UN_eq [pred_set_conv]: "(\i. (\x. x \ r i)) = (\x. x \ (\i. r i))" by (simp add: SUP_apply fun_eq_iff) diff -r 2848e36e0348 -r ae926869a311 src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Mon Feb 20 21:04:00 2012 +0100 +++ b/src/HOL/Product_Type.thy Tue Feb 21 08:15:42 2012 +0100 @@ -22,7 +22,7 @@ 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 False \ \ P" and [code]: "HOL.equal P True \ P" and [code nbe]: "HOL.equal P P \ True" by (simp_all add: equal) diff -r 2848e36e0348 -r ae926869a311 src/HOL/ZF/Games.thy --- a/src/HOL/ZF/Games.thy Mon Feb 20 21:04:00 2012 +0100 +++ b/src/HOL/ZF/Games.thy Tue Feb 21 08:15:42 2012 +0100 @@ -19,7 +19,11 @@ "games_gfp \ gfp fixgames" lemma mono_fixgames: "mono (fixgames)" - by (auto simp add: mono_def fixgames_def) + apply (auto simp add: mono_def fixgames_def) + apply (rule_tac x=l in exI) + apply (rule_tac x=r in exI) + apply auto + done lemma games_lfp_unfold: "games_lfp = fixgames games_lfp" by (auto simp add: def_lfp_unfold games_lfp_def mono_fixgames) @@ -970,4 +974,3 @@ qed end -