# HG changeset patch # User wenzelm # Date 1331469556 -3600 # Node ID 0883804b67bb9842e4125c1c6d5c1b8f15ddd374 # Parent b190913c3c4174b795c1d8a302b94b35a84ed29e modernized locale expression, with some fluctuation of arguments; diff -r b190913c3c41 -r 0883804b67bb src/HOL/Hahn_Banach/Function_Norm.thy --- a/src/HOL/Hahn_Banach/Function_Norm.thy Sat Mar 10 23:45:47 2012 +0100 +++ b/src/HOL/Hahn_Banach/Function_Norm.thy Sun Mar 11 13:39:16 2012 +0100 @@ -21,7 +21,8 @@ linear forms: *} -locale continuous = var_V + norm_syntax + linearform + +locale continuous = linearform + + fixes norm :: "_ \ real" ("\_\") assumes bounded: "\c. \x \ V. \f x\ \ c * \x\" declare continuous.intro [intro?] continuous_axioms.intro [intro?] @@ -30,11 +31,11 @@ fixes norm :: "_ \ real" ("\_\") assumes "linearform V f" assumes r: "\x. x \ V \ \f x\ \ c * \x\" - shows "continuous V norm f" + shows "continuous V f norm" proof show "linearform V f" by fact from r have "\c. \x\V. \f x\ \ c * \x\" by blast - then show "continuous_axioms V norm f" .. + then show "continuous_axioms V f norm" .. qed @@ -71,7 +72,8 @@ supremum exists (otherwise it is undefined). *} -locale fn_norm = norm_syntax + +locale fn_norm = + fixes norm :: "_ \ real" ("\_\") fixes B defines "B V f \ {0} \ {\f x\ / \x\ | x. x \ 0 \ x \ V}" fixes fn_norm ("\_\\_" [0, 1000] 999) defines "\f\\V \ \(B V f)" @@ -87,10 +89,10 @@ *} lemma (in normed_vectorspace_with_fn_norm) fn_norm_works: - assumes "continuous V norm f" + assumes "continuous V f norm" shows "lub (B V f) (\f\\V)" proof - - interpret continuous V norm f by fact + interpret continuous V f norm by fact txt {* The existence of the supremum is shown using the completeness of the reals. Completeness means, that every non-empty bounded set of reals has a supremum. *} @@ -154,39 +156,39 @@ qed lemma (in normed_vectorspace_with_fn_norm) fn_norm_ub [iff?]: - assumes "continuous V norm f" + assumes "continuous V f norm" assumes b: "b \ B V f" shows "b \ \f\\V" proof - - interpret continuous V norm f by fact + interpret continuous V f norm by fact have "lub (B V f) (\f\\V)" - using `continuous V norm f` by (rule fn_norm_works) + using `continuous V f norm` by (rule fn_norm_works) from this and b show ?thesis .. qed lemma (in normed_vectorspace_with_fn_norm) fn_norm_leastB: - assumes "continuous V norm f" + assumes "continuous V f norm" assumes b: "\b. b \ B V f \ b \ y" shows "\f\\V \ y" proof - - interpret continuous V norm f by fact + interpret continuous V f norm by fact have "lub (B V f) (\f\\V)" - using `continuous V norm f` by (rule fn_norm_works) + using `continuous V f norm` by (rule fn_norm_works) from this and b show ?thesis .. qed text {* The norm of a continuous function is always @{text "\ 0"}. *} lemma (in normed_vectorspace_with_fn_norm) fn_norm_ge_zero [iff]: - assumes "continuous V norm f" + assumes "continuous V f norm" shows "0 \ \f\\V" proof - - interpret continuous V norm f by fact + interpret continuous V f norm by fact txt {* The function norm is defined as the supremum of @{text B}. So it is @{text "\ 0"} if all elements in @{text B} are @{text "\ 0"}, provided the supremum exists and @{text B} is not empty. *} have "lub (B V f) (\f\\V)" - using `continuous V norm f` by (rule fn_norm_works) + using `continuous V f norm` by (rule fn_norm_works) moreover have "0 \ B V f" .. ultimately show ?thesis .. qed @@ -199,11 +201,11 @@ *} lemma (in normed_vectorspace_with_fn_norm) fn_norm_le_cong: - assumes "continuous V norm f" "linearform V f" + assumes "continuous V f norm" "linearform V f" assumes x: "x \ V" shows "\f x\ \ \f\\V * \x\" proof - - interpret continuous V norm f by fact + interpret continuous V f norm by fact interpret linearform V f by fact show ?thesis proof cases @@ -212,7 +214,7 @@ also have "f 0 = 0" by rule unfold_locales also have "\\\ = 0" by simp also have a: "0 \ \f\\V" - using `continuous V norm f` by (rule fn_norm_ge_zero) + using `continuous V f norm` by (rule fn_norm_ge_zero) from x have "0 \ norm x" .. with a have "0 \ \f\\V * \x\" by (simp add: zero_le_mult_iff) finally show "\f x\ \ \f\\V * \x\" . @@ -225,7 +227,7 @@ from x show "0 \ \x\" .. from x and neq have "\f x\ * inverse \x\ \ B V f" by (auto simp add: B_def divide_inverse) - with `continuous V norm f` show "\f x\ * inverse \x\ \ \f\\V" + with `continuous V f norm` show "\f x\ * inverse \x\ \ \f\\V" by (rule fn_norm_ub) qed finally show ?thesis . @@ -241,11 +243,11 @@ *} lemma (in normed_vectorspace_with_fn_norm) fn_norm_least [intro?]: - assumes "continuous V norm f" + assumes "continuous V f norm" assumes ineq: "\x \ V. \f x\ \ c * \x\" and ge: "0 \ c" shows "\f\\V \ c" proof - - interpret continuous V norm f by fact + interpret continuous V f norm by fact show ?thesis proof (rule fn_norm_leastB [folded B_def fn_norm_def]) fix b assume b: "b \ B V f" @@ -272,7 +274,7 @@ qed finally show ?thesis . qed - qed (insert `continuous V norm f`, simp_all add: continuous_def) + qed (insert `continuous V f norm`, simp_all add: continuous_def) qed end diff -r b190913c3c41 -r 0883804b67bb src/HOL/Hahn_Banach/Hahn_Banach.thy --- a/src/HOL/Hahn_Banach/Hahn_Banach.thy Sat Mar 10 23:45:47 2012 +0100 +++ b/src/HOL/Hahn_Banach/Hahn_Banach.thy Sun Mar 11 13:39:16 2012 +0100 @@ -356,9 +356,9 @@ fixes fn_norm ("\_\\_" [0, 1000] 999) defines "\V f. \f\\V \ \(B V f)" assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E" - and linearform: "linearform F f" and "continuous F norm f" + and linearform: "linearform F f" and "continuous F f norm" shows "\g. linearform E g - \ continuous E norm g + \ continuous E g norm \ (\x \ F. g x = f x) \ \g\\E = \f\\F" proof - @@ -367,7 +367,7 @@ by (auto simp: B_def fn_norm_def) intro_locales interpret subspace F E by fact interpret linearform F f by fact - interpret continuous F norm f by fact + interpret continuous F f norm by fact have E: "vectorspace E" by intro_locales have F: "vectorspace F" by rule intro_locales have F_norm: "normed_vectorspace F norm" @@ -375,7 +375,7 @@ have ge_zero: "0 \ \f\\F" by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero [OF normed_vectorspace_with_fn_norm.intro, - OF F_norm `continuous F norm f` , folded B_def fn_norm_def]) + OF F_norm `continuous F f norm` , folded B_def fn_norm_def]) txt {* We define a function @{text p} on @{text E} as follows: @{text "p x = \f\ \ \x\"} *} def p \ "\x. \f\\F * \x\" @@ -422,7 +422,7 @@ have "\x \ F. \f x\ \ p x" proof fix x assume "x \ F" - with `continuous F norm f` and linearform + with `continuous F f norm` and linearform show "\f x\ \ p x" unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong [OF normed_vectorspace_with_fn_norm.intro, @@ -442,7 +442,7 @@ txt {* We furthermore have to show that @{text g} is also continuous: *} - have g_cont: "continuous E norm g" using linearformE + have g_cont: "continuous E g norm" using linearformE proof fix x assume "x \ E" with b show "\g x\ \ \f\\F * \x\" @@ -500,7 +500,7 @@ show "0 \ \g\\E" using g_cont by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def]) - show "continuous F norm f" by fact + show "continuous F f norm" by fact qed qed with linearformE a g_cont show ?thesis by blast diff -r b190913c3c41 -r 0883804b67bb src/HOL/Hahn_Banach/Normed_Space.thy --- a/src/HOL/Hahn_Banach/Normed_Space.thy Sat Mar 10 23:45:47 2012 +0100 +++ b/src/HOL/Hahn_Banach/Normed_Space.thy Sun Mar 11 13:39:16 2012 +0100 @@ -16,11 +16,9 @@ definite, absolute homogenous and subadditive. *} -locale norm_syntax = +locale seminorm = + fixes V :: "'a\{minus, plus, zero, uminus} set" fixes norm :: "'a \ real" ("\_\") - -locale seminorm = var_V + norm_syntax + - constrains V :: "'a\{minus, plus, zero, uminus} set" assumes ge_zero [iff?]: "x \ V \ 0 \ \x\" and abs_homogenous [iff?]: "x \ V \ \a \ x\ = \a\ * \x\" and subadditive [iff?]: "x \ V \ y \ V \ \x + y\ \ \x\ + \y\" diff -r b190913c3c41 -r 0883804b67bb src/HOL/Hahn_Banach/Vector_Space.thy --- a/src/HOL/Hahn_Banach/Vector_Space.thy Sat Mar 10 23:45:47 2012 +0100 +++ b/src/HOL/Hahn_Banach/Vector_Space.thy Sun Mar 11 13:39:16 2012 +0100 @@ -38,9 +38,8 @@ the neutral element of scalar multiplication. *} -locale var_V = fixes V - -locale vectorspace = var_V + +locale vectorspace = + fixes V assumes non_empty [iff, intro?]: "V \ {}" and add_closed [iff]: "x \ V \ y \ V \ x + y \ V" and mult_closed [iff]: "x \ V \ a \ x \ V"