# HG changeset patch # User wenzelm # Date 1354896040 -3600 # Node ID 79858bd9f5ef731148cf5577150fea70c8619796 # Parent eb7b59cc8e081424a3d91e65a61320d45c1079b3# Parent 7c8ce63a3c0021dbf09a588abd9356ac68284dda merged diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/IMP/HoareT.thy --- a/src/HOL/IMP/HoareT.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/IMP/HoareT.thy Fri Dec 07 17:00:40 2012 +0100 @@ -1,7 +1,9 @@ -header{* Hoare Logic for Total Correctness *} +(* Author: Tobias Nipkow *) theory HoareT imports Hoare_Sound_Complete begin +subsection "Hoare Logic for Total Correctness" + text{* Note that this definition of total validity @{text"\\<^sub>t"} only works if execution is deterministic (which it is in our case). *} diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/IMP/VC.thy --- a/src/HOL/IMP/VC.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/IMP/VC.thy Fri Dec 07 17:00:40 2012 +0100 @@ -1,8 +1,8 @@ -header "Verification Conditions" +(* Author: Tobias Nipkow *) theory VC imports Hoare begin -subsection "VCG via Weakest Preconditions" +subsection "Verification Conditions" text{* Annotated commands: commands where loops are annotated with invariants. *} @@ -46,7 +46,7 @@ "strip (Aif b c\<^isub>1 c\<^isub>2) = (IF b THEN strip c\<^isub>1 ELSE strip c\<^isub>2)" | "strip (Awhile I b c) = (WHILE b DO strip c)" -subsection "Soundness" +text {* Soundness: *} lemma vc_sound: "\s. vc c Q s \ \ {pre c Q} strip c {Q}" proof(induction c arbitrary: Q) @@ -68,7 +68,7 @@ by (metis strengthen_pre vc_sound) -subsection "Completeness" +text{* Completeness: *} lemma pre_mono: "\s. P s \ P' s \ pre c P s \ pre c P' s" @@ -122,7 +122,7 @@ qed -subsection "An Optimization" +text{* An Optimization: *} fun vcpre :: "acom \ assn \ assn \ assn" where "vcpre ASKIP Q = (\s. True, Q)" | diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Limits.thy --- a/src/HOL/Limits.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Limits.thy Fri Dec 07 17:00:40 2012 +0100 @@ -732,6 +732,9 @@ "filterlim f F2 F1 \ F2 \ F2' \ F1' \ F1 \ filterlim f F2' F1'" unfolding filterlim_def by (metis filtermap_mono order_trans) +lemma filterlim_ident: "LIM x F. x :> F" + by (simp add: filterlim_def filtermap_ident) + lemma filterlim_cong: "F1 = F1' \ F2 = F2' \ eventually (\x. f x = g x) F2 \ filterlim f F1 F2 = filterlim g F1' F2'" by (auto simp: filterlim_def le_filter_def eventually_filtermap elim: eventually_elim2) @@ -1560,6 +1563,10 @@ by (auto intro: order_trans simp: filterlim_def filtermap_filtermap) qed +lemma tendsto_inverse_0_at_top: + "LIM x F. f x :> at_top \ ((\x. inverse (f x) :: real) ---> 0) F" + by (metis at_top_le_at_infinity filterlim_at filterlim_inverse_at_iff filterlim_mono order_refl) + text {* We only show rules for multiplication and addition when the functions are either against a real @@ -1610,6 +1617,12 @@ qed qed +lemma filterlim_tendsto_pos_mult_at_bot: + assumes "(f ---> c) F" "0 < (c::real)" "filterlim g at_bot F" + shows "LIM x F. f x * g x :> at_bot" + using filterlim_tendsto_pos_mult_at_top[OF assms(1,2), of "\x. - g x"] assms(3) + unfolding filterlim_uminus_at_bot by simp + lemma filterlim_tendsto_add_at_top: assumes f: "(f ---> c) F" assumes g: "LIM x F. g x :> at_top" diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Fri Dec 07 17:00:40 2012 +0100 @@ -1698,6 +1698,13 @@ by auto qed +lemma has_vector_derivative_withinI_DERIV: + assumes f: "DERIV f x :> y" shows "(f has_vector_derivative y) (at x within s)" + unfolding has_vector_derivative_def real_scaleR_def + apply (rule has_derivative_at_within) + using DERIV_conv_has_derivative[THEN iffD1, OF f] + apply (subst mult_commute) . + lemma vector_derivative_unique_at: assumes "(f has_vector_derivative f') (at x)" assumes "(f has_vector_derivative f'') (at x)" diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Probability/Borel_Space.thy Fri Dec 07 17:00:40 2012 +0100 @@ -816,6 +816,10 @@ "f \ borel_measurable M \ g \ borel_measurable M \ (\x. log (g x) (f x)) \ borel_measurable M" unfolding log_def by auto +lemma borel_measurable_exp[measurable]: "exp \ borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI + continuous_isCont[THEN iffD1] isCont_exp) + lemma measurable_count_space_eq2_countable: fixes f :: "'a => 'c::countable" shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Probability/Distributions.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Distributions.thy Fri Dec 07 17:00:40 2012 +0100 @@ -0,0 +1,392 @@ +theory Distributions + imports Probability_Measure +begin + +subsection {* Exponential distribution *} + +definition exponential_density :: "real \ real \ real" where + "exponential_density l x = (if x < 0 then 0 else l * exp (- x * l))" + +lemma borel_measurable_exponential_density[measurable]: "exponential_density l \ borel_measurable borel" + by (auto simp add: exponential_density_def[abs_def]) + +lemma (in prob_space) exponential_distributed_params: + assumes D: "distributed M lborel X (exponential_density l)" + shows "0 < l" +proof (cases l "0 :: real" rule: linorder_cases) + assume "l < 0" + have "emeasure lborel {0 <.. 1::real} \ + emeasure lborel {x :: real \ space lborel. 0 < x}" + by (rule emeasure_mono) (auto simp: greaterThan_def[symmetric]) + also have "emeasure lborel {x :: real \ space lborel. 0 < x} = 0" + proof - + have "AE x in lborel. 0 \ exponential_density l x" + using assms by (auto simp: distributed_real_AE) + then have "AE x in lborel. x \ (0::real)" + apply eventually_elim + using `l < 0` + apply (auto simp: exponential_density_def zero_le_mult_iff split: split_if_asm) + done + then show "emeasure lborel {x :: real \ space lborel. 0 < x} = 0" + by (subst (asm) AE_iff_measurable[OF _ refl]) (auto simp: not_le greaterThan_def[symmetric]) + qed + finally show "0 < l" by simp +next + assume "l = 0" + then have [simp]: "\x. ereal (exponential_density l x) = 0" + by (simp add: exponential_density_def) + interpret X: prob_space "distr M lborel X" + using distributed_measurable[OF D] by (rule prob_space_distr) + from X.emeasure_space_1 + show "0 < l" + by (simp add: emeasure_density distributed_distr_eq_density[OF D]) +qed assumption + +lemma + assumes [arith]: "0 < l" + shows emeasure_exponential_density_le0: "0 \ a \ emeasure (density lborel (exponential_density l)) {.. a} = 1 - exp (- a * l)" + and prob_space_exponential_density: "prob_space (density lborel (exponential_density l))" + (is "prob_space ?D") +proof - + let ?f = "\x. l * exp (- x * l)" + let ?F = "\x. - exp (- x * l)" + + have deriv: "\x. DERIV ?F x :> ?f x" "\x. 0 \ l * exp (- x * l)" + by (auto intro!: DERIV_intros simp: zero_le_mult_iff) + + have "emeasure ?D (space ?D) = (\\<^isup>+ x. ereal (?f x) * indicator {0..} x \lborel)" + by (auto simp: emeasure_density exponential_density_def + intro!: positive_integral_cong split: split_indicator) + also have "\ = ereal (0 - ?F 0)" + proof (rule positive_integral_FTC_atLeast) + have "((\x. exp (l * x)) ---> 0) at_bot" + by (rule filterlim_compose[OF exp_at_bot filterlim_tendsto_pos_mult_at_bot[of _ l]]) + (simp_all add: tendsto_const filterlim_ident) + then show "((\x. - exp (- x * l)) ---> 0) at_top" + unfolding filterlim_at_top_mirror + by (simp add: tendsto_minus_cancel_left[symmetric] ac_simps) + qed (insert deriv, auto) + also have "\ = 1" by (simp add: one_ereal_def) + finally have "emeasure ?D (space ?D) = 1" . + then show "prob_space ?D" by rule + + assume "0 \ a" + have "emeasure ?D {..a} = (\\<^isup>+x. ereal (?f x) * indicator {0..a} x \lborel)" + by (auto simp add: emeasure_density intro!: positive_integral_cong split: split_indicator) + (auto simp: exponential_density_def) + also have "(\\<^isup>+x. ereal (?f x) * indicator {0..a} x \lborel) = ereal (?F a) - ereal (?F 0)" + using `0 \ a` deriv by (intro positive_integral_FTC_atLeastAtMost) auto + also have "\ = 1 - exp (- a * l)" + by simp + finally show "emeasure ?D {.. a} = 1 - exp (- a * l)" . +qed + + +lemma (in prob_space) exponential_distributedD_le: + assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \ a" + shows "\

(x in M. X x \ a) = 1 - exp (- a * l)" +proof - + have "emeasure M {x \ space M. X x \ a } = emeasure (distr M lborel X) {.. a}" + using distributed_measurable[OF D] + by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure]) + also have "\ = emeasure (density lborel (exponential_density l)) {.. a}" + unfolding distributed_distr_eq_density[OF D] .. + also have "\ = 1 - exp (- a * l)" + using emeasure_exponential_density_le0[OF exponential_distributed_params[OF D] a] . + finally show ?thesis + by (auto simp: measure_def) +qed + +lemma (in prob_space) exponential_distributedD_gt: + assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \ a" + shows "\

(x in M. a < X x ) = exp (- a * l)" +proof - + have "exp (- a * l) = 1 - \

(x in M. X x \ a)" + unfolding exponential_distributedD_le[OF D a] by simp + also have "\ = prob (space M - {x \ space M. X x \ a })" + using distributed_measurable[OF D] + by (subst prob_compl) auto + also have "\ = \

(x in M. a < X x )" + by (auto intro!: arg_cong[where f=prob] simp: not_le) + finally show ?thesis by simp +qed + +lemma (in prob_space) exponential_distributed_memoryless: + assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \ a"and t: "0 \ t" + shows "\

(x in M. a + t < X x \ a < X x) = \

(x in M. t < X x)" +proof - + have "\

(x in M. a + t < X x \ a < X x) = \

(x in M. a + t < X x) / \

(x in M. a < X x)" + using `0 \ t` by (auto simp: cond_prob_def intro!: arg_cong[where f=prob] arg_cong2[where f="op /"]) + also have "\ = exp (- (a + t) * l) / exp (- a * l)" + using a t by (simp add: exponential_distributedD_gt[OF D]) + also have "\ = exp (- t * l)" + using exponential_distributed_params[OF D] by (auto simp: field_simps exp_add[symmetric]) + finally show ?thesis + using t by (simp add: exponential_distributedD_gt[OF D]) +qed + +lemma exponential_distributedI: + assumes X[measurable]: "X \ borel_measurable M" and [arith]: "0 < l" + and X_distr: "\a. 0 \ a \ emeasure M {x\space M. X x \ a} = 1 - exp (- a * l)" + shows "distributed M lborel X (exponential_density l)" +proof (rule distributedI_borel_atMost) + fix a :: real + + { assume "a \ 0" + with X have "emeasure M {x\space M. X x \ a} \ emeasure M {x\space M. X x \ 0}" + by (intro emeasure_mono) auto + then have "emeasure M {x\space M. X x \ a} = 0" + using X_distr[of 0] by (simp add: one_ereal_def emeasure_le_0_iff) } + note eq_0 = this + + have "\x. \ 0 \ a \ ereal (exponential_density l x) * indicator {..a} x = 0" + by (simp add: exponential_density_def) + then show "(\\<^isup>+ x. exponential_density l x * indicator {..a} x \lborel) = ereal (if 0 \ a then 1 - exp (- a * l) else 0)" + using emeasure_exponential_density_le0[of l a] + by (auto simp: emeasure_density times_ereal.simps[symmetric] ereal_indicator + simp del: times_ereal.simps ereal_zero_times) + show "emeasure M {x\space M. X x \ a} = ereal (if 0 \ a then 1 - exp (- a * l) else 0)" + using X_distr[of a] eq_0 by (auto simp: one_ereal_def) + show "AE x in lborel. 0 \ exponential_density l x " + by (auto simp: exponential_density_def intro!: AE_I2 mult_nonneg_nonneg) +qed simp_all + +lemma (in prob_space) exponential_distributed_iff: + "distributed M lborel X (exponential_density l) \ + (X \ borel_measurable M \ 0 < l \ (\a\0. \

(x in M. X x \ a) = 1 - exp (- a * l)))" + using + distributed_measurable[of M lborel X "exponential_density l"] + exponential_distributed_params[of X l] + emeasure_exponential_density_le0[of l] + exponential_distributedD_le[of X l] + by (auto intro!: exponential_distributedI simp: one_ereal_def emeasure_eq_measure) + +lemma borel_integral_x_exp: + "(\x. x * exp (- x) * indicator {0::real ..} x \lborel) = 1" +proof (rule integral_monotone_convergence) + let ?f = "\i x. x * exp (- x) * indicator {0::real .. i} x" + have "eventually (\b::real. 0 \ b) at_top" + by (rule eventually_ge_at_top) + then have "eventually (\b. 1 - (inverse (exp b) + b / exp b) = integral\<^isup>L lborel (?f b)) at_top" + proof eventually_elim + fix b :: real assume [simp]: "0 \ b" + have "(\x. (exp (-x)) * indicator {0 .. b} x \lborel) - (integral\<^isup>L lborel (?f b)) = + (\x. (exp (-x) - x * exp (-x)) * indicator {0 .. b} x \lborel)" + by (subst integral_diff(2)[symmetric]) + (auto intro!: borel_integrable_atLeastAtMost integral_cong split: split_indicator) + also have "\ = b * exp (-b) - 0 * exp (- 0)" + proof (rule integral_FTC_atLeastAtMost) + fix x assume "0 \ x" "x \ b" + show "DERIV (\x. x * exp (-x)) x :> exp (-x) - x * exp (-x)" + by (auto intro!: DERIV_intros) + show "isCont (\x. exp (- x) - x * exp (- x)) x " + by (intro isCont_intros isCont_exp') + qed fact + also have "(\x. (exp (-x)) * indicator {0 .. b} x \lborel) = - exp (- b) - - exp (- 0)" + by (rule integral_FTC_atLeastAtMost) (auto intro!: DERIV_intros) + finally show "1 - (inverse (exp b) + b / exp b) = integral\<^isup>L lborel (?f b)" + by (auto simp: field_simps exp_minus inverse_eq_divide) + qed + then have "((\i. integral\<^isup>L lborel (?f i)) ---> 1 - (0 + 0)) at_top" + proof (rule Lim_transform_eventually) + show "((\i. 1 - (inverse (exp i) + i / exp i)) ---> 1 - (0 + 0 :: real)) at_top" + using tendsto_power_div_exp_0[of 1] + by (intro tendsto_intros tendsto_inverse_0_at_top exp_at_top) simp + qed + then show "(\i. integral\<^isup>L lborel (?f i)) ----> 1" + by (intro filterlim_compose[OF _ filterlim_real_sequentially]) simp + show "AE x in lborel. mono (\n::nat. x * exp (- x) * indicator {0..real n} x)" + by (auto simp: mono_def mult_le_0_iff zero_le_mult_iff split: split_indicator) + show "\i::nat. integrable lborel (\x. x * exp (- x) * indicator {0..real i} x)" + by (rule borel_integrable_atLeastAtMost) auto + show "AE x in lborel. (\i. x * exp (- x) * indicator {0..real i} x) ----> x * exp (- x) * indicator {0..} x" + apply (intro AE_I2 Lim_eventually ) + apply (rule_tac c="natfloor x + 1" in eventually_sequentiallyI) + apply (auto simp add: not_le dest!: ge_natfloor_plus_one_imp_gt[simplified] split: split_indicator) + done +qed (auto intro!: borel_measurable_times borel_measurable_exp) + +lemma (in prob_space) exponential_distributed_expectation: + assumes D: "distributed M lborel X (exponential_density l)" + shows "expectation X = 1 / l" +proof (subst distributed_integral[OF D, of "\x. x", symmetric]) + have "0 < l" + using exponential_distributed_params[OF D] . + have [simp]: "\x. x * (l * (exp (- (x * l)) * indicator {0..} (x * l))) = + x * exponential_density l x" + using `0 < l` + by (auto split: split_indicator simp: zero_le_mult_iff exponential_density_def) + from borel_integral_x_exp `0 < l` + show "(\ x. exponential_density l x * x \lborel) = 1 / l" + by (subst (asm) lebesgue_integral_real_affine[of "l" _ 0]) + (simp_all add: borel_measurable_exp nonzero_eq_divide_eq ac_simps) +qed simp + +subsection {* Uniform distribution *} + +lemma uniform_distrI: + assumes X: "X \ measurable M M'" + and A: "A \ sets M'" "emeasure M' A \ \" "emeasure M' A \ 0" + assumes distr: "\B. B \ sets M' \ emeasure M (X -` B \ space M) = emeasure M' (A \ B) / emeasure M' A" + shows "distr M M' X = uniform_measure M' A" + unfolding uniform_measure_def +proof (intro measure_eqI) + let ?f = "\x. indicator A x / emeasure M' A" + fix B assume B: "B \ sets (distr M M' X)" + with X have "emeasure M (X -` B \ space M) = emeasure M' (A \ B) / emeasure M' A" + by (simp add: distr[of B] measurable_sets) + also have "\ = (1 / emeasure M' A) * emeasure M' (A \ B)" + by simp + also have "\ = (\\<^isup>+ x. (1 / emeasure M' A) * indicator (A \ B) x \M')" + using A B + by (intro positive_integral_cmult_indicator[symmetric]) (auto intro!: zero_le_divide_ereal) + also have "\ = (\\<^isup>+ x. ?f x * indicator B x \M')" + by (rule positive_integral_cong) (auto split: split_indicator) + finally show "emeasure (distr M M' X) B = emeasure (density M' ?f) B" + using A B X by (auto simp add: emeasure_distr emeasure_density) +qed simp + +lemma uniform_distrI_borel: + fixes A :: "real set" + assumes X[measurable]: "X \ borel_measurable M" and A: "emeasure lborel A = ereal r" "0 < r" + and [measurable]: "A \ sets borel" + assumes distr: "\a. emeasure M {x\space M. X x \ a} = emeasure lborel (A \ {.. a}) / r" + shows "distributed M lborel X (\x. indicator A x / measure lborel A)" +proof (rule distributedI_borel_atMost) + let ?f = "\x. 1 / r * indicator A x" + fix a + have "emeasure lborel (A \ {..a}) \ emeasure lborel A" + using A by (intro emeasure_mono) auto + also have "\ < \" + using A by simp + finally have fin: "emeasure lborel (A \ {..a}) \ \" + by simp + from emeasure_eq_ereal_measure[OF this] + have fin_eq: "emeasure lborel (A \ {..a}) / r = ereal (measure lborel (A \ {..a}) / r)" + using A by simp + then show "emeasure M {x\space M. X x \ a} = ereal (measure lborel (A \ {..a}) / r)" + using distr by simp + + have "(\\<^isup>+ x. ereal (indicator A x / measure lborel A * indicator {..a} x) \lborel) = + (\\<^isup>+ x. ereal (1 / measure lborel A) * indicator (A \ {..a}) x \lborel)" + by (auto intro!: positive_integral_cong split: split_indicator) + also have "\ = ereal (1 / measure lborel A) * emeasure lborel (A \ {..a})" + using `A \ sets borel` + by (intro positive_integral_cmult_indicator) (auto simp: measure_nonneg) + also have "\ = ereal (measure lborel (A \ {..a}) / r)" + unfolding emeasure_eq_ereal_measure[OF fin] using A by (simp add: measure_def) + finally show "(\\<^isup>+ x. ereal (indicator A x / measure lborel A * indicator {..a} x) \lborel) = + ereal (measure lborel (A \ {..a}) / r)" . +qed (auto intro!: divide_nonneg_nonneg measure_nonneg) + +lemma (in prob_space) uniform_distrI_borel_atLeastAtMost: + fixes a b :: real + assumes X: "X \ borel_measurable M" and "a < b" + assumes distr: "\t. a \ t \ t \ b \ \

(x in M. X x \ t) = (t - a) / (b - a)" + shows "distributed M lborel X (\x. indicator {a..b} x / measure lborel {a..b})" +proof (rule uniform_distrI_borel) + fix t + have "t < a \ (a \ t \ t \ b) \ b < t" + by auto + then show "emeasure M {x\space M. X x \ t} = emeasure lborel ({a .. b} \ {..t}) / (b - a)" + proof (elim disjE conjE) + assume "t < a" + then have "emeasure M {x\space M. X x \ t} \ emeasure M {x\space M. X x \ a}" + using X by (auto intro!: emeasure_mono measurable_sets) + also have "\ = 0" + using distr[of a] `a < b` by (simp add: emeasure_eq_measure) + finally have "emeasure M {x\space M. X x \ t} = 0" + by (simp add: antisym measure_nonneg emeasure_le_0_iff) + with `t < a` show ?thesis by simp + next + assume bnds: "a \ t" "t \ b" + have "{a..b} \ {..t} = {a..t}" + using bnds by auto + then show ?thesis using `a \ t` `a < b` + using distr[OF bnds] by (simp add: emeasure_eq_measure) + next + assume "b < t" + have "1 = emeasure M {x\space M. X x \ b}" + using distr[of b] `a < b` by (simp add: one_ereal_def emeasure_eq_measure) + also have "\ \ emeasure M {x\space M. X x \ t}" + using X `b < t` by (auto intro!: emeasure_mono measurable_sets) + finally have "emeasure M {x\space M. X x \ t} = 1" + by (simp add: antisym emeasure_eq_measure one_ereal_def) + with `b < t` `a < b` show ?thesis by (simp add: measure_def one_ereal_def) + qed +qed (insert X `a < b`, auto) + +lemma (in prob_space) uniform_distributed_measure: + fixes a b :: real + assumes D: "distributed M lborel X (\x. indicator {a .. b} x / measure lborel {a .. b})" + assumes " a \ t" "t \ b" + shows "\

(x in M. X x \ t) = (t - a) / (b - a)" +proof - + have "emeasure M {x \ space M. X x \ t} = emeasure (distr M lborel X) {.. t}" + using distributed_measurable[OF D] + by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure]) + also have "\ = (\\<^isup>+x. ereal (1 / (b - a)) * indicator {a .. t} x \lborel)" + using distributed_borel_measurable[OF D] `a \ t` `t \ b` + unfolding distributed_distr_eq_density[OF D] + by (subst emeasure_density) + (auto intro!: positive_integral_cong simp: measure_def split: split_indicator) + also have "\ = ereal (1 / (b - a)) * (t - a)" + using `a \ t` `t \ b` + by (subst positive_integral_cmult_indicator) auto + finally show ?thesis + by (simp add: measure_def) +qed + +lemma (in prob_space) uniform_distributed_bounds: + fixes a b :: real + assumes D: "distributed M lborel X (\x. indicator {a .. b} x / measure lborel {a .. b})" + shows "a < b" +proof (rule ccontr) + assume "\ a < b" + then have "{a .. b} = {} \ {a .. b} = {a .. a}" by simp + with uniform_distributed_params[OF D] show False + by (auto simp: measure_def) +qed + +lemma (in prob_space) uniform_distributed_iff: + fixes a b :: real + shows "distributed M lborel X (\x. indicator {a..b} x / measure lborel {a..b}) \ + (X \ borel_measurable M \ a < b \ (\t\{a .. b}. \

(x in M. X x \ t)= (t - a) / (b - a)))" + using + uniform_distributed_bounds[of X a b] + uniform_distributed_measure[of X a b] + distributed_measurable[of M lborel X] + by (auto intro!: uniform_distrI_borel_atLeastAtMost simp: one_ereal_def emeasure_eq_measure) + +lemma (in prob_space) uniform_distributed_expectation: + fixes a b :: real + assumes D: "distributed M lborel X (\x. indicator {a .. b} x / measure lborel {a .. b})" + shows "expectation X = (a + b) / 2" +proof (subst distributed_integral[OF D, of "\x. x", symmetric]) + have "a < b" + using uniform_distributed_bounds[OF D] . + + have "(\ x. indicator {a .. b} x / measure lborel {a .. b} * x \lborel) = + (\ x. (x / measure lborel {a .. b}) * indicator {a .. b} x \lborel)" + by (intro integral_cong) auto + also have "(\ x. (x / measure lborel {a .. b}) * indicator {a .. b} x \lborel) = (a + b) / 2" + proof (subst integral_FTC_atLeastAtMost) + fix x + show "DERIV (\x. x ^ 2 / (2 * measure lborel {a..b})) x :> x / measure lborel {a..b}" + using uniform_distributed_params[OF D] + by (auto intro!: DERIV_intros) + show "isCont (\x. x / Sigma_Algebra.measure lborel {a..b}) x" + using uniform_distributed_params[OF D] + by (auto intro!: isCont_divide) + have *: "b\ / (2 * measure lborel {a..b}) - a\ / (2 * measure lborel {a..b}) = + (b*b - a * a) / (2 * (b - a))" + using `a < b` + by (auto simp: measure_def power2_eq_square diff_divide_distrib[symmetric]) + show "b\ / (2 * measure lborel {a..b}) - a\ / (2 * measure lborel {a..b}) = (a + b) / 2" + using `a < b` + unfolding * square_diff_square_factored by (auto simp: field_simps) + qed (insert `a < b`, simp) + finally show "(\ x. indicator {a .. b} x / measure lborel {a .. b} * x \lborel) = (a + b) / 2" . +qed auto + +end diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Probability/Information.thy Fri Dec 07 17:00:40 2012 +0100 @@ -932,24 +932,6 @@ finally show ?thesis . qed -lemma (in prob_space) uniform_distributed_params: - assumes X: "distributed M MX X (\x. indicator A x / measure MX A)" - shows "A \ sets MX" "measure MX A \ 0" -proof - - interpret X: prob_space "distr M MX X" - using distributed_measurable[OF X] by (rule prob_space_distr) - - show "measure MX A \ 0" - proof - assume "measure MX A = 0" - with X.emeasure_space_1 X.prob_space distributed_distr_eq_density[OF X] - show False - by (simp add: emeasure_density zero_ereal_def[symmetric]) - qed - with measure_notin_sets[of A MX] show "A \ sets MX" - by blast -qed - lemma (in information_space) entropy_uniform: assumes X: "distributed M MX X (\x. indicator A x / measure MX A)" (is "distributed _ _ _ ?f") shows "entropy b MX X = log b (measure MX A)" diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Fri Dec 07 17:00:40 2012 +0100 @@ -1020,4 +1020,159 @@ shows "integral\<^isup>L lborel f = \x. f (p2e x) \((\\<^isub>M i\{.. real" + assumes f_borel: "f \ borel_measurable borel" and nonneg: "\x. 0 \ f x" + assumes f: "f integrable_on UNIV" + shows "integrable lborel f" +proof - + have "(\\<^isup>+ x. ereal (f x) \lborel) \ \" + using has_integral_iff_positive_integral_lborel[OF f_borel nonneg] f + by (auto simp: integrable_on_def) + moreover have "(\\<^isup>+ x. ereal (- f x) \lborel) = 0" + using f_borel nonneg by (subst positive_integral_0_iff_AE) auto + ultimately show ?thesis + using f_borel by (auto simp: integrable_def) +qed + +subsection {* Fundamental Theorem of Calculus for the Lebesgue integral *} + +lemma borel_integrable_atLeastAtMost: + fixes a b :: real + assumes f: "\x. a \ x \ x \ b \ isCont f x" + shows "integrable lborel (\x. f x * indicator {a .. b} x)" (is "integrable _ ?f") +proof cases + assume "a \ b" + + from isCont_Lb_Ub[OF `a \ b`, of f] f + obtain M L where + bounds: "\x. a \ x \ x \ b \ f x \ M" "\x. a \ x \ x \ b \ L \ f x" + by metis + + show ?thesis + proof (rule integrable_bound) + show "integrable lborel (\x. max \M\ \L\ * indicator {a..b} x)" + by (rule integral_cmul_indicator) simp_all + show "AE x in lborel. \?f x\ \ max \M\ \L\ * indicator {a..b} x" + proof (rule AE_I2) + fix x show "\?f x\ \ max \M\ \L\ * indicator {a..b} x" + using bounds[of x] by (auto split: split_indicator) + qed + + let ?g = "\x. if x = a then f a else if x = b then f b else if x \ {a <..< b} then f x else 0" + from f have "continuous_on {a <..< b} f" + by (subst continuous_on_eq_continuous_at) (auto simp add: continuous_isCont) + then have "?g \ borel_measurable borel" + using borel_measurable_continuous_on_open[of "{a <..< b }" f "\x. x" borel 0] + by (auto intro!: measurable_If[where P="\x. x = a"] measurable_If[where P="\x. x = b"]) + also have "?g = ?f" + using `a \ b` by (auto intro!: ext split: split_indicator) + finally show "?f \ borel_measurable lborel" + by simp + qed +qed simp + +lemma integral_FTC_atLeastAtMost: + fixes a b :: real + assumes "a \ b" + and F: "\x. a \ x \ x \ b \ DERIV F x :> f x" + and f: "\x. a \ x \ x \ b \ isCont f x" + shows "integral\<^isup>L lborel (\x. f x * indicator {a .. b} x) = F b - F a" +proof - + let ?f = "\x. f x * indicator {a .. b} x" + have "(?f has_integral (\x. ?f x \lborel)) UNIV" + using borel_integrable_atLeastAtMost[OF f] + by (rule borel_integral_has_integral) + moreover + have "(f has_integral F b - F a) {a .. b}" + by (intro fundamental_theorem_of_calculus has_vector_derivative_withinI_DERIV ballI assms) auto + then have "(?f has_integral F b - F a) {a .. b}" + by (subst has_integral_eq_eq[where g=f]) auto + then have "(?f has_integral F b - F a) UNIV" + by (intro has_integral_on_superset[where t=UNIV and s="{a..b}"]) auto + ultimately show "integral\<^isup>L lborel ?f = F b - F a" + by (rule has_integral_unique) +qed + +text {* + +For the positive integral we replace continuity with Borel-measurability. + +*} + +lemma positive_integral_FTC_atLeastAtMost: + assumes f_borel: "f \ borel_measurable borel" + assumes f: "\x. x \ {a..b} \ DERIV F x :> f x" "\x. x \ {a..b} \ 0 \ f x" and "a \ b" + shows "(\\<^isup>+x. f x * indicator {a .. b} x \lborel) = F b - F a" +proof - + have i: "(f has_integral F b - F a) {a..b}" + by (intro fundamental_theorem_of_calculus ballI has_vector_derivative_withinI_DERIV assms) + have i: "((\x. f x * indicator {a..b} x) has_integral F b - F a) {a..b}" + by (rule has_integral_eq[OF _ i]) auto + have i: "((\x. f x * indicator {a..b} x) has_integral F b - F a) UNIV" + by (rule has_integral_on_superset[OF _ _ i]) auto + then have "(\\<^isup>+x. ereal (f x * indicator {a .. b} x) \lborel) = F b - F a" + using f f_borel + by (subst has_integral_iff_positive_integral_lborel[symmetric]) (auto split: split_indicator) + also have "(\\<^isup>+x. ereal (f x * indicator {a .. b} x) \lborel) = (\\<^isup>+x. ereal (f x) * indicator {a .. b} x \lborel)" + by (auto intro!: positive_integral_cong simp: indicator_def) + finally show ?thesis by simp +qed + +lemma positive_integral_FTC_atLeast: + fixes f :: "real \ real" + assumes f_borel: "f \ borel_measurable borel" + assumes f: "\x. a \ x \ DERIV F x :> f x" + assumes nonneg: "\x. a \ x \ 0 \ f x" + assumes lim: "(F ---> T) at_top" + shows "(\\<^isup>+x. ereal (f x) * indicator {a ..} x \lborel) = T - F a" +proof - + let ?f = "\(i::nat) (x::real). ereal (f x) * indicator {a..a + real i} x" + let ?fR = "\x. ereal (f x) * indicator {a ..} x" + have "\x. (SUP i::nat. ?f i x) = ?fR x" + proof (rule SUP_Lim_ereal) + show "\x. incseq (\i. ?f i x)" + using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator) + + fix x + from reals_Archimedean2[of "x - a"] guess n .. + then have "eventually (\n. ?f n x = ?fR x) sequentially" + by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator) + then show "(\n. ?f n x) ----> ?fR x" + by (rule Lim_eventually) + qed + then have "integral\<^isup>P lborel ?fR = (\\<^isup>+ x. (SUP i::nat. ?f i x) \lborel)" + by simp + also have "\ = (SUP i::nat. (\\<^isup>+ x. ?f i x \lborel))" + proof (rule positive_integral_monotone_convergence_SUP) + show "incseq ?f" + using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator) + show "\i. (?f i) \ borel_measurable lborel" + using f_borel by auto + show "\i x. 0 \ ?f i x" + using nonneg by (auto split: split_indicator) + qed + also have "\ = (SUP i::nat. F (a + real i) - F a)" + by (subst positive_integral_FTC_atLeastAtMost[OF f_borel f nonneg]) auto + also have "\ = T - F a" + proof (rule SUP_Lim_ereal) + show "incseq (\n. ereal (F (a + real n) - F a))" + proof (simp add: incseq_def, safe) + fix m n :: nat assume "m \ n" + with f nonneg show "F (a + real m) \ F (a + real n)" + by (intro DERIV_nonneg_imp_nondecreasing[where f=F]) + (simp, metis add_increasing2 order_refl order_trans real_of_nat_ge_zero) + qed + have "(\x. F (a + real x)) ----> T" + apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top]) + apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl]) + apply (rule filterlim_real_sequentially) + done + then show "(\n. ereal (F (a + real n) - F a)) ----> ereal (T - F a)" + unfolding lim_ereal + by (intro tendsto_diff) auto + qed + finally show ?thesis . +qed + end diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Probability/Measure_Space.thy --- a/src/HOL/Probability/Measure_Space.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Probability/Measure_Space.thy Fri Dec 07 17:00:40 2012 +0100 @@ -443,6 +443,12 @@ lemma emeasure_not_MInf[simp]: "emeasure M A \ - \" using emeasure_nonneg[of M A] by auto + +lemma emeasure_le_0_iff: "emeasure M A \ 0 \ emeasure M A = 0" + using emeasure_nonneg[of M A] by auto + +lemma emeasure_less_0_iff: "emeasure M A < 0 \ False" + using emeasure_nonneg[of M A] by auto lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Probability/Probability.thy --- a/src/HOL/Probability/Probability.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Probability/Probability.thy Fri Dec 07 17:00:40 2012 +0100 @@ -7,6 +7,7 @@ Projective_Limit Independent_Family Information + Distributions begin end diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Probability/Probability_Measure.thy Fri Dec 07 17:00:40 2012 +0100 @@ -843,6 +843,81 @@ done qed +lemma distributedI_real: + fixes f :: "'a \ real" + assumes gen: "sets M1 = sigma_sets (space M1) E" and "Int_stable E" + and A: "range A \ E" "(\i::nat. A i) = space M1" "\i. emeasure (distr M M1 X) (A i) \ \" + and X: "X \ measurable M M1" + and f: "f \ borel_measurable M1" "AE x in M1. 0 \ f x" + and eq: "\A. A \ E \ emeasure M (X -` A \ space M) = (\\<^isup>+ x. f x * indicator A x \M1)" + shows "distributed M M1 X f" + unfolding distributed_def +proof (intro conjI) + show "distr M M1 X = density M1 f" + proof (rule measure_eqI_generator_eq[where A=A]) + { fix A assume A: "A \ E" + then have "A \ sigma_sets (space M1) E" by auto + then have "A \ sets M1" + using gen by simp + with f A eq[of A] X show "emeasure (distr M M1 X) A = emeasure (density M1 f) A" + by (simp add: emeasure_distr emeasure_density borel_measurable_ereal + times_ereal.simps[symmetric] ereal_indicator + del: times_ereal.simps) } + note eq_E = this + show "Int_stable E" by fact + { fix e assume "e \ E" + then have "e \ sigma_sets (space M1) E" by auto + then have "e \ sets M1" unfolding gen . + then have "e \ space M1" by (rule sets.sets_into_space) } + then show "E \ Pow (space M1)" by auto + show "sets (distr M M1 X) = sigma_sets (space M1) E" + "sets (density M1 (\x. ereal (f x))) = sigma_sets (space M1) E" + unfolding gen[symmetric] by auto + qed fact+ +qed (insert X f, auto) + +lemma distributedI_borel_atMost: + fixes f :: "real \ real" + assumes [measurable]: "X \ borel_measurable M" + and [measurable]: "f \ borel_measurable borel" and f[simp]: "AE x in lborel. 0 \ f x" + and g_eq: "\a. (\\<^isup>+x. f x * indicator {..a} x \lborel) = ereal (g a)" + and M_eq: "\a. emeasure M {x\space M. X x \ a} = ereal (g a)" + shows "distributed M lborel X f" +proof (rule distributedI_real) + show "sets lborel = sigma_sets (space lborel) (range atMost)" + by (simp add: borel_eq_atMost) + show "Int_stable (range atMost :: real set set)" + by (auto simp: Int_stable_def) + have vimage_eq: "\a. (X -` {..a} \ space M) = {x\space M. X x \ a}" by auto + def A \ "\i::nat. {.. real i}" + then show "range A \ range atMost" "(\i. A i) = space lborel" + "\i. emeasure (distr M lborel X) (A i) \ \" + by (auto simp: real_arch_simple emeasure_distr vimage_eq M_eq) + + fix A :: "real set" assume "A \ range atMost" + then obtain a where A: "A = {..a}" by auto + show "emeasure M (X -` A \ space M) = (\\<^isup>+x. f x * indicator A x \lborel)" + unfolding vimage_eq A M_eq g_eq .. +qed auto + +lemma (in prob_space) uniform_distributed_params: + assumes X: "distributed M MX X (\x. indicator A x / measure MX A)" + shows "A \ sets MX" "measure MX A \ 0" +proof - + interpret X: prob_space "distr M MX X" + using distributed_measurable[OF X] by (rule prob_space_distr) + + show "measure MX A \ 0" + proof + assume "measure MX A = 0" + with X.emeasure_space_1 X.prob_space distributed_distr_eq_density[OF X] + show False + by (simp add: emeasure_density zero_ereal_def[symmetric]) + qed + with measure_notin_sets[of A MX] show "A \ sets MX" + by blast +qed + lemma prob_space_uniform_measure: assumes A: "emeasure M A \ 0" "emeasure M A \ \" shows "prob_space (uniform_measure M A)" diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Relation.thy --- a/src/HOL/Relation.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Relation.thy Fri Dec 07 17:00:40 2012 +0100 @@ -919,7 +919,7 @@ subsubsection {* Image of a set under a relation *} -definition Image :: "('a \ 'b) set \ 'a set \ 'b set" (infixl "``" 90) +definition Image :: "('a \ 'b) set \ 'a set \ 'b set" (infixr "``" 90) where "r `` s = {y. \x\s. (x, y) \ r}" diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Rings.thy --- a/src/HOL/Rings.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Rings.thy Fri Dec 07 17:00:40 2012 +0100 @@ -95,7 +95,7 @@ class dvd = times begin -definition dvd :: "'a \ 'a \ bool" (infixl "dvd" 50) where +definition dvd :: "'a \ 'a \ bool" (infix "dvd" 50) where "b dvd a \ (\k. a = b * k)" lemma dvdI [intro?]: "a = b * k \ b dvd a" diff -r 7c8ce63a3c00 -r 79858bd9f5ef src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Fri Dec 07 16:53:35 2012 +0100 +++ b/src/HOL/Set_Interval.thy Fri Dec 07 17:00:40 2012 +0100 @@ -351,6 +351,9 @@ lemma Int_greaterThanLessThan[simp]: "{a<.. {..b} = {.. min a b}" + by (auto simp: min_def) + end