src/HOL/Probability/Information.thy
Mon, 19 May 2014 14:26:58 +0200 hoelzl renamed positive_integral to nn_integral
Mon, 19 May 2014 13:44:13 +0200 hoelzl fixed document generation for HOL-Probability
Mon, 19 May 2014 12:04:45 +0200 hoelzl introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
Mon, 14 Apr 2014 13:08:17 +0200 hoelzl added divide_nonneg_nonneg and co; made it a simp rule
Sat, 12 Apr 2014 17:26:27 +0200 nipkow made mult_pos_pos a simp rule
Fri, 11 Apr 2014 22:53:33 +0200 nipkow made divide_pos_pos a simp rule
Fri, 11 Apr 2014 13:36:57 +0200 nipkow made mult_nonneg_nonneg a simp rule
Wed, 09 Apr 2014 09:37:47 +0200 hoelzl revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
Thu, 03 Apr 2014 23:51:52 +0100 paulson removing simprule status for divide_minus_left and divide_minus_right
Sat, 15 Mar 2014 08:31:33 +0100 haftmann more complete set of lemmas wrt. image and composition
Tue, 03 Sep 2013 01:12:40 +0200 wenzelm tuned proofs -- clarified flow of facts wrt. calculation;
Tue, 13 Aug 2013 16:25:47 +0200 wenzelm standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
Fri, 07 Dec 2012 14:29:09 +0100 hoelzl add exponential and uniform distributions
Tue, 27 Nov 2012 11:29:47 +0100 immler qualified interpretation of sigma_algebra, to avoid name clashes
Fri, 02 Nov 2012 14:23:54 +0100 hoelzl use measurability prover
Fri, 02 Nov 2012 14:23:40 +0100 hoelzl add measurability prover; add support for Borel sets
Fri, 02 Nov 2012 14:00:39 +0100 hoelzl for the product measure it is enough if only one measure is sigma-finite
Thu, 11 Oct 2012 14:38:58 +0200 hoelzl cleanup borel_measurable_positive_integral_(fst|snd)
Wed, 10 Oct 2012 12:12:36 +0200 hoelzl add finite entropy
Wed, 10 Oct 2012 12:12:36 +0200 hoelzl continuous version of mutual_information_eq_entropy_conditional_entropy
Wed, 10 Oct 2012 12:12:29 +0200 hoelzl remove unnecessary assumption from conditional_entropy_eq
Wed, 10 Oct 2012 12:12:28 +0200 hoelzl alternative definition of conditional entropy
Wed, 10 Oct 2012 12:12:27 +0200 hoelzl remove unneeded assumption from conditional_entropy_generic_eq
Wed, 10 Oct 2012 12:12:26 +0200 hoelzl show and use distributed_swap and distributed_jointI
Wed, 10 Oct 2012 12:12:25 +0200 hoelzl rule to show that conditional mutual information is non-negative in the continuous case
Wed, 10 Oct 2012 12:12:25 +0200 hoelzl continuous version of entropy_le
Wed, 10 Oct 2012 12:12:24 +0200 hoelzl simplified entropy_uniform
Wed, 10 Oct 2012 12:12:18 +0200 hoelzl tuned product measurability
Mon, 23 Apr 2012 12:14:35 +0200 hoelzl reworked Probability theory
Tue, 13 Mar 2012 16:56:56 +0100 wenzelm prefer abs_def over def_raw;
Tue, 28 Feb 2012 21:53:36 +0100 wenzelm avoid undeclared variables in let bindings;
Wed, 07 Dec 2011 15:10:29 +0100 hoelzl remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
Thu, 01 Dec 2011 14:03:57 +0100 hoelzl moved theorems about distribution to the definition; removed oopsed-lemma
Thu, 01 Dec 2011 14:03:57 +0100 hoelzl remove duplicate theorem setsum_real_distribution
Mon, 12 Sep 2011 07:55:43 +0200 nipkow new fastforce replacing fastsimp - less confusing name
Tue, 19 Jul 2011 14:36:12 +0200 hoelzl Rename extreal => ereal
Mon, 27 Jun 2011 09:42:46 +0200 hoelzl move conditional expectation to its own theory file
Thu, 09 Jun 2011 14:04:34 +0200 hoelzl lemma: independence is equal to mutual information = 0
Tue, 29 Mar 2011 14:27:42 +0200 hoelzl rename Probability_Space to Probability_Measure
Tue, 22 Mar 2011 20:06:10 +0100 hoelzl standardized headers
Mon, 14 Mar 2011 14:37:49 +0100 hoelzl reworked Probability theory: measures are not type restricted to positive extended reals
Wed, 23 Feb 2011 11:40:18 +0100 hoelzl add lemma KL_divergence_vimage, mutual_information_generic
Wed, 02 Feb 2011 12:34:45 +0100 hoelzl the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
Mon, 24 Jan 2011 22:29:50 +0100 hoelzl use pre-image measure, instead of image
Wed, 29 Dec 2010 17:34:41 +0100 wenzelm explicit file specifications -- avoid secondary load path;
Wed, 08 Dec 2010 16:15:14 +0100 hoelzl cleanup bijectivity btw. product spaces and pairs
Fri, 03 Dec 2010 15:25:14 +0100 hoelzl it is known as the extended reals, not the infinite reals
Wed, 01 Dec 2010 19:20:30 +0100 hoelzl Support product spaces on sigma finite measures.
Mon, 13 Sep 2010 11:13:15 +0200 nipkow renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
Tue, 07 Sep 2010 10:05:19 +0200 nipkow expand_fun_eq -> ext_iff
Thu, 02 Sep 2010 19:51:53 +0200 hoelzl Moved lemmas to appropriate locations
Thu, 02 Sep 2010 17:12:40 +0200 hoelzl move lemmas to correct theory files
Mon, 23 Aug 2010 19:35:57 +0200 hoelzl Rewrite the Probability theory.
Tue, 04 May 2010 18:19:24 +0200 hoelzl Corrected imports; better approximation of dependencies.
Mon, 03 May 2010 14:35:10 +0200 hoelzl Cleanup information theory
Mon, 03 May 2010 14:35:10 +0200 hoelzl Moved Convex theory to library.
Wed, 07 Apr 2010 17:24:44 +0200 hoelzl Added Information theory and Example: dining cryptographers
less more (0) tip