diff -r e4207dfa36b5 -r 4b132ea7d575 src/HOL/Analysis/Measurable.thy --- a/src/HOL/Analysis/Measurable.thy Fri Apr 18 14:19:41 2025 +0200 +++ b/src/HOL/Analysis/Measurable.thy Tue Apr 22 17:35:02 2025 +0100 @@ -290,7 +290,7 @@ {x\space M. if (\i. \n\i. P n x) then the None = n else if (\i. P i x) then P n x \ (\i>n. \ P i x) else Max {} = n}" - by (intro arg_cong[where f=Collect] ext conj_cong) + by (intro arg_cong[where f=Collect] ext) (auto simp add: 1 2 3 not_le[symmetric] intro!: Max_eqI) also have "\ \ sets M" by measurable @@ -317,7 +317,7 @@ {x\space M. if (\i. \n\i. P n x) then the None = n else if (\i. P i x) then P n x \ (\i P i x) else Min {} = n}" - by (intro arg_cong[where f=Collect] ext conj_cong) + by (intro arg_cong[where f=Collect] ext) (auto simp add: 1 2 3 not_le[symmetric] intro!: Min_eqI) also have "\ \ sets M" by measurable @@ -378,7 +378,7 @@ assumes [measurable]: "\i. i \ I \ F i \ measurable M (count_space UNIV)" shows "(\x. SUP i\I. F i x) \ measurable M (count_space UNIV)" unfolding measurable_count_space_eq2_countable -proof (safe intro!: UNIV_I) +proof (intro conjI strip) fix a have "(\x. SUP i\I. F i x) -` {a} \ space M = {x\space M. (\i\I. F i x \ a) \ (\b. (\i\I. F i x \ b) \ a \ b)}" @@ -386,7 +386,7 @@ also have "\ \ sets M" by measurable finally show "(\x. SUP i\I. F i x) -` {a} \ space M \ sets M" . -qed +qed auto lemma measurable_INF[measurable]: fixes F :: "'i \ 'a \ 'b::{complete_lattice, countable}" @@ -394,7 +394,7 @@ assumes [measurable]: "\i. i \ I \ F i \ measurable M (count_space UNIV)" shows "(\x. INF i\I. F i x) \ measurable M (count_space UNIV)" unfolding measurable_count_space_eq2_countable -proof (safe intro!: UNIV_I) +proof (intro conjI strip) fix a have "(\x. INF i\I. F i x) -` {a} \ space M = {x\space M. (\i\I. a \ F i x) \ (\b. (\i\I. b \ F i x) \ b \ a)}" @@ -402,7 +402,7 @@ also have "\ \ sets M" by measurable finally show "(\x. INF i\I. F i x) -` {a} \ space M \ sets M" . -qed +qed auto lemma measurable_lfp_coinduct[consumes 1, case_names continuity step]: fixes F :: "('a \ 'b) \ ('a \ 'b::{complete_lattice, countable})" @@ -628,10 +628,10 @@ lemma measurable_case_enat[measurable (raw)]: assumes f: "f \ M \\<^sub>M count_space UNIV" and g: "\i. g i \ M \\<^sub>M N" and h: "h \ M \\<^sub>M N" shows "(\x. case f x of enat i \ g i x | \ \ h x) \ M \\<^sub>M N" - apply (rule measurable_compose_countable[OF _ f]) - subgoal for i +proof (rule measurable_compose_countable[OF _ f]) + show "(\x. case i of enat i \ g i x | \ \ h x) \ M \\<^sub>M N" for i by (cases i) (auto intro: g h) - done +qed hide_const (open) pred