# HG changeset patch # User krauss # Date 1334264821 -7200 # Node ID 69e96e5500df55b406fa8f424a41f8d262212a9b # Parent d21c95af2df776e6111f9fb864a8f9047b4cec0b Set_Algebras: removed syntax \ and \, in favour of plain + and * diff -r d21c95af2df7 -r 69e96e5500df src/HOL/Hahn_Banach/Hahn_Banach.thy --- a/src/HOL/Hahn_Banach/Hahn_Banach.thy Thu Apr 12 22:55:11 2012 +0200 +++ b/src/HOL/Hahn_Banach/Hahn_Banach.thy Thu Apr 12 23:07:01 2012 +0200 @@ -151,12 +151,12 @@ qed qed - def H' \ "H \ lin x'" + def H' \ "H + lin x'" -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *} have HH': "H \ H'" proof (unfold H'_def) from x'E have "vectorspace (lin x')" .. - with H show "H \ H \ lin x'" .. + with H show "H \ H + lin x'" .. qed obtain xi where diff -r d21c95af2df7 -r 69e96e5500df src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy --- a/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy Thu Apr 12 22:55:11 2012 +0200 +++ b/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy Thu Apr 12 23:07:01 2012 +0200 @@ -90,7 +90,7 @@ lemma h'_lf: assumes h'_def: "h' \ \x. let (y, a) = SOME (y, a). x = y + a \ x0 \ y \ H in h y + a * xi" - and H'_def: "H' \ H \ lin x0" + and H'_def: "H' \ H + lin x0" and HE: "H \ E" assumes "linearform H h" assumes x0: "x0 \ H" "x0 \ E" "x0 \ 0" @@ -106,7 +106,7 @@ proof (unfold H'_def) from `x0 \ E` have "lin x0 \ E" .. - with HE show "vectorspace (H \ lin x0)" using E .. + with HE show "vectorspace (H + lin x0)" using E .. qed { fix x1 x2 assume x1: "x1 \ H'" and x2: "x2 \ H'" @@ -194,7 +194,7 @@ lemma h'_norm_pres: assumes h'_def: "h' \ \x. let (y, a) = SOME (y, a). x = y + a \ x0 \ y \ H in h y + a * xi" - and H'_def: "H' \ H \ lin x0" + and H'_def: "H' \ H + lin x0" and x0: "x0 \ H" "x0 \ E" "x0 \ 0" assumes E: "vectorspace E" and HE: "subspace H E" and "seminorm E p" and "linearform H h" diff -r d21c95af2df7 -r 69e96e5500df src/HOL/Hahn_Banach/Subspace.thy --- a/src/HOL/Hahn_Banach/Subspace.thy Thu Apr 12 22:55:11 2012 +0200 +++ b/src/HOL/Hahn_Banach/Subspace.thy Thu Apr 12 23:07:01 2012 +0200 @@ -228,38 +228,38 @@ set of all sums of elements from @{text U} and @{text V}. *} -lemma sum_def: "U \ V = {u + v | u v. u \ U \ v \ V}" +lemma sum_def: "U + V = {u + v | u v. u \ U \ v \ V}" unfolding set_plus_def by auto lemma sumE [elim]: - "x \ U \ V \ (\u v. x = u + v \ u \ U \ v \ V \ C) \ C" + "x \ U + V \ (\u v. x = u + v \ u \ U \ v \ V \ C) \ C" unfolding sum_def by blast lemma sumI [intro]: - "u \ U \ v \ V \ x = u + v \ x \ U \ V" + "u \ U \ v \ V \ x = u + v \ x \ U + V" unfolding sum_def by blast lemma sumI' [intro]: - "u \ U \ v \ V \ u + v \ U \ V" + "u \ U \ v \ V \ u + v \ U + V" unfolding sum_def by blast -text {* @{text U} is a subspace of @{text "U \ V"}. *} +text {* @{text U} is a subspace of @{text "U + V"}. *} lemma subspace_sum1 [iff]: assumes "vectorspace U" "vectorspace V" - shows "U \ U \ V" + shows "U \ U + V" proof - interpret vectorspace U by fact interpret vectorspace V by fact show ?thesis proof show "U \ {}" .. - show "U \ U \ V" + show "U \ U + V" proof fix x assume x: "x \ U" moreover have "0 \ V" .. - ultimately have "x + 0 \ U \ V" .. - with x show "x \ U \ V" by simp + ultimately have "x + 0 \ U + V" .. + with x show "x \ U + V" by simp qed fix x y assume x: "x \ U" and "y \ U" then show "x + y \ U" by simp @@ -271,30 +271,30 @@ lemma sum_subspace [intro?]: assumes "subspace U E" "vectorspace E" "subspace V E" - shows "U \ V \ E" + shows "U + V \ E" proof - interpret subspace U E by fact interpret vectorspace E by fact interpret subspace V E by fact show ?thesis proof - have "0 \ U \ V" + have "0 \ U + V" proof show "0 \ U" using `vectorspace E` .. show "0 \ V" using `vectorspace E` .. show "(0::'a) = 0 + 0" by simp qed - then show "U \ V \ {}" by blast - show "U \ V \ E" + then show "U + V \ {}" by blast + show "U + V \ E" proof - fix x assume "x \ U \ V" + fix x assume "x \ U + V" then obtain u v where "x = u + v" and "u \ U" and "v \ V" .. then show "x \ E" by simp qed next - fix x y assume x: "x \ U \ V" and y: "y \ U \ V" - show "x + y \ U \ V" + fix x y assume x: "x \ U + V" and y: "y \ U + V" + show "x + y \ U + V" proof - from x obtain ux vx where "x = ux + vx" and "ux \ U" and "vx \ V" .. moreover @@ -306,7 +306,7 @@ using x y by (simp_all add: add_ac) then show ?thesis .. qed - fix a show "a \ x \ U \ V" + fix a show "a \ x \ U + V" proof - from x obtain u v where "x = u + v" and "u \ U" and "v \ V" .. then have "a \ u \ U" and "a \ v \ V" @@ -319,7 +319,7 @@ text{* The sum of two subspaces is a vectorspace. *} lemma sum_vs [intro?]: - "U \ E \ V \ E \ vectorspace E \ vectorspace (U \ V)" + "U \ E \ V \ E \ vectorspace E \ vectorspace (U + V)" by (rule subspace.vectorspace) (rule sum_subspace) @@ -481,7 +481,7 @@ proof - interpret vectorspace E by fact interpret subspace H E by fact - from x y x' have "x \ H \ lin x'" by auto + from x y x' have "x \ H + lin x'" by auto have "\!p. (\(y, a). x = y + a \ x' \ y \ H) p" (is "\!p. ?P p") proof (rule ex_ex1I) from x y show "\p. ?P p" by blast diff -r d21c95af2df7 -r 69e96e5500df src/HOL/Library/BigO.thy --- a/src/HOL/Library/BigO.thy Thu Apr 12 22:55:11 2012 +0200 +++ b/src/HOL/Library/BigO.thy Thu Apr 12 23:07:01 2012 +0200 @@ -92,7 +92,7 @@ by (auto simp add: bigo_def) lemma bigo_plus_self_subset [intro]: - "O(f) \ O(f) <= O(f)" + "O(f) + O(f) <= O(f)" apply (auto simp add: bigo_alt_def set_plus_def) apply (rule_tac x = "c + ca" in exI) apply auto @@ -104,14 +104,14 @@ apply force done -lemma bigo_plus_idemp [simp]: "O(f) \ O(f) = O(f)" +lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)" apply (rule equalityI) apply (rule bigo_plus_self_subset) apply (rule set_zero_plus2) apply (rule bigo_zero) done -lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \ O(g)" +lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) + O(g)" apply (rule subsetI) apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def) apply (subst bigo_pos_const [symmetric])+ @@ -153,15 +153,15 @@ apply simp done -lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A \ B <= O(f)" - apply (subgoal_tac "A \ B <= O(f) \ O(f)") +lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A + B <= O(f)" + apply (subgoal_tac "A + B <= O(f) + O(f)") apply (erule order_trans) apply simp apply (auto del: subsetI simp del: bigo_plus_idemp) done lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==> - O(f + g) = O(f) \ O(g)" + O(f + g) = O(f) + O(g)" apply (rule equalityI) apply (rule bigo_plus_subset) apply (simp add: bigo_alt_def set_plus_def func_plus) @@ -273,12 +273,12 @@ lemma bigo_abs5: "f =o O(g) ==> (%x. abs(f x)) =o O(g)" by (unfold bigo_def, auto) -lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) \ O(h)" +lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) + O(h)" proof - assume "f : g +o O(h)" - also have "... <= O(g) \ O(h)" + also have "... <= O(g) + O(h)" by (auto del: subsetI) - also have "... = O(%x. abs(g x)) \ O(%x. abs(h x))" + also have "... = O(%x. abs(g x)) + O(%x. abs(h x))" apply (subst bigo_abs3 [symmetric])+ apply (rule refl) done @@ -287,13 +287,13 @@ finally have "f : ...". then have "O(f) <= ..." by (elim bigo_elt_subset) - also have "... = O(%x. abs(g x)) \ O(%x. abs(h x))" + also have "... = O(%x. abs(g x)) + O(%x. abs(h x))" by (rule bigo_plus_eq, auto) finally show ?thesis by (simp add: bigo_abs3 [symmetric]) qed -lemma bigo_mult [intro]: "O(f)\O(g) <= O(f * g)" +lemma bigo_mult [intro]: "O(f)*O(g) <= O(f * g)" apply (rule subsetI) apply (subst bigo_def) apply (auto simp add: bigo_alt_def set_times_def func_times) @@ -369,7 +369,7 @@ done lemma bigo_mult7: "ALL x. f x ~= 0 ==> - O(f * g) <= O(f::'a => ('b::linordered_field)) \ O(g)" + O(f * g) <= O(f::'a => ('b::linordered_field)) * O(g)" apply (subst bigo_mult6) apply assumption apply (rule set_times_mono3) @@ -377,7 +377,7 @@ done lemma bigo_mult8: "ALL x. f x ~= 0 ==> - O(f * g) = O(f::'a => ('b::linordered_field)) \ O(g)" + O(f * g) = O(f::'a => ('b::linordered_field)) * O(g)" apply (rule equalityI) apply (erule bigo_mult7) apply (rule bigo_mult) @@ -402,9 +402,9 @@ show "f +o O(g) <= O(g)" proof - have "f : O(f)" by auto - then have "f +o O(g) <= O(f) \ O(g)" + then have "f +o O(g) <= O(f) + O(g)" by (auto del: subsetI) - also have "... <= O(g) \ O(g)" + also have "... <= O(g) + O(g)" proof - from a have "O(f) <= O(g)" by (auto del: subsetI) thus ?thesis by (auto del: subsetI) @@ -656,7 +656,7 @@ subsection {* Misc useful stuff *} lemma bigo_useful_intro: "A <= O(f) ==> B <= O(f) ==> - A \ B <= O(f)" + A + B <= O(f)" apply (subst bigo_plus_idemp [symmetric]) apply (rule set_plus_mono2) apply assumption+ diff -r d21c95af2df7 -r 69e96e5500df src/HOL/Library/Set_Algebras.thy --- a/src/HOL/Library/Set_Algebras.thy Thu Apr 12 22:55:11 2012 +0200 +++ b/src/HOL/Library/Set_Algebras.thy Thu Apr 12 23:07:01 2012 +0200 @@ -34,14 +34,6 @@ end - -text {* Legacy syntax: *} - -abbreviation (input) set_plus :: "'a::plus set \ 'a set \ 'a set" (infixl "\" 65) where - "A \ B \ A + B" -abbreviation (input) set_times :: "'a::times set \ 'a set \ 'a set" (infixl "\" 70) where - "A \ B \ A * B" - instantiation set :: (zero) zero begin @@ -95,14 +87,14 @@ instance set :: (comm_monoid_mult) comm_monoid_mult by default (simp_all add: set_times_def) -lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C \ D" +lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C + D" by (auto simp add: set_plus_def) lemma set_plus_intro2 [intro]: "b : C ==> a + b : a +o C" by (auto simp add: elt_set_plus_def) -lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) \ - (b +o D) = (a + b) +o (C \ D)" +lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) + + (b +o D) = (a + b) +o (C + D)" apply (auto simp add: elt_set_plus_def set_plus_def add_ac) apply (rule_tac x = "ba + bb" in exI) apply (auto simp add: add_ac) @@ -114,8 +106,8 @@ (a + b) +o C" by (auto simp add: elt_set_plus_def add_assoc) -lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) \ C = - a +o (B \ C)" +lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) + C = + a +o (B + C)" apply (auto simp add: elt_set_plus_def set_plus_def) apply (blast intro: add_ac) apply (rule_tac x = "a + aa" in exI) @@ -126,8 +118,8 @@ apply (auto simp add: add_ac) done -theorem set_plus_rearrange4: "C \ ((a::'a::comm_monoid_add) +o D) = - a +o (C \ D)" +theorem set_plus_rearrange4: "C + ((a::'a::comm_monoid_add) +o D) = + a +o (C + D)" apply (auto simp add: elt_set_plus_def set_plus_def add_ac) apply (rule_tac x = "aa + ba" in exI) apply (auto simp add: add_ac) @@ -140,17 +132,17 @@ by (auto simp add: elt_set_plus_def) lemma set_plus_mono2 [intro]: "(C::('a::plus) set) <= D ==> E <= F ==> - C \ E <= D \ F" + C + E <= D + F" by (auto simp add: set_plus_def) -lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C \ D" +lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C + D" by (auto simp add: elt_set_plus_def set_plus_def) lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) : C ==> - a +o D <= D \ C" + a +o D <= D + C" by (auto simp add: elt_set_plus_def set_plus_def add_ac) -lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C \ D" +lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C + D" apply (subgoal_tac "a +o B <= a +o D") apply (erule order_trans) apply (erule set_plus_mono3) @@ -163,21 +155,21 @@ apply auto done -lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C \ E ==> - x : D \ F" +lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C + E ==> + x : D + F" apply (frule set_plus_mono2) prefer 2 apply force apply assumption done -lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C \ D" +lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C + D" apply (frule set_plus_mono3) apply auto done lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C ==> - x : a +o D ==> x : D \ C" + x : a +o D ==> x : D + C" apply (frule set_plus_mono4) apply auto done @@ -185,7 +177,7 @@ lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C" by (auto simp add: elt_set_plus_def) -lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A \ B" +lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A + B" apply (auto simp add: set_plus_def) apply (rule_tac x = 0 in bexI) apply (rule_tac x = x in bexI) @@ -206,14 +198,14 @@ by (rule iffI, rule set_minus_imp_plus, assumption, rule set_plus_imp_minus, assumption) -lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C \ D" +lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C * D" by (auto simp add: set_times_def) lemma set_times_intro2 [intro!]: "b : C ==> a * b : a *o C" by (auto simp add: elt_set_times_def) -lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) \ - (b *o D) = (a * b) *o (C \ D)" +lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) * + (b *o D) = (a * b) *o (C * D)" apply (auto simp add: elt_set_times_def set_times_def) apply (rule_tac x = "ba * bb" in exI) apply (auto simp add: mult_ac) @@ -225,8 +217,8 @@ (a * b) *o C" by (auto simp add: elt_set_times_def mult_assoc) -lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) \ C = - a *o (B \ C)" +lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) * C = + a *o (B * C)" apply (auto simp add: elt_set_times_def set_times_def) apply (blast intro: mult_ac) apply (rule_tac x = "a * aa" in exI) @@ -237,8 +229,8 @@ apply (auto simp add: mult_ac) done -theorem set_times_rearrange4: "C \ ((a::'a::comm_monoid_mult) *o D) = - a *o (C \ D)" +theorem set_times_rearrange4: "C * ((a::'a::comm_monoid_mult) *o D) = + a *o (C * D)" apply (auto simp add: elt_set_times_def set_times_def mult_ac) apply (rule_tac x = "aa * ba" in exI) @@ -252,17 +244,17 @@ by (auto simp add: elt_set_times_def) lemma set_times_mono2 [intro]: "(C::('a::times) set) <= D ==> E <= F ==> - C \ E <= D \ F" + C * E <= D * F" by (auto simp add: set_times_def) -lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C \ D" +lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C * D" by (auto simp add: elt_set_times_def set_times_def) lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C ==> - a *o D <= D \ C" + a *o D <= D * C" by (auto simp add: elt_set_times_def set_times_def mult_ac) -lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C \ D" +lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C * D" apply (subgoal_tac "a *o B <= a *o D") apply (erule order_trans) apply (erule set_times_mono3) @@ -275,21 +267,21 @@ apply auto done -lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C \ E ==> - x : D \ F" +lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C * E ==> + x : D * F" apply (frule set_times_mono2) prefer 2 apply force apply assumption done -lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C \ D" +lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C * D" apply (frule set_times_mono3) apply auto done lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) : C ==> - x : a *o D ==> x : D \ C" + x : a *o D ==> x : D * C" apply (frule set_times_mono4) apply auto done @@ -301,16 +293,16 @@ (a * b) +o (a *o C)" by (auto simp add: elt_set_plus_def elt_set_times_def ring_distribs) -lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B \ C) = - (a *o B) \ (a *o C)" +lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B + C) = + (a *o B) + (a *o C)" apply (auto simp add: set_plus_def elt_set_times_def ring_distribs) apply blast apply (rule_tac x = "b + bb" in exI) apply (auto simp add: ring_distribs) done -lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) \ D <= - a *o D \ C \ D" +lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) * D <= + a *o D + C * D" apply (auto simp add: elt_set_plus_def elt_set_times_def set_times_def set_plus_def ring_distribs) @@ -330,7 +322,7 @@ by (auto simp add: elt_set_times_def) lemma set_plus_image: - fixes S T :: "'n::semigroup_add set" shows "S \ T = (\(x, y). x + y) ` (S \ T)" + fixes S T :: "'n::semigroup_add set" shows "S + T = (\(x, y). x + y) ` (S \ T)" unfolding set_plus_def by (fastforce simp: image_iff) lemma set_setsum_alt: @@ -339,7 +331,7 @@ (is "_ = ?setsum I") using fin proof induct case (insert x F) - have "setsum S (insert x F) = S x \ ?setsum F" + have "setsum S (insert x F) = S x + ?setsum F" using insert.hyps by auto also have "...= {s x + setsum s F |s. \ i\insert x F. s i \ S i}" unfolding set_plus_def @@ -355,8 +347,8 @@ lemma setsum_set_cond_linear: fixes f :: "('a::comm_monoid_add) set \ ('b::comm_monoid_add) set" - assumes [intro!]: "\A B. P A \ P B \ P (A \ B)" "P {0}" - and f: "\A B. P A \ P B \ f (A \ B) = f A \ f B" "f {0} = {0}" + assumes [intro!]: "\A B. P A \ P B \ P (A + B)" "P {0}" + and f: "\A B. P A \ P B \ f (A + B) = f A + f B" "f {0} = {0}" assumes all: "\i. i \ I \ P (S i)" shows "f (setsum S I) = setsum (f \ S) I" proof cases @@ -372,7 +364,7 @@ lemma setsum_set_linear: fixes f :: "('a::comm_monoid_add) set => ('b::comm_monoid_add) set" - assumes "\A B. f(A) \ f(B) = f(A \ B)" "f {0} = {0}" + assumes "\A B. f(A) + f(B) = f(A + B)" "f {0} = {0}" shows "f (setsum S I) = setsum (f \ S) I" using setsum_set_cond_linear[of "\x. True" f I S] assms by auto diff -r d21c95af2df7 -r 69e96e5500df src/HOL/Metis_Examples/Big_O.thy --- a/src/HOL/Metis_Examples/Big_O.thy Thu Apr 12 22:55:11 2012 +0200 +++ b/src/HOL/Metis_Examples/Big_O.thy Thu Apr 12 23:07:01 2012 +0200 @@ -146,17 +146,17 @@ by (auto simp add: bigo_def) lemma bigo_plus_self_subset [intro]: - "O(f) \ O(f) <= O(f)" + "O(f) + O(f) <= O(f)" apply (auto simp add: bigo_alt_def set_plus_def) apply (rule_tac x = "c + ca" in exI) apply auto apply (simp add: ring_distribs func_plus) by (metis order_trans abs_triangle_ineq add_mono) -lemma bigo_plus_idemp [simp]: "O(f) \ O(f) = O(f)" +lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)" by (metis bigo_plus_self_subset bigo_zero set_eq_subset set_zero_plus2) -lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \ O(g)" +lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) + O(g)" apply (rule subsetI) apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def) apply (subst bigo_pos_const [symmetric])+ @@ -187,10 +187,10 @@ apply (metis abs_triangle_ineq mult_le_cancel_left_pos) by (metis abs_ge_zero abs_of_pos zero_le_mult_iff) -lemma bigo_plus_subset2 [intro]: "A <= O(f) \ B <= O(f) \ A \ B <= O(f)" +lemma bigo_plus_subset2 [intro]: "A <= O(f) \ B <= O(f) \ A + B <= O(f)" by (metis bigo_plus_idemp set_plus_mono2) -lemma bigo_plus_eq: "\x. 0 <= f x \ \x. 0 <= g x \ O(f + g) = O(f) \ O(g)" +lemma bigo_plus_eq: "\x. 0 <= f x \ \x. 0 <= g x \ O(f + g) = O(f) + O(g)" apply (rule equalityI) apply (rule bigo_plus_subset) apply (simp add: bigo_alt_def set_plus_def func_plus) @@ -284,25 +284,25 @@ lemma bigo_abs5: "f =o O(g) \ (\x. abs(f x)) =o O(g)" by (unfold bigo_def, auto) -lemma bigo_elt_subset2 [intro]: "f : g +o O(h) \ O(f) <= O(g) \ O(h)" +lemma bigo_elt_subset2 [intro]: "f : g +o O(h) \ O(f) <= O(g) + O(h)" proof - assume "f : g +o O(h)" - also have "... <= O(g) \ O(h)" + also have "... <= O(g) + O(h)" by (auto del: subsetI) - also have "... = O(\x. abs(g x)) \ O(\x. abs(h x))" + also have "... = O(\x. abs(g x)) + O(\x. abs(h x))" by (metis bigo_abs3) also have "... = O((\x. abs(g x)) + (\x. abs(h x)))" by (rule bigo_plus_eq [symmetric], auto) finally have "f : ...". then have "O(f) <= ..." by (elim bigo_elt_subset) - also have "... = O(\x. abs(g x)) \ O(\x. abs(h x))" + also have "... = O(\x. abs(g x)) + O(\x. abs(h x))" by (rule bigo_plus_eq, auto) finally show ?thesis by (simp add: bigo_abs3 [symmetric]) qed -lemma bigo_mult [intro]: "O(f) \ O(g) <= O(f * g)" +lemma bigo_mult [intro]: "O(f) * O(g) <= O(f * g)" apply (rule subsetI) apply (subst bigo_def) apply (auto simp del: abs_mult mult_ac @@ -358,14 +358,14 @@ declare bigo_mult6 [simp] lemma bigo_mult7: -"\x. f x \ 0 \ O(f * g) \ O(f\'a \ ('b\linordered_field)) \ O(g)" +"\x. f x \ 0 \ O(f * g) \ O(f\'a \ ('b\linordered_field)) * O(g)" by (metis bigo_refl bigo_mult6 set_times_mono3) declare bigo_mult6 [simp del] declare bigo_mult7 [intro!] lemma bigo_mult8: -"\x. f x \ 0 \ O(f * g) = O(f\'a \ ('b\linordered_field)) \ O(g)" +"\x. f x \ 0 \ O(f * g) = O(f\'a \ ('b\linordered_field)) * O(g)" by (metis bigo_mult bigo_mult7 order_antisym_conv) lemma bigo_minus [intro]: "f : O(g) \ - f : O(g)" @@ -575,7 +575,7 @@ subsection {* Misc useful stuff *} lemma bigo_useful_intro: "A <= O(f) \ B <= O(f) \ - A \ B <= O(f)" + A + B <= O(f)" apply (subst bigo_plus_idemp [symmetric]) apply (rule set_plus_mono2) apply assumption+ diff -r d21c95af2df7 -r 69e96e5500df src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Thu Apr 12 22:55:11 2012 +0200 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Thu Apr 12 23:07:01 2012 +0200 @@ -5428,13 +5428,13 @@ lemma closure_sum: fixes S T :: "('n::euclidean_space) set" - shows "closure S \ closure T \ closure (S \ T)" + shows "closure S + closure T \ closure (S + T)" proof- - have "(closure S) \ (closure T) = (\(x,y). x + y) ` (closure S \ closure T)" + have "(closure S) + (closure T) = (\(x,y). x + y) ` (closure S \ closure T)" by (simp add: set_plus_image) also have "... = (\(x,y). x + y) ` closure (S \ T)" using closure_direct_sum by auto - also have "... \ closure (S \ T)" + also have "... \ closure (S + T)" using fst_snd_linear closure_linear_image[of "(\(x,y). x + y)" "S \ T"] by (auto simp: set_plus_image) finally show ?thesis @@ -5444,7 +5444,7 @@ lemma convex_oplus: fixes S T :: "('n::euclidean_space) set" assumes "convex S" "convex T" -shows "convex (S \ T)" +shows "convex (S + T)" proof- have "{x + y |x y. x : S & y : T} = {c. EX a:S. EX b:T. c = a + b}" by auto thus ?thesis unfolding set_plus_def using convex_sums[of S T] assms by auto @@ -5452,13 +5452,13 @@ lemma convex_hull_sum: fixes S T :: "('n::euclidean_space) set" -shows "convex hull (S \ T) = (convex hull S) \ (convex hull T)" +shows "convex hull (S + T) = (convex hull S) + (convex hull T)" proof- -have "(convex hull S) \ (convex hull T) = +have "(convex hull S) + (convex hull T) = (%(x,y). x + y) ` ((convex hull S) <*> (convex hull T))" by (simp add: set_plus_image) also have "... = (%(x,y). x + y) ` (convex hull (S <*> T))" using convex_hull_direct_sum by auto -also have "...= convex hull (S \ T)" using fst_snd_linear linear_conv_bounded_linear +also have "...= convex hull (S + T)" using fst_snd_linear linear_conv_bounded_linear convex_hull_linear_image[of "(%(x,y). x + y)" "S <*> T"] by (auto simp add: set_plus_image) finally show ?thesis by auto qed @@ -5466,12 +5466,12 @@ lemma rel_interior_sum: fixes S T :: "('n::euclidean_space) set" assumes "convex S" "convex T" -shows "rel_interior (S \ T) = (rel_interior S) \ (rel_interior T)" +shows "rel_interior (S + T) = (rel_interior S) + (rel_interior T)" proof- -have "(rel_interior S) \ (rel_interior T) = (%(x,y). x + y) ` (rel_interior S <*> rel_interior T)" +have "(rel_interior S) + (rel_interior T) = (%(x,y). x + y) ` (rel_interior S <*> rel_interior T)" by (simp add: set_plus_image) also have "... = (%(x,y). x + y) ` rel_interior (S <*> T)" using rel_interior_direct_sum assms by auto -also have "...= rel_interior (S \ T)" using fst_snd_linear convex_direct_sum assms +also have "...= rel_interior (S + T)" using fst_snd_linear convex_direct_sum assms rel_interior_convex_linear_image[of "(%(x,y). x + y)" "S <*> T"] by (auto simp add: set_plus_image) finally show ?thesis by auto qed @@ -5507,7 +5507,7 @@ lemma convex_rel_open_sum: fixes S T :: "('n::euclidean_space) set" assumes "convex S" "rel_open S" "convex T" "rel_open T" -shows "convex (S \ T) & rel_open (S \ T)" +shows "convex (S + T) & rel_open (S + T)" by (metis assms convex_oplus rel_interior_sum rel_open_def) lemma convex_hull_finite_union_cones: @@ -5547,7 +5547,7 @@ fixes S T :: "('m::euclidean_space) set" assumes "convex S" "cone S" "S ~= {}" assumes "convex T" "cone T" "T ~= {}" -shows "convex hull (S Un T) = S \ T" +shows "convex hull (S Un T) = S + T" proof- def I == "{(1::nat),2}" def A == "(%i. (if i=(1::nat) then S else T))" @@ -5556,7 +5556,7 @@ moreover have "convex hull Union (A ` I) = setsum A I" apply (subst convex_hull_finite_union_cones[of I A]) using assms A_def I_def by auto moreover have - "setsum A I = S \ T" using A_def I_def + "setsum A I = S + T" using A_def I_def unfolding set_plus_def apply auto unfolding set_plus_def by auto ultimately show ?thesis by auto qed