merged
authorpaulson
Mon, 04 Apr 2011 12:40:00 +0100
changeset 42210 8de8c38d503b
parent 42208 02513eb26eb7 (diff)
parent 42209 bc7d938991e0 (current diff)
child 42212 b33d871675bb
merged
--- a/Admin/mira.py	Mon Apr 04 12:39:14 2011 +0100
+++ b/Admin/mira.py	Mon Apr 04 12:40:00 2011 +0100
@@ -10,6 +10,9 @@
 
 import util
 
+from mira import schedule
+from mira.environment import scheduler
+
 
 # build and evaluation tools
 
@@ -63,10 +66,10 @@
     def to_secs(h, m, s):
         return (int(h) * 60 + int(m)) * 60 + int(s)
     pat = r'Finished (\S+) \((\d+):(\d+):(\d+) elapsed time, (\d+):(\d+):(\d+) cpu time'
-    pat2 = r'Timing (\S+) \((\d+) threads, (\d+\.\d+)s elapsed time, (\d+\.\d+)s cpu time, (\d+\.\d+)s GC time\)'
+    pat2 = r'Timing (\S+) \((\d+) threads, (\d+\.\d+)s elapsed time, (\d+\.\d+)s cpu time, (\d+\.\d+)s GC time, factor (\d+\.\d+)\)'
     t = dict((name, {'elapsed': to_secs(eh,em,es), 'cpu': to_secs(ch,cm,cs)})
              for name, eh, em, es, ch, cm, cs in re.findall(pat, logdata))
-    for name, threads, elapsed, cpu, gc in re.findall(pat2, logdata):
+    for name, threads, elapsed, cpu, gc, factor in re.findall(pat2, logdata):
 
         if name not in t:
             t[name] = {}
@@ -75,6 +78,7 @@
         t[name]['elapsed_inner'] = elapsed
         t[name]['cpu_inner'] = cpu
         t[name]['gc'] = gc
+        t[name]['factor'] = factor
 
     return t
 
@@ -331,6 +335,14 @@
     return judgement_day('Isabelle/src/HOL/Proofs/Lambda', 'StrongNorm', 'prover_timeout=10', *args)
 
 
+JD_confs = 'JD_NS JD_FTA JD_Hoare JD_SN JD_Arrow JD_FFT JD_Jinja JD_QE JD_S2S'.split(' ')
+
+@scheduler()
+def judgement_day_scheduler(env):
+    """Scheduler for Judgement Day."""
+    return schedule.age_scheduler(env, 'Isabelle', JD_confs)
+
+
 # SML/NJ
 
 smlnj_settings = '''
--- a/NEWS	Mon Apr 04 12:39:14 2011 +0100
+++ b/NEWS	Mon Apr 04 12:40:00 2011 +0100
@@ -49,13 +49,25 @@
 * Sledgehammer:
   - sledgehammer available_provers ~> sledgehammer supported_provers
     INCOMPATIBILITY.
+  - Added "monomorphize" and "monomorphize_limit" options.
 
 * "try":
   - Added "simp:", "intro:", and "elim:" options.
 
+* Quickcheck:
+  - Added "eval" option to evaluate terms for the found counterexample
+    (currently only supported by the default (exhaustive) tester)
+  - Added post-processing of terms to obtain readable counterexamples
+    (currently only supported by the default (exhaustive) tester)
+
 * Function package: discontinued option "tailrec".
 INCOMPATIBILITY. Use partial_function instead.
 
+* HOL-Probability:
+  - Caratheodory's extension lemma is now proved for ring_of_sets.
+  - Infinite products of probability measures are now available.
+  - Use extended reals instead of positive extended reals.
+    INCOMPATIBILITY.
 
 *** Document preparation ***
 
--- a/doc-src/Sledgehammer/sledgehammer.tex	Mon Apr 04 12:39:14 2011 +0100
+++ b/doc-src/Sledgehammer/sledgehammer.tex	Mon Apr 04 12:40:00 2011 +0100
@@ -555,6 +555,19 @@
 filter. If the option is set to \textit{smart}, it is set to a value that was
 empirically found to be appropriate for the prover. A typical value would be
 300.
+
+\opfalse{monomorphize}{dont\_monomorphize}
+Specifies whether the relevant facts should be monomorphized---i.e., whether
+their type variables should be instantiated with relevant ground types.
+Monomorphization is always performed for SMT solvers, irrespective of this
+option. Monomorphization can simplify reasoning but also leads to larger fact
+bases, which can slow down the ATPs.
+
+\opdefault{monomorphize\_limit}{int}{\upshape 4}
+Specifies the maximum number of iterations for the monomorphization fixpoint
+construction. The higher this limit is, the more monomorphic instances are
+potentially generated.
+
 \end{enum}
 
 \subsection{Output Format}
--- a/src/CCL/CCL.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/CCL.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -85,6 +85,8 @@
 
   apply_def:     "f ` t == case(f,bot,bot,%x y. bot,%u. u(t))"
   bot_def:         "bot == (lam x. x`x)`(lam x. x`x)"
+
+defs
   fix_def:      "fix(f) == (lam x. f(x`x))`(lam x. f(x`x))"
 
   (*  The pre-order ([=) is defined as a simulation, and behavioural equivalence (=) *)
@@ -106,6 +108,7 @@
 
   (** Partial Order **)
 
+axioms
   po_refl:        "a [= a"
   po_trans:       "[| a [= b;  b [= c |] ==> a [= c"
   po_cong:        "a [= b ==> f(a) [= f(b)"
@@ -136,6 +139,7 @@
 
   (*** Definitions of Termination and Divergence ***)
 
+defs
   Dvg_def:  "Dvg(t) == t = bot"
   Trm_def:  "Trm(t) == ~ Dvg(t)"
 
--- a/src/CCL/Fix.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/Fix.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -9,17 +9,12 @@
 imports Type
 begin
 
-consts
-  idgen      ::       "[i]=>i"
-  INCL      :: "[i=>o]=>o"
+definition idgen :: "i => i"
+  where "idgen(f) == lam t. case(t,true,false,%x y.<f`x, f`y>,%u. lam x. f ` u(x))"
 
-defs
-  idgen_def:
-  "idgen(f) == lam t. case(t,true,false,%x y.<f`x, f`y>,%u. lam x. f ` u(x))"
-
-axioms
-  INCL_def:   "INCL(%x. P(x)) == (ALL f.(ALL n:Nat. P(f^n`bot)) --> P(fix(f)))"
-  po_INCL:    "INCL(%x. a(x) [= b(x))"
+axiomatization INCL :: "[i=>o]=>o" where
+  INCL_def: "INCL(%x. P(x)) == (ALL f.(ALL n:Nat. P(f^n`bot)) --> P(fix(f)))" and
+  po_INCL: "INCL(%x. a(x) [= b(x))" and
   INCL_subst: "INCL(P) ==> INCL(%x. P((g::i=>i)(x)))"
 
 
--- a/src/CCL/Hered.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/Hered.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -15,18 +15,13 @@
   is.  Not so useful for functions!
 *}
 
-consts
-      (*** Predicates ***)
-  HTTgen     ::       "i set => i set"
-  HTT        ::       "i set"
+definition HTTgen :: "i set => i set" where
+  "HTTgen(R) ==
+    {t. t=true | t=false | (EX a b. t= <a, b> & a : R & b : R) |
+      (EX f. t = lam x. f(x) & (ALL x. f(x) : R))}"
 
-axioms
-  (*** Definitions of Hereditary Termination ***)
-
-  HTTgen_def:
-  "HTTgen(R) == {t. t=true | t=false | (EX a b. t=<a,b>      & a : R & b : R) |
-                                      (EX f.  t=lam x. f(x) & (ALL x. f(x) : R))}"
-  HTT_def:       "HTT == gfp(HTTgen)"
+definition HTT :: "i set"
+  where "HTT == gfp(HTTgen)"
 
 
 subsection {* Hereditary Termination *}
--- a/src/CCL/Set.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/Set.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -46,9 +46,9 @@
   "ALL x:A. P"  == "CONST Ball(A, %x. P)"
   "EX x:A. P"   == "CONST Bex(A, %x. P)"
 
-axioms
-  mem_Collect_iff:       "(a : {x. P(x)}) <-> P(a)"
-  set_extension:         "A=B <-> (ALL x. x:A <-> x:B)"
+axiomatization where
+  mem_Collect_iff: "(a : {x. P(x)}) <-> P(a)" and
+  set_extension: "A = B <-> (ALL x. x:A <-> x:B)"
 
 defs
   Ball_def:      "Ball(A, P)  == ALL x. x:A --> P(x)"
--- a/src/CCL/Term.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/Term.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -111,8 +111,7 @@
 consts
   napply     :: "[i=>i,i,i]=>i"            ("(_ ^ _ ` _)" [56,56,56] 56)
 
-axioms
-
+defs
   one_def:                    "one == true"
   if_def:     "if b then t else u  == case(b,t,u,% x y. bot,%v. bot)"
   inl_def:                 "inl(a) == <true,a>"
--- a/src/CCL/Trancl.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/Trancl.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -9,21 +9,20 @@
 imports CCL
 begin
 
-consts
-  trans   :: "i set => o"                   (*transitivity predicate*)
-  id      :: "i set"
-  rtrancl :: "i set => i set"               ("(_^*)" [100] 100)
-  trancl  :: "i set => i set"               ("(_^+)" [100] 100)
-  relcomp :: "[i set,i set] => i set"       (infixr "O" 60)
+definition trans :: "i set => o"  (*transitivity predicate*)
+  where "trans(r) == (ALL x y z. <x,y>:r --> <y,z>:r --> <x,z>:r)"
+
+definition id :: "i set"  (*the identity relation*)
+  where "id == {p. EX x. p = <x,x>}"
 
-axioms
-  trans_def:       "trans(r) == (ALL x y z. <x,y>:r --> <y,z>:r --> <x,z>:r)"
-  relcomp_def:     (*composition of relations*)
-                   "r O s == {xz. EX x y z. xz = <x,z> & <x,y>:s & <y,z>:r}"
-  id_def:          (*the identity relation*)
-                   "id == {p. EX x. p = <x,x>}"
-  rtrancl_def:     "r^* == lfp(%s. id Un (r O s))"
-  trancl_def:      "r^+ == r O rtrancl(r)"
+definition relcomp :: "[i set,i set] => i set"  (infixr "O" 60)  (*composition of relations*)
+  where "r O s == {xz. EX x y z. xz = <x,z> & <x,y>:s & <y,z>:r}"
+
+definition rtrancl :: "i set => i set"  ("(_^*)" [100] 100)
+  where "r^* == lfp(%s. id Un (r O s))"
+
+definition trancl :: "i set => i set"  ("(_^+)" [100] 100)
+  where "r^+ == r O rtrancl(r)"
 
 
 subsection {* Natural deduction for @{text "trans(r)"} *}
--- a/src/CCL/Type.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/Type.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -50,7 +50,7 @@
   (@{const_syntax Sigma}, dependent_tr' (@{syntax_const "_Sigma"}, @{syntax_const "_star"}))]
 *}
 
-axioms
+defs
   Subtype_def: "{x:A. P(x)} == {x. x:A & P(x)}"
   Unit_def:          "Unit == {x. x=one}"
   Bool_def:          "Bool == {x. x=true | x=false}"
--- a/src/CCL/ex/Flag.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/ex/Flag.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -10,24 +10,24 @@
 imports List
 begin
 
-consts
-  Colour             :: "i set"
-  red                :: "i"
-  white              :: "i"
-  blue               :: "i"
-  ccase              :: "[i,i,i,i]=>i"
-  flag               :: "i"
+definition Colour :: "i set"
+  where "Colour == Unit + Unit + Unit"
+
+definition red :: "i"
+  where "red == inl(one)"
+
+definition white :: "i"
+  where "white == inr(inl(one))"
 
-axioms
+definition blue :: "i"
+  where "blue == inr(inr(one))"
 
-  Colour_def:  "Colour == Unit + Unit + Unit"
-  red_def:        "red == inl(one)"
-  white_def:    "white == inr(inl(one))"
-  blue_def:     "blue == inr(inr(one))"
+definition ccase :: "[i,i,i,i]=>i"
+  where "ccase(c,r,w,b) == when(c,%x. r,%wb. when(wb,%x. w,%x. b))"
 
-  ccase_def:   "ccase(c,r,w,b) == when(c,%x. r,%wb. when(wb,%x. w,%x. b))"
-
-  flag_def:    "flag == lam l. letrec
+definition flag :: "i"
+  where
+    "flag == lam l. letrec
       flagx l be lcase(l,<[],<[],[]>>,
                        %h t. split(flagx(t),%lr p. split(p,%lw lb.
                             ccase(h, <red$lr,<lw,lb>>,
@@ -35,13 +35,14 @@
                                      <lr,<lw,blue$lb>>))))
       in flagx(l)"
 
-  Flag_def:
-     "Flag(l,x) == ALL lr:List(Colour).ALL lw:List(Colour).ALL lb:List(Colour).
-                    x = <lr,<lw,lb>> -->
-                  (ALL c:Colour.(c mem lr = true --> c=red) &
-                                (c mem lw = true --> c=white) &
-                                (c mem lb = true --> c=blue)) &
-                  Perm(l,lr @ lw @ lb)"
+axiomatization Perm :: "i => i => o"
+definition Flag :: "i => i => o" where
+  "Flag(l,x) == ALL lr:List(Colour).ALL lw:List(Colour).ALL lb:List(Colour).
+                x = <lr,<lw,lb>> -->
+              (ALL c:Colour.(c mem lr = true --> c=red) &
+                            (c mem lw = true --> c=white) &
+                            (c mem lb = true --> c=blue)) &
+              Perm(l,lr @ lw @ lb)"
 
 
 lemmas flag_defs = Colour_def red_def white_def blue_def ccase_def
@@ -68,7 +69,7 @@
   apply assumption
   done
 
-lemma "flag : PROD l:List(Colour).{x:List(Colour)*List(Colour)*List(Colour).FLAG(x,l)}"
+lemma "flag : PROD l:List(Colour).{x:List(Colour)*List(Colour)*List(Colour).Flag(x,l)}"
   apply (unfold flag_def)
   apply (tactic {* gen_ccs_tac @{context}
     [@{thm redT}, @{thm whiteT}, @{thm blueT}, @{thm ccaseT}] 1 *})
--- a/src/CCL/ex/List.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/ex/List.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -9,46 +9,47 @@
 imports Nat
 begin
 
-consts
-  map       :: "[i=>i,i]=>i"
-  comp      :: "[i=>i,i=>i]=>i=>i"    (infixr "o" 55)
-  append    :: "[i,i]=>i"             (infixr "@" 55)
-  member    :: "[i,i]=>i"             (infixr "mem" 55)
-  filter    :: "[i,i]=>i"
-  flat      :: "i=>i"
-  partition :: "[i,i]=>i"
-  insert    :: "[i,i,i]=>i"
-  isort     :: "i=>i"
-  qsort     :: "i=>i"
+definition map :: "[i=>i,i]=>i"
+  where "map(f,l) == lrec(l,[],%x xs g. f(x)$g)"
+
+definition comp :: "[i=>i,i=>i]=>i=>i"  (infixr "\<circ>" 55)
+  where "f \<circ> g == (%x. f(g(x)))"
+
+definition append :: "[i,i]=>i"  (infixr "@" 55)
+  where "l @ m == lrec(l,m,%x xs g. x$g)"
 
-axioms
+axiomatization member :: "[i,i]=>i"  (infixr "mem" 55)  (* FIXME dangling eq *)
+  where "a mem l == lrec(l,false,%h t g. if eq(a,h) then true else g)"
 
-  map_def:     "map(f,l)   == lrec(l,[],%x xs g. f(x)$g)"
-  comp_def:    "f o g == (%x. f(g(x)))"
-  append_def:  "l @ m == lrec(l,m,%x xs g. x$g)"
-  member_def:  "a mem l == lrec(l,false,%h t g. if eq(a,h) then true else g)"
-  filter_def:  "filter(f,l) == lrec(l,[],%x xs g. if f`x then x$g else g)"
-  flat_def:    "flat(l) == lrec(l,[],%h t g. h @ g)"
+definition filter :: "[i,i]=>i"
+  where "filter(f,l) == lrec(l,[],%x xs g. if f`x then x$g else g)"
 
-  insert_def:  "insert(f,a,l) == lrec(l,a$[],%h t g. if f`a`h then a$h$t else h$g)"
-  isort_def:   "isort(f) == lam l. lrec(l,[],%h t g. insert(f,h,g))"
+definition flat :: "i=>i"
+  where "flat(l) == lrec(l,[],%h t g. h @ g)"
 
-  partition_def:
+definition partition :: "[i,i]=>i" where
   "partition(f,l) == letrec part l a b be lcase(l,<a,b>,%x xs.
                             if f`x then part(xs,x$a,b) else part(xs,a,x$b))
                     in part(l,[],[])"
-  qsort_def:   "qsort(f) == lam l. letrec qsortx l be lcase(l,[],%h t.
+
+definition insert :: "[i,i,i]=>i"
+  where "insert(f,a,l) == lrec(l,a$[],%h t g. if f`a`h then a$h$t else h$g)"
+
+definition isort :: "i=>i"
+  where "isort(f) == lam l. lrec(l,[],%h t g. insert(f,h,g))"
+
+definition qsort :: "i=>i" where
+  "qsort(f) == lam l. letrec qsortx l be lcase(l,[],%h t.
                                    let p be partition(f`h,t)
                                    in split(p,%x y. qsortx(x) @ h$qsortx(y)))
                           in qsortx(l)"
 
-
 lemmas list_defs = map_def comp_def append_def filter_def flat_def
   insert_def isort_def partition_def qsort_def
 
 lemma listBs [simp]:
-  "!!f g. (f o g) = (%a. f(g(a)))"
-  "!!a f g. (f o g)(a) = f(g(a))"
+  "!!f g. (f \<circ> g) = (%a. f(g(a)))"
+  "!!a f g. (f \<circ> g)(a) = f(g(a))"
   "!!f. map(f,[]) = []"
   "!!f x xs. map(f,x$xs) = f(x)$map(f,xs)"
   "!!m. [] @ m = m"
@@ -85,7 +86,7 @@
 
 lemma appendTS:
   "[| l : {l:List(A). m : {m:List(A).P(l @ m)}} |] ==> l @ m : {x:List(A). P(x)}"
-  by (blast intro!: SubtypeI appendT elim!: SubtypeE)
+  by (blast intro!: appendT)
 
 lemma filterT: "[| f:A->Bool;   l : List(A) |] ==> filter(f,l) : List(A)"
   apply (unfold filter_def)
@@ -105,7 +106,7 @@
 lemma insertTS:
   "[| f : {f:A->A->Bool. a : {a:A. l : {l:List(A).P(insert(f,a,l))}}} |] ==>  
    insert(f,a,l)  : {x:List(A). P(x)}"
-  by (blast intro!: SubtypeI insertT elim!: SubtypeE)
+  by (blast intro!: insertT)
 
 lemma partitionT:
   "[| f:A->Bool;  l : List(A) |] ==> partition(f,l) : List(A)*List(A)"
--- a/src/CCL/ex/Nat.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/ex/Nat.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -9,37 +9,44 @@
 imports Wfd
 begin
 
-consts
+definition not :: "i=>i"
+  where "not(b) == if b then false else true"
+
+definition add :: "[i,i]=>i"  (infixr "#+" 60)
+  where "a #+ b == nrec(a,b,%x g. succ(g))"
 
-  not   :: "i=>i"
-  add   :: "[i,i]=>i"            (infixr "#+" 60)
-  mult  :: "[i,i]=>i"            (infixr "#*" 60)
-  sub   :: "[i,i]=>i"            (infixr "#-" 60)
-  div   :: "[i,i]=>i"            (infixr "##" 60)
-  lt    :: "[i,i]=>i"            (infixr "#<" 60)
-  le    :: "[i,i]=>i"            (infixr "#<=" 60)
-  ackermann :: "[i,i]=>i"
+definition mult :: "[i,i]=>i"  (infixr "#*" 60)
+  where "a #* b == nrec(a,zero,%x g. b #+ g)"
 
-defs
-
-  not_def:     "not(b) == if b then false else true"
+definition sub :: "[i,i]=>i"  (infixr "#-" 60)
+  where
+    "a #- b ==
+      letrec sub x y be ncase(y,x,%yy. ncase(x,zero,%xx. sub(xx,yy)))
+      in sub(a,b)"
 
-  add_def:     "a #+ b == nrec(a,b,%x g. succ(g))"
-  mult_def:    "a #* b == nrec(a,zero,%x g. b #+ g)"
-  sub_def:     "a #- b == letrec sub x y be ncase(y,x,%yy. ncase(x,zero,%xx. sub(xx,yy)))
-                        in sub(a,b)"
-  le_def:     "a #<= b == letrec le x y be ncase(x,true,%xx. ncase(y,false,%yy. le(xx,yy)))
-                        in le(a,b)"
-  lt_def:     "a #< b == not(b #<= a)"
+definition le :: "[i,i]=>i"  (infixr "#<=" 60)
+  where
+    "a #<= b ==
+      letrec le x y be ncase(x,true,%xx. ncase(y,false,%yy. le(xx,yy)))
+      in le(a,b)"
+
+definition lt :: "[i,i]=>i"  (infixr "#<" 60)
+  where "a #< b == not(b #<= a)"
 
-  div_def:    "a ## b == letrec div x y be if x #< y then zero else succ(div(x#-y,y))
-                       in div(a,b)"
-  ack_def:
-  "ackermann(a,b) == letrec ack n m be ncase(n,succ(m),%x.
-                          ncase(m,ack(x,succ(zero)),%y. ack(x,ack(succ(x),y))))
-                    in ack(a,b)"
+definition div :: "[i,i]=>i"  (infixr "##" 60)
+  where
+    "a ## b ==
+      letrec div x y be if x #< y then zero else succ(div(x#-y,y))
+      in div(a,b)"
 
-lemmas nat_defs = not_def add_def mult_def sub_def le_def lt_def ack_def napply_def
+definition ackermann :: "[i,i]=>i"
+  where
+    "ackermann(a,b) ==
+      letrec ack n m be ncase(n,succ(m),%x.
+        ncase(m,ack(x,succ(zero)),%y. ack(x,ack(succ(x),y))))
+      in ack(a,b)"
+
+lemmas nat_defs = not_def add_def mult_def sub_def le_def lt_def ackermann_def napply_def
 
 lemma natBs [simp]:
   "not(true) = false"
@@ -94,7 +101,7 @@
 lemmas relI = NatPR_wf [THEN NatPR_wf [THEN lex_wf, THEN wfI]]
 
 lemma "[| a:Nat;  b:Nat |] ==> ackermann(a,b) : Nat"
-  apply (unfold ack_def)
+  apply (unfold ackermann_def)
   apply (tactic {* gen_ccs_tac @{context} [] 1 *})
   apply (erule NatPRI [THEN lexI1 [THEN relI]] NatPRI [THEN lexI2 [THEN relI]])+
   done
--- a/src/CCL/ex/Stream.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/CCL/ex/Stream.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -9,15 +9,11 @@
 imports List
 begin
 
-consts
-  iter1   ::  "[i=>i,i]=>i"
-  iter2   ::  "[i=>i,i]=>i"
+definition iter1 :: "[i=>i,i]=>i"
+  where "iter1(f,a) == letrec iter x be x$iter(f(x)) in iter(a)"
 
-defs
-
-  iter1_def:   "iter1(f,a) == letrec iter x be x$iter(f(x)) in iter(a)"
-  iter2_def:   "iter2(f,a) == letrec iter x be x$map(f,iter(x)) in iter(a)"
-
+definition iter2 :: "[i=>i,i]=>i"
+  where "iter2(f,a) == letrec iter x be x$map(f,iter(x)) in iter(a)"
 
 (*
 Proving properties about infinite lists using coinduction:
@@ -30,9 +26,9 @@
 
 lemma map_comp:
   assumes 1: "l:Lists(A)"
-  shows "map(f o g,l) = map(f,map(g,l))"
+  shows "map(f \<circ> g,l) = map(f,map(g,l))"
   apply (tactic {* eq_coinduct3_tac @{context}
-    "{p. EX x y. p=<x,y> & (EX l:Lists (A) .x=map (f o g,l) & y=map (f,map (g,l)))}" 1 *})
+    "{p. EX x y. p=<x,y> & (EX l:Lists (A) .x=map (f \<circ> g,l) & y=map (f,map (g,l)))}" 1 *})
    apply (blast intro: 1)
   apply safe
   apply (drule ListsXH [THEN iffD1])
--- a/src/HOL/DSequence.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/DSequence.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -7,7 +7,7 @@
 imports Lazy_Sequence Code_Numeral
 begin
 
-types 'a dseq = "code_numeral => bool => 'a Lazy_Sequence.lazy_sequence option"
+type_synonym 'a dseq = "code_numeral => bool => 'a Lazy_Sequence.lazy_sequence option"
 
 definition empty :: "'a dseq"
 where
--- a/src/HOL/Datatype.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Datatype.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -28,8 +28,8 @@
 
 text{*Datatypes will be represented by sets of type @{text node}*}
 
-types 'a item        = "('a, unit) node set"
-      ('a, 'b) dtree = "('a, 'b) node set"
+type_synonym 'a item        = "('a, unit) node set"
+type_synonym ('a, 'b) dtree = "('a, 'b) node set"
 
 consts
   Push      :: "[('b + nat), nat => ('b + nat)] => (nat => ('b + nat))"
--- a/src/HOL/Finite_Set.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Finite_Set.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -246,24 +246,20 @@
   by (simp add: image_Collect [symmetric])
 
 lemma finite_imageD:
-  "finite (f ` A) \<Longrightarrow> inj_on f A \<Longrightarrow> finite A"
-proof -
-  have aux: "!!A. finite (A - {}) = finite A" by simp
-  fix B :: "'a set"
-  assume "finite B"
-  thus "!!A. f`A = B ==> inj_on f A ==> finite A"
-    apply induct
-     apply simp
-    apply (subgoal_tac "EX y:A. f y = x & F = f ` (A - {y})")
-     apply clarify
-     apply (simp (no_asm_use) add: inj_on_def)
-     apply (blast dest!: aux [THEN iffD1], atomize)
-    apply (erule_tac V = "ALL A. ?PP (A)" in thin_rl)
-    apply (frule subsetD [OF equalityD2 insertI1], clarify)
-    apply (rule_tac x = xa in bexI)
-     apply (simp_all add: inj_on_image_set_diff)
-    done
-qed (rule refl)
+  assumes "finite (f ` A)" and "inj_on f A"
+  shows "finite A"
+using assms proof (induct "f ` A" arbitrary: A)
+  case empty then show ?case by simp
+next
+  case (insert x B)
+  then have B_A: "insert x B = f ` A" by simp
+  then obtain y where "x = f y" and "y \<in> A" by blast
+  from B_A `x \<notin> B` have "B = f ` A - {x}" by blast
+  with B_A `x \<notin> B` `x = f y` `inj_on f A` `y \<in> A` have "B = f ` (A - {y})" by (simp add: inj_on_image_set_diff)
+  moreover from `inj_on f A` have "inj_on f (A - {y})" by (rule inj_on_diff)
+  ultimately have "finite (A - {y})" by (rule insert.hyps)
+  then show "finite A" by simp
+qed
 
 lemma finite_surj:
   "finite A \<Longrightarrow> B \<subseteq> f ` A \<Longrightarrow> finite B"
@@ -359,34 +355,32 @@
   by (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product)
 
 lemma finite_cartesian_productD1:
-  "finite (A \<times> B) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> finite A"
-apply (auto simp add: finite_conv_nat_seg_image) 
-apply (drule_tac x=n in spec) 
-apply (drule_tac x="fst o f" in spec) 
-apply (auto simp add: o_def) 
- prefer 2 apply (force dest!: equalityD2) 
-apply (drule equalityD1) 
-apply (rename_tac y x)
-apply (subgoal_tac "\<exists>k. k<n & f k = (x,y)") 
- prefer 2 apply force
-apply clarify
-apply (rule_tac x=k in image_eqI, auto)
-done
+  assumes "finite (A \<times> B)" and "B \<noteq> {}"
+  shows "finite A"
+proof -
+  from assms obtain n f where "A \<times> B = f ` {i::nat. i < n}"
+    by (auto simp add: finite_conv_nat_seg_image)
+  then have "fst ` (A \<times> B) = fst ` f ` {i::nat. i < n}" by simp
+  with `B \<noteq> {}` have "A = (fst \<circ> f) ` {i::nat. i < n}"
+    by (simp add: image_compose)
+  then have "\<exists>n f. A = f ` {i::nat. i < n}" by blast
+  then show ?thesis
+    by (auto simp add: finite_conv_nat_seg_image)
+qed
 
 lemma finite_cartesian_productD2:
-     "[| finite (A <*> B); A \<noteq> {} |] ==> finite B"
-apply (auto simp add: finite_conv_nat_seg_image) 
-apply (drule_tac x=n in spec) 
-apply (drule_tac x="snd o f" in spec) 
-apply (auto simp add: o_def) 
- prefer 2 apply (force dest!: equalityD2) 
-apply (drule equalityD1)
-apply (rename_tac x y)
-apply (subgoal_tac "\<exists>k. k<n & f k = (x,y)") 
- prefer 2 apply force
-apply clarify
-apply (rule_tac x=k in image_eqI, auto)
-done
+  assumes "finite (A \<times> B)" and "A \<noteq> {}"
+  shows "finite B"
+proof -
+  from assms obtain n f where "A \<times> B = f ` {i::nat. i < n}"
+    by (auto simp add: finite_conv_nat_seg_image)
+  then have "snd ` (A \<times> B) = snd ` f ` {i::nat. i < n}" by simp
+  with `A \<noteq> {}` have "B = (snd \<circ> f) ` {i::nat. i < n}"
+    by (simp add: image_compose)
+  then have "\<exists>n f. B = f ` {i::nat. i < n}" by blast
+  then show ?thesis
+    by (auto simp add: finite_conv_nat_seg_image)
+qed
 
 lemma finite_Pow_iff [iff]:
   "finite (Pow A) \<longleftrightarrow> finite A"
--- a/src/HOL/HOL.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOL.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1923,7 +1923,7 @@
     and infixl 3 "&&" and infixl 2 "||"
     and "!(if (_)/ then (_)/ else (_))")
   (Haskell "True" and "False" and "not"
-    and infixl 3 "&&" and infixl 2 "||"
+    and infixr 3 "&&" and infixr 2 "||"
     and "!(if (_)/ then (_)/ else (_))")
   (Scala "true" and "false" and "'! _"
     and infixl 3 "&&" and infixl 1 "||"
--- a/src/HOL/HOLCF/Bifinite.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Bifinite.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Bifinite.thy
+(*  Title:      HOL/HOLCF/Bifinite.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Cfun.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Cfun.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Cfun.thy
+(*  Title:      HOL/HOLCF/Cfun.thy
     Author:     Franz Regensburger
     Author:     Brian Huffman
 *)
--- a/src/HOL/HOLCF/Compact_Basis.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Compact_Basis.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Compact_Basis.thy
+(*  Title:      HOL/HOLCF/Compact_Basis.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Completion.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Completion.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Completion.thy
+(*  Title:      HOL/HOLCF/Completion.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Cont.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Cont.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Cont.thy
+(*  Title:      HOL/HOLCF/Cont.thy
     Author:     Franz Regensburger
     Author:     Brian Huffman
 *)
--- a/src/HOL/HOLCF/ConvexPD.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ConvexPD.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ConvexPD.thy
+(*  Title:      HOL/HOLCF/ConvexPD.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Cpodef.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Cpodef.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Pcpodef.thy
+(*  Title:      HOL/HOLCF/Cpodef.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Cprod.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Cprod.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Cprod.thy
+(*  Title:      HOL/HOLCF/Cprod.thy
     Author:     Franz Regensburger
 *)
 
--- a/src/HOL/HOLCF/Deflation.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Deflation.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Deflation.thy
+(*  Title:      HOL/HOLCF/Deflation.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Discrete.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Discrete.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Discrete.thy
+(*  Title:      HOL/HOLCF/Discrete.thy
     Author:     Tobias Nipkow
 *)
 
--- a/src/HOL/HOLCF/Domain.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Domain.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Domain.thy
+(*  Title:      HOL/HOLCF/Domain.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Domain_Aux.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Domain_Aux.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Domain_Aux.thy
+(*  Title:      HOL/HOLCF/Domain_Aux.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/FOCUS/Buffer.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/FOCUS/Buffer.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/FOCUS/Buffer.thy
+(*  Title:      HOL/HOLCF/FOCUS/Buffer.thy
     Author:     David von Oheimb, TU Muenchen
 
 Formalization of section 4 of
--- a/src/HOL/HOLCF/FOCUS/Buffer_adm.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/FOCUS/Buffer_adm.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/FOCUS/Buffer_adm.thy
+(*  Title:      HOL/HOLCF/FOCUS/Buffer_adm.thy
     Author:     David von Oheimb, TU Muenchen
 *)
 
--- a/src/HOL/HOLCF/FOCUS/FOCUS.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/FOCUS/FOCUS.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/FOCUS/FOCUS.thy
+(*  Title:      HOL/HOLCF/FOCUS/FOCUS.thy
     Author:     David von Oheimb, TU Muenchen
 *)
 
--- a/src/HOL/HOLCF/FOCUS/Fstream.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/FOCUS/Fstream.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/FOCUS/Fstream.thy
+(*  Title:      HOL/HOLCF/FOCUS/Fstream.thy
     Author:     David von Oheimb, TU Muenchen
 
 FOCUS streams (with lifted elements).
--- a/src/HOL/HOLCF/FOCUS/Fstreams.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/FOCUS/Fstreams.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/FOCUS/Fstreams.thy
+(*  Title:      HOL/HOLCF/FOCUS/Fstreams.thy
     Author:     Borislav Gajanovic
 
 FOCUS flat streams (with lifted elements).
--- a/src/HOL/HOLCF/FOCUS/Stream_adm.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/FOCUS/Stream_adm.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Stream_adm.thy
+(*  Title:      HOL/HOLCF/FOCUS/Stream_adm.thy
     Author:     David von Oheimb, TU Muenchen
 *)
 
--- a/src/HOL/HOLCF/Fix.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Fix.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Fix.thy
+(*  Title:      HOL/HOLCF/Fix.thy
     Author:     Franz Regensburger
     Author:     Brian Huffman
 *)
--- a/src/HOL/HOLCF/Fixrec.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Fixrec.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Fixrec.thy
+(*  Title:      HOL/HOLCF/Fixrec.thy
     Author:     Amber Telfer and Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Fun_Cpo.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Fun_Cpo.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Fun_Cpo.thy
+(*  Title:      HOL/HOLCF/Fun_Cpo.thy
     Author:     Franz Regensburger
     Author:     Brian Huffman
 *)
--- a/src/HOL/HOLCF/HOLCF.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/HOLCF.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/HOLCF.thy
+(*  Title:      HOL/HOLCF/HOLCF.thy
     Author:     Franz Regensburger
 
 HOLCF -- a semantic extension of HOL by the LCF logic.
--- a/src/HOL/HOLCF/IMP/Denotational.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IMP/Denotational.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IMP/Denotational.thy
+(*  Title:      HOL/HOLCF/IMP/Denotational.thy
     Author:     Tobias Nipkow and Robert Sandner, TUM
     Copyright   1996 TUM
 *)
--- a/src/HOL/HOLCF/IMP/HoareEx.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IMP/HoareEx.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IMP/HoareEx.thy
+(*  Title:      HOL/HOLCF/IMP/HoareEx.thy
     Author:     Tobias Nipkow, TUM
     Copyright   1997 TUM
 *)
--- a/src/HOL/HOLCF/IOA/ABP/Abschannel.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Abschannel.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Abschannel.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Abschannel.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/Abschannel_finite.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Abschannel_finite.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Abschannels.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Abschannel_finite.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/Action.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Action.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Action.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Action.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/Check.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Check.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Check.ML
+(*  Title:      HOL/HOLCF/IOA/ABP/Check.ML
     Author:     Olaf Mueller
 
 The Model Checker.
--- a/src/HOL/HOLCF/IOA/ABP/Correctness.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Correctness.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Correctness.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Correctness.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/Env.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Env.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Impl.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Env.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/Impl.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Impl.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Impl.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Impl.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/Impl_finite.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Impl_finite.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Impl.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Impl_finite.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/Lemmas.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Lemmas.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Lemmas.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Lemmas.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/Packet.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Packet.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Packet.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Packet.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/ROOT.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/ROOT.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,7 +1,8 @@
-(*  Title:      HOLCF/IOA/ABP/ROOT.ML
+(*  Title:      HOL/HOLCF/IOA/ABP/ROOT.ML
     Author:     Olaf Mueller
 
 This is the ROOT file for the Alternating Bit Protocol performed in
 I/O-Automata.
 *)
+
 use_thys ["Correctness"];
--- a/src/HOL/HOLCF/IOA/ABP/Receiver.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Receiver.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Receiver.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Receiver.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/Sender.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Sender.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Sender.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Sender.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ABP/Spec.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Spec.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Spec.thy
+(*  Title:      HOL/HOLCF/IOA/ABP/Spec.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/NTP/Abschannel.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Abschannel.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/NTP/Abschannel.thy
+(*  Title:      HOL/HOLCF/IOA/NTP/Abschannel.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/NTP/Action.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Action.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/NTP/Action.thy
+(*  Title:      HOL/HOLCF/IOA/NTP/Action.thy
     Author:     Tobias Nipkow & Konrad Slind
 *)
 
--- a/src/HOL/HOLCF/IOA/NTP/Correctness.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Correctness.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/NTP/Correctness.thy
+(*  Title:      HOL/HOLCF/IOA/NTP/Correctness.thy
     Author:     Tobias Nipkow & Konrad Slind
 *)
 
--- a/src/HOL/HOLCF/IOA/NTP/Impl.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Impl.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/NTP/Impl.thy
+(*  Title:      HOL/HOLCF/IOA/NTP/Impl.thy
     Author:     Tobias Nipkow & Konrad Slind
 *)
 
--- a/src/HOL/HOLCF/IOA/NTP/Lemmas.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Lemmas.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/NTP/Lemmas.thy
+(*  Title:      HOL/HOLCF/IOA/NTP/Lemmas.thy
     Author:     Tobias Nipkow & Konrad Slind
 *)
 
--- a/src/HOL/HOLCF/IOA/NTP/Multiset.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Multiset.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/NTP/Multiset.thy
+(*  Title:      HOL/HOLCF/IOA/NTP/Multiset.thy
     Author:     Tobias Nipkow & Konrad Slind
 *)
 
--- a/src/HOL/HOLCF/IOA/NTP/Packet.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Packet.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/NTP/Packet.thy
+(*  Title:      HOL/HOLCF/IOA/NTP/Packet.thy
     Author:     Tobias Nipkow & Konrad Slind
 *)
 
--- a/src/HOL/HOLCF/IOA/NTP/ROOT.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/ROOT.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/NTP/ROOT.ML
+(*  Title:      HOL/HOLCF/IOA/NTP/ROOT.ML
     Author:     Tobias Nipkow & Konrad Slind
 
 This is the ROOT file for a network transmission protocol (NTP
--- a/src/HOL/HOLCF/IOA/NTP/Receiver.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Receiver.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/NTP/Receiver.thy
+(*  Title:      HOL/HOLCF/IOA/NTP/Receiver.thy
     Author:     Tobias Nipkow & Konrad Slind
 *)
 
--- a/src/HOL/HOLCF/IOA/NTP/Sender.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Sender.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/NTP/Sender.thy
+(*  Title:      HOL/HOLCF/IOA/NTP/Sender.thy
     Author:     Tobias Nipkow & Konrad Slind
 *)
 
--- a/src/HOL/HOLCF/IOA/NTP/Spec.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Spec.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/NTP/Spec.thy
+(*  Title:      HOL/HOLCF/IOA/NTP/Spec.thy
     Author:     Tobias Nipkow & Konrad Slind
 *)
 
--- a/src/HOL/HOLCF/IOA/ROOT.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ROOT.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ROOT.ML
+(*  Title:      HOL/HOLCF/IOA/ROOT.ML
     Author:     Olaf Mueller
 
 Formalization of a semantic model of I/O-Automata.  See README.html
--- a/src/HOL/HOLCF/IOA/Storage/Action.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/Storage/Action.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ABP/Action.thy
+(*  Title:      HOL/HOLCF/IOA/Storage/Action.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/Storage/Correctness.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/Storage/Correctness.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/example/Correctness.thy
+(*  Title:      HOL/HOLCF/IOA/Storage/Correctness.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/Storage/Impl.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/Storage/Impl.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/example/Spec.thy
+(*  Title:      HOL/HOLCF/IOA/Storage/Impl.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/Storage/ROOT.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/Storage/ROOT.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/Storage/ROOT.ML
+(*  Title:      HOL/HOLCF/IOA/Storage/ROOT.ML
     Author:     Olaf Mueller
 
 Memory storage case study.
--- a/src/HOL/HOLCF/IOA/Storage/Spec.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/Storage/Spec.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/example/Spec.thy
+(*  Title:      HOL/HOLCF/IOA/Storage/Spec.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ex/ROOT.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ex/ROOT.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/ex/ROOT.ML
+(*  Title:      HOL/HOLCF/IOA/ex/ROOT.ML
     Author:     Olaf Mueller
 *)
 
--- a/src/HOL/HOLCF/IOA/ex/TrivEx.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ex/TrivEx.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/TrivEx.thy
+(*  Title:      HOL/HOLCF/IOA/ex/TrivEx.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/ex/TrivEx2.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/ex/TrivEx2.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/TrivEx.thy
+(*  Title:      HOL/HOLCF/IOA/ex/TrivEx2.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Abstraction.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Abstraction.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/Abstraction.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/Abstraction.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Asig.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Asig.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/IOA/meta_theory/Asig.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/Asig.thy
     Author:     Olaf Müller, Tobias Nipkow & Konrad Slind
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Automata.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Automata.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/Automata.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/Automata.thy
     Author:     Olaf Müller, Konrad Slind, Tobias Nipkow
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/CompoExecs.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/CompoExecs.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/CompoExecs.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/CompoExecs.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/CompoScheds.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/CompoScheds.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/CompoScheds.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/CompoScheds.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/CompoTraces.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/CompoTraces.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/CompoTraces.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/CompoTraces.thy
     Author:     Olaf Müller
 *) 
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Compositionality.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Compositionality.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/Compositionality.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/Compositionality.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Deadlock.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Deadlock.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/Deadlock.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/Deadlock.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/IOA.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/IOA.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/IOA.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/IOA.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/LiveIOA.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/LiveIOA.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/LiveIOA.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/LiveIOA.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Pred.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Pred.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/Pred.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/Pred.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/RefCorrectness.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/RefCorrectness.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/RefCorrectness.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/RefCorrectness.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/RefMappings.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/RefMappings.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/RefMappings.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/RefMappings.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Seq.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Seq.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/Seq.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/Seq.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Sequence.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Sequence.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/Sequence.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/Sequence.thy
     Author:     Olaf Müller
 
 Sequences over flat domains with lifted elements.
--- a/src/HOL/HOLCF/IOA/meta_theory/ShortExecutions.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/ShortExecutions.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/ShortExecutions.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/ShortExecutions.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/SimCorrectness.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/SimCorrectness.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/SimCorrectness.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/SimCorrectness.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Simulations.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Simulations.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/Simulations.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/Simulations.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/TL.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/TL.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/TLS.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/TL.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/TLS.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/TLS.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/TLS.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/TLS.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Traces.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Traces.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/IOA/meta_theory/Traces.thy
+(*  Title:      HOL/HOLCF/IOA/meta_theory/Traces.thy
     Author:     Olaf Müller
 *)
 
--- a/src/HOL/HOLCF/Library/Bool_Discrete.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/Bool_Discrete.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Library/Bool_Discrete.thy
+(*  Title:      HOL/HOLCF/Library/Bool_Discrete.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Library/Char_Discrete.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/Char_Discrete.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Library/Char_Discrete.thy
+(*  Title:      HOL/HOLCF/Library/Char_Discrete.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Library/Defl_Bifinite.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/Defl_Bifinite.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Library/Defl_Bifinite.thy
+(*  Title:      HOL/HOLCF/Library/Defl_Bifinite.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Library/HOL_Cpo.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/HOL_Cpo.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/HOL_Cpo.thy
+(*  Title:      HOL/HOLCF/Library/HOL_Cpo.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Library/Int_Discrete.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/Int_Discrete.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Library/Int_Discrete.thy
+(*  Title:      HOL/HOLCF/Library/Int_Discrete.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Library/List_Cpo.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/List_Cpo.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Library/List_Cpo.thy
+(*  Title:      HOL/HOLCF/Library/List_Cpo.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Library/List_Predomain.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/List_Predomain.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Library/List_Predomain.thy
+(*  Title:      HOL/HOLCF/Library/List_Predomain.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Library/Nat_Discrete.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/Nat_Discrete.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Library/Nat_Discrete.thy
+(*  Title:      HOL/HOLCF/Library/Nat_Discrete.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Library/Option_Cpo.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/Option_Cpo.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Option_Cpo.thy
+(*  Title:      HOL/HOLCF/Library/Option_Cpo.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Library/Stream.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/Stream.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Stream.thy
+(*  Title:      HOL/HOLCF/Library/Stream.thy
     Author:     Franz Regensburger, David von Oheimb, Borislav Gajanovic
 *)
 
--- a/src/HOL/HOLCF/Library/Sum_Cpo.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Library/Sum_Cpo.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Sum_Cpo.thy
+(*  Title:      HOL/HOLCF/Library/Sum_Cpo.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Lift.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Lift.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Lift.thy
+(*  Title:      HOL/HOLCF/Lift.thy
     Author:     Olaf Mueller
 *)
 
--- a/src/HOL/HOLCF/LowerPD.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/LowerPD.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/LowerPD.thy
+(*  Title:      HOL/HOLCF/LowerPD.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Map_Functions.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Map_Functions.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Map_Functions.thy
+(*  Title:      HOL/HOLCF/Map_Functions.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/One.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/One.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/One.thy
+(*  Title:      HOL/HOLCF/One.thy
     Author:     Oscar Slotosch
 *)
 
--- a/src/HOL/HOLCF/Pcpo.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Pcpo.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Pcpo.thy
+(*  Title:      HOL/HOLCF/Pcpo.thy
     Author:     Franz Regensburger
 *)
 
--- a/src/HOL/HOLCF/Plain_HOLCF.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Plain_HOLCF.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Plain_HOLCF.thy
+(*  Title:      HOL/HOLCF/Plain_HOLCF.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Porder.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Porder.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Porder.thy
+(*  Title:      HOL/HOLCF/Porder.thy
     Author:     Franz Regensburger and Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Powerdomains.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Powerdomains.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Powerdomains.thy
+(*  Title:      HOL/HOLCF/Powerdomains.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Product_Cpo.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Product_Cpo.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Product_Cpo.thy
+(*  Title:      HOL/HOLCF/Product_Cpo.thy
     Author:     Franz Regensburger
 *)
 
--- a/src/HOL/HOLCF/ROOT.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ROOT.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ROOT.ML
+(*  Title:      HOL/HOLCF/ROOT.ML
     Author:     Franz Regensburger
 
 HOLCF -- a semantic extension of HOL by the LCF logic.
--- a/src/HOL/HOLCF/Representable.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Representable.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Representable.thy
+(*  Title:      HOL/HOLCF/Representable.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Sfun.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Sfun.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Sfun.thy
+(*  Title:      HOL/HOLCF/Sfun.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Sprod.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Sprod.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Sprod.thy
+(*  Title:      HOL/HOLCF/Sprod.thy
     Author:     Franz Regensburger
     Author:     Brian Huffman
 *)
--- a/src/HOL/HOLCF/Ssum.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Ssum.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Ssum.thy
+(*  Title:      HOL/HOLCF/Ssum.thy
     Author:     Franz Regensburger
     Author:     Brian Huffman
 *)
--- a/src/HOL/HOLCF/Tools/Domain/domain.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/Domain/domain.ML
+(*  Title:      HOL/HOLCF/Tools/Domain/domain.ML
     Author:     David von Oheimb
     Author:     Brian Huffman
 
--- a/src/HOL/HOLCF/Tools/Domain/domain_axioms.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain_axioms.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/Domain/domain_axioms.ML
+(*  Title:      HOL/HOLCF/Tools/Domain/domain_axioms.ML
     Author:     David von Oheimb
     Author:     Brian Huffman
 
--- a/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/Domain/domain_constructors.ML
+(*  Title:      HOL/HOLCF/Tools/Domain/domain_constructors.ML
     Author:     Brian Huffman
 
 Defines constructor functions for a given domain isomorphism
--- a/src/HOL/HOLCF/Tools/Domain/domain_induction.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain_induction.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/Domain/domain_induction.ML
+(*  Title:      HOL/HOLCF/Tools/Domain/domain_induction.ML
     Author:     David von Oheimb
     Author:     Brian Huffman
 
--- a/src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/Domain/domain_isomorphism.ML
+(*  Title:      HOL/HOLCF/Tools/Domain/domain_isomorphism.ML
     Author:     Brian Huffman
 
 Defines new types satisfying the given domain equations.
--- a/src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/Domain/domain_take_proofs.ML
+(*  Title:      HOL/HOLCF/Tools/Domain/domain_take_proofs.ML
     Author:     Brian Huffman
 
 Defines take functions for the given domain equation
--- a/src/HOL/HOLCF/Tools/cont_consts.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/cont_consts.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/cont_consts.ML
+(*  Title:      HOL/HOLCF/Tools/cont_consts.ML
     Author:     Tobias Mayr, David von Oheimb, and Markus Wenzel
 
 HOLCF version of consts: handle continuous function types in mixfix
--- a/src/HOL/HOLCF/Tools/cont_proc.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/cont_proc.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/cont_proc.ML
+(*  Title:      HOL/HOLCF/Tools/cont_proc.ML
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Tools/cpodef.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/cpodef.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/cpodef.ML
+(*  Title:      HOL/HOLCF/Tools/cpodef.ML
     Author:     Brian Huffman
 
 Primitive domain definitions for HOLCF, similar to Gordon/HOL-style
--- a/src/HOL/HOLCF/Tools/domaindef.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/domaindef.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/repdef.ML
+(*  Title:      HOL/HOLCF/Tools/domaindef.ML
     Author:     Brian Huffman
 
 Defining representable domains using algebraic deflations.
--- a/src/HOL/HOLCF/Tools/fixrec.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/fixrec.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/fixrec.ML
+(*  Title:      HOL/HOLCF/Tools/fixrec.ML
     Author:     Amber Telfer and Brian Huffman
 
 Recursive function definition package for HOLCF.
--- a/src/HOL/HOLCF/Tools/holcf_library.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tools/holcf_library.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tools/holcf_library.ML
+(*  Title:      HOL/HOLCF/Tools/holcf_library.ML
     Author:     Brian Huffman
 
 Functions for constructing HOLCF types and terms.
--- a/src/HOL/HOLCF/Tr.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tr.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Tr.thy
+(*  Title:      HOL/HOLCF/Tr.thy
     Author:     Franz Regensburger
 *)
 
--- a/src/HOL/HOLCF/Tutorial/Domain_ex.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tutorial/Domain_ex.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Domain_ex.thy
+(*  Title:      HOL/HOLCF/Tutorial/Domain_ex.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Tutorial/Fixrec_ex.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tutorial/Fixrec_ex.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Fixrec_ex.thy
+(*  Title:      HOL/HOLCF/Tutorial/Fixrec_ex.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Tutorial/New_Domain.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Tutorial/New_Domain.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/New_Domain.thy
+(*  Title:      HOL/HOLCF/Tutorial/New_Domain.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Universal.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Universal.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Universal.thy
+(*  Title:      HOL/HOLCF/Universal.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/Up.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/Up.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Up.thy
+(*  Title:      HOL/HOLCF/Up.thy
     Author:     Franz Regensburger
     Author:     Brian Huffman
 *)
--- a/src/HOL/HOLCF/UpperPD.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/UpperPD.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/UpperPD.thy
+(*  Title:      HOL/HOLCF/UpperPD.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/ex/Dnat.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ex/Dnat.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/Dnat.thy
+(*  Title:      HOL/HOLCF/ex/Dnat.thy
     Author:     Franz Regensburger
 
 Theory for the domain of natural numbers  dnat = one ++ dnat
--- a/src/HOL/HOLCF/ex/Domain_Proofs.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ex/Domain_Proofs.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Domain_Proofs.thy
+(*  Title:      HOL/HOLCF/ex/Domain_Proofs.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/ex/Fix2.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ex/Fix2.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Fix2.thy
+(*  Title:      HOL/HOLCF/ex/Fix2.thy
     Author:     Franz Regensburger
 
 Show that fix is the unique least fixed-point operator.
--- a/src/HOL/HOLCF/ex/Hoare.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ex/Hoare.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/hoare.thy
+(*  Title:      HOL/HOLCF/ex/Hoare.thy
     Author:     Franz Regensburger
 
 Theory for an example by C.A.R. Hoare
--- a/src/HOL/HOLCF/ex/Letrec.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ex/Letrec.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Letrec.thy
+(*  Title:      HOL/HOLCF/ex/Letrec.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/ex/Loop.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ex/Loop.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Loop.thy
+(*  Title:      HOL/HOLCF/ex/Loop.thy
     Author:     Franz Regensburger
 *)
 
--- a/src/HOL/HOLCF/ex/Pattern_Match.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ex/Pattern_Match.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Pattern_Match.thy
+(*  Title:      HOL/HOLCF/ex/Pattern_Match.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/ex/Powerdomain_ex.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ex/Powerdomain_ex.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Powerdomain_ex.thy
+(*  Title:      HOL/HOLCF/ex/Powerdomain_ex.thy
     Author:     Brian Huffman
 *)
 
--- a/src/HOL/HOLCF/ex/ROOT.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/HOLCF/ex/ROOT.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/ROOT.ML
+(*  Title:      HOL/HOLCF/ex/ROOT.ML
 
 Misc HOLCF examples.
 *)
--- a/src/HOL/Hoare/Examples.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare/Examples.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -90,7 +90,7 @@
  {c = A^B}"
 apply vcg_simp
 apply(case_tac "b")
- apply(simp add: mod_less)
+ apply simp
 apply simp
 done
 
--- a/src/HOL/Hoare/ExamplesAbort.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare/ExamplesAbort.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -14,8 +14,7 @@
 lemma
  "VARS a i j
  {k <= length a & i < k & j < k} j < length a \<rightarrow> a[i] := a!j {True}"
-apply vcg_simp
-done
+by vcg_simp
 
 lemma "VARS (a::int list) i
  {True}
@@ -24,7 +23,6 @@
  INV {i <= length a}
  DO a[i] := 7; i := i+1 OD
  {True}"
-apply vcg_simp
-done
+by vcg_simp
 
 end
--- a/src/HOL/Hoare/Hoare_Logic.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare/Hoare_Logic.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -10,12 +10,11 @@
 
 theory Hoare_Logic
 imports Main
-uses ("hoare_tac.ML")
+uses ("hoare_syntax.ML") ("hoare_tac.ML")
 begin
 
-types
-    'a bexp = "'a set"
-    'a assn = "'a set"
+type_synonym 'a bexp = "'a set"
+type_synonym 'a assn = "'a set"
 
 datatype
  'a com = Basic "'a \<Rightarrow> 'a"
@@ -25,7 +24,7 @@
 
 abbreviation annskip ("SKIP") where "SKIP == Basic id"
 
-types 'a sem = "'a => 'a => bool"
+type_synonym 'a sem = "'a => 'a => bool"
 
 inductive Sem :: "'a com \<Rightarrow> 'a sem"
 where
@@ -45,11 +44,8 @@
   where "Valid p c q \<longleftrightarrow> (!s s'. Sem c s s' --> s : p --> s' : q)"
 
 
-
-(** parse translations **)
-
 syntax
-  "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
+  "_assign" :: "idt => 'b => 'a com"  ("(2_ :=/ _)" [70, 65] 61)
 
 syntax
  "_hoare_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
@@ -57,147 +53,11 @@
 syntax ("" output)
  "_hoare"      :: "['a assn,'a com,'a assn] => bool"
                  ("{_} // _ // {_}" [0,55,0] 50)
-ML {*
-
-local
-
-fun abs((a,T),body) =
-  let val a = absfree(a, dummyT, body)
-  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
-in
-
-fun mk_abstuple [x] body = abs (x, body)
-  | mk_abstuple (x::xs) body =
-      Syntax.const @{const_syntax prod_case} $ abs (x, mk_abstuple xs body);
-
-fun mk_fbody a e [x as (b,_)] = if a=b then e else Syntax.free b
-  | mk_fbody a e ((b,_)::xs) =
-      Syntax.const @{const_syntax Pair} $ (if a=b then e else Syntax.free b) $ mk_fbody a e xs;
-
-fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
-end
-*}
-
-(* bexp_tr & assn_tr *)
-(*all meta-variables for bexp except for TRUE are translated as if they
-  were boolean expressions*)
-ML{*
-fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
-  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
-
-fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
-*}
-(* com_tr *)
-ML{*
-fun com_tr (t as (Const(@{syntax_const "_assign"},_) $ x $ e)) xs =
-      (case Syntax.strip_positions x of
-        Free (a, _) => Syntax.const @{const_syntax Basic} $ mk_fexp a e xs
-      | _ => t)
-  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
-  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
-      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
-  | com_tr t _ = t (* if t is just a Free/Var *)
-*}
-
-(* triple_tr *)    (* FIXME does not handle "_idtdummy" *)
-ML{*
-local
-
-fun var_tr (Free (a, _)) = (a, Bound 0) (* Bound 0 = dummy term *)
-  | var_tr (Const (@{syntax_const "_constrain"}, _) $ Free (a, _) $ T) = (a, T);
-
-fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) =
-      var_tr (Syntax.strip_positions idt) :: vars_tr vars
-  | vars_tr t = [var_tr (Syntax.strip_positions t)]
-
-in
-fun hoare_vars_tr [vars, pre, prg, post] =
-      let val xs = vars_tr vars
-      in Syntax.const @{const_syntax Valid} $
-         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
-      end
-  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
-end
-*}
-
-parse_translation {* [(@{syntax_const "_hoare_vars"}, hoare_vars_tr)] *}
-
 
-(*****************************************************************************)
-
-(*** print translations ***)
-ML{*
-fun dest_abstuple (Const (@{const_syntax prod_case},_) $ (Abs(v,_, body))) =
-                            subst_bound (Syntax.free v, dest_abstuple body)
-  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
-  | dest_abstuple trm = trm;
-
-fun abs2list (Const (@{const_syntax prod_case},_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
-  | abs2list (Abs(x,T,t)) = [Free (x, T)]
-  | abs2list _ = [];
-
-fun mk_ts (Const (@{const_syntax prod_case},_) $ (Abs(x,_,t))) = mk_ts t
-  | mk_ts (Abs(x,_,t)) = mk_ts t
-  | mk_ts (Const (@{const_syntax Pair},_) $ a $ b) = a::(mk_ts b)
-  | mk_ts t = [t];
-
-fun mk_vts (Const (@{const_syntax prod_case},_) $ (Abs(x,_,t))) =
-           ((Syntax.free x)::(abs2list t), mk_ts t)
-  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
-  | mk_vts t = raise Match;
-
-fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))
-  | find_ch ((v,t)::vts) i xs =
-      if t = Bound i then find_ch vts (i-1) xs
-      else (true, (v, subst_bounds (xs, t)));
-
-fun is_f (Const (@{const_syntax prod_case},_) $ (Abs(x,_,t))) = true
-  | is_f (Abs(x,_,t)) = true
-  | is_f t = false;
-*}
+use "hoare_syntax.ML"
+parse_translation {* [(@{syntax_const "_hoare_vars"}, Hoare_Syntax.hoare_vars_tr)] *}
+print_translation {* [(@{const_syntax Valid}, Hoare_Syntax.spec_tr' @{syntax_const "_hoare"})] *}
 
-(* assn_tr' & bexp_tr'*)
-ML{*
-fun assn_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | assn_tr' (Const (@{const_syntax inter}, _) $
-        (Const (@{const_syntax Collect},_) $ T1) $ (Const (@{const_syntax Collect},_) $ T2)) =
-      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
-  | assn_tr' t = t;
-
-fun bexp_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | bexp_tr' t = t;
-*}
-
-(*com_tr' *)
-ML{*
-fun mk_assign f =
-  let val (vs, ts) = mk_vts f;
-      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
-  in
-    if ch then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
-    else Syntax.const @{const_syntax annskip}
-  end;
-
-fun com_tr' (Const (@{const_syntax Basic},_) $ f) =
-      if is_f f then mk_assign f
-      else Syntax.const @{const_syntax Basic} $ f
-  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
-      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) =
-      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax While},_) $ b $ I $ c) =
-      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
-  | com_tr' t = t;
-
-fun spec_tr' [p, c, q] =
-  Syntax.const @{syntax_const "_hoare"} $ assn_tr' p $ com_tr' c $ assn_tr' q
-*}
-
-print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
 
 lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
 by (auto simp:Valid_def)
--- a/src/HOL/Hoare/Hoare_Logic_Abort.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -7,12 +7,11 @@
 
 theory Hoare_Logic_Abort
 imports Main
-uses ("hoare_tac.ML")
+uses ("hoare_syntax.ML") ("hoare_tac.ML")
 begin
 
-types
-    'a bexp = "'a set"
-    'a assn = "'a set"
+type_synonym 'a bexp = "'a set"
+type_synonym 'a assn = "'a set"
 
 datatype
  'a com = Basic "'a \<Rightarrow> 'a"
@@ -23,7 +22,7 @@
 
 abbreviation annskip ("SKIP") where "SKIP == Basic id"
 
-types 'a sem = "'a option => 'a option => bool"
+type_synonym 'a sem = "'a option => 'a option => bool"
 
 inductive Sem :: "'a com \<Rightarrow> 'a sem"
 where
@@ -47,11 +46,8 @@
   "Valid p c q == \<forall>s s'. Sem c s s' \<longrightarrow> s : Some ` p \<longrightarrow> s' : Some ` q"
 
 
-
-(** parse translations **)
-
 syntax
-  "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
+  "_assign" :: "idt => 'b => 'a com"  ("(2_ :=/ _)" [70, 65] 61)
 
 syntax
   "_hoare_abort_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
@@ -59,146 +55,12 @@
 syntax ("" output)
   "_hoare_abort"      :: "['a assn,'a com,'a assn] => bool"
                  ("{_} // _ // {_}" [0,55,0] 50)
-ML {*
-
-local
-fun free a = Free(a,dummyT)
-fun abs((a,T),body) =
-  let val a = absfree(a, dummyT, body)
-  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
-in
-
-fun mk_abstuple [x] body = abs (x, body)
-  | mk_abstuple (x::xs) body =
-      Syntax.const @{const_syntax prod_case} $ abs (x, mk_abstuple xs body);
-
-fun mk_fbody a e [x as (b,_)] = if a=b then e else free b
-  | mk_fbody a e ((b,_)::xs) =
-      Syntax.const @{const_syntax Pair} $ (if a=b then e else free b) $ mk_fbody a e xs;
-
-fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
-end
-*}
-
-(* bexp_tr & assn_tr *)
-(*all meta-variables for bexp except for TRUE are translated as if they
-  were boolean expressions*)
-ML{*
-fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
-  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
-
-fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
-*}
-(* com_tr *)
-ML{*
-fun com_tr (t as (Const (@{syntax_const "_assign"},_) $ x $ e)) xs =
-      (case Syntax.strip_positions x of
-        Free (a, _) => Syntax.const @{const_syntax Basic} $ mk_fexp a e xs
-      | _ => t)
-  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
-  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
-      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
-  | com_tr t _ = t (* if t is just a Free/Var *)
-*}
-
-(* triple_tr *)    (* FIXME does not handle "_idtdummy" *)
-ML{*
-local
-
-fun var_tr (Free (a, _)) = (a, Bound 0) (* Bound 0 = dummy term *)
-  | var_tr (Const (@{syntax_const "_constrain"}, _) $ Free (a, _) $ T) = (a, T);
-
-fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) =
-      var_tr (Syntax.strip_positions idt) :: vars_tr vars
-  | vars_tr t = [var_tr (Syntax.strip_positions t)]
-
-in
-fun hoare_vars_tr [vars, pre, prg, post] =
-      let val xs = vars_tr vars
-      in Syntax.const @{const_syntax Valid} $
-         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
-      end
-  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
-end
-*}
-
-parse_translation {* [(@{syntax_const "_hoare_abort_vars"}, hoare_vars_tr)] *}
-
 
-(*****************************************************************************)
-
-(*** print translations ***)
-ML{*
-fun dest_abstuple (Const (@{const_syntax prod_case},_) $ (Abs(v,_, body))) =
-      subst_bound (Syntax.free v, dest_abstuple body)
-  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
-  | dest_abstuple trm = trm;
-
-fun abs2list (Const (@{const_syntax prod_case},_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
-  | abs2list (Abs(x,T,t)) = [Free (x, T)]
-  | abs2list _ = [];
-
-fun mk_ts (Const (@{const_syntax prod_case},_) $ (Abs(x,_,t))) = mk_ts t
-  | mk_ts (Abs(x,_,t)) = mk_ts t
-  | mk_ts (Const (@{const_syntax Pair},_) $ a $ b) = a::(mk_ts b)
-  | mk_ts t = [t];
-
-fun mk_vts (Const (@{const_syntax prod_case},_) $ (Abs(x,_,t))) =
-           ((Syntax.free x)::(abs2list t), mk_ts t)
-  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
-  | mk_vts t = raise Match;
-
-fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))
-  | find_ch ((v,t)::vts) i xs =
-      if t = Bound i then find_ch vts (i-1) xs
-      else (true, (v, subst_bounds (xs,t)));
-
-fun is_f (Const (@{const_syntax prod_case},_) $ (Abs(x,_,t))) = true
-  | is_f (Abs(x,_,t)) = true
-  | is_f t = false;
-*}
+use "hoare_syntax.ML"
+parse_translation {* [(@{syntax_const "_hoare_abort_vars"}, Hoare_Syntax.hoare_vars_tr)] *}
+print_translation
+  {* [(@{const_syntax Valid}, Hoare_Syntax.spec_tr' @{syntax_const "_hoare_abort"})] *}
 
-(* assn_tr' & bexp_tr'*)
-ML{*
-fun assn_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | assn_tr' (Const (@{const_syntax inter},_) $ (Const (@{const_syntax Collect},_) $ T1) $
-        (Const (@{const_syntax Collect},_) $ T2)) =
-      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
-  | assn_tr' t = t;
-
-fun bexp_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | bexp_tr' t = t;
-*}
-
-(*com_tr' *)
-ML{*
-fun mk_assign f =
-  let val (vs, ts) = mk_vts f;
-      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
-  in
-    if ch then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
-    else Syntax.const @{const_syntax annskip}
-  end;
-
-fun com_tr' (Const (@{const_syntax Basic},_) $ f) =
-      if is_f f then mk_assign f else Syntax.const @{const_syntax Basic} $ f
-  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
-      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) =
-      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax While},_) $ b $ I $ c) =
-      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
-  | com_tr' t = t;
-
-fun spec_tr' [p, c, q] =
-  Syntax.const @{syntax_const "_hoare_abort"} $ assn_tr' p $ com_tr' c $ assn_tr' q
-*}
-
-print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
 
 (*** The proof rules ***)
 
--- a/src/HOL/Hoare/SchorrWaite.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare/SchorrWaite.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -40,7 +40,7 @@
 done
 
 lemma still_reachable: "\<lbrakk>B\<subseteq>Ra\<^sup>*``A; \<forall> (x,y) \<in> Rb-Ra. y\<in> (Ra\<^sup>*``A)\<rbrakk> \<Longrightarrow> Rb\<^sup>* `` B \<subseteq> Ra\<^sup>* `` A "
-apply (clarsimp simp only:Image_iff intro:subsetI)
+apply (clarsimp simp only:Image_iff)
 apply (erule rtrancl_induct)
  apply blast
 apply (subgoal_tac "(y, z) \<in> Ra\<union>(Rb-Ra)")
--- a/src/HOL/Hoare/SepLogHeap.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare/SepLogHeap.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -10,7 +10,7 @@
 imports Main
 begin
 
-types heap = "(nat \<Rightarrow> nat option)"
+type_synonym heap = "(nat \<Rightarrow> nat option)"
 
 text{* @{text "Some"} means allocated, @{text "None"} means
 free. Address @{text "0"} serves as the null reference. *}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare/hoare_syntax.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -0,0 +1,156 @@
+(*  Title:      HOL/Hoare/hoare_syntax.ML
+    Author:     Leonor Prensa Nieto & Tobias Nipkow
+
+Syntax translations for Hoare logic.
+*)
+
+signature HOARE_SYNTAX =
+sig
+  val hoare_vars_tr: term list -> term
+  val spec_tr': string -> term list -> term
+end;
+
+structure Hoare_Syntax: HOARE_SYNTAX =
+struct
+
+(** parse translation **)
+
+local
+
+fun idt_name (Free (x, _)) = SOME x
+  | idt_name (Const ("_constrain", _) $ t $ _) = idt_name t
+  | idt_name _ = NONE;
+
+fun eq_idt tu =
+  (case pairself idt_name tu of
+    (SOME x, SOME y) => x = y
+  | _ => false);
+
+fun mk_abstuple [x] body = Syntax.abs_tr [x, body]
+  | mk_abstuple (x :: xs) body =
+      Syntax.const @{const_syntax prod_case} $ Syntax.abs_tr [x, mk_abstuple xs body];
+
+fun mk_fbody x e [y] = if eq_idt (x, y) then e else y
+  | mk_fbody x e (y :: xs) =
+      Syntax.const @{const_syntax Pair} $
+        (if eq_idt (x, y) then e else y) $ mk_fbody x e xs;
+
+fun mk_fexp x e xs = mk_abstuple xs (mk_fbody x e xs);
+
+
+(* bexp_tr & assn_tr *)
+(*all meta-variables for bexp except for TRUE are translated as if they
+  were boolean expressions*)
+
+fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
+  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
+
+fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
+
+
+(* com_tr *)
+
+fun com_tr (Const (@{syntax_const "_assign"}, _) $ x $ e) xs =
+      Syntax.const @{const_syntax Basic} $ mk_fexp x e xs
+  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
+  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
+      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
+  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
+      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
+  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
+      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
+  | com_tr t _ = t;
+
+fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) = idt :: vars_tr vars
+  | vars_tr t = [t];
+
+in
+
+fun hoare_vars_tr [vars, pre, prg, post] =
+      let val xs = vars_tr vars
+      in Syntax.const @{const_syntax Valid} $
+         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
+      end
+  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
+
+end;
+
+
+
+(** print translation **)
+
+local
+
+fun dest_abstuple
+      (Const (@{const_syntax prod_case}, _) $ Abs (v, _, body)) =
+        subst_bound (Syntax.free v, dest_abstuple body)
+  | dest_abstuple (Abs (v,_, body)) = subst_bound (Syntax.free v, body)
+  | dest_abstuple tm = tm;
+
+fun abs2list (Const (@{const_syntax prod_case}, _) $ Abs (x, T, t)) = Free (x, T) :: abs2list t
+  | abs2list (Abs (x, T, t)) = [Free (x, T)]
+  | abs2list _ = [];
+
+fun mk_ts (Const (@{const_syntax prod_case}, _) $ Abs (x, _, t)) = mk_ts t
+  | mk_ts (Abs (x, _, t)) = mk_ts t
+  | mk_ts (Const (@{const_syntax Pair}, _) $ a $ b) = a :: mk_ts b
+  | mk_ts t = [t];
+
+fun mk_vts (Const (@{const_syntax prod_case},_) $ Abs (x, _, t)) =
+      (Syntax.free x :: abs2list t, mk_ts t)
+  | mk_vts (Abs (x, _, t)) = ([Syntax.free x], [t])
+  | mk_vts t = raise Match;
+
+fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))  (* FIXME no_ch!? *)
+  | find_ch ((v, t) :: vts) i xs =
+      if t = Bound i then find_ch vts (i - 1) xs
+      else (true, (v, subst_bounds (xs, t)));
+
+fun is_f (Const (@{const_syntax prod_case}, _) $ Abs (x, _, t)) = true
+  | is_f (Abs (x, _, t)) = true
+  | is_f t = false;
+
+
+(* assn_tr' & bexp_tr'*)
+
+fun assn_tr' (Const (@{const_syntax Collect}, _) $ T) = dest_abstuple T
+  | assn_tr' (Const (@{const_syntax inter}, _) $
+        (Const (@{const_syntax Collect}, _) $ T1) $ (Const (@{const_syntax Collect}, _) $ T2)) =
+      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
+  | assn_tr' t = t;
+
+fun bexp_tr' (Const (@{const_syntax Collect}, _) $ T) = dest_abstuple T
+  | bexp_tr' t = t;
+
+
+(* com_tr' *)
+
+fun mk_assign f =
+  let
+    val (vs, ts) = mk_vts f;
+    val (ch, which) = find_ch (vs ~~ ts) (length vs - 1) (rev vs);
+  in
+    if ch
+    then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
+    else Syntax.const @{const_syntax annskip}
+  end;
+
+fun com_tr' (Const (@{const_syntax Basic}, _) $ f) =
+      if is_f f then mk_assign f
+      else Syntax.const @{const_syntax Basic} $ f
+  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
+      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
+  | com_tr' (Const (@{const_syntax Cond}, _) $ b $ c1 $ c2) =
+      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
+  | com_tr' (Const (@{const_syntax While}, _) $ b $ I $ c) =
+      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
+  | com_tr' t = t;
+
+in
+
+fun spec_tr' syn [p, c, q] = Syntax.const syn $ assn_tr' p $ com_tr' c $ assn_tr' q
+
+end;
+
+end;
+
--- a/src/HOL/Hoare_Parallel/Graph.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare_Parallel/Graph.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -6,10 +6,9 @@
 
 datatype node = Black | White
 
-types 
-  nodes = "node list"
-  edge  = "nat \<times> nat"
-  edges = "edge list"
+type_synonym nodes = "node list"
+type_synonym edge = "nat \<times> nat"
+type_synonym edges = "edge list"
 
 consts Roots :: "nat set"
 
--- a/src/HOL/Hoare_Parallel/OG_Com.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Com.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -7,9 +7,8 @@
 
 text {* Type abbreviations for boolean expressions and assertions: *}
 
-types
-    'a bexp = "'a set"
-    'a assn = "'a set"
+type_synonym 'a bexp = "'a set"
+type_synonym 'a assn = "'a set"
 
 text {* The syntax of commands is defined by two mutually recursive
 datatypes: @{text "'a ann_com"} for annotated commands and @{text "'a
--- a/src/HOL/Hoare_Parallel/OG_Tran.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Tran.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -3,9 +3,8 @@
 
 theory OG_Tran imports OG_Com begin
 
-types
-  'a ann_com_op = "('a ann_com) option"
-  'a ann_triple_op = "('a ann_com_op \<times> 'a assn)"
+type_synonym 'a ann_com_op = "('a ann_com) option"
+type_synonym 'a ann_triple_op = "('a ann_com_op \<times> 'a assn)"
   
 primrec com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op" where
   "com (c, q) = c"
--- a/src/HOL/Hoare_Parallel/RG_Com.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Com.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -10,8 +10,7 @@
 of states.  Syntax of commands @{text com} and parallel commands
 @{text par_com}. *}
 
-types
-  'a bexp = "'a set"
+type_synonym 'a bexp = "'a set"
 
 datatype 'a com = 
     Basic "'a \<Rightarrow>'a"
@@ -20,6 +19,6 @@
   | While "'a bexp" "'a com"       
   | Await "'a bexp" "'a com"                 
 
-types 'a par_com = "(('a com) option) list"
+type_synonym 'a par_com = "'a com option list"
 
 end
\ No newline at end of file
--- a/src/HOL/Hoare_Parallel/RG_Hoare.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -55,7 +55,8 @@
 
 subsection {* Proof System for Parallel Programs *}
 
-types 'a par_rgformula = "('a rgformula) list \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
+type_synonym 'a par_rgformula =
+  "('a rgformula) list \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
 
 inductive
   par_rghoare :: "('a rgformula) list \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool"
--- a/src/HOL/Hoare_Parallel/RG_Tran.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Tran.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -8,7 +8,7 @@
 
 subsubsection {* Environment transitions *}
 
-types 'a conf = "(('a com) option) \<times> 'a"
+type_synonym 'a conf = "(('a com) option) \<times> 'a"
 
 inductive_set
   etran :: "('a conf \<times> 'a conf) set" 
@@ -48,7 +48,7 @@
 
 subsection {* Semantics of Parallel Programs *}
 
-types 'a par_conf = "('a par_com) \<times> 'a"
+type_synonym 'a par_conf = "('a par_com) \<times> 'a"
 
 inductive_set
   par_etran :: "('a par_conf \<times> 'a par_conf) set"
@@ -73,9 +73,9 @@
 
 subsubsection {* Sequential computations *}
 
-types 'a confs = "('a conf) list"
+type_synonym 'a confs = "'a conf list"
 
-inductive_set cptn :: "('a confs) set"
+inductive_set cptn :: "'a confs set"
 where
   CptnOne: "[(P,s)] \<in> cptn"
 | CptnEnv: "(P, t)#xs \<in> cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> cptn"
@@ -86,9 +86,9 @@
 
 subsubsection {* Parallel computations *}
 
-types  'a par_confs = "('a par_conf) list"
+type_synonym 'a par_confs = "'a par_conf list"
 
-inductive_set par_cptn :: "('a par_confs) set"
+inductive_set par_cptn :: "'a par_confs set"
 where
   ParCptnOne: "[(P,s)] \<in> par_cptn"
 | ParCptnEnv: "(P,t)#xs \<in> par_cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> par_cptn"
@@ -363,7 +363,8 @@
 
 subsection {* Validity for Component Programs. *}
 
-types 'a rgformula = "'a com \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
+type_synonym 'a rgformula =
+  "'a com \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
 
 definition assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a confs) set" where
   "assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow> 
--- a/src/HOL/IMP/Com.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/IMP/Com.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -11,11 +11,10 @@
   -- "an unspecified (arbitrary) type of locations 
       (adresses/names) for variables"
       
-types 
-  val   = nat -- "or anything else, @{text nat} used in examples"
-  state = "loc \<Rightarrow> val"
-  aexp  = "state \<Rightarrow> val"  
-  bexp  = "state \<Rightarrow> bool"
+type_synonym val = nat -- "or anything else, @{text nat} used in examples"
+type_synonym state = "loc \<Rightarrow> val"
+type_synonym aexp = "state \<Rightarrow> val"
+type_synonym bexp = "state \<Rightarrow> bool"
   -- "arithmetic and boolean expressions are not modelled explicitly here,"
   -- "they are just functions on states"
 
--- a/src/HOL/IMP/Denotation.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/IMP/Denotation.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -6,7 +6,7 @@
 
 theory Denotation imports Natural begin
 
-types com_den = "(state\<times>state)set"
+type_synonym com_den = "(state \<times> state) set"
 
 definition
   Gamma :: "[bexp,com_den] => (com_den => com_den)" where
--- a/src/HOL/IMP/Expr.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/IMP/Expr.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -14,8 +14,7 @@
 subsection "Arithmetic expressions"
 typedecl loc
 
-types
-  state = "loc => nat"
+type_synonym state = "loc => nat"
 
 datatype
   aexp = N nat
--- a/src/HOL/IMP/Hoare.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/IMP/Hoare.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -6,7 +6,7 @@
 
 theory Hoare imports Natural begin
 
-types assn = "state => bool"
+type_synonym assn = "state => bool"
 
 inductive
   hoare :: "assn => com => assn => bool" ("|- ({(1_)}/ (_)/ {(1_)})" 50)
--- a/src/HOL/IMP/Machines.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/IMP/Machines.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -14,7 +14,7 @@
 text {* There are only three instructions: *}
 datatype instr = SET loc aexp | JMPF bexp nat | JMPB nat
 
-types instrs = "instr list"
+type_synonym instrs = "instr list"
 
 subsection "M0 with PC"
 
@@ -47,7 +47,7 @@
   an operational (small step) semantics:
 *}
 
-types config = "instrs \<times> instrs \<times> state"
+type_synonym config = "instrs \<times> instrs \<times> state"
 
 
 inductive_set
--- a/src/HOL/IMPP/Com.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/IMPP/Com.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -8,8 +8,9 @@
 imports Main
 begin
 
-types    val = nat   (* for the meta theory, this may be anything, but with
-                        current Isabelle, types cannot be refined later *)
+type_synonym val = nat
+  (* for the meta theory, this may be anything, but types cannot be refined later *)
+
 typedecl glb
 typedecl loc
 
@@ -18,15 +19,15 @@
   Res :: loc
 
 datatype vname  = Glb glb | Loc loc
-types    globs  = "glb => val"
-         locals = "loc => val"
+type_synonym globs = "glb => val"
+type_synonym locals = "loc => val"
 datatype state  = st globs locals
 (* for the meta theory, the following would be sufficient:
 typedecl state
 consts   st :: "[globs , locals] => state"
 *)
-types    aexp   = "state => val"
-         bexp   = "state => bool"
+type_synonym aexp = "state => val"
+type_synonym bexp = "state => bool"
 
 typedecl pname
 
--- a/src/HOL/IMPP/Hoare.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/IMPP/Hoare.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -16,7 +16,7 @@
   vs. simultaneous recursion in call rule
 *}
 
-types 'a assn = "'a => state => bool"
+type_synonym 'a assn = "'a => state => bool"
 translations
   (type) "'a assn" <= (type) "'a => state => bool"
 
--- a/src/HOL/IOA/Asig.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/IOA/Asig.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -9,7 +9,7 @@
 imports Main
 begin
 
-types
+type_synonym
   'a signature = "('a set * 'a set * 'a set)"
 
 consts
--- a/src/HOL/IOA/IOA.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/IOA/IOA.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -9,12 +9,11 @@
 imports Asig
 begin
 
-types
-   'a seq            =   "nat => 'a"
-   'a oseq           =   "nat => 'a option"
-   ('a,'b)execution  =   "'a oseq * 'b seq"
-   ('a,'s)transition =   "('s * 'a * 's)"
-   ('a,'s)ioa        =   "'a signature * 's set * ('a,'s)transition set"
+type_synonym 'a seq = "nat => 'a"
+type_synonym 'a oseq = "nat => 'a option"
+type_synonym ('a, 'b) execution = "'a oseq * 'b seq"
+type_synonym ('a, 's) transition = "('s * 'a * 's)"
+type_synonym ('a,'s) ioa = "'a signature * 's set * ('a, 's) transition set"
 
 consts
 
--- a/src/HOL/IsaMakefile	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/IsaMakefile	Mon Apr 04 12:40:00 2011 +0100
@@ -650,12 +650,12 @@
 HOL-Hoare: HOL $(LOG)/HOL-Hoare.gz
 
 $(LOG)/HOL-Hoare.gz: $(OUT)/HOL Hoare/Arith2.thy Hoare/Examples.thy	\
-  Hoare/Hoare.thy Hoare/hoare_tac.ML Hoare/Heap.thy			\
-  Hoare/Hoare_Logic.thy	Hoare/Hoare_Logic_Abort.thy			\
+  Hoare/Hoare.thy Hoare/hoare_syntax.ML Hoare/hoare_tac.ML		\
+  Hoare/Heap.thy Hoare/Hoare_Logic.thy Hoare/Hoare_Logic_Abort.thy	\
   Hoare/HeapSyntax.thy Hoare/Pointer_Examples.thy Hoare/ROOT.ML		\
   Hoare/ExamplesAbort.thy Hoare/HeapSyntaxAbort.thy			\
-  Hoare/SchorrWaite.thy Hoare/Separation.thy				\
-  Hoare/SepLogHeap.thy Hoare/document/root.tex Hoare/document/root.bib
+  Hoare/SchorrWaite.thy Hoare/Separation.thy Hoare/SepLogHeap.thy	\
+  Hoare/document/root.tex Hoare/document/root.bib
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Hoare
 
 
@@ -1049,7 +1049,7 @@
   ex/Multiquote.thy ex/NatSum.thy ex/Normalization_by_Evaluation.thy	\
   ex/Numeral.thy ex/PER.thy ex/PresburgerEx.thy ex/Primrec.thy		\
   ex/Quickcheck_Examples.thy ex/Quickcheck_Lattice_Examples.thy		\
-  ex/Quickcheck_Narrowing_Examples.thy \
+  ex/Quickcheck_Narrowing_Examples.thy ex/SML_Quickcheck_Examples.thy   \
   ex/Quicksort.thy ex/ROOT.ML ex/Recdefs.thy ex/Records.thy		\
   ex/ReflectionEx.thy ex/Refute_Examples.thy ex/SAT_Examples.thy	\
   ex/SVC_Oracle.thy ex/Serbian.thy ex/Set_Algebras.thy			\
@@ -1186,17 +1186,17 @@
 HOL-Probability: HOL-Multivariate_Analysis $(OUT)/HOL-Probability
 
 $(OUT)/HOL-Probability: $(OUT)/HOL-Multivariate_Analysis		\
-  Probability/Borel_Space.thy Probability/Caratheodory.thy		\
-  Probability/Complete_Measure.thy					\
+  Probability/Binary_Product_Measure.thy Probability/Borel_Space.thy	\
+  Probability/Caratheodory.thy Probability/Complete_Measure.thy		\
   Probability/ex/Dining_Cryptographers.thy				\
   Probability/ex/Koepf_Duermuth_Countermeasure.thy			\
-  Probability/Information.thy Probability/Lebesgue_Integration.thy	\
-  Probability/Lebesgue_Measure.thy Probability/Measure.thy		\
-  Probability/Probability_Space.thy Probability/Probability.thy		\
-  Probability/Product_Measure.thy Probability/Radon_Nikodym.thy		\
+  Probability/Finite_Product_Measure.thy				\
+  Probability/Infinite_Product_Measure.thy Probability/Information.thy	\
+  Probability/Lebesgue_Integration.thy Probability/Lebesgue_Measure.thy \
+  Probability/Measure.thy Probability/Probability_Measure.thy		\
+  Probability/Probability.thy Probability/Radon_Nikodym.thy		\
   Probability/ROOT.ML Probability/Sigma_Algebra.thy			\
-  Library/Countable.thy Library/FuncSet.thy				\
-  Library/Nat_Bijection.thy
+  Library/Countable.thy Library/FuncSet.thy Library/Nat_Bijection.thy
 	@cd Probability; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Multivariate_Analysis HOL-Probability
 
 
--- a/src/HOL/Lazy_Sequence.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Lazy_Sequence.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -208,7 +208,7 @@
 subsection {* With Hit Bound Value *}
 text {* assuming in negative context *}
 
-types 'a hit_bound_lazy_sequence = "'a option lazy_sequence"
+type_synonym 'a hit_bound_lazy_sequence = "'a option lazy_sequence"
 
 definition hit_bound :: "'a hit_bound_lazy_sequence"
 where
--- a/src/HOL/Library/SML_Quickcheck.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Library/SML_Quickcheck.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -8,9 +8,10 @@
 setup {*
   Inductive_Codegen.quickcheck_setup #>
   Context.theory_map (Quickcheck.add_generator ("SML",
-    fn ctxt => fn (t, eval_terms) =>
+    fn ctxt => fn [(t, eval_terms)] =>
       let
-        val test_fun = Codegen.test_term ctxt t 
+        val prop = list_abs_free (Term.add_frees t [], t)
+        val test_fun = Codegen.test_term ctxt prop 
         val iterations = Config.get ctxt Quickcheck.iterations
         fun iterate size 0 = NONE
           | iterate size j =
@@ -21,7 +22,7 @@
             in
               case result of NONE => iterate size (j - 1) | SOME q => SOME q
             end
-     in fn size => (iterate size iterations, NONE) end))
+     in fn [_, size] => (iterate size iterations, NONE) end))
 *}
 
 end
--- a/src/HOL/List.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/List.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -341,16 +341,17 @@
 mangled). In such cases it can be advisable to introduce separate
 definitions for the list comprehensions in question.  *}
 
-nonterminal lc_qual and lc_quals
+nonterminal lc_gen and lc_qual and lc_quals
 
 syntax
 "_listcompr" :: "'a \<Rightarrow> lc_qual \<Rightarrow> lc_quals \<Rightarrow> 'a list"  ("[_ . __")
-"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ <- _")
+"_lc_gen" :: "lc_gen \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ <- _")
 "_lc_test" :: "bool \<Rightarrow> lc_qual" ("_")
 (*"_lc_let" :: "letbinds => lc_qual"  ("let _")*)
 "_lc_end" :: "lc_quals" ("]")
 "_lc_quals" :: "lc_qual \<Rightarrow> lc_quals \<Rightarrow> lc_quals" (", __")
 "_lc_abs" :: "'a => 'b list => 'b list"
+"_strip_positions" :: "'a \<Rightarrow> lc_gen"  ("_")
 
 (* These are easier than ML code but cannot express the optimized
    translation of [e. p<-xs]
@@ -366,9 +367,9 @@
 *)
 
 syntax (xsymbols)
-"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
+"_lc_gen" :: "lc_gen \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
 syntax (HTML output)
-"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
+"_lc_gen" :: "lc_gen \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
 
 parse_translation (advanced) {*
 let
@@ -425,26 +426,45 @@
 in [(@{syntax_const "_listcompr"}, lc_tr)] end
 *}
 
-term "[(x,y,z). b]"
-term "[(x,y,z). x\<leftarrow>xs]"
-term "[e x y. x\<leftarrow>xs, y\<leftarrow>ys]"
-term "[(x,y,z). x<a, x>b]"
-term "[(x,y,z). x\<leftarrow>xs, x>b]"
-term "[(x,y,z). x<a, x\<leftarrow>xs]"
-term "[(x,y). Cons True x \<leftarrow> xs]"
-term "[(x,y,z). Cons x [] \<leftarrow> xs]"
-term "[(x,y,z). x<a, x>b, x=d]"
-term "[(x,y,z). x<a, x>b, y\<leftarrow>ys]"
-term "[(x,y,z). x<a, x\<leftarrow>xs,y>b]"
-term "[(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys]"
-term "[(x,y,z). x\<leftarrow>xs, x>b, y<a]"
-term "[(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys]"
-term "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x]"
-term "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs]"
+ML {*
+  let
+    val read = Syntax.read_term @{context};
+    fun check s1 s2 = read s1 aconv read s2 orelse error ("Check failed: " ^ quote s1);
+  in
+    check "[(x,y,z). b]" "if b then [(x, y, z)] else []";
+    check "[(x,y,z). x\<leftarrow>xs]" "map (\<lambda>x. (x, y, z)) xs";
+    check "[e x y. x\<leftarrow>xs, y\<leftarrow>ys]" "concat (map (\<lambda>x. map (\<lambda>y. e x y) ys) xs)";
+    check "[(x,y,z). x<a, x>b]" "if x < a then if b < x then [(x, y, z)] else [] else []";
+    check "[(x,y,z). x\<leftarrow>xs, x>b]" "concat (map (\<lambda>x. if b < x then [(x, y, z)] else []) xs)";
+    check "[(x,y,z). x<a, x\<leftarrow>xs]" "if x < a then map (\<lambda>x. (x, y, z)) xs else []";
+    check "[(x,y). Cons True x \<leftarrow> xs]"
+      "concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | True # x \<Rightarrow> [(x, y)] | False # x \<Rightarrow> []) xs)";
+    check "[(x,y,z). Cons x [] \<leftarrow> xs]"
+      "concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | [x] \<Rightarrow> [(x, y, z)] | x # aa # lista \<Rightarrow> []) xs)";
+    check "[(x,y,z). x<a, x>b, x=d]"
+      "if x < a then if b < x then if x = d then [(x, y, z)] else [] else [] else []";
+    check "[(x,y,z). x<a, x>b, y\<leftarrow>ys]"
+      "if x < a then if b < x then map (\<lambda>y. (x, y, z)) ys else [] else []";
+    check "[(x,y,z). x<a, x\<leftarrow>xs,y>b]"
+      "if x < a then concat (map (\<lambda>x. if b < y then [(x, y, z)] else []) xs) else []";
+    check "[(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys]"
+      "if x < a then concat (map (\<lambda>x. map (\<lambda>y. (x, y, z)) ys) xs) else []";
+    check "[(x,y,z). x\<leftarrow>xs, x>b, y<a]"
+      "concat (map (\<lambda>x. if b < x then if y < a then [(x, y, z)] else [] else []) xs)";
+    check "[(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys]"
+      "concat (map (\<lambda>x. if b < x then map (\<lambda>y. (x, y, z)) ys else []) xs)";
+    check "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x]"
+      "concat (map (\<lambda>x. concat (map (\<lambda>y. if x < y then [(x, y, z)] else []) ys)) xs)";
+    check "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs]"
+      "concat (map (\<lambda>x. concat (map (\<lambda>y. map (\<lambda>z. (x, y, z)) zs) ys)) xs)"
+  end;
+*}
+
 (*
 term "[(x,y). x\<leftarrow>xs, let xx = x+x, y\<leftarrow>ys, y \<noteq> xx]"
 *)
 
+
 use "Tools/list_to_set_comprehension.ML"
 
 simproc_setup list_to_set_comprehension ("set xs") = {* K List_to_Set_Comprehension.simproc *}
--- a/src/HOL/Map.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Map.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -11,7 +11,7 @@
 imports List
 begin
 
-types ('a,'b) "map" = "'a => 'b option" (infixr "~=>" 0)
+type_synonym ('a,'b) "map" = "'a => 'b option" (infixr "~=>" 0)
 
 type_notation (xsymbols)
   "map" (infixr "\<rightharpoonup>" 0)
--- a/src/HOL/MicroJava/BV/BVNoTypeError.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/BV/BVNoTypeError.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/BVNoTypeErrors.thy
+(*  Title:      HOL/MicroJava/BV/BVNoTypeError.thy
     Author:     Gerwin Klein
 *)
 
--- a/src/HOL/MicroJava/BV/EffectMono.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/BV/EffectMono.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/EffMono.thy
+(*  Title:      HOL/MicroJava/BV/EffectMono.thy
     Author:     Gerwin Klein
     Copyright   2000 Technische Universitaet Muenchen
 *)
--- a/src/HOL/MicroJava/BV/JVMType.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/BV/JVMType.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/JVM.thy
+(*  Title:      HOL/MicroJava/BV/JVMType.thy
     Author:     Gerwin Klein
     Copyright   2000 TUM
 *)
--- a/src/HOL/MicroJava/BV/LBVJVM.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/BV/LBVJVM.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/JVM.thy
+(*  Title:      HOL/MicroJava/BV/LBVJVM.thy
     Author:     Tobias Nipkow, Gerwin Klein
     Copyright   2000 TUM
 *)
--- a/src/HOL/MicroJava/BV/Typing_Framework_JVM.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/BV/Typing_Framework_JVM.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/JVM.thy
+(*  Title:      HOL/MicroJava/BV/Typing_Framework_JVM.thy
     Author:     Tobias Nipkow, Gerwin Klein
     Copyright   2000 TUM
 *)
--- a/src/HOL/MicroJava/DFA/Abstract_BV.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/Abstract_BV.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/Semilat.thy
+(*  Title:      HOL/MicroJava/DFA/Abstract_BV.thy
     Author:     Gerwin Klein
     Copyright   2003 TUM
 *)
--- a/src/HOL/MicroJava/DFA/Err.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/Err.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/Err.thy
+(*  Title:      HOL/MicroJava/DFA/Err.thy
     Author:     Tobias Nipkow
     Copyright   2000 TUM
 *)
--- a/src/HOL/MicroJava/DFA/Kildall.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/Kildall.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/Kildall.thy
+(*  Title:      HOL/MicroJava/DFA/Kildall.thy
     Author:     Tobias Nipkow, Gerwin Klein
     Copyright   2000 TUM
 *)
--- a/src/HOL/MicroJava/DFA/LBVComplete.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/LBVComplete.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/LBVComplete.thy
+(*  Title:      HOL/MicroJava/DFA/LBVComplete.thy
     Author:     Gerwin Klein
     Copyright   2000 Technische Universitaet Muenchen
 *)
--- a/src/HOL/MicroJava/DFA/LBVSpec.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/LBVSpec.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/LBVSpec.thy
+(*  Title:      HOL/MicroJava/DFA/LBVSpec.thy
     Author:     Gerwin Klein
     Copyright   1999 Technische Universitaet Muenchen
 *)
--- a/src/HOL/MicroJava/DFA/Listn.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/Listn.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/Listn.thy
+(*  Title:      HOL/MicroJava/DFA/Listn.thy
     Author:     Tobias Nipkow
     Copyright   2000 TUM
 *)
--- a/src/HOL/MicroJava/DFA/Opt.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/Opt.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/Opt.thy
+(*  Title:      HOL/MicroJava/DFA/Opt.thy
     Author:     Tobias Nipkow
     Copyright   2000 TUM
 *)
--- a/src/HOL/MicroJava/DFA/Product.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/Product.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/Product.thy
+(*  Title:      HOL/MicroJava/DFA/Product.thy
     Author:     Tobias Nipkow
     Copyright   2000 TUM
 *)
--- a/src/HOL/MicroJava/DFA/Semilat.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/Semilat.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/Semilat.thy
+(*  Title:      HOL/MicroJava/DFA/Semilat.thy
     Author:     Tobias Nipkow
     Copyright   2000 TUM
 *)
--- a/src/HOL/MicroJava/DFA/SemilatAlg.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/SemilatAlg.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/SemilatAlg.thy
+(*  Title:      HOL/MicroJava/DFA/SemilatAlg.thy
     Author:     Gerwin Klein
     Copyright   2002 Technische Universitaet Muenchen
 *)
--- a/src/HOL/MicroJava/DFA/Semilattices.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/Semilattices.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/Semilat.thy
+(*  Title:      HOL/MicroJava/DFA/Semilattices.thy
     Author:     Gerwin Klein
     Copyright   2003 TUM
 *)
--- a/src/HOL/MicroJava/DFA/Typing_Framework.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/Typing_Framework.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/Typing_Framework.thy
+(*  Title:      HOL/MicroJava/DFA/Typing_Framework.thy
     Author:     Tobias Nipkow
     Copyright   2000 TUM
 *)
--- a/src/HOL/MicroJava/DFA/Typing_Framework_err.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/MicroJava/DFA/Typing_Framework_err.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/MicroJava/BV/Typing_Framework_err.thy
+(*  Title:      HOL/MicroJava/DFA/Typing_Framework_err.thy
     Author:     Gerwin Klein
     Copyright   2000 TUM
 *)
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -558,7 +558,7 @@
         Unsynchronized.ref (NONE : ((string * locality) * thm list) list option)
       val minimize = AList.defined (op =) args minimizeK
       val metis_ft = AList.defined (op =) args metis_ftK
-      val trivial = Try.invoke_try (SOME try_timeout) ([], [], []) pre
+      val trivial = Try.invoke_try (SOME try_timeout) ([], [], [], []) pre
         handle TimeLimit.TimeOut => false
       fun apply_reconstructor m1 m2 =
         if metis_ft
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -3829,7 +3829,8 @@
              ==> continuous net (\<lambda>x. c(x) *\<^sub>R f x) "
   unfolding continuous_def by (intro tendsto_intros)
 
-lemmas continuous_intros = continuous_add continuous_vmul continuous_cmul continuous_const continuous_sub continuous_at_id continuous_within_id continuous_mul
+lemmas continuous_intros = continuous_add continuous_vmul continuous_cmul
+  continuous_const continuous_sub continuous_at_id continuous_within_id continuous_mul
 
 lemma continuous_on_vmul:
   fixes c :: "'a::metric_space \<Rightarrow> real" and v :: "'b::real_normed_vector"
@@ -3843,9 +3844,20 @@
              ==> continuous_on s (\<lambda>x. c(x) *\<^sub>R f x)"
   unfolding continuous_on_eq_continuous_within using continuous_mul[of _ c] by auto
 
-lemmas continuous_on_intros = continuous_on_add continuous_on_const continuous_on_id continuous_on_compose continuous_on_cmul continuous_on_neg continuous_on_sub
-  uniformly_continuous_on_add uniformly_continuous_on_const uniformly_continuous_on_id uniformly_continuous_on_compose uniformly_continuous_on_cmul uniformly_continuous_on_neg uniformly_continuous_on_sub
-  continuous_on_mul continuous_on_vmul
+lemma continuous_on_mul_real:
+  fixes f :: "'a::metric_space \<Rightarrow> real"
+  fixes g :: "'a::metric_space \<Rightarrow> real"
+  shows "continuous_on s f \<Longrightarrow> continuous_on s g
+             ==> continuous_on s (\<lambda>x. f x * g x)"
+  using continuous_on_mul[of s f g] unfolding real_scaleR_def .
+
+lemmas continuous_on_intros = continuous_on_add continuous_on_const
+  continuous_on_id continuous_on_compose continuous_on_cmul continuous_on_neg
+  continuous_on_sub continuous_on_mul continuous_on_vmul continuous_on_mul_real
+  uniformly_continuous_on_add uniformly_continuous_on_const
+  uniformly_continuous_on_id uniformly_continuous_on_compose
+  uniformly_continuous_on_cmul uniformly_continuous_on_neg
+  uniformly_continuous_on_sub
 
 text{* And so we have continuity of inverse.                                     *}
 
--- a/src/HOL/Mutabelle/mutabelle_extra.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -154,7 +154,7 @@
   let
     val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (ProofContext.init_global thy)
   in
-    case Try.invoke_try (SOME (seconds 5.0)) ([], [], []) state of
+    case Try.invoke_try (SOME (seconds 5.0)) ([], [], [], []) state of
       true => (Solved, [])
     | false => (Unsolved, [])
   end
--- a/src/HOL/New_DSequence.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/New_DSequence.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -9,7 +9,7 @@
 
 subsection {* Positive Depth-Limited Sequence *}
 
-types 'a pos_dseq = "code_numeral => 'a Lazy_Sequence.lazy_sequence"
+type_synonym 'a pos_dseq = "code_numeral => 'a Lazy_Sequence.lazy_sequence"
 
 definition pos_empty :: "'a pos_dseq"
 where
@@ -49,7 +49,7 @@
 
 subsection {* Negative Depth-Limited Sequence *}
 
-types 'a neg_dseq = "code_numeral => 'a Lazy_Sequence.hit_bound_lazy_sequence"
+type_synonym 'a neg_dseq = "code_numeral => 'a Lazy_Sequence.hit_bound_lazy_sequence"
 
 definition neg_empty :: "'a neg_dseq"
 where
--- a/src/HOL/New_Random_Sequence.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/New_Random_Sequence.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -5,8 +5,8 @@
 imports Quickcheck New_DSequence
 begin
 
-types 'a pos_random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> 'a New_DSequence.pos_dseq"
-types 'a neg_random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> 'a New_DSequence.neg_dseq"
+type_synonym 'a pos_random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> 'a New_DSequence.pos_dseq"
+type_synonym 'a neg_random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> 'a New_DSequence.neg_dseq"
 
 definition pos_empty :: "'a pos_random_dseq"
 where
--- a/src/HOL/Nitpick_Examples/Core_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Core_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -12,7 +12,7 @@
 begin
 
 nitpick_params [verbose, card = 1\<midarrow>6, unary_ints, max_potential = 0,
-                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60]
+                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 240]
 
 subsection {* Curry in a Hurry *}
 
--- a/src/HOL/Nitpick_Examples/Datatype_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Datatype_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -12,7 +12,7 @@
 begin
 
 nitpick_params [verbose, card = 1\<midarrow>8, max_potential = 0,
-                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60]
+                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 240]
 
 primrec rot where
 "rot Nibble0 = Nibble1" | "rot Nibble1 = Nibble2" | "rot Nibble2 = Nibble3" |
--- a/src/HOL/Nitpick_Examples/Hotel_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Hotel_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -13,7 +13,7 @@
 begin
 
 nitpick_params [verbose, max_potential = 0, sat_solver = MiniSat_JNI,
-                max_threads = 1, timeout = 120]
+                max_threads = 1, timeout = 240]
 
 typedecl guest
 typedecl key
--- a/src/HOL/Nitpick_Examples/Induct_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Induct_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -12,7 +12,7 @@
 begin
 
 nitpick_params [verbose, card = 1\<midarrow>8, unary_ints,
-                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60]
+                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 240]
 
 inductive p1 :: "nat \<Rightarrow> bool" where
 "p1 0" |
--- a/src/HOL/Nitpick_Examples/Integer_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Integer_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -12,7 +12,7 @@
 begin
 
 nitpick_params [verbose, card = 1\<midarrow>5, bits = 1,2,3,4,6,
-                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60]
+                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 240]
 
 lemma "Suc x = x + 1"
 nitpick [unary_ints, expect = none]
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -18,7 +18,7 @@
 chapter {* 3. First Steps *}
 
 nitpick_params [verbose, sat_solver = MiniSat_JNI, max_threads = 1,
-                timeout = 60]
+                timeout = 240]
 
 subsection {* 3.1. Propositional Logic *}
 
--- a/src/HOL/Nitpick_Examples/Pattern_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Pattern_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -12,7 +12,7 @@
 begin
 
 nitpick_params [verbose, card = 10, max_potential = 0, sat_solver = MiniSat_JNI,
-                max_threads = 1, timeout = 60]
+                max_threads = 1, timeout = 240]
 
 lemma "x = (case u of () \<Rightarrow> y)"
 nitpick [expect = genuine]
--- a/src/HOL/Nitpick_Examples/Record_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Record_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -12,7 +12,7 @@
 begin
 
 nitpick_params [verbose, card = 1\<midarrow>6, max_potential = 0,
-                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60]
+                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 240]
 
 record point2d =
   xc :: int
--- a/src/HOL/Nitpick_Examples/Refute_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -12,7 +12,7 @@
 begin
 
 nitpick_params [verbose, card = 1\<midarrow>6, max_potential = 0,
-                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60]
+                sat_solver = MiniSat_JNI, max_threads = 1, timeout = 240]
 
 lemma "P \<and> Q"
 apply (rule conjI)
--- a/src/HOL/Nitpick_Examples/Special_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Special_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -12,7 +12,7 @@
 begin
 
 nitpick_params [verbose, card = 4, sat_solver = MiniSat_JNI, max_threads = 1,
-                timeout = 60]
+                timeout = 240]
 
 fun f1 :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
 "f1 a b c d e = a + b + c + d + e"
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -12,7 +12,7 @@
 begin
 
 nitpick_params [verbose, card = 1\<midarrow>4, sat_solver = MiniSat_JNI, max_threads = 1,
-                timeout = 60]
+                timeout = 240]
 
 typedef three = "{0\<Colon>nat, 1, 2}"
 by blast
--- a/src/HOL/Predicate_Compile_Examples/Examples.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Predicate_Compile_Examples/Examples.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -2,6 +2,8 @@
 imports Main "~~/src/HOL/Library/Predicate_Compile_Alternative_Defs"
 begin
 
+declare [[values_timeout = 480.0]]
+
 section {* Formal Languages *}
 
 subsection {* General Context Free Grammars *}
--- a/src/HOL/Predicate_Compile_Examples/Hotel_Example_Small_Generator.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example_Small_Generator.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -50,8 +50,8 @@
 
 lemma
   "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
-(*quickcheck[tester = small_generators_depth_14, finite_types = false, iterations = 1, size = 1, timeout = 60.0, expect = no_counterexample, timeout = 60]*)
-quickcheck[tester = small_generators_depth_15, finite_types = false, iterations = 1, size = 1, timeout = 60.0, expect = counterexample, timeout = 60]
+(*quickcheck[tester = small_generators_depth_14, finite_types = false, iterations = 1, size = 1, timeout = 600.0, expect = no_counterexample]*)
+quickcheck[tester = small_generators_depth_15, finite_types = false, iterations = 1, size = 1, timeout = 600.0, expect = counterexample]
 oops
 
 
--- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -2,6 +2,8 @@
 imports "~~/src/HOL/Library/Predicate_Compile_Alternative_Defs"
 begin
 
+declare [[values_timeout = 480.0]]
+
 subsection {* Basic predicates *}
 
 inductive False' :: "bool"
--- a/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -2,6 +2,8 @@
 imports Main "~~/src/HOL/Library/Predicate_Compile_Alternative_Defs"
 begin
 
+declare [[values_timeout = 480.0]]
+
 section {* Specialisation Examples *}
 
 primrec nth_el'
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Binary_Product_Measure.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -0,0 +1,975 @@
+(*  Title:      HOL/Probability/Binary_Product_Measure.thy
+    Author:     Johannes Hölzl, TU München
+*)
+
+header {*Binary product measures*}
+
+theory Binary_Product_Measure
+imports Lebesgue_Integration
+begin
+
+lemma times_Int_times: "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)"
+  by auto
+
+lemma Pair_vimage_times[simp]: "\<And>A B x. Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
+  by auto
+
+lemma rev_Pair_vimage_times[simp]: "\<And>A B y. (\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
+  by auto
+
+lemma case_prod_distrib: "f (case x of (x, y) \<Rightarrow> g x y) = (case x of (x, y) \<Rightarrow> f (g x y))"
+  by (cases x) simp
+
+lemma split_beta': "(\<lambda>(x,y). f x y) = (\<lambda>x. f (fst x) (snd x))"
+  by (auto simp: fun_eq_iff)
+
+section "Binary products"
+
+definition
+  "pair_measure_generator A B =
+    \<lparr> space = space A \<times> space B,
+      sets = {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B},
+      measure = \<lambda>X. \<integral>\<^isup>+x. (\<integral>\<^isup>+y. indicator X (x,y) \<partial>B) \<partial>A \<rparr>"
+
+definition pair_measure (infixr "\<Otimes>\<^isub>M" 80) where
+  "A \<Otimes>\<^isub>M B = sigma (pair_measure_generator A B)"
+
+locale pair_sigma_algebra = M1: sigma_algebra M1 + M2: sigma_algebra M2
+  for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
+
+abbreviation (in pair_sigma_algebra)
+  "E \<equiv> pair_measure_generator M1 M2"
+
+abbreviation (in pair_sigma_algebra)
+  "P \<equiv> M1 \<Otimes>\<^isub>M M2"
+
+lemma sigma_algebra_pair_measure:
+  "sets M1 \<subseteq> Pow (space M1) \<Longrightarrow> sets M2 \<subseteq> Pow (space M2) \<Longrightarrow> sigma_algebra (pair_measure M1 M2)"
+  by (force simp: pair_measure_def pair_measure_generator_def intro!: sigma_algebra_sigma)
+
+sublocale pair_sigma_algebra \<subseteq> sigma_algebra P
+  using M1.space_closed M2.space_closed
+  by (rule sigma_algebra_pair_measure)
+
+lemma pair_measure_generatorI[intro, simp]:
+  "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (pair_measure_generator A B)"
+  by (auto simp add: pair_measure_generator_def)
+
+lemma pair_measureI[intro, simp]:
+  "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^isub>M B)"
+  by (auto simp add: pair_measure_def)
+
+lemma space_pair_measure:
+  "space (A \<Otimes>\<^isub>M B) = space A \<times> space B"
+  by (simp add: pair_measure_def pair_measure_generator_def)
+
+lemma sets_pair_measure_generator:
+  "sets (pair_measure_generator N M) = (\<lambda>(x, y). x \<times> y) ` (sets N \<times> sets M)"
+  unfolding pair_measure_generator_def by auto
+
+lemma pair_measure_generator_sets_into_space:
+  assumes "sets M \<subseteq> Pow (space M)" "sets N \<subseteq> Pow (space N)"
+  shows "sets (pair_measure_generator M N) \<subseteq> Pow (space (pair_measure_generator M N))"
+  using assms by (auto simp: pair_measure_generator_def)
+
+lemma pair_measure_generator_Int_snd:
+  assumes "sets S1 \<subseteq> Pow (space S1)"
+  shows "sets (pair_measure_generator S1 (algebra.restricted_space S2 A)) =
+         sets (algebra.restricted_space (pair_measure_generator S1 S2) (space S1 \<times> A))"
+  (is "?L = ?R")
+  apply (auto simp: pair_measure_generator_def image_iff)
+  using assms
+  apply (rule_tac x="a \<times> xa" in exI)
+  apply force
+  using assms
+  apply (rule_tac x="a" in exI)
+  apply (rule_tac x="b \<inter> A" in exI)
+  apply auto
+  done
+
+lemma (in pair_sigma_algebra)
+  shows measurable_fst[intro!, simp]:
+    "fst \<in> measurable P M1" (is ?fst)
+  and measurable_snd[intro!, simp]:
+    "snd \<in> measurable P M2" (is ?snd)
+proof -
+  { fix X assume "X \<in> sets M1"
+    then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. fst -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2"
+      apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"])
+      using M1.sets_into_space by force+ }
+  moreover
+  { fix X assume "X \<in> sets M2"
+    then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. snd -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2"
+      apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X])
+      using M2.sets_into_space by force+ }
+  ultimately have "?fst \<and> ?snd"
+    by (fastsimp simp: measurable_def sets_sigma space_pair_measure
+                 intro!: sigma_sets.Basic)
+  then show ?fst ?snd by auto
+qed
+
+lemma (in pair_sigma_algebra) measurable_pair_iff:
+  assumes "sigma_algebra M"
+  shows "f \<in> measurable M P \<longleftrightarrow>
+    (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
+proof -
+  interpret M: sigma_algebra M by fact
+  from assms show ?thesis
+  proof (safe intro!: measurable_comp[where b=P])
+    assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2"
+    show "f \<in> measurable M P" unfolding pair_measure_def
+    proof (rule M.measurable_sigma)
+      show "sets (pair_measure_generator M1 M2) \<subseteq> Pow (space E)"
+        unfolding pair_measure_generator_def using M1.sets_into_space M2.sets_into_space by auto
+      show "f \<in> space M \<rightarrow> space E"
+        using f s by (auto simp: mem_Times_iff measurable_def comp_def space_sigma pair_measure_generator_def)
+      fix A assume "A \<in> sets E"
+      then obtain B C where "B \<in> sets M1" "C \<in> sets M2" "A = B \<times> C"
+        unfolding pair_measure_generator_def by auto
+      moreover have "(fst \<circ> f) -` B \<inter> space M \<in> sets M"
+        using f `B \<in> sets M1` unfolding measurable_def by auto
+      moreover have "(snd \<circ> f) -` C \<inter> space M \<in> sets M"
+        using s `C \<in> sets M2` unfolding measurable_def by auto
+      moreover have "f -` A \<inter> space M = ((fst \<circ> f) -` B \<inter> space M) \<inter> ((snd \<circ> f) -` C \<inter> space M)"
+        unfolding `A = B \<times> C` by (auto simp: vimage_Times)
+      ultimately show "f -` A \<inter> space M \<in> sets M" by auto
+    qed
+  qed
+qed
+
+lemma (in pair_sigma_algebra) measurable_pair:
+  assumes "sigma_algebra M"
+  assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2"
+  shows "f \<in> measurable M P"
+  unfolding measurable_pair_iff[OF assms(1)] using assms(2,3) by simp
+
+lemma pair_measure_generatorE:
+  assumes "X \<in> sets (pair_measure_generator M1 M2)"
+  obtains A B where "X = A \<times> B" "A \<in> sets M1" "B \<in> sets M2"
+  using assms unfolding pair_measure_generator_def by auto
+
+lemma (in pair_sigma_algebra) pair_measure_generator_swap:
+  "(\<lambda>X. (\<lambda>(x,y). (y,x)) -` X \<inter> space M2 \<times> space M1) ` sets E = sets (pair_measure_generator M2 M1)"
+proof (safe elim!: pair_measure_generatorE)
+  fix A B assume "A \<in> sets M1" "B \<in> sets M2"
+  moreover then have "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 = B \<times> A"
+    using M1.sets_into_space M2.sets_into_space by auto
+  ultimately show "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 \<in> sets (pair_measure_generator M2 M1)"
+    by (auto intro: pair_measure_generatorI)
+next
+  fix A B assume "A \<in> sets M1" "B \<in> sets M2"
+  then show "B \<times> A \<in> (\<lambda>X. (\<lambda>(x, y). (y, x)) -` X \<inter> space M2 \<times> space M1) ` sets E"
+    using M1.sets_into_space M2.sets_into_space
+    by (auto intro!: image_eqI[where x="A \<times> B"] pair_measure_generatorI)
+qed
+
+lemma (in pair_sigma_algebra) sets_pair_sigma_algebra_swap:
+  assumes Q: "Q \<in> sets P"
+  shows "(\<lambda>(x,y). (y, x)) -` Q \<in> sets (M2 \<Otimes>\<^isub>M M1)" (is "_ \<in> sets ?Q")
+proof -
+  let "?f Q" = "(\<lambda>(x,y). (y, x)) -` Q \<inter> space M2 \<times> space M1"
+  have *: "(\<lambda>(x,y). (y, x)) -` Q = ?f Q"
+    using sets_into_space[OF Q] by (auto simp: space_pair_measure)
+  have "sets (M2 \<Otimes>\<^isub>M M1) = sets (sigma (pair_measure_generator M2 M1))"
+    unfolding pair_measure_def ..
+  also have "\<dots> = sigma_sets (space M2 \<times> space M1) (?f ` sets E)"
+    unfolding sigma_def pair_measure_generator_swap[symmetric]
+    by (simp add: pair_measure_generator_def)
+  also have "\<dots> = ?f ` sigma_sets (space M1 \<times> space M2) (sets E)"
+    using M1.sets_into_space M2.sets_into_space
+    by (intro sigma_sets_vimage) (auto simp: pair_measure_generator_def)
+  also have "\<dots> = ?f ` sets P"
+    unfolding pair_measure_def pair_measure_generator_def sigma_def by simp
+  finally show ?thesis
+    using Q by (subst *) auto
+qed
+
+lemma (in pair_sigma_algebra) pair_sigma_algebra_swap_measurable:
+  shows "(\<lambda>(x,y). (y, x)) \<in> measurable P (M2 \<Otimes>\<^isub>M M1)"
+    (is "?f \<in> measurable ?P ?Q")
+  unfolding measurable_def
+proof (intro CollectI conjI Pi_I ballI)
+  fix x assume "x \<in> space ?P" then show "(case x of (x, y) \<Rightarrow> (y, x)) \<in> space ?Q"
+    unfolding pair_measure_generator_def pair_measure_def by auto
+next
+  fix A assume "A \<in> sets (M2 \<Otimes>\<^isub>M M1)"
+  interpret Q: pair_sigma_algebra M2 M1 by default
+  with Q.sets_pair_sigma_algebra_swap[OF `A \<in> sets (M2 \<Otimes>\<^isub>M M1)`]
+  show "?f -` A \<inter> space ?P \<in> sets ?P" by simp
+qed
+
+lemma (in pair_sigma_algebra) measurable_cut_fst[simp,intro]:
+  assumes "Q \<in> sets P" shows "Pair x -` Q \<in> sets M2"
+proof -
+  let ?Q' = "{Q. Q \<subseteq> space P \<and> Pair x -` Q \<in> sets M2}"
+  let ?Q = "\<lparr> space = space P, sets = ?Q' \<rparr>"
+  interpret Q: sigma_algebra ?Q
+    proof qed (auto simp: vimage_UN vimage_Diff space_pair_measure)
+  have "sets E \<subseteq> sets ?Q"
+    using M1.sets_into_space M2.sets_into_space
+    by (auto simp: pair_measure_generator_def space_pair_measure)
+  then have "sets P \<subseteq> sets ?Q"
+    apply (subst pair_measure_def, intro Q.sets_sigma_subset)
+    by (simp add: pair_measure_def)
+  with assms show ?thesis by auto
+qed
+
+lemma (in pair_sigma_algebra) measurable_cut_snd:
+  assumes Q: "Q \<in> sets P" shows "(\<lambda>x. (x, y)) -` Q \<in> sets M1" (is "?cut Q \<in> sets M1")
+proof -
+  interpret Q: pair_sigma_algebra M2 M1 by default
+  with Q.measurable_cut_fst[OF sets_pair_sigma_algebra_swap[OF Q], of y]
+  show ?thesis by (simp add: vimage_compose[symmetric] comp_def)
+qed
+
+lemma (in pair_sigma_algebra) measurable_pair_image_snd:
+  assumes m: "f \<in> measurable P M" and "x \<in> space M1"
+  shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
+  unfolding measurable_def
+proof (intro CollectI conjI Pi_I ballI)
+  fix y assume "y \<in> space M2" with `f \<in> measurable P M` `x \<in> space M1`
+  show "f (x, y) \<in> space M"
+    unfolding measurable_def pair_measure_generator_def pair_measure_def by auto
+next
+  fix A assume "A \<in> sets M"
+  then have "Pair x -` (f -` A \<inter> space P) \<in> sets M2" (is "?C \<in> _")
+    using `f \<in> measurable P M`
+    by (intro measurable_cut_fst) (auto simp: measurable_def)
+  also have "?C = (\<lambda>y. f (x, y)) -` A \<inter> space M2"
+    using `x \<in> space M1` by (auto simp: pair_measure_generator_def pair_measure_def)
+  finally show "(\<lambda>y. f (x, y)) -` A \<inter> space M2 \<in> sets M2" .
+qed
+
+lemma (in pair_sigma_algebra) measurable_pair_image_fst:
+  assumes m: "f \<in> measurable P M" and "y \<in> space M2"
+  shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"
+proof -
+  interpret Q: pair_sigma_algebra M2 M1 by default
+  from Q.measurable_pair_image_snd[OF measurable_comp `y \<in> space M2`,
+                                      OF Q.pair_sigma_algebra_swap_measurable m]
+  show ?thesis by simp
+qed
+
+lemma (in pair_sigma_algebra) Int_stable_pair_measure_generator: "Int_stable E"
+  unfolding Int_stable_def
+proof (intro ballI)
+  fix A B assume "A \<in> sets E" "B \<in> sets E"
+  then obtain A1 A2 B1 B2 where "A = A1 \<times> A2" "B = B1 \<times> B2"
+    "A1 \<in> sets M1" "A2 \<in> sets M2" "B1 \<in> sets M1" "B2 \<in> sets M2"
+    unfolding pair_measure_generator_def by auto
+  then show "A \<inter> B \<in> sets E"
+    by (auto simp add: times_Int_times pair_measure_generator_def)
+qed
+
+lemma finite_measure_cut_measurable:
+  fixes M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
+  assumes "sigma_finite_measure M1" "finite_measure M2"
+  assumes "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+  shows "(\<lambda>x. measure M2 (Pair x -` Q)) \<in> borel_measurable M1"
+    (is "?s Q \<in> _")
+proof -
+  interpret M1: sigma_finite_measure M1 by fact
+  interpret M2: finite_measure M2 by fact
+  interpret pair_sigma_algebra M1 M2 by default
+  have [intro]: "sigma_algebra M1" by fact
+  have [intro]: "sigma_algebra M2" by fact
+  let ?D = "\<lparr> space = space P, sets = {A\<in>sets P. ?s A \<in> borel_measurable M1}  \<rparr>"
+  note space_pair_measure[simp]
+  interpret dynkin_system ?D
+  proof (intro dynkin_systemI)
+    fix A assume "A \<in> sets ?D" then show "A \<subseteq> space ?D"
+      using sets_into_space by simp
+  next
+    from top show "space ?D \<in> sets ?D"
+      by (auto simp add: if_distrib intro!: M1.measurable_If)
+  next
+    fix A assume "A \<in> sets ?D"
+    with sets_into_space have "\<And>x. measure M2 (Pair x -` (space M1 \<times> space M2 - A)) =
+        (if x \<in> space M1 then measure M2 (space M2) - ?s A x else 0)"
+      by (auto intro!: M2.measure_compl simp: vimage_Diff)
+    with `A \<in> sets ?D` top show "space ?D - A \<in> sets ?D"
+      by (auto intro!: Diff M1.measurable_If M1.borel_measurable_extreal_diff)
+  next
+    fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> sets ?D"
+    moreover then have "\<And>x. measure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
+      by (intro M2.measure_countably_additive[symmetric])
+         (auto simp: disjoint_family_on_def)
+    ultimately show "(\<Union>i. F i) \<in> sets ?D"
+      by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf)
+  qed
+  have "sets P = sets ?D" apply (subst pair_measure_def)
+  proof (intro dynkin_lemma)
+    show "Int_stable E" by (rule Int_stable_pair_measure_generator)
+    from M1.sets_into_space have "\<And>A. A \<in> sets M1 \<Longrightarrow> {x \<in> space M1. x \<in> A} = A"
+      by auto
+    then show "sets E \<subseteq> sets ?D"
+      by (auto simp: pair_measure_generator_def sets_sigma if_distrib
+               intro: sigma_sets.Basic intro!: M1.measurable_If)
+  qed (auto simp: pair_measure_def)
+  with `Q \<in> sets P` have "Q \<in> sets ?D" by simp
+  then show "?s Q \<in> borel_measurable M1" by simp
+qed
+
+subsection {* Binary products of $\sigma$-finite measure spaces *}
+
+locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2
+  for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
+
+sublocale pair_sigma_finite \<subseteq> pair_sigma_algebra M1 M2
+  by default
+
+lemma times_eq_iff: "A \<times> B = C \<times> D \<longleftrightarrow> A = C \<and> B = D \<or> ((A = {} \<or> B = {}) \<and> (C = {} \<or> D = {}))"
+  by auto
+
+lemma sigma_sets_subseteq: assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
+proof
+  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
+    by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros)
+qed
+
+lemma (in pair_sigma_finite) measure_cut_measurable_fst:
+  assumes "Q \<in> sets P" shows "(\<lambda>x. measure M2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _")
+proof -
+  have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+
+  have M1: "sigma_finite_measure M1" by default
+  from M2.disjoint_sigma_finite guess F .. note F = this
+  then have F_sets: "\<And>i. F i \<in> sets M2" by auto
+  let "?C x i" = "F i \<inter> Pair x -` Q"
+  { fix i
+    let ?R = "M2.restricted_space (F i)"
+    have [simp]: "space M1 \<times> F i \<inter> space M1 \<times> space M2 = space M1 \<times> F i"
+      using F M2.sets_into_space by auto
+    let ?R2 = "M2.restricted_space (F i)"
+    have "(\<lambda>x. measure ?R2 (Pair x -` (space M1 \<times> space ?R2 \<inter> Q))) \<in> borel_measurable M1"
+    proof (intro finite_measure_cut_measurable[OF M1])
+      show "finite_measure ?R2"
+        using F by (intro M2.restricted_to_finite_measure) auto
+      have "(space M1 \<times> space ?R2) \<inter> Q \<in> (op \<inter> (space M1 \<times> F i)) ` sets P"
+        using `Q \<in> sets P` by (auto simp: image_iff)
+      also have "\<dots> = sigma_sets (space M1 \<times> F i) ((op \<inter> (space M1 \<times> F i)) ` sets E)"
+        unfolding pair_measure_def pair_measure_generator_def sigma_def
+        using `F i \<in> sets M2` M2.sets_into_space
+        by (auto intro!: sigma_sets_Int sigma_sets.Basic)
+      also have "\<dots> \<subseteq> sets (M1 \<Otimes>\<^isub>M ?R2)"
+        using M1.sets_into_space
+        apply (auto simp: times_Int_times pair_measure_def pair_measure_generator_def sigma_def
+                    intro!: sigma_sets_subseteq)
+        apply (rule_tac x="a" in exI)
+        apply (rule_tac x="b \<inter> F i" in exI)
+        by auto
+      finally show "(space M1 \<times> space ?R2) \<inter> Q \<in> sets (M1 \<Otimes>\<^isub>M ?R2)" .
+    qed
+    moreover have "\<And>x. Pair x -` (space M1 \<times> F i \<inter> Q) = ?C x i"
+      using `Q \<in> sets P` sets_into_space by (auto simp: space_pair_measure)
+    ultimately have "(\<lambda>x. measure M2 (?C x i)) \<in> borel_measurable M1"
+      by simp }
+  moreover
+  { fix x
+    have "(\<Sum>i. measure M2 (?C x i)) = measure M2 (\<Union>i. ?C x i)"
+    proof (intro M2.measure_countably_additive)
+      show "range (?C x) \<subseteq> sets M2"
+        using F `Q \<in> sets P` by (auto intro!: M2.Int)
+      have "disjoint_family F" using F by auto
+      show "disjoint_family (?C x)"
+        by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto
+    qed
+    also have "(\<Union>i. ?C x i) = Pair x -` Q"
+      using F sets_into_space `Q \<in> sets P`
+      by (auto simp: space_pair_measure)
+    finally have "measure M2 (Pair x -` Q) = (\<Sum>i. measure M2 (?C x i))"
+      by simp }
+  ultimately show ?thesis using `Q \<in> sets P` F_sets
+    by (auto intro!: M1.borel_measurable_psuminf M2.Int)
+qed
+
+lemma (in pair_sigma_finite) measure_cut_measurable_snd:
+  assumes "Q \<in> sets P" shows "(\<lambda>y. M1.\<mu> ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
+proof -
+  interpret Q: pair_sigma_finite M2 M1 by default
+  note sets_pair_sigma_algebra_swap[OF assms]
+  from Q.measure_cut_measurable_fst[OF this]
+  show ?thesis by (simp add: vimage_compose[symmetric] comp_def)
+qed
+
+lemma (in pair_sigma_algebra) pair_sigma_algebra_measurable:
+  assumes "f \<in> measurable P M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^isub>M M1) M"
+proof -
+  interpret Q: pair_sigma_algebra M2 M1 by default
+  have *: "(\<lambda>(x,y). f (y, x)) = f \<circ> (\<lambda>(x,y). (y, x))" by (simp add: fun_eq_iff)
+  show ?thesis
+    using Q.pair_sigma_algebra_swap_measurable assms
+    unfolding * by (rule measurable_comp)
+qed
+
+lemma (in pair_sigma_finite) pair_measure_alt:
+  assumes "A \<in> sets P"
+  shows "measure (M1 \<Otimes>\<^isub>M M2) A = (\<integral>\<^isup>+ x. measure M2 (Pair x -` A) \<partial>M1)"
+  apply (simp add: pair_measure_def pair_measure_generator_def)
+proof (rule M1.positive_integral_cong)
+  fix x assume "x \<in> space M1"
+  have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: extreal)"
+    unfolding indicator_def by auto
+  show "(\<integral>\<^isup>+ y. indicator A (x, y) \<partial>M2) = measure M2 (Pair x -` A)"
+    unfolding *
+    apply (subst M2.positive_integral_indicator)
+    apply (rule measurable_cut_fst[OF assms])
+    by simp
+qed
+
+lemma (in pair_sigma_finite) pair_measure_times:
+  assumes A: "A \<in> sets M1" and "B \<in> sets M2"
+  shows "measure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = M1.\<mu> A * measure M2 B"
+proof -
+  have "measure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = (\<integral>\<^isup>+ x. measure M2 B * indicator A x \<partial>M1)"
+    using assms by (auto intro!: M1.positive_integral_cong simp: pair_measure_alt)
+  with assms show ?thesis
+    by (simp add: M1.positive_integral_cmult_indicator ac_simps)
+qed
+
+lemma (in measure_space) measure_not_negative[simp,intro]:
+  assumes A: "A \<in> sets M" shows "\<mu> A \<noteq> - \<infinity>"
+  using positive_measure[OF A] by auto
+
+lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
+  "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets E \<and> incseq F \<and> (\<Union>i. F i) = space E \<and>
+    (\<forall>i. measure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
+proof -
+  obtain F1 :: "nat \<Rightarrow> 'a set" and F2 :: "nat \<Rightarrow> 'b set" where
+    F1: "range F1 \<subseteq> sets M1" "incseq F1" "(\<Union>i. F1 i) = space M1" "\<And>i. M1.\<mu> (F1 i) \<noteq> \<infinity>" and
+    F2: "range F2 \<subseteq> sets M2" "incseq F2" "(\<Union>i. F2 i) = space M2" "\<And>i. M2.\<mu> (F2 i) \<noteq> \<infinity>"
+    using M1.sigma_finite_up M2.sigma_finite_up by auto
+  then have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto
+  let ?F = "\<lambda>i. F1 i \<times> F2 i"
+  show ?thesis unfolding space_pair_measure
+  proof (intro exI[of _ ?F] conjI allI)
+    show "range ?F \<subseteq> sets E" using F1 F2
+      by (fastsimp intro!: pair_measure_generatorI)
+  next
+    have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"
+    proof (intro subsetI)
+      fix x assume "x \<in> space M1 \<times> space M2"
+      then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j"
+        by (auto simp: space)
+      then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)"
+        using `incseq F1` `incseq F2` unfolding incseq_def
+        by (force split: split_max)+
+      then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)"
+        by (intro SigmaI) (auto simp add: min_max.sup_commute)
+      then show "x \<in> (\<Union>i. ?F i)" by auto
+    qed
+    then show "(\<Union>i. ?F i) = space E"
+      using space by (auto simp: space pair_measure_generator_def)
+  next
+    fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)"
+      using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto
+  next
+    fix i
+    from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
+    with F1 F2 M1.positive_measure[OF this(1)] M2.positive_measure[OF this(2)]
+    show "measure P (F1 i \<times> F2 i) \<noteq> \<infinity>"
+      by (simp add: pair_measure_times)
+  qed
+qed
+
+sublocale pair_sigma_finite \<subseteq> sigma_finite_measure P
+proof
+  show "positive P (measure P)"
+    unfolding pair_measure_def pair_measure_generator_def sigma_def positive_def
+    by (auto intro: M1.positive_integral_positive M2.positive_integral_positive)
+
+  show "countably_additive P (measure P)"
+    unfolding countably_additive_def
+  proof (intro allI impI)
+    fix F :: "nat \<Rightarrow> ('a \<times> 'b) set"
+    assume F: "range F \<subseteq> sets P" "disjoint_family F"
+    from F have *: "\<And>i. F i \<in> sets P" "(\<Union>i. F i) \<in> sets P" by auto
+    moreover from F have "\<And>i. (\<lambda>x. measure M2 (Pair x -` F i)) \<in> borel_measurable M1"
+      by (intro measure_cut_measurable_fst) auto
+    moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"
+      by (intro disjoint_family_on_bisimulation[OF F(2)]) auto
+    moreover have "\<And>x. x \<in> space M1 \<Longrightarrow> range (\<lambda>i. Pair x -` F i) \<subseteq> sets M2"
+      using F by auto
+    ultimately show "(\<Sum>n. measure P (F n)) = measure P (\<Union>i. F i)"
+      by (simp add: pair_measure_alt vimage_UN M1.positive_integral_suminf[symmetric]
+                    M2.measure_countably_additive
+               cong: M1.positive_integral_cong)
+  qed
+
+  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
+  show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets P \<and> (\<Union>i. F i) = space P \<and> (\<forall>i. measure P (F i) \<noteq> \<infinity>)"
+  proof (rule exI[of _ F], intro conjI)
+    show "range F \<subseteq> sets P" using F by (auto simp: pair_measure_def)
+    show "(\<Union>i. F i) = space P"
+      using F by (auto simp: pair_measure_def pair_measure_generator_def)
+    show "\<forall>i. measure P (F i) \<noteq> \<infinity>" using F by auto
+  qed
+qed
+
+lemma (in pair_sigma_algebra) sets_swap:
+  assumes "A \<in> sets P"
+  shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
+    (is "_ -` A \<inter> space ?Q \<in> sets ?Q")
+proof -
+  have *: "(\<lambda>(x, y). (y, x)) -` A \<inter> space ?Q = (\<lambda>(x, y). (y, x)) -` A"
+    using `A \<in> sets P` sets_into_space by (auto simp: space_pair_measure)
+  show ?thesis
+    unfolding * using assms by (rule sets_pair_sigma_algebra_swap)
+qed
+
+lemma (in pair_sigma_finite) pair_measure_alt2:
+  assumes A: "A \<in> sets P"
+  shows "\<mu> A = (\<integral>\<^isup>+y. M1.\<mu> ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
+    (is "_ = ?\<nu> A")
+proof -
+  interpret Q: pair_sigma_finite M2 M1 by default
+  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
+  have [simp]: "\<And>m. \<lparr> space = space E, sets = sets (sigma E), measure = m \<rparr> = P\<lparr> measure := m \<rparr>"
+    unfolding pair_measure_def by simp
+
+  have "\<mu> A = Q.\<mu> ((\<lambda>(y, x). (x, y)) -` A \<inter> space Q.P)"
+  proof (rule measure_unique_Int_stable_vimage[OF Int_stable_pair_measure_generator])
+    show "measure_space P" "measure_space Q.P" by default
+    show "(\<lambda>(y, x). (x, y)) \<in> measurable Q.P P" by (rule Q.pair_sigma_algebra_swap_measurable)
+    show "sets (sigma E) = sets P" "space E = space P" "A \<in> sets (sigma E)"
+      using assms unfolding pair_measure_def by auto
+    show "range F \<subseteq> sets E" "incseq F" "(\<Union>i. F i) = space E" "\<And>i. \<mu> (F i) \<noteq> \<infinity>"
+      using F `A \<in> sets P` by (auto simp: pair_measure_def)
+    fix X assume "X \<in> sets E"
+    then obtain A B where X[simp]: "X = A \<times> B" and AB: "A \<in> sets M1" "B \<in> sets M2"
+      unfolding pair_measure_def pair_measure_generator_def by auto
+    then have "(\<lambda>(y, x). (x, y)) -` X \<inter> space Q.P = B \<times> A"
+      using M1.sets_into_space M2.sets_into_space by (auto simp: space_pair_measure)
+    then show "\<mu> X = Q.\<mu> ((\<lambda>(y, x). (x, y)) -` X \<inter> space Q.P)"
+      using AB by (simp add: pair_measure_times Q.pair_measure_times ac_simps)
+  qed
+  then show ?thesis
+    using sets_into_space[OF A] Q.pair_measure_alt[OF sets_swap[OF A]]
+    by (auto simp add: Q.pair_measure_alt space_pair_measure
+             intro!: M2.positive_integral_cong arg_cong[where f="M1.\<mu>"])
+qed
+
+lemma pair_sigma_algebra_sigma:
+  assumes 1: "incseq S1" "(\<Union>i. S1 i) = space E1" "range S1 \<subseteq> sets E1" and E1: "sets E1 \<subseteq> Pow (space E1)"
+  assumes 2: "decseq S2" "(\<Union>i. S2 i) = space E2" "range S2 \<subseteq> sets E2" and E2: "sets E2 \<subseteq> Pow (space E2)"
+  shows "sets (sigma (pair_measure_generator (sigma E1) (sigma E2))) = sets (sigma (pair_measure_generator E1 E2))"
+    (is "sets ?S = sets ?E")
+proof -
+  interpret M1: sigma_algebra "sigma E1" using E1 by (rule sigma_algebra_sigma)
+  interpret M2: sigma_algebra "sigma E2" using E2 by (rule sigma_algebra_sigma)
+  have P: "sets (pair_measure_generator E1 E2) \<subseteq> Pow (space E1 \<times> space E2)"
+    using E1 E2 by (auto simp add: pair_measure_generator_def)
+  interpret E: sigma_algebra ?E unfolding pair_measure_generator_def
+    using E1 E2 by (intro sigma_algebra_sigma) auto
+  { fix A assume "A \<in> sets E1"
+    then have "fst -` A \<inter> space ?E = A \<times> (\<Union>i. S2 i)"
+      using E1 2 unfolding pair_measure_generator_def by auto
+    also have "\<dots> = (\<Union>i. A \<times> S2 i)" by auto
+    also have "\<dots> \<in> sets ?E" unfolding pair_measure_generator_def sets_sigma
+      using 2 `A \<in> sets E1`
+      by (intro sigma_sets.Union)
+         (force simp: image_subset_iff intro!: sigma_sets.Basic)
+    finally have "fst -` A \<inter> space ?E \<in> sets ?E" . }
+  moreover
+  { fix B assume "B \<in> sets E2"
+    then have "snd -` B \<inter> space ?E = (\<Union>i. S1 i) \<times> B"
+      using E2 1 unfolding pair_measure_generator_def by auto
+    also have "\<dots> = (\<Union>i. S1 i \<times> B)" by auto
+    also have "\<dots> \<in> sets ?E"
+      using 1 `B \<in> sets E2` unfolding pair_measure_generator_def sets_sigma
+      by (intro sigma_sets.Union)
+         (force simp: image_subset_iff intro!: sigma_sets.Basic)
+    finally have "snd -` B \<inter> space ?E \<in> sets ?E" . }
+  ultimately have proj:
+    "fst \<in> measurable ?E (sigma E1) \<and> snd \<in> measurable ?E (sigma E2)"
+    using E1 E2 by (subst (1 2) E.measurable_iff_sigma)
+                   (auto simp: pair_measure_generator_def sets_sigma)
+  { fix A B assume A: "A \<in> sets (sigma E1)" and B: "B \<in> sets (sigma E2)"
+    with proj have "fst -` A \<inter> space ?E \<in> sets ?E" "snd -` B \<inter> space ?E \<in> sets ?E"
+      unfolding measurable_def by simp_all
+    moreover have "A \<times> B = (fst -` A \<inter> space ?E) \<inter> (snd -` B \<inter> space ?E)"
+      using A B M1.sets_into_space M2.sets_into_space
+      by (auto simp: pair_measure_generator_def)
+    ultimately have "A \<times> B \<in> sets ?E" by auto }
+  then have "sigma_sets (space ?E) (sets (pair_measure_generator (sigma E1) (sigma E2))) \<subseteq> sets ?E"
+    by (intro E.sigma_sets_subset) (auto simp add: pair_measure_generator_def sets_sigma)
+  then have subset: "sets ?S \<subseteq> sets ?E"
+    by (simp add: sets_sigma pair_measure_generator_def)
+  show "sets ?S = sets ?E"
+  proof (intro set_eqI iffI)
+    fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S"
+      unfolding sets_sigma
+    proof induct
+      case (Basic A) then show ?case
+        by (auto simp: pair_measure_generator_def sets_sigma intro: sigma_sets.Basic)
+    qed (auto intro: sigma_sets.intros simp: pair_measure_generator_def)
+  next
+    fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
+  qed
+qed
+
+section "Fubinis theorem"
+
+lemma (in pair_sigma_finite) simple_function_cut:
+  assumes f: "simple_function P f" "\<And>x. 0 \<le> f x"
+  shows "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
+    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
+proof -
+  have f_borel: "f \<in> borel_measurable P"
+    using f(1) by (rule borel_measurable_simple_function)
+  let "?F z" = "f -` {z} \<inter> space P"
+  let "?F' x z" = "Pair x -` ?F z"
+  { fix x assume "x \<in> space M1"
+    have [simp]: "\<And>z y. indicator (?F z) (x, y) = indicator (?F' x z) y"
+      by (auto simp: indicator_def)
+    have "\<And>y. y \<in> space M2 \<Longrightarrow> (x, y) \<in> space P" using `x \<in> space M1`
+      by (simp add: space_pair_measure)
+    moreover have "\<And>x z. ?F' x z \<in> sets M2" using f_borel
+      by (intro borel_measurable_vimage measurable_cut_fst)
+    ultimately have "simple_function M2 (\<lambda> y. f (x, y))"
+      apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _])
+      apply (rule simple_function_indicator_representation[OF f(1)])
+      using `x \<in> space M1` by (auto simp del: space_sigma) }
+  note M2_sf = this
+  { fix x assume x: "x \<in> space M1"
+    then have "(\<integral>\<^isup>+y. f (x, y) \<partial>M2) = (\<Sum>z\<in>f ` space P. z * M2.\<mu> (?F' x z))"
+      unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)]
+      unfolding simple_integral_def
+    proof (safe intro!: setsum_mono_zero_cong_left)
+      from f(1) show "finite (f ` space P)" by (rule simple_functionD)
+    next
+      fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space P"
+        using `x \<in> space M1` by (auto simp: space_pair_measure)
+    next
+      fix x' y assume "(x', y) \<in> space P"
+        "f (x', y) \<notin> (\<lambda>y. f (x, y)) ` space M2"
+      then have *: "?F' x (f (x', y)) = {}"
+        by (force simp: space_pair_measure)
+      show  "f (x', y) * M2.\<mu> (?F' x (f (x', y))) = 0"
+        unfolding * by simp
+    qed (simp add: vimage_compose[symmetric] comp_def
+                   space_pair_measure) }
+  note eq = this
+  moreover have "\<And>z. ?F z \<in> sets P"
+    by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma)
+  moreover then have "\<And>z. (\<lambda>x. M2.\<mu> (?F' x z)) \<in> borel_measurable M1"
+    by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int)
+  moreover have *: "\<And>i x. 0 \<le> M2.\<mu> (Pair x -` (f -` {i} \<inter> space P))"
+    using f(1)[THEN simple_functionD(2)] f(2) by (intro M2.positive_measure measurable_cut_fst)
+  moreover { fix i assume "i \<in> f`space P"
+    with * have "\<And>x. 0 \<le> i * M2.\<mu> (Pair x -` (f -` {i} \<inter> space P))"
+      using f(2) by auto }
+  ultimately
+  show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
+    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f" using f(2)
+    by (auto simp del: vimage_Int cong: measurable_cong
+             intro!: M1.borel_measurable_extreal_setsum setsum_cong
+             simp add: M1.positive_integral_setsum simple_integral_def
+                       M1.positive_integral_cmult
+                       M1.positive_integral_cong[OF eq]
+                       positive_integral_eq_simple_integral[OF f]
+                       pair_measure_alt[symmetric])
+qed
+
+lemma (in pair_sigma_finite) positive_integral_fst_measurable:
+  assumes f: "f \<in> borel_measurable P"
+  shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
+      (is "?C f \<in> borel_measurable M1")
+    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
+proof -
+  from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this
+  then have F_borel: "\<And>i. F i \<in> borel_measurable P"
+    by (auto intro: borel_measurable_simple_function)
+  note sf = simple_function_cut[OF F(1,5)]
+  then have "(\<lambda>x. SUP i. ?C (F i) x) \<in> borel_measurable M1"
+    using F(1) by auto
+  moreover
+  { fix x assume "x \<in> space M1"
+    from F measurable_pair_image_snd[OF F_borel`x \<in> space M1`]
+    have "(\<integral>\<^isup>+y. (SUP i. F i (x, y)) \<partial>M2) = (SUP i. ?C (F i) x)"
+      by (intro M2.positive_integral_monotone_convergence_SUP)
+         (auto simp: incseq_Suc_iff le_fun_def)
+    then have "(SUP i. ?C (F i) x) = ?C f x"
+      unfolding F(4) positive_integral_max_0 by simp }
+  note SUPR_C = this
+  ultimately show "?C f \<in> borel_measurable M1"
+    by (simp cong: measurable_cong)
+  have "(\<integral>\<^isup>+x. (SUP i. F i x) \<partial>P) = (SUP i. integral\<^isup>P P (F i))"
+    using F_borel F
+    by (intro positive_integral_monotone_convergence_SUP) auto
+  also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1)"
+    unfolding sf(2) by simp
+  also have "\<dots> = \<integral>\<^isup>+ x. (SUP i. \<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1" using F sf(1)
+    by (intro M1.positive_integral_monotone_convergence_SUP[symmetric])
+       (auto intro!: M2.positive_integral_mono M2.positive_integral_positive
+                simp: incseq_Suc_iff le_fun_def)
+  also have "\<dots> = \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. (SUP i. F i (x, y)) \<partial>M2) \<partial>M1"
+    using F_borel F(2,5)
+    by (auto intro!: M1.positive_integral_cong M2.positive_integral_monotone_convergence_SUP[symmetric]
+             simp: incseq_Suc_iff le_fun_def measurable_pair_image_snd)
+  finally show "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
+    using F by (simp add: positive_integral_max_0)
+qed
+
+lemma (in pair_sigma_finite) measure_preserving_swap:
+  "(\<lambda>(x,y). (y, x)) \<in> measure_preserving (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
+proof
+  interpret Q: pair_sigma_finite M2 M1 by default
+  show *: "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
+    using pair_sigma_algebra_swap_measurable .
+  fix X assume "X \<in> sets (M2 \<Otimes>\<^isub>M M1)"
+  from measurable_sets[OF * this] this Q.sets_into_space[OF this]
+  show "measure (M1 \<Otimes>\<^isub>M M2) ((\<lambda>(x, y). (y, x)) -` X \<inter> space P) = measure (M2 \<Otimes>\<^isub>M M1) X"
+    by (auto intro!: M1.positive_integral_cong arg_cong[where f="M2.\<mu>"]
+      simp: pair_measure_alt Q.pair_measure_alt2 space_pair_measure)
+qed
+
+lemma (in pair_sigma_finite) positive_integral_product_swap:
+  assumes f: "f \<in> borel_measurable P"
+  shows "(\<integral>\<^isup>+x. f (case x of (x,y)\<Rightarrow>(y,x)) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P P f"
+proof -
+  interpret Q: pair_sigma_finite M2 M1 by default
+  have "sigma_algebra P" by default
+  with f show ?thesis
+    by (subst Q.positive_integral_vimage[OF _ Q.measure_preserving_swap]) auto
+qed
+
+lemma (in pair_sigma_finite) positive_integral_snd_measurable:
+  assumes f: "f \<in> borel_measurable P"
+  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P P f"
+proof -
+  interpret Q: pair_sigma_finite M2 M1 by default
+  note pair_sigma_algebra_measurable[OF f]
+  from Q.positive_integral_fst_measurable[OF this]
+  have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>Q.P)"
+    by simp
+  also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>Q.P) = integral\<^isup>P P f"
+    unfolding positive_integral_product_swap[OF f, symmetric]
+    by (auto intro!: Q.positive_integral_cong)
+  finally show ?thesis .
+qed
+
+lemma (in pair_sigma_finite) Fubini:
+  assumes f: "f \<in> borel_measurable P"
+  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
+  unfolding positive_integral_snd_measurable[OF assms]
+  unfolding positive_integral_fst_measurable[OF assms] ..
+
+lemma (in pair_sigma_finite) AE_pair:
+  assumes "AE x in P. Q x"
+  shows "AE x in M1. (AE y in M2. Q (x, y))"
+proof -
+  obtain N where N: "N \<in> sets P" "\<mu> N = 0" "{x\<in>space P. \<not> Q x} \<subseteq> N"
+    using assms unfolding almost_everywhere_def by auto
+  show ?thesis
+  proof (rule M1.AE_I)
+    from N measure_cut_measurable_fst[OF `N \<in> sets P`]
+    show "M1.\<mu> {x\<in>space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} = 0"
+      by (auto simp: pair_measure_alt M1.positive_integral_0_iff)
+    show "{x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} \<in> sets M1"
+      by (intro M1.borel_measurable_extreal_neq_const measure_cut_measurable_fst N)
+    { fix x assume "x \<in> space M1" "M2.\<mu> (Pair x -` N) = 0"
+      have "M2.almost_everywhere (\<lambda>y. Q (x, y))"
+      proof (rule M2.AE_I)
+        show "M2.\<mu> (Pair x -` N) = 0" by fact
+        show "Pair x -` N \<in> sets M2" by (intro measurable_cut_fst N)
+        show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
+          using N `x \<in> space M1` unfolding space_sigma space_pair_measure by auto
+      qed }
+    then show "{x \<in> space M1. \<not> M2.almost_everywhere (\<lambda>y. Q (x, y))} \<subseteq> {x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0}"
+      by auto
+  qed
+qed
+
+lemma (in pair_sigma_algebra) measurable_product_swap:
+  "f \<in> measurable (M2 \<Otimes>\<^isub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable P M"
+proof -
+  interpret Q: pair_sigma_algebra M2 M1 by default
+  show ?thesis
+    using pair_sigma_algebra_measurable[of "\<lambda>(x,y). f (y, x)"]
+    by (auto intro!: pair_sigma_algebra_measurable Q.pair_sigma_algebra_measurable iffI)
+qed
+
+lemma (in pair_sigma_finite) integrable_product_swap:
+  assumes "integrable P f"
+  shows "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x))"
+proof -
+  interpret Q: pair_sigma_finite M2 M1 by default
+  have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
+  show ?thesis unfolding *
+    using assms unfolding integrable_def
+    apply (subst (1 2) positive_integral_product_swap)
+    using `integrable P f` unfolding integrable_def
+    by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric])
+qed
+
+lemma (in pair_sigma_finite) integrable_product_swap_iff:
+  "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable P f"
+proof -
+  interpret Q: pair_sigma_finite M2 M1 by default
+  from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f]
+  show ?thesis by auto
+qed
+
+lemma (in pair_sigma_finite) integral_product_swap:
+  assumes "integrable P f"
+  shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>L P f"
+proof -
+  interpret Q: pair_sigma_finite M2 M1 by default
+  have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
+  show ?thesis
+    unfolding lebesgue_integral_def *
+    apply (subst (1 2) positive_integral_product_swap)
+    using `integrable P f` unfolding integrable_def
+    by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric])
+qed
+
+lemma (in pair_sigma_finite) integrable_fst_measurable:
+  assumes f: "integrable P f"
+  shows "M1.almost_everywhere (\<lambda>x. integrable M2 (\<lambda> y. f (x, y)))" (is "?AE")
+    and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L P f" (is "?INT")
+proof -
+  let "?pf x" = "extreal (f x)" and "?nf x" = "extreal (- f x)"
+  have
+    borel: "?nf \<in> borel_measurable P""?pf \<in> borel_measurable P" and
+    int: "integral\<^isup>P P ?nf \<noteq> \<infinity>" "integral\<^isup>P P ?pf \<noteq> \<infinity>"
+    using assms by auto
+  have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
+     "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
+    using borel[THEN positive_integral_fst_measurable(1)] int
+    unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all
+  with borel[THEN positive_integral_fst_measurable(1)]
+  have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
+    "AE x in M1. (\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
+    by (auto intro!: M1.positive_integral_PInf_AE )
+  then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
+    "AE x in M1. \<bar>\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
+    by (auto simp: M2.positive_integral_positive)
+  from AE_pos show ?AE using assms
+    by (simp add: measurable_pair_image_snd integrable_def)
+  { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. extreal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
+      using M2.positive_integral_positive
+      by (intro M1.positive_integral_cong_pos) (auto simp: extreal_uminus_le_reorder)
+    then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. extreal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
+  note this[simp]
+  { fix f assume borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable P"
+      and int: "integral\<^isup>P P (\<lambda>x. extreal (f x)) \<noteq> \<infinity>"
+      and AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<noteq> \<infinity>)"
+    have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
+    proof (intro integrable_def[THEN iffD2] conjI)
+      show "?f \<in> borel_measurable M1"
+        using borel by (auto intro!: M1.borel_measurable_real_of_extreal positive_integral_fst_measurable)
+      have "(\<integral>\<^isup>+x. extreal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (f (x, y))  \<partial>M2) \<partial>M1)"
+        using AE M2.positive_integral_positive
+        by (auto intro!: M1.positive_integral_cong_AE simp: extreal_real)
+      then show "(\<integral>\<^isup>+x. extreal (?f x) \<partial>M1) \<noteq> \<infinity>"
+        using positive_integral_fst_measurable[OF borel] int by simp
+      have "(\<integral>\<^isup>+x. extreal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
+        by (intro M1.positive_integral_cong_pos)
+           (simp add: M2.positive_integral_positive real_of_extreal_pos)
+      then show "(\<integral>\<^isup>+x. extreal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
+    qed }
+  with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)]
+  show ?INT
+    unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2]
+      borel[THEN positive_integral_fst_measurable(2), symmetric]
+    using AE[THEN M1.integral_real]
+    by simp
+qed
+
+lemma (in pair_sigma_finite) integrable_snd_measurable:
+  assumes f: "integrable P f"
+  shows "M2.almost_everywhere (\<lambda>y. integrable M1 (\<lambda>x. f (x, y)))" (is "?AE")
+    and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>L P f" (is "?INT")
+proof -
+  interpret Q: pair_sigma_finite M2 M1 by default
+  have Q_int: "integrable Q.P (\<lambda>(x, y). f (y, x))"
+    using f unfolding integrable_product_swap_iff .
+  show ?INT
+    using Q.integrable_fst_measurable(2)[OF Q_int]
+    using integral_product_swap[OF f] by simp
+  show ?AE
+    using Q.integrable_fst_measurable(1)[OF Q_int]
+    by simp
+qed
+
+lemma (in pair_sigma_finite) Fubini_integral:
+  assumes f: "integrable P f"
+  shows "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1)"
+  unfolding integrable_snd_measurable[OF assms]
+  unfolding integrable_fst_measurable[OF assms] ..
+
+section "Products on finite spaces"
+
+lemma sigma_sets_pair_measure_generator_finite:
+  assumes "finite A" and "finite B"
+  shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<in> Pow A \<and> b \<in> Pow B} = Pow (A \<times> B)"
+  (is "sigma_sets ?prod ?sets = _")
+proof safe
+  have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
+  fix x assume subset: "x \<subseteq> A \<times> B"
+  hence "finite x" using fin by (rule finite_subset)
+  from this subset show "x \<in> sigma_sets ?prod ?sets"
+  proof (induct x)
+    case empty show ?case by (rule sigma_sets.Empty)
+  next
+    case (insert a x)
+    hence "{a} \<in> sigma_sets ?prod ?sets"
+      by (auto simp: pair_measure_generator_def intro!: sigma_sets.Basic)
+    moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
+    ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)
+  qed
+next
+  fix x a b
+  assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x"
+  from sigma_sets_into_sp[OF _ this(1)] this(2)
+  show "a \<in> A" and "b \<in> B" by auto
+qed
+
+locale pair_finite_sigma_algebra = M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2
+  for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
+
+sublocale pair_finite_sigma_algebra \<subseteq> pair_sigma_algebra by default
+
+lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra:
+  shows "P = \<lparr> space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2), \<dots> = algebra.more P \<rparr>"
+proof -
+  show ?thesis
+    using sigma_sets_pair_measure_generator_finite[OF M1.finite_space M2.finite_space]
+    by (intro algebra.equality) (simp_all add: pair_measure_def pair_measure_generator_def sigma_def)
+qed
+
+sublocale pair_finite_sigma_algebra \<subseteq> finite_sigma_algebra P
+  apply default
+  using M1.finite_space M2.finite_space
+  apply (subst finite_pair_sigma_algebra) apply simp
+  apply (subst (1 2) finite_pair_sigma_algebra) apply simp
+  done
+
+locale pair_finite_space = M1: finite_measure_space M1 + M2: finite_measure_space M2
+  for M1 M2
+
+sublocale pair_finite_space \<subseteq> pair_finite_sigma_algebra
+  by default
+
+sublocale pair_finite_space \<subseteq> pair_sigma_finite
+  by default
+
+lemma (in pair_finite_space) pair_measure_Pair[simp]:
+  assumes "a \<in> space M1" "b \<in> space M2"
+  shows "\<mu> {(a, b)} = M1.\<mu> {a} * M2.\<mu> {b}"
+proof -
+  have "\<mu> ({a}\<times>{b}) = M1.\<mu> {a} * M2.\<mu> {b}"
+    using M1.sets_eq_Pow M2.sets_eq_Pow assms
+    by (subst pair_measure_times) auto
+  then show ?thesis by simp
+qed
+
+lemma (in pair_finite_space) pair_measure_singleton[simp]:
+  assumes "x \<in> space M1 \<times> space M2"
+  shows "\<mu> {x} = M1.\<mu> {fst x} * M2.\<mu> {snd x}"
+  using pair_measure_Pair assms by (cases x) auto
+
+sublocale pair_finite_space \<subseteq> finite_measure_space P
+  by default (auto simp: space_pair_measure)
+
+end
\ No newline at end of file
--- a/src/HOL/Probability/Borel_Space.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Probability/Borel_Space.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -1,4 +1,4 @@
-(*  Title:      HOL/Probability/Lebesgue_Integration.thy
+(*  Title:      HOL/Probability/Borel_Space.thy
     Author:     Johannes Hölzl, TU München
     Author:     Armin Heller, TU München
 *)
--- a/src/HOL/Probability/Caratheodory.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Probability/Caratheodory.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -258,7 +258,21 @@
     by (simp add: Un)
 qed
 
-lemma (in algebra) countably_subadditive_subadditive:
+lemma (in ring_of_sets) disjointed_additive:
+  assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> sets M" "incseq A"
+  shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
+proof (induct n)
+  case (Suc n)
+  then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
+    by simp
+  also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
+    using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_incseq)
+  also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
+    using `incseq A` by (auto dest: incseq_SucD simp: disjointed_incseq)
+  finally show ?case .
+qed simp
+
+lemma (in ring_of_sets) countably_subadditive_subadditive:
   assumes f: "positive M f" and cs: "countably_subadditive M f"
   shows  "subadditive M f"
 proof (auto simp add: subadditive_def)
@@ -277,7 +291,7 @@
     by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
 qed
 
-lemma (in algebra) additive_sum:
+lemma (in ring_of_sets) additive_sum:
   fixes A:: "nat \<Rightarrow> 'a set"
   assumes f: "positive M f" and ad: "additive M f" and "finite S"
       and A: "range A \<subseteq> sets M"
@@ -785,17 +799,17 @@
   moreover
   have "Inf (measure_set M f s)
        \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
-    proof -
+  proof -
     have "Inf (measure_set M f s) = Inf (measure_set M f ((s\<inter>x) \<union> (s-x)))"
       by (metis Un_Diff_Int Un_commute)
     also have "... \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
       apply (rule subadditiveD)
-      apply (rule algebra.countably_subadditive_subadditive[OF algebra_Pow])
+      apply (rule ring_of_sets.countably_subadditive_subadditive [OF ring_of_sets_Pow])
       apply (simp add: positive_def inf_measure_empty[OF posf] inf_measure_pos[OF posf])
       apply (rule inf_measure_countably_subadditive)
       using s by (auto intro!: posf inc)
     finally show ?thesis .
-    qed
+  qed
   ultimately
   show "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
         = Inf (measure_set M f s)"
@@ -837,4 +851,183 @@
     by (intro exI[of _ ?infm]) auto
 qed
 
+subsubsection {*Alternative instances of caratheodory*}
+
+lemma sums_def2:
+  "f sums x \<longleftrightarrow> (\<lambda>n. (\<Sum>i\<le>n. f i)) ----> x"
+  unfolding sums_def
+  apply (subst LIMSEQ_Suc_iff[symmetric])
+  unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost ..
+
+lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
+  assumes f: "positive M f" "additive M f"
+  shows "countably_additive M f \<longleftrightarrow>
+    (\<forall>A. range A \<subseteq> sets M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i))"
+  unfolding countably_additive_def
+proof safe
+  assume count_sum: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> UNION UNIV A \<in> sets M \<longrightarrow> (\<Sum>i. f (A i)) = f (UNION UNIV A)"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "incseq A" "(\<Union>i. A i) \<in> sets M"
+  then have dA: "range (disjointed A) \<subseteq> sets M" by (auto simp: range_disjointed_sets)
+  with count_sum[THEN spec, of "disjointed A"] A(3)
+  have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)"
+    by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
+  moreover have "(\<lambda>n. (\<Sum>i=0..<n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
+    using f(1)[unfolded positive_def] dA
+    by (auto intro!: summable_sumr_LIMSEQ_suminf summable_extreal_pos)
+  from LIMSEQ_Suc[OF this]
+  have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
+    unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost .
+  moreover have "\<And>n. (\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
+    using disjointed_additive[OF f A(1,2)] .
+  ultimately show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)" by simp
+next
+  assume cont: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family A" "(\<Union>i. A i) \<in> sets M"
+  have *: "(\<Union>n. (\<Union>i\<le>n. A i)) = (\<Union>i. A i)" by auto
+  have "(\<lambda>n. f (\<Union>i\<le>n. A i)) ----> f (\<Union>i. A i)"
+  proof (unfold *[symmetric], intro cont[rule_format])
+    show "range (\<lambda>i. \<Union> i\<le>i. A i) \<subseteq> sets M" "(\<Union>i. \<Union> i\<le>i. A i) \<in> sets M"
+      using A * by auto
+  qed (force intro!: incseq_SucI)
+  moreover have "\<And>n. f (\<Union>i\<le>n. A i) = (\<Sum>i\<le>n. f (A i))"
+    using A
+    by (intro additive_sum[OF f, of _ A, symmetric])
+       (auto intro: disjoint_family_on_mono[where B=UNIV])
+  ultimately
+  have "(\<lambda>i. f (A i)) sums f (\<Union>i. A i)"
+    unfolding sums_def2 by simp
+  from sums_unique[OF this]
+  show "(\<Sum>i. f (A i)) = f (\<Union>i. A i)" by simp
+qed
+
+lemma uminus_extreal_add_uminus_uminus:
+  fixes a b :: extreal shows "a \<noteq> \<infinity> \<Longrightarrow> b \<noteq> \<infinity> \<Longrightarrow> - (- a + - b) = a + b"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma INFI_extreal_add:
+  assumes "decseq f" "decseq g" and fin: "\<And>i. f i \<noteq> \<infinity>" "\<And>i. g i \<noteq> \<infinity>"
+  shows "(INF i. f i + g i) = INFI UNIV f + INFI UNIV g"
+proof -
+  have INF_less: "(INF i. f i) < \<infinity>" "(INF i. g i) < \<infinity>"
+    using assms unfolding INF_less_iff by auto
+  { fix i from fin[of i] have "- ((- f i) + (- g i)) = f i + g i"
+      by (rule uminus_extreal_add_uminus_uminus) }
+  then have "(INF i. f i + g i) = (INF i. - ((- f i) + (- g i)))"
+    by simp
+  also have "\<dots> = INFI UNIV f + INFI UNIV g"
+    unfolding extreal_INFI_uminus
+    using assms INF_less
+    by (subst SUPR_extreal_add)
+       (auto simp: extreal_SUPR_uminus intro!: uminus_extreal_add_uminus_uminus)
+  finally show ?thesis .
+qed
+
+lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous:
+  assumes f: "positive M f" "additive M f"
+  shows "(\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> sets M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))
+     \<longleftrightarrow> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0)"
+proof safe
+  assume cont: "(\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> sets M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
+  with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) ----> 0"
+    using `positive M f`[unfolded positive_def] by auto
+next
+  assume cont: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "decseq A" "(\<Inter>i. A i) \<in> sets M" "\<forall>i. f (A i) \<noteq> \<infinity>"
+
+  have f_mono: "\<And>a b. a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b"
+    using additive_increasing[OF f] unfolding increasing_def by simp
+
+  have decseq_fA: "decseq (\<lambda>i. f (A i))"
+    using A by (auto simp: decseq_def intro!: f_mono)
+  have decseq: "decseq (\<lambda>i. A i - (\<Inter>i. A i))"
+    using A by (auto simp: decseq_def)
+  then have decseq_f: "decseq (\<lambda>i. f (A i - (\<Inter>i. A i)))"
+    using A unfolding decseq_def by (auto intro!: f_mono Diff)
+  have "f (\<Inter>x. A x) \<le> f (A 0)"
+    using A by (auto intro!: f_mono)
+  then have f_Int_fin: "f (\<Inter>x. A x) \<noteq> \<infinity>"
+    using A by auto
+  { fix i
+    have "f (A i - (\<Inter>i. A i)) \<le> f (A i)" using A by (auto intro!: f_mono)
+    then have "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>"
+      using A by auto }
+  note f_fin = this
+  have "(\<lambda>i. f (A i - (\<Inter>i. A i))) ----> 0"
+  proof (intro cont[rule_format, OF _ decseq _ f_fin])
+    show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> sets M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
+      using A by auto
+  qed
+  from INF_Lim_extreal[OF decseq_f this]
+  have "(INF n. f (A n - (\<Inter>i. A i))) = 0" .
+  moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)"
+    by auto
+  ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)"
+    using A(4) f_fin f_Int_fin
+    by (subst INFI_extreal_add) (auto simp: decseq_f)
+  moreover {
+    fix n
+    have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f ((A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i))"
+      using A by (subst f(2)[THEN additiveD]) auto
+    also have "(A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i) = A n"
+      by auto
+    finally have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f (A n)" . }
+  ultimately have "(INF n. f (A n)) = f (\<Inter>i. A i)"
+    by simp
+  with LIMSEQ_extreal_INFI[OF decseq_fA]
+  show "(\<lambda>i. f (A i)) ----> f (\<Inter>i. A i)" by simp
+qed
+
+lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
+lemma positiveD2: "positive M f \<Longrightarrow> A \<in> sets M \<Longrightarrow> 0 \<le> f A" by (auto simp: positive_def)
+
+lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
+  assumes f: "positive M f" "additive M f" "\<forall>A\<in>sets M. f A \<noteq> \<infinity>"
+  assumes cont: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
+  assumes A: "range A \<subseteq> sets M" "incseq A" "(\<Union>i. A i) \<in> sets M"
+  shows "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
+proof -
+  have "\<forall>A\<in>sets M. \<exists>x. f A = extreal x"
+  proof
+    fix A assume "A \<in> sets M" with f show "\<exists>x. f A = extreal x"
+      unfolding positive_def by (cases "f A") auto
+  qed
+  from bchoice[OF this] guess f' .. note f' = this[rule_format]
+  from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) ----> 0"
+    by (intro cont[rule_format]) (auto simp: decseq_def incseq_def)
+  moreover
+  { fix i
+    have "f ((\<Union>i. A i) - A i) + f (A i) = f ((\<Union>i. A i) - A i \<union> A i)"
+      using A by (intro f(2)[THEN additiveD, symmetric]) auto
+    also have "(\<Union>i. A i) - A i \<union> A i = (\<Union>i. A i)"
+      by auto
+    finally have "f' (\<Union>i. A i) - f' (A i) = f' ((\<Union>i. A i) - A i)"
+      using A by (subst (asm) (1 2 3) f') auto
+    then have "f ((\<Union>i. A i) - A i) = extreal (f' (\<Union>i. A i) - f' (A i))"
+      using A f' by auto }
+  ultimately have "(\<lambda>i. f' (\<Union>i. A i) - f' (A i)) ----> 0"
+    by (simp add: zero_extreal_def)
+  then have "(\<lambda>i. f' (A i)) ----> f' (\<Union>i. A i)"
+    by (rule LIMSEQ_diff_approach_zero2[OF LIMSEQ_const])
+  then show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
+    using A by (subst (1 2) f') auto
+qed
+
+lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
+  assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>sets M. f A \<noteq> \<infinity>"
+  assumes cont: "\<And>A. range A \<subseteq> sets M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
+  shows "countably_additive M f"
+  using countably_additive_iff_continuous_from_below[OF f]
+  using empty_continuous_imp_continuous_from_below[OF f fin] cont
+  by blast
+
+lemma (in ring_of_sets) caratheodory_empty_continuous:
+  assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> sets M \<Longrightarrow> f A \<noteq> \<infinity>"
+  assumes cont: "\<And>A. range A \<subseteq> sets M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
+  shows "\<exists>\<mu> :: 'a set \<Rightarrow> extreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
+            measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
+proof (intro caratheodory empty_continuous_imp_countably_additive f)
+  show "\<forall>A\<in>sets M. f A \<noteq> \<infinity>" using fin by auto
+qed (rule cont)
+
 end
--- a/src/HOL/Probability/Complete_Measure.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Probability/Complete_Measure.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -3,7 +3,7 @@
 *)
 
 theory Complete_Measure
-imports Product_Measure
+imports Lebesgue_Integration
 begin
 
 locale completeable_measure_space = measure_space
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Finite_Product_Measure.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -0,0 +1,1014 @@
+(*  Title:      HOL/Probability/Finite_Product_Measure.thy
+    Author:     Johannes Hölzl, TU München
+*)
+
+header {*Finite product measures*}
+
+theory Finite_Product_Measure
+imports Binary_Product_Measure
+begin
+
+lemma Pi_iff: "f \<in> Pi I X \<longleftrightarrow> (\<forall>i\<in>I. f i \<in> X i)"
+  unfolding Pi_def by auto
+
+abbreviation
+  "Pi\<^isub>E A B \<equiv> Pi A B \<inter> extensional A"
+
+syntax
+  "_PiE"  :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3PIE _:_./ _)" 10)
+
+syntax (xsymbols)
+  "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>\<^isub>E _\<in>_./ _)"   10)
+
+syntax (HTML output)
+  "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>\<^isub>E _\<in>_./ _)"   10)
+
+translations
+  "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)"
+
+abbreviation
+  funcset_extensional :: "['a set, 'b set] => ('a => 'b) set"
+    (infixr "->\<^isub>E" 60) where
+  "A ->\<^isub>E B \<equiv> Pi\<^isub>E A (%_. B)"
+
+notation (xsymbols)
+  funcset_extensional  (infixr "\<rightarrow>\<^isub>E" 60)
+
+lemma extensional_empty[simp]: "extensional {} = {\<lambda>x. undefined}"
+  by safe (auto simp add: extensional_def fun_eq_iff)
+
+lemma extensional_insert[intro, simp]:
+  assumes "a \<in> extensional (insert i I)"
+  shows "a(i := b) \<in> extensional (insert i I)"
+  using assms unfolding extensional_def by auto
+
+lemma extensional_Int[simp]:
+  "extensional I \<inter> extensional I' = extensional (I \<inter> I')"
+  unfolding extensional_def by auto
+
+definition
+  "merge I x J y = (\<lambda>i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
+
+lemma merge_apply[simp]:
+  "I \<inter> J = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I x J y i = x i"
+  "I \<inter> J = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I x J y i = y i"
+  "J \<inter> I = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I x J y i = x i"
+  "J \<inter> I = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I x J y i = y i"
+  "i \<notin> I \<Longrightarrow> i \<notin> J \<Longrightarrow> merge I x J y i = undefined"
+  unfolding merge_def by auto
+
+lemma merge_commute:
+  "I \<inter> J = {} \<Longrightarrow> merge I x J y = merge J y I x"
+  by (auto simp: merge_def intro!: ext)
+
+lemma Pi_cancel_merge_range[simp]:
+  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge I A J B) \<longleftrightarrow> x \<in> Pi I A"
+  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge J B I A) \<longleftrightarrow> x \<in> Pi I A"
+  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge I A J B) \<longleftrightarrow> x \<in> Pi I A"
+  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge J B I A) \<longleftrightarrow> x \<in> Pi I A"
+  by (auto simp: Pi_def)
+
+lemma Pi_cancel_merge[simp]:
+  "I \<inter> J = {} \<Longrightarrow> merge I x J y \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
+  "J \<inter> I = {} \<Longrightarrow> merge I x J y \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
+  "I \<inter> J = {} \<Longrightarrow> merge I x J y \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
+  "J \<inter> I = {} \<Longrightarrow> merge I x J y \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
+  by (auto simp: Pi_def)
+
+lemma extensional_merge[simp]: "merge I x J y \<in> extensional (I \<union> J)"
+  by (auto simp: extensional_def)
+
+lemma restrict_Pi_cancel: "restrict x I \<in> Pi I A \<longleftrightarrow> x \<in> Pi I A"
+  by (auto simp: restrict_def Pi_def)
+
+lemma restrict_merge[simp]:
+  "I \<inter> J = {} \<Longrightarrow> restrict (merge I x J y) I = restrict x I"
+  "I \<inter> J = {} \<Longrightarrow> restrict (merge I x J y) J = restrict y J"
+  "J \<inter> I = {} \<Longrightarrow> restrict (merge I x J y) I = restrict x I"
+  "J \<inter> I = {} \<Longrightarrow> restrict (merge I x J y) J = restrict y J"
+  by (auto simp: restrict_def intro!: ext)
+
+lemma extensional_insert_undefined[intro, simp]:
+  assumes "a \<in> extensional (insert i I)"
+  shows "a(i := undefined) \<in> extensional I"
+  using assms unfolding extensional_def by auto
+
+lemma extensional_insert_cancel[intro, simp]:
+  assumes "a \<in> extensional I"
+  shows "a \<in> extensional (insert i I)"
+  using assms unfolding extensional_def by auto
+
+lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I x {i} y = restrict (x(i := y i)) (insert i I)"
+  unfolding merge_def by (auto simp: fun_eq_iff)
+
+lemma Pi_Int: "Pi I E \<inter> Pi I F = (\<Pi> i\<in>I. E i \<inter> F i)"
+  by auto
+
+lemma PiE_Int: "(Pi\<^isub>E I A) \<inter> (Pi\<^isub>E I B) = Pi\<^isub>E I (\<lambda>x. A x \<inter> B x)"
+  by auto
+
+lemma Pi_cancel_fupd_range[simp]: "i \<notin> I \<Longrightarrow> x \<in> Pi I (B(i := b)) \<longleftrightarrow> x \<in> Pi I B"
+  by (auto simp: Pi_def)
+
+lemma Pi_split_insert_domain[simp]: "x \<in> Pi (insert i I) X \<longleftrightarrow> x \<in> Pi I X \<and> x i \<in> X i"
+  by (auto simp: Pi_def)
+
+lemma Pi_split_domain[simp]: "x \<in> Pi (I \<union> J) X \<longleftrightarrow> x \<in> Pi I X \<and> x \<in> Pi J X"
+  by (auto simp: Pi_def)
+
+lemma Pi_cancel_fupd[simp]: "i \<notin> I \<Longrightarrow> x(i := a) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
+  by (auto simp: Pi_def)
+
+lemma restrict_vimage:
+  assumes "I \<inter> J = {}"
+  shows "(\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \<times> Pi\<^isub>E J F) = Pi (I \<union> J) (merge I E J F)"
+  using assms by (auto simp: restrict_Pi_cancel)
+
+lemma merge_vimage:
+  assumes "I \<inter> J = {}"
+  shows "(\<lambda>(x,y). merge I x J y) -` Pi\<^isub>E (I \<union> J) E = Pi I E \<times> Pi J E"
+  using assms by (auto simp: restrict_Pi_cancel)
+
+lemma restrict_fupd[simp]: "i \<notin> I \<Longrightarrow> restrict (f (i := x)) I = restrict f I"
+  by (auto simp: restrict_def intro!: ext)
+
+lemma merge_restrict[simp]:
+  "merge I (restrict x I) J y = merge I x J y"
+  "merge I x J (restrict y J) = merge I x J y"
+  unfolding merge_def by (auto intro!: ext)
+
+lemma merge_x_x_eq_restrict[simp]:
+  "merge I x J x = restrict x (I \<union> J)"
+  unfolding merge_def by (auto intro!: ext)
+
+lemma Pi_fupd_iff: "i \<in> I \<Longrightarrow> f \<in> Pi I (B(i := A)) \<longleftrightarrow> f \<in> Pi (I - {i}) B \<and> f i \<in> A"
+  apply auto
+  apply (drule_tac x=x in Pi_mem)
+  apply (simp_all split: split_if_asm)
+  apply (drule_tac x=i in Pi_mem)
+  apply (auto dest!: Pi_mem)
+  done
+
+lemma Pi_UN:
+  fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set"
+  assumes "finite I" and mono: "\<And>i n m. i \<in> I \<Longrightarrow> n \<le> m \<Longrightarrow> A n i \<subseteq> A m i"
+  shows "(\<Union>n. Pi I (A n)) = (\<Pi> i\<in>I. \<Union>n. A n i)"
+proof (intro set_eqI iffI)
+  fix f assume "f \<in> (\<Pi> i\<in>I. \<Union>n. A n i)"
+  then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" by auto
+  from bchoice[OF this] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto
+  obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k"
+    using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto
+  have "f \<in> Pi I (A k)"
+  proof (intro Pi_I)
+    fix i assume "i \<in> I"
+    from mono[OF this, of "n i" k] k[OF this] n[OF this]
+    show "f i \<in> A k i" by auto
+  qed
+  then show "f \<in> (\<Union>n. Pi I (A n))" by auto
+qed auto
+
+lemma PiE_cong:
+  assumes "\<And>i. i\<in>I \<Longrightarrow> A i = B i"
+  shows "Pi\<^isub>E I A = Pi\<^isub>E I B"
+  using assms by (auto intro!: Pi_cong)
+
+lemma restrict_upd[simp]:
+  "i \<notin> I \<Longrightarrow> (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)"
+  by (auto simp: fun_eq_iff)
+
+lemma Pi_eq_subset:
+  assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
+  assumes eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and "i \<in> I"
+  shows "F i \<subseteq> F' i"
+proof
+  fix x assume "x \<in> F i"
+  with ne have "\<forall>j. \<exists>y. ((j \<in> I \<longrightarrow> y \<in> F j \<and> (i = j \<longrightarrow> x = y)) \<and> (j \<notin> I \<longrightarrow> y = undefined))" by auto
+  from choice[OF this] guess f .. note f = this
+  then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def)
+  then have "f \<in> Pi\<^isub>E I F'" using assms by simp
+  then show "x \<in> F' i" using f `i \<in> I` by auto
+qed
+
+lemma Pi_eq_iff_not_empty:
+  assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
+  shows "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i)"
+proof (intro iffI ballI)
+  fix i assume eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and i: "i \<in> I"
+  show "F i = F' i"
+    using Pi_eq_subset[of I F F', OF ne eq i]
+    using Pi_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i]
+    by auto
+qed auto
+
+lemma Pi_eq_empty_iff:
+  "Pi\<^isub>E I F = {} \<longleftrightarrow> (\<exists>i\<in>I. F i = {})"
+proof
+  assume "Pi\<^isub>E I F = {}"
+  show "\<exists>i\<in>I. F i = {}"
+  proof (rule ccontr)
+    assume "\<not> ?thesis"
+    then have "\<forall>i. \<exists>y. (i \<in> I \<longrightarrow> y \<in> F i) \<and> (i \<notin> I \<longrightarrow> y = undefined)" by auto
+    from choice[OF this] guess f ..
+    then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def)
+    with `Pi\<^isub>E I F = {}` show False by auto
+  qed
+qed auto
+
+lemma Pi_eq_iff:
+  "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i) \<or> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
+proof (intro iffI disjCI)
+  assume eq[simp]: "Pi\<^isub>E I F = Pi\<^isub>E I F'"
+  assume "\<not> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
+  then have "(\<forall>i\<in>I. F i \<noteq> {}) \<and> (\<forall>i\<in>I. F' i \<noteq> {})"
+    using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto
+  with Pi_eq_iff_not_empty[of I F F'] show "\<forall>i\<in>I. F i = F' i" by auto
+next
+  assume "(\<forall>i\<in>I. F i = F' i) \<or> (\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {})"
+  then show "Pi\<^isub>E I F = Pi\<^isub>E I F'"
+    using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto
+qed
+
+section "Finite product spaces"
+
+section "Products"
+
+locale product_sigma_algebra =
+  fixes M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme"
+  assumes sigma_algebras: "\<And>i. sigma_algebra (M i)"
+
+locale finite_product_sigma_algebra = product_sigma_algebra M
+  for M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme" +
+  fixes I :: "'i set"
+  assumes finite_index: "finite I"
+
+definition
+  "product_algebra_generator I M = \<lparr> space = (\<Pi>\<^isub>E i \<in> I. space (M i)),
+    sets = Pi\<^isub>E I ` (\<Pi> i \<in> I. sets (M i)),
+    measure = \<lambda>A. (\<Prod>i\<in>I. measure (M i) ((SOME F. A = Pi\<^isub>E I F) i)) \<rparr>"
+
+definition product_algebra_def:
+  "Pi\<^isub>M I M = sigma (product_algebra_generator I M)
+    \<lparr> measure := (SOME \<mu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<mu> \<rparr>) \<and>
+      (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))))\<rparr>"
+
+syntax
+  "_PiM"  :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
+              ('i => 'a, 'b) measure_space_scheme"  ("(3PIM _:_./ _)" 10)
+
+syntax (xsymbols)
+  "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
+             ('i => 'a, 'b) measure_space_scheme"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"   10)
+
+syntax (HTML output)
+  "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
+             ('i => 'a, 'b) measure_space_scheme"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"   10)
+
+translations
+  "PIM x:I. M" == "CONST Pi\<^isub>M I (%x. M)"
+
+abbreviation (in finite_product_sigma_algebra) "G \<equiv> product_algebra_generator I M"
+abbreviation (in finite_product_sigma_algebra) "P \<equiv> Pi\<^isub>M I M"
+
+sublocale product_sigma_algebra \<subseteq> M: sigma_algebra "M i" for i by (rule sigma_algebras)
+
+lemma sigma_into_space:
+  assumes "sets M \<subseteq> Pow (space M)"
+  shows "sets (sigma M) \<subseteq> Pow (space M)"
+  using sigma_sets_into_sp[OF assms] unfolding sigma_def by auto
+
+lemma (in product_sigma_algebra) product_algebra_generator_into_space:
+  "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))"
+  using M.sets_into_space unfolding product_algebra_generator_def
+  by auto blast
+
+lemma (in product_sigma_algebra) product_algebra_into_space:
+  "sets (Pi\<^isub>M I M) \<subseteq> Pow (space (Pi\<^isub>M I M))"
+  using product_algebra_generator_into_space
+  by (auto intro!: sigma_into_space simp add: product_algebra_def)
+
+lemma (in product_sigma_algebra) sigma_algebra_product_algebra: "sigma_algebra (Pi\<^isub>M I M)"
+  using product_algebra_generator_into_space unfolding product_algebra_def
+  by (rule sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) simp_all
+
+sublocale finite_product_sigma_algebra \<subseteq> sigma_algebra P
+  using sigma_algebra_product_algebra .
+
+lemma product_algebraE:
+  assumes "A \<in> sets (product_algebra_generator I M)"
+  obtains E where "A = Pi\<^isub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
+  using assms unfolding product_algebra_generator_def by auto
+
+lemma product_algebra_generatorI[intro]:
+  assumes "E \<in> (\<Pi> i\<in>I. sets (M i))"
+  shows "Pi\<^isub>E I E \<in> sets (product_algebra_generator I M)"
+  using assms unfolding product_algebra_generator_def by auto
+
+lemma space_product_algebra_generator[simp]:
+  "space (product_algebra_generator I M) = Pi\<^isub>E I (\<lambda>i. space (M i))"
+  unfolding product_algebra_generator_def by simp
+
+lemma space_product_algebra[simp]:
+  "space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E i\<in>I. space (M i))"
+  unfolding product_algebra_def product_algebra_generator_def by simp
+
+lemma sets_product_algebra:
+  "sets (Pi\<^isub>M I M) = sets (sigma (product_algebra_generator I M))"
+  unfolding product_algebra_def sigma_def by simp
+
+lemma product_algebra_generator_sets_into_space:
+  assumes "\<And>i. i\<in>I \<Longrightarrow> sets (M i) \<subseteq> Pow (space (M i))"
+  shows "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))"
+  using assms by (auto simp: product_algebra_generator_def) blast
+
+lemma (in finite_product_sigma_algebra) in_P[simp, intro]:
+  "\<lbrakk> \<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i) \<rbrakk> \<Longrightarrow> Pi\<^isub>E I A \<in> sets P"
+  by (auto simp: sets_product_algebra)
+
+section "Generating set generates also product algebra"
+
+lemma sigma_product_algebra_sigma_eq:
+  assumes "finite I"
+  assumes mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)"
+  assumes union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (E i)"
+  assumes sets_into: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> sets (E i)"
+  and E: "\<And>i. sets (E i) \<subseteq> Pow (space (E i))"
+  shows "sets (\<Pi>\<^isub>M i\<in>I. sigma (E i)) = sets (\<Pi>\<^isub>M i\<in>I. E i)"
+    (is "sets ?S = sets ?E")
+proof cases
+  assume "I = {}" then show ?thesis
+    by (simp add: product_algebra_def product_algebra_generator_def)
+next
+  assume "I \<noteq> {}"
+  interpret E: sigma_algebra "sigma (E i)" for i
+    using E by (rule sigma_algebra_sigma)
+  have into_space[intro]: "\<And>i x A. A \<in> sets (E i) \<Longrightarrow> x i \<in> A \<Longrightarrow> x i \<in> space (E i)"
+    using E by auto
+  interpret G: sigma_algebra ?E
+    unfolding product_algebra_def product_algebra_generator_def using E
+    by (intro sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) (auto dest: Pi_mem)
+  { fix A i assume "i \<in> I" and A: "A \<in> sets (E i)"
+    then have "(\<lambda>x. x i) -` A \<inter> space ?E = (\<Pi>\<^isub>E j\<in>I. if j = i then A else \<Union>n. S j n) \<inter> space ?E"
+      using mono union unfolding incseq_Suc_iff space_product_algebra
+      by (auto dest: Pi_mem)
+    also have "\<dots> = (\<Union>n. (\<Pi>\<^isub>E j\<in>I. if j = i then A else S j n))"
+      unfolding space_product_algebra
+      apply simp
+      apply (subst Pi_UN[OF `finite I`])
+      using mono[THEN incseqD] apply simp
+      apply (simp add: PiE_Int)
+      apply (intro PiE_cong)
+      using A sets_into by (auto intro!: into_space)
+    also have "\<dots> \<in> sets ?E"
+      using sets_into `A \<in> sets (E i)`
+      unfolding sets_product_algebra sets_sigma
+      by (intro sigma_sets.Union)
+         (auto simp: image_subset_iff intro!: sigma_sets.Basic)
+    finally have "(\<lambda>x. x i) -` A \<inter> space ?E \<in> sets ?E" . }
+  then have proj:
+    "\<And>i. i\<in>I \<Longrightarrow> (\<lambda>x. x i) \<in> measurable ?E (sigma (E i))"
+    using E by (subst G.measurable_iff_sigma)
+               (auto simp: sets_product_algebra sets_sigma)
+  { fix A assume A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (sigma (E i))"
+    with proj have basic: "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. x i) -` (A i) \<inter> space ?E \<in> sets ?E"
+      unfolding measurable_def by simp
+    have "Pi\<^isub>E I A = (\<Inter>i\<in>I. (\<lambda>x. x i) -` (A i) \<inter> space ?E)"
+      using A E.sets_into_space `I \<noteq> {}` unfolding product_algebra_def by auto blast
+    then have "Pi\<^isub>E I A \<in> sets ?E"
+      using G.finite_INT[OF `finite I` `I \<noteq> {}` basic, of "\<lambda>i. i"] by simp }
+  then have "sigma_sets (space ?E) (sets (product_algebra_generator I (\<lambda>i. sigma (E i)))) \<subseteq> sets ?E"
+    by (intro G.sigma_sets_subset) (auto simp add: product_algebra_generator_def)
+  then have subset: "sets ?S \<subseteq> sets ?E"
+    by (simp add: sets_sigma sets_product_algebra)
+  show "sets ?S = sets ?E"
+  proof (intro set_eqI iffI)
+    fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S"
+      unfolding sets_sigma sets_product_algebra
+    proof induct
+      case (Basic A) then show ?case
+        by (auto simp: sets_sigma product_algebra_generator_def intro: sigma_sets.Basic)
+    qed (auto intro: sigma_sets.intros simp: product_algebra_generator_def)
+  next
+    fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
+  qed
+qed
+
+lemma product_algebraI[intro]:
+    "E \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> Pi\<^isub>E I E \<in> sets (Pi\<^isub>M I M)"
+  using assms unfolding product_algebra_def by (auto intro: product_algebra_generatorI)
+
+lemma (in product_sigma_algebra) measurable_component_update:
+  assumes "x \<in> space (Pi\<^isub>M I M)" and "i \<notin> I"
+  shows "(\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^isub>M (insert i I) M)" (is "?f \<in> _")
+  unfolding product_algebra_def apply simp
+proof (intro measurable_sigma)
+  let ?G = "product_algebra_generator (insert i I) M"
+  show "sets ?G \<subseteq> Pow (space ?G)" using product_algebra_generator_into_space .
+  show "?f \<in> space (M i) \<rightarrow> space ?G"
+    using M.sets_into_space assms by auto
+  fix A assume "A \<in> sets ?G"
+  from product_algebraE[OF this] guess E . note E = this
+  then have "?f -` A \<inter> space (M i) = E i \<or> ?f -` A \<inter> space (M i) = {}"
+    using M.sets_into_space assms by auto
+  then show "?f -` A \<inter> space (M i) \<in> sets (M i)"
+    using E by (auto intro!: product_algebraI)
+qed
+
+lemma (in product_sigma_algebra) measurable_add_dim:
+  assumes "i \<notin> I"
+  shows "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M)"
+proof -
+  let ?f = "(\<lambda>(f, y). f(i := y))" and ?G = "product_algebra_generator (insert i I) M"
+  interpret Ii: pair_sigma_algebra "Pi\<^isub>M I M" "M i"
+    unfolding pair_sigma_algebra_def
+    by (intro sigma_algebra_product_algebra sigma_algebras conjI)
+  have "?f \<in> measurable Ii.P (sigma ?G)"
+  proof (rule Ii.measurable_sigma)
+    show "sets ?G \<subseteq> Pow (space ?G)"
+      using product_algebra_generator_into_space .
+    show "?f \<in> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<rightarrow> space ?G"
+      by (auto simp: space_pair_measure)
+  next
+    fix A assume "A \<in> sets ?G"
+    then obtain F where "A = Pi\<^isub>E (insert i I) F"
+      and F: "\<And>j. j \<in> I \<Longrightarrow> F j \<in> sets (M j)" "F i \<in> sets (M i)"
+      by (auto elim!: product_algebraE)
+    then have "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = Pi\<^isub>E I F \<times> (F i)"
+      using sets_into_space `i \<notin> I`
+      by (auto simp add: space_pair_measure) blast+
+    then show "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M M i)"
+      using F by (auto intro!: pair_measureI)
+  qed
+  then show ?thesis
+    by (simp add: product_algebra_def)
+qed
+
+lemma (in product_sigma_algebra) measurable_merge:
+  assumes [simp]: "I \<inter> J = {}"
+  shows "(\<lambda>(x, y). merge I x J y) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
+proof -
+  let ?I = "Pi\<^isub>M I M" and ?J = "Pi\<^isub>M J M"
+  interpret P: sigma_algebra "?I \<Otimes>\<^isub>M ?J"
+    by (intro sigma_algebra_pair_measure product_algebra_into_space)
+  let ?f = "\<lambda>(x, y). merge I x J y"
+  let ?G = "product_algebra_generator (I \<union> J) M"
+  have "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (sigma ?G)"
+  proof (rule P.measurable_sigma)
+    fix A assume "A \<in> sets ?G"
+    from product_algebraE[OF this]
+    obtain E where E: "A = Pi\<^isub>E (I \<union> J) E" "E \<in> (\<Pi> i\<in>I \<union> J. sets (M i))" .
+    then have *: "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) = Pi\<^isub>E I E \<times> Pi\<^isub>E J E"
+      using sets_into_space `I \<inter> J = {}`
+      by (auto simp add: space_pair_measure) fast+
+    then show "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) \<in> sets (?I \<Otimes>\<^isub>M ?J)"
+      using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI
+        simp: product_algebra_def)
+  qed (insert product_algebra_generator_into_space, auto simp: space_pair_measure)
+  then show "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (Pi\<^isub>M (I \<union> J) M)"
+    unfolding product_algebra_def[of "I \<union> J"] by simp
+qed
+
+lemma (in product_sigma_algebra) measurable_component_singleton:
+  assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>M I M) (M i)"
+proof (unfold measurable_def, intro CollectI conjI ballI)
+  fix A assume "A \<in> sets (M i)"
+  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E j\<in>I. if i = j then A else space (M j))"
+    using M.sets_into_space `i \<in> I` by (fastsimp dest: Pi_mem split: split_if_asm)
+  then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) \<in> sets (Pi\<^isub>M I M)"
+    using `A \<in> sets (M i)` by (auto intro!: product_algebraI)
+qed (insert `i \<in> I`, auto)
+
+locale product_sigma_finite =
+  fixes M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme"
+  assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)"
+
+locale finite_product_sigma_finite = product_sigma_finite M
+  for M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" +
+  fixes I :: "'i set" assumes finite_index'[intro]: "finite I"
+
+sublocale product_sigma_finite \<subseteq> M: sigma_finite_measure "M i" for i
+  by (rule sigma_finite_measures)
+
+sublocale product_sigma_finite \<subseteq> product_sigma_algebra
+  by default
+
+sublocale finite_product_sigma_finite \<subseteq> finite_product_sigma_algebra
+  by default (fact finite_index')
+
+lemma setprod_extreal_0:
+  fixes f :: "'a \<Rightarrow> extreal"
+  shows "(\<Prod>i\<in>A. f i) = 0 \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. f i = 0))"
+proof cases
+  assume "finite A"
+  then show ?thesis by (induct A) auto
+qed auto
+
+lemma setprod_extreal_pos:
+  fixes f :: "'a \<Rightarrow> extreal" assumes pos: "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i" shows "0 \<le> (\<Prod>i\<in>I. f i)"
+proof cases
+  assume "finite I" from this pos show ?thesis by induct auto
+qed simp
+
+lemma setprod_PInf:
+  assumes "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i"
+  shows "(\<Prod>i\<in>I. f i) = \<infinity> \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = \<infinity>) \<and> (\<forall>i\<in>I. f i \<noteq> 0)"
+proof cases
+  assume "finite I" from this assms show ?thesis
+  proof (induct I)
+    case (insert i I)
+    then have pos: "0 \<le> f i" "0 \<le> setprod f I" by (auto intro!: setprod_extreal_pos)
+    from insert have "(\<Prod>j\<in>insert i I. f j) = \<infinity> \<longleftrightarrow> setprod f I * f i = \<infinity>" by auto
+    also have "\<dots> \<longleftrightarrow> (setprod f I = \<infinity> \<or> f i = \<infinity>) \<and> f i \<noteq> 0 \<and> setprod f I \<noteq> 0"
+      using setprod_extreal_pos[of I f] pos
+      by (cases rule: extreal2_cases[of "f i" "setprod f I"]) auto
+    also have "\<dots> \<longleftrightarrow> finite (insert i I) \<and> (\<exists>j\<in>insert i I. f j = \<infinity>) \<and> (\<forall>j\<in>insert i I. f j \<noteq> 0)"
+      using insert by (auto simp: setprod_extreal_0)
+    finally show ?case .
+  qed simp
+qed simp
+
+lemma setprod_extreal: "(\<Prod>i\<in>A. extreal (f i)) = extreal (setprod f A)"
+proof cases
+  assume "finite A" then show ?thesis
+    by induct (auto simp: one_extreal_def)
+qed (simp add: one_extreal_def)
+
+lemma (in finite_product_sigma_finite) product_algebra_generator_measure:
+  assumes "Pi\<^isub>E I F \<in> sets G"
+  shows "measure G (Pi\<^isub>E I F) = (\<Prod>i\<in>I. M.\<mu> i (F i))"
+proof cases
+  assume ne: "\<forall>i\<in>I. F i \<noteq> {}"
+  have "\<forall>i\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') i = F i"
+    by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
+       (insert ne, auto simp: Pi_eq_iff)
+  then show ?thesis
+    unfolding product_algebra_generator_def by simp
+next
+  assume empty: "\<not> (\<forall>j\<in>I. F j \<noteq> {})"
+  then have "(\<Prod>j\<in>I. M.\<mu> j (F j)) = 0"
+    by (auto simp: setprod_extreal_0 intro!: bexI)
+  moreover
+  have "\<exists>j\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}"
+    by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
+       (insert empty, auto simp: Pi_eq_empty_iff[symmetric])
+  then have "(\<Prod>j\<in>I. M.\<mu> j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0"
+    by (auto simp: setprod_extreal_0 intro!: bexI)
+  ultimately show ?thesis
+    unfolding product_algebra_generator_def by simp
+qed
+
+lemma (in finite_product_sigma_finite) sigma_finite_pairs:
+  "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
+    (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
+    (\<forall>k. \<forall>i\<in>I. \<mu> i (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<and>
+    (\<Union>k. \<Pi>\<^isub>E i\<in>I. F i k) = space G"
+proof -
+  have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. \<mu> i (F k) \<noteq> \<infinity>)"
+    using M.sigma_finite_up by simp
+  from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
+  then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. \<mu> i (F i k) \<noteq> \<infinity>"
+    by auto
+  let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k"
+  note space_product_algebra[simp]
+  show ?thesis
+  proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI)
+    fix i show "range (F i) \<subseteq> sets (M i)" by fact
+  next
+    fix i k show "\<mu> i (F i k) \<noteq> \<infinity>" by fact
+  next
+    fix A assume "A \<in> (\<Union>i. ?F i)" then show "A \<in> space G"
+      using `\<And>i. range (F i) \<subseteq> sets (M i)` M.sets_into_space
+      by (force simp: image_subset_iff)
+  next
+    fix f assume "f \<in> space G"
+    with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F
+    show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def)
+  next
+    fix i show "?F i \<subseteq> ?F (Suc i)"
+      using `\<And>i. incseq (F i)`[THEN incseq_SucD] by auto
+  qed
+qed
+
+lemma sets_pair_cancel_measure[simp]:
+  "sets (M1\<lparr>measure := m1\<rparr> \<Otimes>\<^isub>M M2) = sets (M1 \<Otimes>\<^isub>M M2)"
+  "sets (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m2\<rparr>) = sets (M1 \<Otimes>\<^isub>M M2)"
+  unfolding pair_measure_def pair_measure_generator_def sets_sigma
+  by simp_all
+
+lemma measurable_pair_cancel_measure[simp]:
+  "measurable (M1\<lparr>measure := m1\<rparr> \<Otimes>\<^isub>M M2) M = measurable (M1 \<Otimes>\<^isub>M M2) M"
+  "measurable (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m2\<rparr>) M = measurable (M1 \<Otimes>\<^isub>M M2) M"
+  "measurable M (M1\<lparr>measure := m3\<rparr> \<Otimes>\<^isub>M M2) = measurable M (M1 \<Otimes>\<^isub>M M2)"
+  "measurable M (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m4\<rparr>) = measurable M (M1 \<Otimes>\<^isub>M M2)"
+  unfolding measurable_def by (auto simp add: space_pair_measure)
+
+lemma (in product_sigma_finite) product_measure_exists:
+  assumes "finite I"
+  shows "\<exists>\<nu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<nu> \<rparr>) \<and>
+    (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i)))"
+using `finite I` proof induct
+  case empty
+  interpret finite_product_sigma_finite M "{}" by default simp
+  let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> extreal"
+  show ?case
+  proof (intro exI conjI ballI)
+    have "sigma_algebra (sigma G \<lparr>measure := ?\<nu>\<rparr>)"
+      by (rule sigma_algebra_cong) (simp_all add: product_algebra_def)
+    then have "measure_space (sigma G\<lparr>measure := ?\<nu>\<rparr>)"
+      by (rule finite_additivity_sufficient)
+         (simp_all add: positive_def additive_def sets_sigma
+                        product_algebra_generator_def image_constant)
+    then show "sigma_finite_measure (sigma G\<lparr>measure := ?\<nu>\<rparr>)"
+      by (auto intro!: exI[of _ "\<lambda>i. {\<lambda>_. undefined}"]
+               simp: sigma_finite_measure_def sigma_finite_measure_axioms_def
+                     product_algebra_generator_def)
+  qed auto
+next
+  case (insert i I)
+  interpret finite_product_sigma_finite M I by default fact
+  have "finite (insert i I)" using `finite I` by auto
+  interpret I': finite_product_sigma_finite M "insert i I" by default fact
+  from insert obtain \<nu> where
+    prod: "\<And>A. A \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" and
+    "sigma_finite_measure (sigma G\<lparr> measure := \<nu> \<rparr>)" by auto
+  then interpret I: sigma_finite_measure "P\<lparr> measure := \<nu>\<rparr>" unfolding product_algebra_def by simp
+  interpret P: pair_sigma_finite "P\<lparr> measure := \<nu>\<rparr>" "M i" ..
+  let ?h = "(\<lambda>(f, y). f(i := y))"
+  let ?\<nu> = "\<lambda>A. P.\<mu> (?h -` A \<inter> space P.P)"
+  have I': "sigma_algebra (I'.P\<lparr> measure := ?\<nu> \<rparr>)"
+    by (rule I'.sigma_algebra_cong) simp_all
+  interpret I'': measure_space "I'.P\<lparr> measure := ?\<nu> \<rparr>"
+    using measurable_add_dim[OF `i \<notin> I`]
+    by (intro P.measure_space_vimage[OF I']) (auto simp add: measure_preserving_def)
+  show ?case
+  proof (intro exI[of _ ?\<nu>] conjI ballI)
+    let "?m A" = "measure (Pi\<^isub>M I M\<lparr>measure := \<nu>\<rparr> \<Otimes>\<^isub>M M i) (?h -` A \<inter> space P.P)"
+    { fix A assume A: "A \<in> (\<Pi> i\<in>insert i I. sets (M i))"
+      then have *: "?h -` Pi\<^isub>E (insert i I) A \<inter> space P.P = Pi\<^isub>E I A \<times> A i"
+        using `i \<notin> I` M.sets_into_space by (auto simp: space_pair_measure space_product_algebra) blast
+      show "?m (Pi\<^isub>E (insert i I) A) = (\<Prod>i\<in>insert i I. M.\<mu> i (A i))"
+        unfolding * using A
+        apply (subst P.pair_measure_times)
+        using A apply fastsimp
+        using A apply fastsimp
+        using `i \<notin> I` `finite I` prod[of A] A by (auto simp: ac_simps) }
+    note product = this
+    have *: "sigma I'.G\<lparr> measure := ?\<nu> \<rparr> = I'.P\<lparr> measure := ?\<nu> \<rparr>"
+      by (simp add: product_algebra_def)
+    show "sigma_finite_measure (sigma I'.G\<lparr> measure := ?\<nu> \<rparr>)"
+    proof (unfold *, default, simp)
+      from I'.sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
+      then have F: "\<And>j. j \<in> insert i I \<Longrightarrow> range (F j) \<subseteq> sets (M j)"
+        "incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k)"
+        "(\<Union>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) = space I'.G"
+        "\<And>k. \<And>j. j \<in> insert i I \<Longrightarrow> \<mu> j (F j k) \<noteq> \<infinity>"
+        by blast+
+      let "?F k" = "\<Pi>\<^isub>E j \<in> insert i I. F j k"
+      show "\<exists>F::nat \<Rightarrow> ('i \<Rightarrow> 'a) set. range F \<subseteq> sets I'.P \<and>
+          (\<Union>i. F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i)) \<and> (\<forall>i. ?m (F i) \<noteq> \<infinity>)"
+      proof (intro exI[of _ ?F] conjI allI)
+        show "range ?F \<subseteq> sets I'.P" using F(1) by auto
+      next
+        from F(3) show "(\<Union>i. ?F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i))" by simp
+      next
+        fix j
+        have "\<And>k. k \<in> insert i I \<Longrightarrow> 0 \<le> measure (M k) (F k j)"
+          using F(1) by auto
+        with F `finite I` setprod_PInf[of "insert i I", OF this] show "?\<nu> (?F j) \<noteq> \<infinity>"
+          by (subst product) auto
+      qed
+    qed
+  qed
+qed
+
+sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure P
+  unfolding product_algebra_def
+  using product_measure_exists[OF finite_index]
+  by (rule someI2_ex) auto
+
+lemma (in finite_product_sigma_finite) measure_times:
+  assumes "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)"
+  shows "\<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
+  unfolding product_algebra_def
+  using product_measure_exists[OF finite_index]
+  proof (rule someI2_ex, elim conjE)
+    fix \<nu> assume *: "\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
+    have "Pi\<^isub>E I A = Pi\<^isub>E I (\<lambda>i\<in>I. A i)" by (auto dest: Pi_mem)
+    then have "\<nu> (Pi\<^isub>E I A) = \<nu> (Pi\<^isub>E I (\<lambda>i\<in>I. A i))" by simp
+    also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i ((\<lambda>i\<in>I. A i) i))" using assms * by auto
+    finally show "measure (sigma G\<lparr>measure := \<nu>\<rparr>) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
+      by (simp add: sigma_def)
+  qed
+
+lemma (in product_sigma_finite) product_measure_empty[simp]:
+  "measure (Pi\<^isub>M {} M) {\<lambda>x. undefined} = 1"
+proof -
+  interpret finite_product_sigma_finite M "{}" by default auto
+  from measure_times[of "\<lambda>x. {}"] show ?thesis by simp
+qed
+
+lemma (in finite_product_sigma_algebra) P_empty:
+  assumes "I = {}"
+  shows "space P = {\<lambda>k. undefined}" "sets P = { {}, {\<lambda>k. undefined} }"
+  unfolding product_algebra_def product_algebra_generator_def `I = {}`
+  by (simp_all add: sigma_def image_constant)
+
+lemma (in product_sigma_finite) positive_integral_empty:
+  assumes pos: "0 \<le> f (\<lambda>k. undefined)"
+  shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)"
+proof -
+  interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI)
+  have "\<And>A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1"
+    using assms by (subst measure_times) auto
+  then show ?thesis
+    unfolding positive_integral_def simple_function_def simple_integral_def_raw
+  proof (simp add: P_empty, intro antisym)
+    show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined))"
+      by (intro le_SUPI) (auto simp: le_fun_def split: split_max)
+    show "(SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)" using pos
+      by (intro SUP_leI) (auto simp: le_fun_def simp: max_def split: split_if_asm)
+  qed
+qed
+
+lemma (in product_sigma_finite) measure_fold:
+  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
+  assumes A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
+  shows "measure (Pi\<^isub>M (I \<union> J) M) A =
+    measure (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) ((\<lambda>(x,y). merge I x J y) -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M))"
+proof -
+  interpret I: finite_product_sigma_finite M I by default fact
+  interpret J: finite_product_sigma_finite M J by default fact
+  have "finite (I \<union> J)" using fin by auto
+  interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
+  interpret P: pair_sigma_finite I.P J.P by default
+  let ?g = "\<lambda>(x,y). merge I x J y"
+  let "?X S" = "?g -` S \<inter> space P.P"
+  from IJ.sigma_finite_pairs obtain F where
+    F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)"
+       "incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k)"
+       "(\<Union>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) = space IJ.G"
+       "\<And>k. \<forall>i\<in>I\<union>J. \<mu> i (F i k) \<noteq> \<infinity>"
+    by auto
+  let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k"
+  show "IJ.\<mu> A = P.\<mu> (?X A)"
+  proof (rule measure_unique_Int_stable_vimage)
+    show "measure_space IJ.P" "measure_space P.P" by default
+    show "sets (sigma IJ.G) = sets IJ.P" "space IJ.G = space IJ.P" "A \<in> sets (sigma IJ.G)"
+      using A unfolding product_algebra_def by auto
+  next
+    show "Int_stable IJ.G"
+      by (simp add: PiE_Int Int_stable_def product_algebra_def
+                    product_algebra_generator_def)
+          auto
+    show "range ?F \<subseteq> sets IJ.G" using F
+      by (simp add: image_subset_iff product_algebra_def
+                    product_algebra_generator_def)
+    show "incseq ?F" "(\<Union>i. ?F i) = space IJ.G " by fact+
+  next
+    fix k
+    have "\<And>j. j \<in> I \<union> J \<Longrightarrow> 0 \<le> measure (M j) (F j k)"
+      using F(1) by auto
+    with F `finite I` setprod_PInf[of "I \<union> J", OF this]
+    show "IJ.\<mu> (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto
+  next
+    fix A assume "A \<in> sets IJ.G"
+    then obtain F where A: "A = Pi\<^isub>E (I \<union> J) F"
+      and F: "\<And>i. i \<in> I \<union> J \<Longrightarrow> F i \<in> sets (M i)"
+      by (auto simp: product_algebra_generator_def)
+    then have X: "?X A = (Pi\<^isub>E I F \<times> Pi\<^isub>E J F)"
+      using sets_into_space by (auto simp: space_pair_measure) blast+
+    then have "P.\<mu> (?X A) = (\<Prod>i\<in>I. \<mu> i (F i)) * (\<Prod>i\<in>J. \<mu> i (F i))"
+      using `finite J` `finite I` F
+      by (simp add: P.pair_measure_times I.measure_times J.measure_times)
+    also have "\<dots> = (\<Prod>i\<in>I \<union> J. \<mu> i (F i))"
+      using `finite J` `finite I` `I \<inter> J = {}`  by (simp add: setprod_Un_one)
+    also have "\<dots> = IJ.\<mu> A"
+      using `finite J` `finite I` F unfolding A
+      by (intro IJ.measure_times[symmetric]) auto
+    finally show "IJ.\<mu> A = P.\<mu> (?X A)" by (rule sym)
+  qed (rule measurable_merge[OF IJ])
+qed
+
+lemma (in product_sigma_finite) measure_preserving_merge:
+  assumes IJ: "I \<inter> J = {}" and "finite I" "finite J"
+  shows "(\<lambda>(x,y). merge I x J y) \<in> measure_preserving (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
+  by (intro measure_preservingI measurable_merge[OF IJ] measure_fold[symmetric] assms)
+
+lemma (in product_sigma_finite) product_positive_integral_fold:
+  assumes IJ[simp]: "I \<inter> J = {}" "finite I" "finite J"
+  and f: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
+  shows "integral\<^isup>P (Pi\<^isub>M (I \<union> J) M) f =
+    (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (merge I x J y) \<partial>(Pi\<^isub>M J M)) \<partial>(Pi\<^isub>M I M))"
+proof -
+  interpret I: finite_product_sigma_finite M I by default fact
+  interpret J: finite_product_sigma_finite M J by default fact
+  interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
+  interpret IJ: finite_product_sigma_finite M "I \<union> J" by default simp
+  have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable P.P"
+    using measurable_comp[OF measurable_merge[OF IJ(1)] f] by (simp add: comp_def)
+  show ?thesis
+    unfolding P.positive_integral_fst_measurable[OF P_borel, simplified]
+  proof (rule P.positive_integral_vimage)
+    show "sigma_algebra IJ.P" by default
+    show "(\<lambda>(x, y). merge I x J y) \<in> measure_preserving P.P IJ.P"
+      using IJ by (rule measure_preserving_merge)
+    show "f \<in> borel_measurable IJ.P" using f by simp
+  qed
+qed
+
+lemma (in product_sigma_finite) measure_preserving_component_singelton:
+  "(\<lambda>x. x i) \<in> measure_preserving (Pi\<^isub>M {i} M) (M i)"
+proof (intro measure_preservingI measurable_component_singleton)
+  interpret I: finite_product_sigma_finite M "{i}" by default simp
+  fix A let ?P = "(\<lambda>x. x i) -` A \<inter> space I.P"
+  assume A: "A \<in> sets (M i)"
+  then have *: "?P = {i} \<rightarrow>\<^isub>E A" using sets_into_space by auto
+  show "I.\<mu> ?P = M.\<mu> i A" unfolding *
+    using A I.measure_times[of "\<lambda>_. A"] by auto
+qed auto
+
+lemma (in product_sigma_finite) product_positive_integral_singleton:
+  assumes f: "f \<in> borel_measurable (M i)"
+  shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\<lambda>x. f (x i)) = integral\<^isup>P (M i) f"
+proof -
+  interpret I: finite_product_sigma_finite M "{i}" by default simp
+  show ?thesis
+  proof (rule I.positive_integral_vimage[symmetric])
+    show "sigma_algebra (M i)" by (rule sigma_algebras)
+    show "(\<lambda>x. x i) \<in> measure_preserving (Pi\<^isub>M {i} M) (M i)"
+      by (rule measure_preserving_component_singelton)
+    show "f \<in> borel_measurable (M i)" by fact
+  qed
+qed
+
+lemma (in product_sigma_finite) product_positive_integral_insert:
+  assumes [simp]: "finite I" "i \<notin> I"
+    and f: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)"
+  shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^isub>M I M))"
+proof -
+  interpret I: finite_product_sigma_finite M I by default auto
+  interpret i: finite_product_sigma_finite M "{i}" by default auto
+  interpret P: pair_sigma_algebra I.P i.P ..
+  have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I"
+    using f by auto
+  show ?thesis
+    unfolding product_positive_integral_fold[OF IJ, unfolded insert, simplified, OF f]
+  proof (rule I.positive_integral_cong, subst product_positive_integral_singleton)
+    fix x assume x: "x \<in> space I.P"
+    let "?f y" = "f (restrict (x(i := y)) (insert i I))"
+    have f'_eq: "\<And>y. ?f y = f (x(i := y))"
+      using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def)
+    show "?f \<in> borel_measurable (M i)" unfolding f'_eq
+      using measurable_comp[OF measurable_component_update f] x `i \<notin> I`
+      by (simp add: comp_def)
+    show "integral\<^isup>P (M i) ?f = \<integral>\<^isup>+ y. f (x(i:=y)) \<partial>M i"
+      unfolding f'_eq by simp
+  qed
+qed
+
+lemma (in product_sigma_finite) product_positive_integral_setprod:
+  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> extreal"
+  assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
+  and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x"
+  shows "(\<integral>\<^isup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>P (M i) (f i))"
+using assms proof induct
+  case empty
+  interpret finite_product_sigma_finite M "{}" by default auto
+  then show ?case by simp
+next
+  case (insert i I)
+  note `finite I`[intro, simp]
+  interpret I: finite_product_sigma_finite M I by default auto
+  have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
+    using insert by (auto intro!: setprod_cong)
+  have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^isub>M J M)"
+    using sets_into_space insert
+    by (intro sigma_algebra.borel_measurable_extreal_setprod sigma_algebra_product_algebra
+              measurable_comp[OF measurable_component_singleton, unfolded comp_def])
+       auto
+  then show ?case
+    apply (simp add: product_positive_integral_insert[OF insert(1,2) prod])
+    apply (simp add: insert * pos borel setprod_extreal_pos M.positive_integral_multc)
+    apply (subst I.positive_integral_cmult)
+    apply (auto simp add: pos borel insert setprod_extreal_pos positive_integral_positive)
+    done
+qed
+
+lemma (in product_sigma_finite) product_integral_singleton:
+  assumes f: "f \<in> borel_measurable (M i)"
+  shows "(\<integral>x. f (x i) \<partial>Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f"
+proof -
+  interpret I: finite_product_sigma_finite M "{i}" by default simp
+  have *: "(\<lambda>x. extreal (f x)) \<in> borel_measurable (M i)"
+    "(\<lambda>x. extreal (- f x)) \<in> borel_measurable (M i)"
+    using assms by auto
+  show ?thesis
+    unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] ..
+qed
+
+lemma (in product_sigma_finite) product_integral_fold:
+  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
+  and f: "integrable (Pi\<^isub>M (I \<union> J) M) f"
+  shows "integral\<^isup>L (Pi\<^isub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I x J y) \<partial>Pi\<^isub>M J M) \<partial>Pi\<^isub>M I M)"
+proof -
+  interpret I: finite_product_sigma_finite M I by default fact
+  interpret J: finite_product_sigma_finite M J by default fact
+  have "finite (I \<union> J)" using fin by auto
+  interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
+  interpret P: pair_sigma_finite I.P J.P by default
+  let ?M = "\<lambda>(x, y). merge I x J y"
+  let ?f = "\<lambda>x. f (?M x)"
+  show ?thesis
+  proof (subst P.integrable_fst_measurable(2)[of ?f, simplified])
+    have 1: "sigma_algebra IJ.P" by default
+    have 2: "?M \<in> measure_preserving P.P IJ.P" using measure_preserving_merge[OF assms(1,2,3)] .
+    have 3: "integrable (Pi\<^isub>M (I \<union> J) M) f" by fact
+    then have 4: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
+      by (simp add: integrable_def)
+    show "integrable P.P ?f"
+      by (rule P.integrable_vimage[where f=f, OF 1 2 3])
+    show "integral\<^isup>L IJ.P f = integral\<^isup>L P.P ?f"
+      by (rule P.integral_vimage[where f=f, OF 1 2 4])
+  qed
+qed
+
+lemma (in product_sigma_finite) product_integral_insert:
+  assumes [simp]: "finite I" "i \<notin> I"
+    and f: "integrable (Pi\<^isub>M (insert i I) M) f"
+  shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^isub>M I M)"
+proof -
+  interpret I: finite_product_sigma_finite M I by default auto
+  interpret I': finite_product_sigma_finite M "insert i I" by default auto
+  interpret i: finite_product_sigma_finite M "{i}" by default auto
+  interpret P: pair_sigma_finite I.P i.P ..
+  have IJ: "I \<inter> {i} = {}" by auto
+  show ?thesis
+    unfolding product_integral_fold[OF IJ, simplified, OF f]
+  proof (rule I.integral_cong, subst product_integral_singleton)
+    fix x assume x: "x \<in> space I.P"
+    let "?f y" = "f (restrict (x(i := y)) (insert i I))"
+    have f'_eq: "\<And>y. ?f y = f (x(i := y))"
+      using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def)
+    have f: "f \<in> borel_measurable I'.P" using f unfolding integrable_def by auto
+    show "?f \<in> borel_measurable (M i)"
+      unfolding measurable_cong[OF f'_eq]
+      using measurable_comp[OF measurable_component_update f] x `i \<notin> I`
+      by (simp add: comp_def)
+    show "integral\<^isup>L (M i) ?f = integral\<^isup>L (M i) (\<lambda>y. f (x(i := y)))"
+      unfolding f'_eq by simp
+  qed
+qed
+
+lemma (in product_sigma_finite) product_integrable_setprod:
+  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
+  assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
+  shows "integrable (Pi\<^isub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f")
+proof -
+  interpret finite_product_sigma_finite M I by default fact
+  have f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
+    using integrable unfolding integrable_def by auto
+  then have borel: "?f \<in> borel_measurable P"
+    using measurable_comp[OF measurable_component_singleton f]
+    by (auto intro!: borel_measurable_setprod simp: comp_def)
+  moreover have "integrable (Pi\<^isub>M I M) (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)"
+  proof (unfold integrable_def, intro conjI)
+    show "(\<lambda>x. abs (?f x)) \<in> borel_measurable P"
+      using borel by auto
+    have "(\<integral>\<^isup>+x. extreal (abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. extreal (abs (f i (x i)))) \<partial>P)"
+      by (simp add: setprod_extreal abs_setprod)
+    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. extreal (abs (f i x)) \<partial>M i))"
+      using f by (subst product_positive_integral_setprod) auto
+    also have "\<dots> < \<infinity>"
+      using integrable[THEN M.integrable_abs]
+      by (simp add: setprod_PInf integrable_def M.positive_integral_positive)
+    finally show "(\<integral>\<^isup>+x. extreal (abs (?f x)) \<partial>P) \<noteq> \<infinity>" by auto
+    have "(\<integral>\<^isup>+x. extreal (- abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. 0 \<partial>P)"
+      by (intro positive_integral_cong_pos) auto
+    then show "(\<integral>\<^isup>+x. extreal (- abs (?f x)) \<partial>P) \<noteq> \<infinity>" by simp
+  qed
+  ultimately show ?thesis
+    by (rule integrable_abs_iff[THEN iffD1])
+qed
+
+lemma (in product_sigma_finite) product_integral_setprod:
+  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
+  assumes "finite I" "I \<noteq> {}" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
+  shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>L (M i) (f i))"
+using assms proof (induct rule: finite_ne_induct)
+  case (singleton i)
+  then show ?case by (simp add: product_integral_singleton integrable_def)
+next
+  case (insert i I)
+  then have iI: "finite (insert i I)" by auto
+  then have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow>
+    integrable (Pi\<^isub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))"
+    by (intro product_integrable_setprod insert(5)) (auto intro: finite_subset)
+  interpret I: finite_product_sigma_finite M I by default fact
+  have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
+    using `i \<notin> I` by (auto intro!: setprod_cong)
+  show ?case
+    unfolding product_integral_insert[OF insert(1,3) prod[OF subset_refl]]
+    by (simp add: * insert integral_multc I.integral_cmult[OF prod] subset_insertI)
+qed
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -0,0 +1,787 @@
+(*  Title:      HOL/Probability/Infinite_Product_Measure.thy
+    Author:     Johannes Hölzl, TU München
+*)
+
+header {*Infinite Product Measure*}
+
+theory Infinite_Product_Measure
+  imports Probability_Measure
+begin
+
+lemma restrict_extensional_sub[intro]: "A \<subseteq> B \<Longrightarrow> restrict f A \<in> extensional B"
+  unfolding restrict_def extensional_def by auto
+
+lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \<inter> B)"
+  unfolding restrict_def by (simp add: fun_eq_iff)
+
+lemma split_merge: "P (merge I x J y i) \<longleftrightarrow> (i \<in> I \<longrightarrow> P (x i)) \<and> (i \<in> J - I \<longrightarrow> P (y i)) \<and> (i \<notin> I \<union> J \<longrightarrow> P undefined)"
+  unfolding merge_def by auto
+
+lemma extensional_merge_sub: "I \<union> J \<subseteq> K \<Longrightarrow> merge I x J y \<in> extensional K"
+  unfolding merge_def extensional_def by auto
+
+lemma injective_vimage_restrict:
+  assumes J: "J \<subseteq> I"
+  and sets: "A \<subseteq> (\<Pi>\<^isub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^isub>E i\<in>J. S i)" and ne: "(\<Pi>\<^isub>E i\<in>I. S i) \<noteq> {}"
+  and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"
+  shows "A = B"
+proof  (intro set_eqI)
+  fix x
+  from ne obtain y where y: "\<And>i. i \<in> I \<Longrightarrow> y i \<in> S i" by auto
+  have "J \<inter> (I - J) = {}" by auto
+  show "x \<in> A \<longleftrightarrow> x \<in> B"
+  proof cases
+    assume x: "x \<in> (\<Pi>\<^isub>E i\<in>J. S i)"
+    have "x \<in> A \<longleftrightarrow> merge J x (I - J) y \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"
+      using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub split: split_merge)
+    then show "x \<in> A \<longleftrightarrow> x \<in> B"
+      using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub eq split: split_merge)
+  next
+    assume "x \<notin> (\<Pi>\<^isub>E i\<in>J. S i)" with sets show "x \<in> A \<longleftrightarrow> x \<in> B" by auto
+  qed
+qed
+
+locale product_prob_space =
+  fixes M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" and I :: "'i set"
+  assumes prob_spaces: "\<And>i. prob_space (M i)"
+  and I_not_empty: "I \<noteq> {}"
+
+locale finite_product_prob_space = product_prob_space M I
+  for M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" and I :: "'i set" +
+  assumes finite_index'[intro]: "finite I"
+
+sublocale product_prob_space \<subseteq> M: prob_space "M i" for i
+  by (rule prob_spaces)
+
+sublocale product_prob_space \<subseteq> product_sigma_finite
+  by default
+
+sublocale finite_product_prob_space \<subseteq> finite_product_sigma_finite
+  by default (fact finite_index')
+
+sublocale finite_product_prob_space \<subseteq> prob_space "Pi\<^isub>M I M"
+proof
+  show "measure P (space P) = 1"
+    by (simp add: measure_times measure_space_1 setprod_1)
+qed
+
+lemma (in product_prob_space) measure_preserving_restrict:
+  assumes "J \<noteq> {}" "J \<subseteq> K" "finite K"
+  shows "(\<lambda>f. restrict f J) \<in> measure_preserving (\<Pi>\<^isub>M i\<in>K. M i) (\<Pi>\<^isub>M i\<in>J. M i)" (is "?R \<in> _")
+proof -
+  interpret K: finite_product_prob_space M K
+    by default (insert assms, auto)
+  have J: "J \<noteq> {}" "finite J" using assms by (auto simp add: finite_subset)
+  interpret J: finite_product_prob_space M J
+    by default (insert J, auto)
+  from J.sigma_finite_pairs guess F .. note F = this
+  then have [simp,intro]: "\<And>k i. k \<in> J \<Longrightarrow> F k i \<in> sets (M k)"
+    by auto
+  let "?F i" = "\<Pi>\<^isub>E k\<in>J. F k i"
+  let ?J = "product_algebra_generator J M \<lparr> measure := measure (Pi\<^isub>M J M) \<rparr>"
+  have "?R \<in> measure_preserving (\<Pi>\<^isub>M i\<in>K. M i) (sigma ?J)"
+  proof (rule K.measure_preserving_Int_stable)
+    show "Int_stable ?J"
+      by (auto simp: Int_stable_def product_algebra_generator_def PiE_Int)
+    show "range ?F \<subseteq> sets ?J" "incseq ?F" "(\<Union>i. ?F i) = space ?J"
+      using F by auto
+    show "\<And>i. measure ?J (?F i) \<noteq> \<infinity>"
+      using F by (simp add: J.measure_times setprod_PInf)
+    have "measure_space (Pi\<^isub>M J M)" by default
+    then show "measure_space (sigma ?J)"
+      by (simp add: product_algebra_def sigma_def)
+    show "?R \<in> measure_preserving (Pi\<^isub>M K M) ?J"
+    proof (simp add: measure_preserving_def measurable_def product_algebra_generator_def del: vimage_Int,
+           safe intro!: restrict_extensional)
+      fix x k assume "k \<in> J" "x \<in> (\<Pi> i\<in>K. space (M i))"
+      then show "x k \<in> space (M k)" using `J \<subseteq> K` by auto
+    next
+      fix E assume "E \<in> (\<Pi> i\<in>J. sets (M i))"
+      then have E: "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)" by auto
+      then have *: "?R -` Pi\<^isub>E J E \<inter> (\<Pi>\<^isub>E i\<in>K. space (M i)) = (\<Pi>\<^isub>E i\<in>K. if i \<in> J then E i else space (M i))"
+        (is "?X = Pi\<^isub>E K ?M")
+        using `J \<subseteq> K` sets_into_space by (auto simp: Pi_iff split: split_if_asm) blast+
+      with E show "?X \<in> sets (Pi\<^isub>M K M)"
+        by (auto intro!: product_algebra_generatorI)
+      have "measure (Pi\<^isub>M J M) (Pi\<^isub>E J E) = (\<Prod>i\<in>J. measure (M i) (?M i))"
+        using E by (simp add: J.measure_times)
+      also have "\<dots> = measure (Pi\<^isub>M K M) ?X"
+        unfolding * using E `finite K` `J \<subseteq> K`
+        by (auto simp: K.measure_times M.measure_space_1
+                 cong del: setprod_cong
+                 intro!: setprod_mono_one_left)
+      finally show "measure (Pi\<^isub>M J M) (Pi\<^isub>E J E) = measure (Pi\<^isub>M K M) ?X" .
+    qed
+  qed
+  then show ?thesis
+    by (simp add: product_algebra_def sigma_def)
+qed
+
+lemma (in product_prob_space) measurable_restrict:
+  assumes *: "J \<noteq> {}" "J \<subseteq> K" "finite K"
+  shows "(\<lambda>f. restrict f J) \<in> measurable (\<Pi>\<^isub>M i\<in>K. M i) (\<Pi>\<^isub>M i\<in>J. M i)"
+  using measure_preserving_restrict[OF *]
+  by (rule measure_preservingD2)
+
+definition (in product_prob_space)
+  "emb J K X = (\<lambda>x. restrict x K) -` X \<inter> space (Pi\<^isub>M J M)"
+
+lemma (in product_prob_space) emb_trans[simp]:
+  "J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> emb L K (emb K J X) = emb L J X"
+  by (auto simp add: Int_absorb1 emb_def)
+
+lemma (in product_prob_space) emb_empty[simp]:
+  "emb K J {} = {}"
+  by (simp add: emb_def)
+
+lemma (in product_prob_space) emb_Pi:
+  assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
+  shows "emb K J (Pi\<^isub>E J X) = (\<Pi>\<^isub>E i\<in>K. if i \<in> J then X i else space (M i))"
+  using assms space_closed
+  by (auto simp: emb_def Pi_iff split: split_if_asm) blast+
+
+lemma (in product_prob_space) emb_injective:
+  assumes "J \<noteq> {}" "J \<subseteq> L" "finite J" and sets: "X \<in> sets (Pi\<^isub>M J M)" "Y \<in> sets (Pi\<^isub>M J M)"
+  assumes "emb L J X = emb L J Y"
+  shows "X = Y"
+proof -
+  interpret J: finite_product_sigma_finite M J by default fact
+  show "X = Y"
+  proof (rule injective_vimage_restrict)
+    show "X \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))" "Y \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))"
+      using J.sets_into_space sets by auto
+    have "\<forall>i\<in>L. \<exists>x. x \<in> space (M i)"
+      using M.not_empty by auto
+    from bchoice[OF this]
+    show "(\<Pi>\<^isub>E i\<in>L. space (M i)) \<noteq> {}" by auto
+    show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i)) = (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i))"
+      using `emb L J X = emb L J Y` by (simp add: emb_def)
+  qed fact
+qed
+
+lemma (in product_prob_space) emb_id:
+  "B \<subseteq> (\<Pi>\<^isub>E i\<in>L. space (M i)) \<Longrightarrow> emb L L B = B"
+  by (auto simp: emb_def Pi_iff subset_eq extensional_restrict)
+
+lemma (in product_prob_space) emb_simps:
+  shows "emb L K (A \<union> B) = emb L K A \<union> emb L K B"
+    and "emb L K (A \<inter> B) = emb L K A \<inter> emb L K B"
+    and "emb L K (A - B) = emb L K A - emb L K B"
+  by (auto simp: emb_def)
+
+lemma (in product_prob_space) measurable_emb[intro,simp]:
+  assumes *: "J \<noteq> {}" "J \<subseteq> L" "finite L" "X \<in> sets (Pi\<^isub>M J M)"
+  shows "emb L J X \<in> sets (Pi\<^isub>M L M)"
+  using measurable_restrict[THEN measurable_sets, OF *] by (simp add: emb_def)
+
+lemma (in product_prob_space) measure_emb[intro,simp]:
+  assumes *: "J \<noteq> {}" "J \<subseteq> L" "finite L" "X \<in> sets (Pi\<^isub>M J M)"
+  shows "measure (Pi\<^isub>M L M) (emb L J X) = measure (Pi\<^isub>M J M) X"
+  using measure_preserving_restrict[THEN measure_preservingD, OF *]
+  by (simp add: emb_def)
+
+definition (in product_prob_space) generator :: "('i \<Rightarrow> 'a) measure_space" where
+  "generator = \<lparr>
+    space = (\<Pi>\<^isub>E i\<in>I. space (M i)),
+    sets = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^isub>M J M)),
+    measure = undefined
+  \<rparr>"
+
+lemma (in product_prob_space) generatorI:
+  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> A = emb I J X \<Longrightarrow> A \<in> sets generator"
+  unfolding generator_def by auto
+
+lemma (in product_prob_space) generatorI':
+  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> emb I J X \<in> sets generator"
+  unfolding generator_def by auto
+
+lemma (in product_sigma_finite)
+  assumes "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
+  shows measure_fold_integral:
+    "measure (Pi\<^isub>M (I \<union> J) M) A = (\<integral>\<^isup>+x. measure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M)) \<partial>Pi\<^isub>M I M)" (is ?I)
+    and measure_fold_measurable:
+    "(\<lambda>x. measure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M))) \<in> borel_measurable (Pi\<^isub>M I M)" (is ?B)
+proof -
+  interpret I: finite_product_sigma_finite M I by default fact
+  interpret J: finite_product_sigma_finite M J by default fact
+  interpret IJ: pair_sigma_finite I.P J.P ..
+  show ?I
+    unfolding measure_fold[OF assms]
+    apply (subst IJ.pair_measure_alt)
+    apply (intro measurable_sets[OF _ A] measurable_merge assms)
+    apply (auto simp: vimage_compose[symmetric] comp_def space_pair_measure
+      intro!: I.positive_integral_cong)
+    done
+
+  have "(\<lambda>(x, y). merge I x J y) -` A \<inter> space (I.P \<Otimes>\<^isub>M J.P) \<in> sets (I.P \<Otimes>\<^isub>M J.P)"
+    by (intro measurable_sets[OF _ A] measurable_merge assms)
+  from IJ.measure_cut_measurable_fst[OF this]
+  show ?B
+    apply (auto simp: vimage_compose[symmetric] comp_def space_pair_measure)
+    apply (subst (asm) measurable_cong)
+    apply auto
+    done
+qed
+
+lemma (in prob_space) measure_le_1: "X \<in> sets M \<Longrightarrow> \<mu> X \<le> 1"
+  unfolding measure_space_1[symmetric]
+  using sets_into_space
+  by (intro measure_mono) auto
+
+definition (in product_prob_space)
+  "\<mu>G A =
+    (THE x. \<forall>J. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> J \<subseteq> I \<longrightarrow> (\<forall>X\<in>sets (Pi\<^isub>M J M). A = emb I J X \<longrightarrow> x = measure (Pi\<^isub>M J M) X))"
+
+lemma (in product_prob_space) \<mu>G_spec:
+  assumes J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
+  shows "\<mu>G A = measure (Pi\<^isub>M J M) X"
+  unfolding \<mu>G_def
+proof (intro the_equality allI impI ballI)
+  fix K Y assume K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "A = emb I K Y" "Y \<in> sets (Pi\<^isub>M K M)"
+  have "measure (Pi\<^isub>M K M) Y = measure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) K Y)"
+    using K J by simp
+  also have "emb (K \<union> J) K Y = emb (K \<union> J) J X"
+    using K J by (simp add: emb_injective[of "K \<union> J" I])
+  also have "measure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) J X) = measure (Pi\<^isub>M J M) X"
+    using K J by simp
+  finally show "measure (Pi\<^isub>M J M) X = measure (Pi\<^isub>M K M) Y" ..
+qed (insert J, force)
+
+lemma (in product_prob_space) \<mu>G_eq:
+  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> \<mu>G (emb I J X) = measure (Pi\<^isub>M J M) X"
+  by (intro \<mu>G_spec) auto
+
+lemma (in product_prob_space) generator_Ex:
+  assumes *: "A \<in> sets generator"
+  shows "\<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A = emb I J X \<and> \<mu>G A = measure (Pi\<^isub>M J M) X"
+proof -
+  from * obtain J X where J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
+    unfolding generator_def by auto
+  with \<mu>G_spec[OF this] show ?thesis by auto
+qed
+
+lemma (in product_prob_space) generatorE:
+  assumes A: "A \<in> sets generator"
+  obtains J X where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A" "\<mu>G A = measure (Pi\<^isub>M J M) X"
+proof -
+  from generator_Ex[OF A] obtain X J where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A"
+    "\<mu>G A = measure (Pi\<^isub>M J M) X" by auto
+  then show thesis by (intro that) auto
+qed
+
+lemma (in product_prob_space) merge_sets:
+  assumes "finite J" "finite K" "J \<inter> K = {}" and A: "A \<in> sets (Pi\<^isub>M (J \<union> K) M)" and x: "x \<in> space (Pi\<^isub>M J M)"
+  shows "merge J x K -` A \<inter> space (Pi\<^isub>M K M) \<in> sets (Pi\<^isub>M K M)"
+proof -
+  interpret J: finite_product_sigma_algebra M J by default fact
+  interpret K: finite_product_sigma_algebra M K by default fact
+  interpret JK: pair_sigma_algebra J.P K.P ..
+
+  from JK.measurable_cut_fst[OF
+    measurable_merge[THEN measurable_sets, OF `J \<inter> K = {}`], OF A, of x] x
+  show ?thesis
+      by (simp add: space_pair_measure comp_def vimage_compose[symmetric])
+qed
+
+lemma (in product_prob_space) merge_emb:
+  assumes "K \<subseteq> I" "J \<subseteq> I" and y: "y \<in> space (Pi\<^isub>M J M)"
+  shows "(merge J y (I - J) -` emb I K X \<inter> space (Pi\<^isub>M I M)) =
+    emb I (K - J) (merge J y (K - J) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M))"
+proof -
+  have [simp]: "\<And>x J K L. merge J y K (restrict x L) = merge J y (K \<inter> L) x"
+    by (auto simp: restrict_def merge_def)
+  have [simp]: "\<And>x J K L. restrict (merge J y K x) L = merge (J \<inter> L) y (K \<inter> L) x"
+    by (auto simp: restrict_def merge_def)
+  have [simp]: "(I - J) \<inter> K = K - J" using `K \<subseteq> I` `J \<subseteq> I` by auto
+  have [simp]: "(K - J) \<inter> (K \<union> J) = K - J" by auto
+  have [simp]: "(K - J) \<inter> K = K - J" by auto
+  from y `K \<subseteq> I` `J \<subseteq> I` show ?thesis
+    by (simp split: split_merge add: emb_def Pi_iff extensional_merge_sub set_eq_iff) auto
+qed
+
+definition (in product_prob_space) infprod_algebra :: "('i \<Rightarrow> 'a) measure_space" where
+  "infprod_algebra = sigma generator \<lparr> measure :=
+    (SOME \<mu>. (\<forall>s\<in>sets generator. \<mu> s = \<mu>G s) \<and>
+       measure_space \<lparr>space = space generator, sets = sets (sigma generator), measure = \<mu>\<rparr>)\<rparr>"
+
+syntax
+  "_PiP"  :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme"  ("(3PIP _:_./ _)" 10)
+
+syntax (xsymbols)
+  "_PiP" :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme"  ("(3\<Pi>\<^isub>P _\<in>_./ _)"   10)
+
+syntax (HTML output)
+  "_PiP" :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme"  ("(3\<Pi>\<^isub>P _\<in>_./ _)"   10)
+
+abbreviation
+  "Pi\<^isub>P I M \<equiv> product_prob_space.infprod_algebra M I"
+
+translations
+  "PIP x:I. M" == "CONST Pi\<^isub>P I (%x. M)"
+
+sublocale product_prob_space \<subseteq> G: algebra generator
+proof
+  let ?G = generator
+  show "sets ?G \<subseteq> Pow (space ?G)"
+    by (auto simp: generator_def emb_def)
+  from I_not_empty
+  obtain i where "i \<in> I" by auto
+  then show "{} \<in> sets ?G"
+    by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\<lambda>i. {}"]
+      simp: product_algebra_def sigma_def sigma_sets.Empty generator_def emb_def)
+  from `i \<in> I` show "space ?G \<in> sets ?G"
+    by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^isub>E {i} (\<lambda>i. space (M i))"]
+      simp: generator_def emb_def)
+  fix A assume "A \<in> sets ?G"
+  then obtain JA XA where XA: "JA \<noteq> {}" "finite JA" "JA \<subseteq> I" "XA \<in> sets (Pi\<^isub>M JA M)" and A: "A = emb I JA XA"
+    by (auto simp: generator_def)
+  fix B assume "B \<in> sets ?G"
+  then obtain JB XB where XB: "JB \<noteq> {}" "finite JB" "JB \<subseteq> I" "XB \<in> sets (Pi\<^isub>M JB M)" and B: "B = emb I JB XB"
+    by (auto simp: generator_def)
+  let ?RA = "emb (JA \<union> JB) JA XA"
+  let ?RB = "emb (JA \<union> JB) JB XB"
+  interpret JAB: finite_product_sigma_algebra M "JA \<union> JB"
+    by default (insert XA XB, auto)
+  have *: "A - B = emb I (JA \<union> JB) (?RA - ?RB)" "A \<union> B = emb I (JA \<union> JB) (?RA \<union> ?RB)"
+    using XA A XB B by (auto simp: emb_simps)
+  then show "A - B \<in> sets ?G" "A \<union> B \<in> sets ?G"
+    using XA XB by (auto intro!: generatorI')
+qed
+
+lemma (in product_prob_space) positive_\<mu>G: "positive generator \<mu>G"
+proof (intro positive_def[THEN iffD2] conjI ballI)
+  from generatorE[OF G.empty_sets] guess J X . note this[simp]
+  interpret J: finite_product_sigma_finite M J by default fact
+  have "X = {}"
+    by (rule emb_injective[of J I]) simp_all
+  then show "\<mu>G {} = 0" by simp
+next
+  fix A assume "A \<in> sets generator"
+  from generatorE[OF this] guess J X . note this[simp]
+  interpret J: finite_product_sigma_finite M J by default fact
+  show "0 \<le> \<mu>G A" by simp
+qed
+
+lemma (in product_prob_space) additive_\<mu>G: "additive generator \<mu>G"
+proof (intro additive_def[THEN iffD2] ballI impI)
+  fix A assume "A \<in> sets generator" with generatorE guess J X . note J = this
+  fix B assume "B \<in> sets generator" with generatorE guess K Y . note K = this
+  assume "A \<inter> B = {}"
+  have JK: "J \<union> K \<noteq> {}" "J \<union> K \<subseteq> I" "finite (J \<union> K)"
+    using J K by auto
+  interpret JK: finite_product_sigma_finite M "J \<union> K" by default fact
+  have JK_disj: "emb (J \<union> K) J X \<inter> emb (J \<union> K) K Y = {}"
+    apply (rule emb_injective[of "J \<union> K" I])
+    apply (insert `A \<inter> B = {}` JK J K)
+    apply (simp_all add: JK.Int emb_simps)
+    done
+  have AB: "A = emb I (J \<union> K) (emb (J \<union> K) J X)" "B = emb I (J \<union> K) (emb (J \<union> K) K Y)"
+    using J K by simp_all
+  then have "\<mu>G (A \<union> B) = \<mu>G (emb I (J \<union> K) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y))"
+    by (simp add: emb_simps)
+  also have "\<dots> = measure (Pi\<^isub>M (J \<union> K) M) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y)"
+    using JK J(1, 4) K(1, 4) by (simp add: \<mu>G_eq JK.Un)
+  also have "\<dots> = \<mu>G A + \<mu>G B"
+    using J K JK_disj by (simp add: JK.measure_additive[symmetric])
+  finally show "\<mu>G (A \<union> B) = \<mu>G A + \<mu>G B" .
+qed
+
+lemma (in product_prob_space) finite_index_eq_finite_product:
+  assumes "finite I"
+  shows "sets (sigma generator) = sets (Pi\<^isub>M I M)"
+proof safe
+  interpret I: finite_product_sigma_algebra M I by default fact
+  have [simp]: "space generator = space (Pi\<^isub>M I M)"
+    by (simp add: generator_def product_algebra_def)
+  { fix A assume "A \<in> sets (sigma generator)"
+    then show "A \<in> sets I.P" unfolding sets_sigma
+    proof induct
+      case (Basic A)
+      from generatorE[OF this] guess J X . note J = this
+      with `finite I` have "emb I J X \<in> sets I.P" by auto
+      with `emb I J X = A` show "A \<in> sets I.P" by simp
+    qed auto }
+  { fix A assume "A \<in> sets I.P"
+    moreover with I.sets_into_space have "emb I I A = A" by (intro emb_id) auto
+    ultimately show "A \<in> sets (sigma generator)"
+      using `finite I` I_not_empty unfolding sets_sigma
+      by (intro sigma_sets.Basic generatorI[of I A]) auto }
+qed
+
+lemma (in product_prob_space) extend_\<mu>G:
+  "\<exists>\<mu>. (\<forall>s\<in>sets generator. \<mu> s = \<mu>G s) \<and>
+       measure_space \<lparr>space = space generator, sets = sets (sigma generator), measure = \<mu>\<rparr>"
+proof cases
+  assume "finite I"
+  interpret I: finite_product_sigma_finite M I by default fact
+  show ?thesis
+  proof (intro exI[of _ "measure (Pi\<^isub>M I M)"] ballI conjI)
+    fix A assume "A \<in> sets generator"
+    from generatorE[OF this] guess J X . note J = this
+    from J(1-4) `finite I` show "measure I.P A = \<mu>G A"
+      unfolding J(6)
+      by (subst J(5)[symmetric]) (simp add: measure_emb)
+  next
+    have [simp]: "space generator = space (Pi\<^isub>M I M)"
+      by (simp add: generator_def product_algebra_def)
+    have "\<lparr>space = space generator, sets = sets (sigma generator), measure = measure I.P\<rparr>
+      = I.P" (is "?P = _")
+      by (auto intro!: measure_space.equality simp: finite_index_eq_finite_product[OF `finite I`])
+    then show "measure_space ?P" by simp default
+  qed
+next
+  let ?G = generator
+  assume "\<not> finite I"
+  note \<mu>G_mono =
+    G.additive_increasing[OF positive_\<mu>G additive_\<mu>G, THEN increasingD]
+
+  { fix Z J assume J: "J \<noteq> {}" "finite J" "J \<subseteq> I" and Z: "Z \<in> sets ?G"
+
+    from `infinite I` `finite J` obtain k where k: "k \<in> I" "k \<notin> J"
+      by (metis rev_finite_subset subsetI)
+    moreover from Z guess K' X' by (rule generatorE)
+    moreover def K \<equiv> "insert k K'"
+    moreover def X \<equiv> "emb K K' X'"
+    ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^isub>M K M)" "Z = emb I K X"
+      "K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = measure (Pi\<^isub>M K M) X"
+      by (auto simp: subset_insertI)
+
+    let "?M y" = "merge J y (K - J) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M)"
+    { fix y assume y: "y \<in> space (Pi\<^isub>M J M)"
+      note * = merge_emb[OF `K \<subseteq> I` `J \<subseteq> I` y, of X]
+      moreover
+      have **: "?M y \<in> sets (Pi\<^isub>M (K - J) M)"
+        using J K y by (intro merge_sets) auto
+      ultimately
+      have ***: "(merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<in> sets ?G"
+        using J K by (intro generatorI) auto
+      have "\<mu>G (merge J y (I - J) -` emb I K X \<inter> space (Pi\<^isub>M I M)) = measure (Pi\<^isub>M (K - J) M) (?M y)"
+        unfolding * using K J by (subst \<mu>G_eq[OF _ _ _ **]) auto
+      note * ** *** this }
+    note merge_in_G = this
+
+    have "finite (K - J)" using K by auto
+
+    interpret J: finite_product_prob_space M J by default fact+
+    interpret KmJ: finite_product_prob_space M "K - J" by default fact+
+
+    have "\<mu>G Z = measure (Pi\<^isub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)"
+      using K J by simp
+    also have "\<dots> = (\<integral>\<^isup>+ x. measure (Pi\<^isub>M (K - J) M) (?M x) \<partial>Pi\<^isub>M J M)"
+      using K J by (subst measure_fold_integral) auto
+    also have "\<dots> = (\<integral>\<^isup>+ y. \<mu>G (merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<partial>Pi\<^isub>M J M)"
+      (is "_ = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)")
+    proof (intro J.positive_integral_cong)
+      fix x assume x: "x \<in> space (Pi\<^isub>M J M)"
+      with K merge_in_G(2)[OF this]
+      show "measure (Pi\<^isub>M (K - J) M) (?M x) = \<mu>G (?MZ x)"
+        unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst \<mu>G_eq) auto
+    qed
+    finally have fold: "\<mu>G Z = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)" .
+
+    { fix x assume x: "x \<in> space (Pi\<^isub>M J M)"
+      then have "\<mu>G (?MZ x) \<le> 1"
+        unfolding merge_in_G(4)[OF x] `Z = emb I K X`
+        by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) }
+    note le_1 = this
+
+    let "?q y" = "\<mu>G (merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M))"
+    have "?q \<in> borel_measurable (Pi\<^isub>M J M)"
+      unfolding `Z = emb I K X` using J K merge_in_G(3)
+      by (simp add: merge_in_G  \<mu>G_eq measure_fold_measurable
+               del: space_product_algebra cong: measurable_cong)
+    note this fold le_1 merge_in_G(3) }
+  note fold = this
+
+  show ?thesis
+  proof (rule G.caratheodory_empty_continuous[OF positive_\<mu>G additive_\<mu>G])
+    fix A assume "A \<in> sets ?G"
+    with generatorE guess J X . note JX = this
+    interpret JK: finite_product_prob_space M J by default fact+
+    with JX show "\<mu>G A \<noteq> \<infinity>" by simp
+  next
+    fix A assume A: "range A \<subseteq> sets ?G" "decseq A" "(\<Inter>i. A i) = {}"
+    then have "decseq (\<lambda>i. \<mu>G (A i))"
+      by (auto intro!: \<mu>G_mono simp: decseq_def)
+    moreover
+    have "(INF i. \<mu>G (A i)) = 0"
+    proof (rule ccontr)
+      assume "(INF i. \<mu>G (A i)) \<noteq> 0" (is "?a \<noteq> 0")
+      moreover have "0 \<le> ?a"
+        using A positive_\<mu>G by (auto intro!: le_INFI simp: positive_def)
+      ultimately have "0 < ?a" by auto
+
+      have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = measure (Pi\<^isub>M J M) X"
+        using A by (intro allI generator_Ex) auto
+      then obtain J' X' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I" "\<And>n. X' n \<in> sets (Pi\<^isub>M (J' n) M)"
+        and A': "\<And>n. A n = emb I (J' n) (X' n)"
+        unfolding choice_iff by blast
+      moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)"
+      moreover def X \<equiv> "\<lambda>n. emb (J n) (J' n) (X' n)"
+      ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. X n \<in> sets (Pi\<^isub>M (J n) M)"
+        by auto
+      with A' have A_eq: "\<And>n. A n = emb I (J n) (X n)" "\<And>n. A n \<in> sets ?G"
+        unfolding J_def X_def by (subst emb_trans) (insert A, auto)
+
+      have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
+        unfolding J_def by force
+
+      interpret J: finite_product_prob_space M "J i" for i by default fact+
+
+      have a_le_1: "?a \<le> 1"
+        using \<mu>G_spec[of "J 0" "A 0" "X 0"] J A_eq
+        by (auto intro!: INF_leI2[of 0] J.measure_le_1)
+
+      let "?M K Z y" = "merge K y (I - K) -` Z \<inter> space (Pi\<^isub>M I M)"
+
+      { fix Z k assume Z: "range Z \<subseteq> sets ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)"
+        then have Z_sets: "\<And>n. Z n \<in> sets ?G" by auto
+        fix J' assume J': "J' \<noteq> {}" "finite J'" "J' \<subseteq> I"
+        interpret J': finite_product_prob_space M J' by default fact+
+
+        let "?q n y" = "\<mu>G (?M J' (Z n) y)"
+        let "?Q n" = "?q n -` {?a / 2^(k+1) ..} \<inter> space (Pi\<^isub>M J' M)"
+        { fix n
+          have "?q n \<in> borel_measurable (Pi\<^isub>M J' M)"
+            using Z J' by (intro fold(1)) auto
+          then have "?Q n \<in> sets (Pi\<^isub>M J' M)"
+            by (rule measurable_sets) auto }
+        note Q_sets = this
+
+        have "?a / 2^(k+1) \<le> (INF n. measure (Pi\<^isub>M J' M) (?Q n))"
+        proof (intro le_INFI)
+          fix n
+          have "?a / 2^k \<le> \<mu>G (Z n)" using Z by auto
+          also have "\<dots> \<le> (\<integral>\<^isup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \<partial>Pi\<^isub>M J' M)"
+            unfolding fold(2)[OF J' `Z n \<in> sets ?G`]
+          proof (intro J'.positive_integral_mono)
+            fix x assume x: "x \<in> space (Pi\<^isub>M J' M)"
+            then have "?q n x \<le> 1 + 0"
+              using J' Z fold(3) Z_sets by auto
+            also have "\<dots> \<le> 1 + ?a / 2^(k+1)"
+              using `0 < ?a` by (intro add_mono) auto
+            finally have "?q n x \<le> 1 + ?a / 2^(k+1)" .
+            with x show "?q n x \<le> indicator (?Q n) x + ?a / 2^(k+1)"
+              by (auto split: split_indicator simp del: power_Suc)
+          qed
+          also have "\<dots> = measure (Pi\<^isub>M J' M) (?Q n) + ?a / 2^(k+1)"
+            using `0 \<le> ?a` Q_sets J'.measure_space_1
+            by (subst J'.positive_integral_add) auto
+          finally show "?a / 2^(k+1) \<le> measure (Pi\<^isub>M J' M) (?Q n)" using `?a \<le> 1`
+            by (cases rule: extreal2_cases[of ?a "measure (Pi\<^isub>M J' M) (?Q n)"])
+               (auto simp: field_simps)
+        qed
+        also have "\<dots> = measure (Pi\<^isub>M J' M) (\<Inter>n. ?Q n)"
+        proof (intro J'.continuity_from_above)
+          show "range ?Q \<subseteq> sets (Pi\<^isub>M J' M)" using Q_sets by auto
+          show "decseq ?Q"
+            unfolding decseq_def
+          proof (safe intro!: vimageI[OF refl])
+            fix m n :: nat assume "m \<le> n"
+            fix x assume x: "x \<in> space (Pi\<^isub>M J' M)"
+            assume "?a / 2^(k+1) \<le> ?q n x"
+            also have "?q n x \<le> ?q m x"
+            proof (rule \<mu>G_mono)
+              from fold(4)[OF J', OF Z_sets x]
+              show "?M J' (Z n) x \<in> sets ?G" "?M J' (Z m) x \<in> sets ?G" by auto
+              show "?M J' (Z n) x \<subseteq> ?M J' (Z m) x"
+                using `decseq Z`[THEN decseqD, OF `m \<le> n`] by auto
+            qed
+            finally show "?a / 2^(k+1) \<le> ?q m x" .
+          qed
+        qed (intro J'.finite_measure Q_sets)
+        finally have "(\<Inter>n. ?Q n) \<noteq> {}"
+          using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
+        then have "\<exists>w\<in>space (Pi\<^isub>M J' M). \<forall>n. ?a / 2 ^ (k + 1) \<le> ?q n w" by auto }
+      note Ex_w = this
+
+      let "?q k n y" = "\<mu>G (?M (J k) (A n) y)"
+
+      have "\<forall>n. ?a / 2 ^ 0 \<le> \<mu>G (A n)" by (auto intro: INF_leI)
+      from Ex_w[OF A(1,2) this J(1-3), of 0] guess w0 .. note w0 = this
+
+      let "?P k wk w" =
+        "w \<in> space (Pi\<^isub>M (J (Suc k)) M) \<and> restrict w (J k) = wk \<and> (\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n w)"
+      def w \<equiv> "nat_rec w0 (\<lambda>k wk. Eps (?P k wk))"
+
+      { fix k have w: "w k \<in> space (Pi\<^isub>M (J k) M) \<and>
+          (\<forall>n. ?a / 2 ^ (k + 1) \<le> ?q k n (w k)) \<and> (k \<noteq> 0 \<longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1))"
+        proof (induct k)
+          case 0 with w0 show ?case
+            unfolding w_def nat_rec_0 by auto
+        next
+          case (Suc k)
+          then have wk: "w k \<in> space (Pi\<^isub>M (J k) M)" by auto
+          have "\<exists>w'. ?P k (w k) w'"
+          proof cases
+            assume [simp]: "J k = J (Suc k)"
+            show ?thesis
+            proof (intro exI[of _ "w k"] conjI allI)
+              fix n
+              have "?a / 2 ^ (Suc k + 1) \<le> ?a / 2 ^ (k + 1)"
+                using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: field_simps)
+              also have "\<dots> \<le> ?q k n (w k)" using Suc by auto
+              finally show "?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n (w k)" by simp
+            next
+              show "w k \<in> space (Pi\<^isub>M (J (Suc k)) M)"
+                using Suc by simp
+              then show "restrict (w k) (J k) = w k"
+                by (simp add: extensional_restrict)
+            qed
+          next
+            assume "J k \<noteq> J (Suc k)"
+            with J_mono[of k "Suc k"] have "J (Suc k) - J k \<noteq> {}" (is "?D \<noteq> {}") by auto
+            have "range (\<lambda>n. ?M (J k) (A n) (w k)) \<subseteq> sets ?G"
+              "decseq (\<lambda>n. ?M (J k) (A n) (w k))"
+              "\<forall>n. ?a / 2 ^ (k + 1) \<le> \<mu>G (?M (J k) (A n) (w k))"
+              using `decseq A` fold(4)[OF J(1-3) A_eq(2), of "w k" k] Suc
+              by (auto simp: decseq_def)
+            from Ex_w[OF this `?D \<noteq> {}`] J[of "Suc k"]
+            obtain w' where w': "w' \<in> space (Pi\<^isub>M ?D M)"
+              "\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> \<mu>G (?M ?D (?M (J k) (A n) (w k)) w')" by auto
+            let ?w = "merge (J k) (w k) ?D w'"
+            have [simp]: "\<And>x. merge (J k) (w k) (I - J k) (merge ?D w' (I - ?D) x) =
+              merge (J (Suc k)) ?w (I - (J (Suc k))) x"
+              using J(3)[of "Suc k"] J(3)[of k] J_mono[of k "Suc k"]
+              by (auto intro!: ext split: split_merge)
+            have *: "\<And>n. ?M ?D (?M (J k) (A n) (w k)) w' = ?M (J (Suc k)) (A n) ?w"
+              using w'(1) J(3)[of "Suc k"]
+              by (auto split: split_merge intro!: extensional_merge_sub) force+
+            show ?thesis
+              apply (rule exI[of _ ?w])
+              using w' J_mono[of k "Suc k"] wk unfolding *
+              apply (auto split: split_merge intro!: extensional_merge_sub ext)
+              apply (force simp: extensional_def)
+              done
+          qed
+          then have "?P k (w k) (w (Suc k))"
+            unfolding w_def nat_rec_Suc unfolding w_def[symmetric]
+            by (rule someI_ex)
+          then show ?case by auto
+        qed
+        moreover
+        then have "w k \<in> space (Pi\<^isub>M (J k) M)" by auto
+        moreover
+        from w have "?a / 2 ^ (k + 1) \<le> ?q k k (w k)" by auto
+        then have "?M (J k) (A k) (w k) \<noteq> {}"
+          using positive_\<mu>G[unfolded positive_def] `0 < ?a` `?a \<le> 1`
+          by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
+        then obtain x where "x \<in> ?M (J k) (A k) (w k)" by auto
+        then have "merge (J k) (w k) (I - J k) x \<in> A k" by auto
+        then have "\<exists>x\<in>A k. restrict x (J k) = w k"
+          using `w k \<in> space (Pi\<^isub>M (J k) M)`
+          by (intro rev_bexI) (auto intro!: ext simp: extensional_def)
+        ultimately have "w k \<in> space (Pi\<^isub>M (J k) M)"
+          "\<exists>x\<in>A k. restrict x (J k) = w k"
+          "k \<noteq> 0 \<Longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1)"
+          by auto }
+      note w = this
+
+      { fix k l i assume "k \<le> l" "i \<in> J k"
+        { fix l have "w k i = w (k + l) i"
+          proof (induct l)
+            case (Suc l)
+            from `i \<in> J k` J_mono[of k "k + l"] have "i \<in> J (k + l)" by auto
+            with w(3)[of "k + Suc l"]
+            have "w (k + l) i = w (k + Suc l) i"
+              by (auto simp: restrict_def fun_eq_iff split: split_if_asm)
+            with Suc show ?case by simp
+          qed simp }
+        from this[of "l - k"] `k \<le> l` have "w l i = w k i" by simp }
+      note w_mono = this
+
+      def w' \<equiv> "\<lambda>i. if i \<in> (\<Union>k. J k) then w (LEAST k. i \<in> J k) i else if i \<in> I then (SOME x. x \<in> space (M i)) else undefined"
+      { fix i k assume k: "i \<in> J k"
+        have "w k i = w (LEAST k. i \<in> J k) i"
+          by (intro w_mono Least_le k LeastI[of _ k])
+        then have "w' i = w k i"
+          unfolding w'_def using k by auto }
+      note w'_eq = this
+      have w'_simps1: "\<And>i. i \<notin> I \<Longrightarrow> w' i = undefined"
+        using J by (auto simp: w'_def)
+      have w'_simps2: "\<And>i. i \<notin> (\<Union>k. J k) \<Longrightarrow> i \<in> I \<Longrightarrow> w' i \<in> space (M i)"
+        using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]])
+      { fix i assume "i \<in> I" then have "w' i \<in> space (M i)"
+          using w(1) by (cases "i \<in> (\<Union>k. J k)") (force simp: w'_simps2 w'_eq)+ }
+      note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this
+
+      have w': "w' \<in> space (Pi\<^isub>M I M)"
+        using w(1) by (auto simp add: Pi_iff extensional_def)
+
+      { fix n
+        have "restrict w' (J n) = w n" using w(1)
+          by (auto simp add: fun_eq_iff restrict_def Pi_iff extensional_def)
+        with w[of n] obtain x where "x \<in> A n" "restrict x (J n) = restrict w' (J n)" by auto
+        then have "w' \<in> A n" unfolding A_eq using w' by (auto simp: emb_def) }
+      then have "w' \<in> (\<Inter>i. A i)" by auto
+      with `(\<Inter>i. A i) = {}` show False by auto
+    qed
+    ultimately show "(\<lambda>i. \<mu>G (A i)) ----> 0"
+      using LIMSEQ_extreal_INFI[of "\<lambda>i. \<mu>G (A i)"] by simp
+  qed
+qed
+
+lemma (in product_prob_space) infprod_spec:
+  shows "(\<forall>s\<in>sets generator. measure (Pi\<^isub>P I M) s = \<mu>G s) \<and> measure_space (Pi\<^isub>P I M)"
+proof -
+  let ?P = "\<lambda>\<mu>. (\<forall>A\<in>sets generator. \<mu> A = \<mu>G A) \<and>
+       measure_space \<lparr>space = space generator, sets = sets (sigma generator), measure = \<mu>\<rparr>"
+  have **: "measure infprod_algebra = (SOME \<mu>. ?P \<mu>)"
+    unfolding infprod_algebra_def by simp
+  have *: "Pi\<^isub>P I M = \<lparr>space = space generator, sets = sets (sigma generator), measure = measure (Pi\<^isub>P I M)\<rparr>"
+    unfolding infprod_algebra_def by auto
+  show ?thesis
+    apply (subst (2) *)
+    apply (unfold **)
+    apply (rule someI_ex[where P="?P"])
+    apply (rule extend_\<mu>G)
+    done
+qed
+
+sublocale product_prob_space \<subseteq> measure_space "Pi\<^isub>P I M"
+  using infprod_spec by auto
+
+lemma (in product_prob_space) measure_infprod_emb:
+  assumes "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)"
+  shows "measure (Pi\<^isub>P I M) (emb I J X) = measure (Pi\<^isub>M J M) X"
+proof -
+  have "emb I J X \<in> sets generator"
+    using assms by (rule generatorI')
+  with \<mu>G_eq[OF assms] infprod_spec show ?thesis by auto
+qed
+
+sublocale product_prob_space \<subseteq> prob_space "Pi\<^isub>P I M"
+proof
+  obtain i where "i \<in> I" using I_not_empty by auto
+  interpret i: finite_product_sigma_finite M "{i}" by default auto
+  let ?X = "\<Pi>\<^isub>E i\<in>{i}. space (M i)"
+  have "?X \<in> sets (Pi\<^isub>M {i} M)"
+    by auto
+  from measure_infprod_emb[OF _ _ _ this] `i \<in> I`
+  have "measure (Pi\<^isub>P I M) (emb I {i} ?X) = measure (M i) (space (M i))"
+    by (simp add: i.measure_times)
+  also have "emb I {i} ?X = space (Pi\<^isub>P I M)"
+    using `i \<in> I` by (auto simp: emb_def infprod_algebra_def generator_def)
+  finally show "measure (Pi\<^isub>P I M) (space (Pi\<^isub>P I M)) = 1"
+    using M.measure_space_1 by simp
+qed
+
+lemma (in product_prob_space) measurable_component:
+  assumes "i \<in> I"
+  shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>P I M) (M i)"
+proof (unfold measurable_def, safe)
+  fix x assume "x \<in> space (Pi\<^isub>P I M)"
+  then show "x i \<in> space (M i)"
+    using `i \<in> I` by (auto simp: infprod_algebra_def generator_def)
+next
+  fix A assume "A \<in> sets (M i)"
+  with `i \<in> I` have
+    "(\<Pi>\<^isub>E x \<in> {i}. A) \<in> sets (Pi\<^isub>M {i} M)"
+    "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M) = emb I {i} (\<Pi>\<^isub>E x \<in> {i}. A)"
+    by (auto simp: infprod_algebra_def generator_def emb_def)
+  from generatorI[OF _ _ _ this] `i \<in> I`
+  show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M) \<in> sets (Pi\<^isub>P I M)"
+    unfolding infprod_algebra_def by auto
+qed
+
+end
\ No newline at end of file
--- a/src/HOL/Probability/Information.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Probability/Information.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -7,7 +7,7 @@
 
 theory Information
 imports
-  Probability_Space
+  Probability_Measure
   "~~/src/HOL/Library/Convex"
 begin
 
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Probability/Lebesgue_Measure.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -6,7 +6,7 @@
 header {* Lebsegue measure *}
 
 theory Lebesgue_Measure
-  imports Product_Measure
+  imports Finite_Product_Measure
 begin
 
 subsection {* Standard Cubes *}
@@ -50,9 +50,6 @@
 lemma cube_subset_Suc[intro]: "cube n \<subseteq> cube (Suc n)"
   unfolding cube_def_raw subset_eq apply safe unfolding mem_interval by auto
 
-lemma Pi_iff: "f \<in> Pi I X \<longleftrightarrow> (\<forall>i\<in>I. f i \<in> X i)"
-  unfolding Pi_def by auto
-
 subsection {* Lebesgue measure *} 
 
 definition lebesgue :: "'a::ordered_euclidean_space measure_space" where
@@ -983,4 +980,97 @@
   shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>(lborel_space.P DIM('a))"
   using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def)
 
+
+lemma Int_stable_atLeastAtMost:
+  "Int_stable \<lparr>space = UNIV, sets = range (\<lambda>(a,b). {a::'a::ordered_euclidean_space .. b}) \<rparr>"
+proof (simp add: Int_stable_def image_iff, intro allI)
+  fix a1 b1 a2 b2 :: 'a
+  have "\<forall>i<DIM('a). \<exists>a b. {a1$$i..b1$$i} \<inter> {a2$$i..b2$$i} = {a..b}" by auto
+  then have "\<exists>a b. \<forall>i<DIM('a). {a1$$i..b1$$i} \<inter> {a2$$i..b2$$i} = {a i..b i}"
+    unfolding choice_iff' .
+  then guess a b by safe
+  then have "{a1..b1} \<inter> {a2..b2} = {(\<chi>\<chi> i. a i) .. (\<chi>\<chi> i. b i)}"
+    by (simp add: set_eq_iff eucl_le[where 'a='a] all_conj_distrib[symmetric]) blast
+  then show "\<exists>a' b'. {a1..b1} \<inter> {a2..b2} = {a'..b'}" by blast
+qed
+
+lemma (in sigma_algebra) borel_measurable_sets:
+  assumes "f \<in> measurable borel M" "A \<in> sets M"
+  shows "f -` A \<in> sets borel"
+  using measurable_sets[OF assms] by simp
+
+lemma (in sigma_algebra) measurable_identity[intro,simp]:
+  "(\<lambda>x. x) \<in> measurable M M"
+  unfolding measurable_def by auto
+
+lemma lebesgue_real_affine:
+  fixes X :: "real set"
+  assumes "X \<in> sets borel" and "c \<noteq> 0"
+  shows "measure lebesgue X = extreal \<bar>c\<bar> * measure lebesgue ((\<lambda>x. t + c * x) -` X)"
+    (is "_ = ?\<nu> X")
+proof -
+  let ?E = "\<lparr>space = UNIV, sets = range (\<lambda>(a,b). {a::real .. b})\<rparr> :: real algebra"
+  let "?M \<nu>" = "\<lparr>space = space ?E, sets = sets (sigma ?E), measure = \<nu>\<rparr> :: real measure_space"
+  have *: "?M (measure lebesgue) = lborel"
+    unfolding borel_eq_atLeastAtMost[symmetric]
+    by (simp add: lborel_def lebesgue_def)
+  have **: "?M ?\<nu> = lborel \<lparr> measure := ?\<nu> \<rparr>"
+    unfolding borel_eq_atLeastAtMost[symmetric]
+    by (simp add: lborel_def lebesgue_def)
+  show ?thesis
+  proof (rule measure_unique_Int_stable[where X=X, OF Int_stable_atLeastAtMost], unfold * **)
+    show "X \<in> sets (sigma ?E)"
+      unfolding borel_eq_atLeastAtMost[symmetric] by fact
+    have "\<And>x. \<exists>xa. x \<in> cube xa" apply(rule_tac x=x in mem_big_cube) by fastsimp
+    then show "(\<Union>i. cube i) = space ?E" by auto
+    show "incseq cube" by (intro incseq_SucI cube_subset_Suc)
+    show "range cube \<subseteq> sets ?E"
+      unfolding cube_def_raw by auto
+    show "\<And>i. measure lebesgue (cube i) \<noteq> \<infinity>"
+      by (simp add: cube_def)
+    show "measure_space lborel" by default
+    then interpret sigma_algebra "lborel\<lparr>measure := ?\<nu>\<rparr>"
+      by (auto simp add: measure_space_def)
+    show "measure_space (lborel\<lparr>measure := ?\<nu>\<rparr>)"
+    proof
+      show "positive (lborel\<lparr>measure := ?\<nu>\<rparr>) (measure (lborel\<lparr>measure := ?\<nu>\<rparr>))"
+        by (auto simp: positive_def intro!: extreal_0_le_mult borel.borel_measurable_sets)
+      show "countably_additive (lborel\<lparr>measure := ?\<nu>\<rparr>) (measure (lborel\<lparr>measure := ?\<nu>\<rparr>))"
+      proof (simp add: countably_additive_def, safe)
+        fix A :: "nat \<Rightarrow> real set" assume A: "range A \<subseteq> sets borel" "disjoint_family A"
+        then have Ai: "\<And>i. A i \<in> sets borel" by auto
+        have "(\<Sum>n. measure lebesgue ((\<lambda>x. t + c * x) -` A n)) = measure lebesgue (\<Union>n. (\<lambda>x. t + c * x) -` A n)"
+        proof (intro lborel.measure_countably_additive)
+          { fix n have "(\<lambda>x. t + c * x) -` A n \<inter> space borel \<in> sets borel"
+              using A borel.measurable_ident unfolding id_def
+              by (intro measurable_sets[where A=borel] borel.borel_measurable_add[OF _ borel.borel_measurable_times]) auto }
+          then show "range (\<lambda>i. (\<lambda>x. t + c * x) -` A i) \<subseteq> sets borel" by auto
+          from `disjoint_family A`
+          show "disjoint_family (\<lambda>i. (\<lambda>x. t + c * x) -` A i)"
+            by (rule disjoint_family_on_bisimulation) auto
+        qed
+        with Ai show "(\<Sum>n. ?\<nu> (A n)) = ?\<nu> (UNION UNIV A)"
+          by (subst suminf_cmult_extreal)
+             (auto simp: vimage_UN borel.borel_measurable_sets)
+      qed
+    qed
+    fix X assume "X \<in> sets ?E"
+    then obtain a b where [simp]: "X = {a .. b}" by auto
+    show "measure lebesgue X = ?\<nu> X"
+    proof cases
+      assume "0 < c"
+      then have "(\<lambda>x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}"
+        by (auto simp: field_simps)
+      with `0 < c` show ?thesis
+        by (cases "a \<le> b") (auto simp: field_simps)
+    next
+      assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto
+      then have *: "(\<lambda>x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}"
+        by (auto simp: field_simps)
+      with `c < 0` show ?thesis
+        by (cases "a \<le> b") (auto simp: field_simps)
+    qed
+  qed
+qed
+
 end
--- a/src/HOL/Probability/Probability.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Probability/Probability.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -2,6 +2,8 @@
 imports
   Complete_Measure
   Lebesgue_Measure
+  Probability_Measure
+  Infinite_Product_Measure
   Information
   "ex/Dining_Cryptographers"
   "ex/Koepf_Duermuth_Countermeasure"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Probability_Measure.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -0,0 +1,945 @@
+(*  Title:      HOL/Probability/Probability_Measure.thy
+    Author:     Johannes Hölzl, TU München
+    Author:     Armin Heller, TU München
+*)
+
+header {*Probability measure*}
+
+theory Probability_Measure
+imports Lebesgue_Integration Radon_Nikodym Finite_Product_Measure
+begin
+
+lemma real_of_extreal_inverse[simp]:
+  fixes X :: extreal
+  shows "real (inverse X) = 1 / real X"
+  by (cases X) (auto simp: inverse_eq_divide)
+
+lemma real_of_extreal_le_0[simp]: "real (X :: extreal) \<le> 0 \<longleftrightarrow> (X \<le> 0 \<or> X = \<infinity>)"
+  by (cases X) auto
+
+lemma abs_real_of_extreal[simp]: "\<bar>real (X :: extreal)\<bar> = real \<bar>X\<bar>"
+  by (cases X) auto
+
+lemma zero_less_real_of_extreal: "0 < real X \<longleftrightarrow> (0 < X \<and> X \<noteq> \<infinity>)"
+  by (cases X) auto
+
+lemma real_of_extreal_le_1: fixes X :: extreal shows "X \<le> 1 \<Longrightarrow> real X \<le> 1"
+  by (cases X) (auto simp: one_extreal_def)
+
+locale prob_space = measure_space +
+  assumes measure_space_1: "measure M (space M) = 1"
+
+sublocale prob_space < finite_measure
+proof
+  from measure_space_1 show "\<mu> (space M) \<noteq> \<infinity>" by simp
+qed
+
+abbreviation (in prob_space) "events \<equiv> sets M"
+abbreviation (in prob_space) "prob \<equiv> \<mu>'"
+abbreviation (in prob_space) "random_variable M' X \<equiv> sigma_algebra M' \<and> X \<in> measurable M M'"
+abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M"
+
+definition (in prob_space)
+  "indep A B \<longleftrightarrow> A \<in> events \<and> B \<in> events \<and> prob (A \<inter> B) = prob A * prob B"
+
+definition (in prob_space)
+  "indep_families F G \<longleftrightarrow> (\<forall> A \<in> F. \<forall> B \<in> G. indep A B)"
+
+definition (in prob_space)
+  "distribution X A = \<mu>' (X -` A \<inter> space M)"
+
+abbreviation (in prob_space)
+  "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
+
+declare (in finite_measure) positive_measure'[intro, simp]
+
+lemma (in prob_space) distribution_cong:
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
+  shows "distribution X = distribution Y"
+  unfolding distribution_def fun_eq_iff
+  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
+
+lemma (in prob_space) joint_distribution_cong:
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
+  shows "joint_distribution X Y = joint_distribution X' Y'"
+  unfolding distribution_def fun_eq_iff
+  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
+
+lemma (in prob_space) distribution_id[simp]:
+  "N \<in> events \<Longrightarrow> distribution (\<lambda>x. x) N = prob N"
+  by (auto simp: distribution_def intro!: arg_cong[where f=prob])
+
+lemma (in prob_space) prob_space: "prob (space M) = 1"
+  using measure_space_1 unfolding \<mu>'_def by (simp add: one_extreal_def)
+
+lemma (in prob_space) prob_le_1[simp, intro]: "prob A \<le> 1"
+  using bounded_measure[of A] by (simp add: prob_space)
+
+lemma (in prob_space) distribution_positive[simp, intro]:
+  "0 \<le> distribution X A" unfolding distribution_def by auto
+
+lemma (in prob_space) joint_distribution_remove[simp]:
+    "joint_distribution X X {(x, x)} = distribution X {x}"
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
+
+lemma (in prob_space) distribution_1:
+  "distribution X A \<le> 1"
+  unfolding distribution_def by simp
+
+lemma (in prob_space) prob_compl:
+  assumes A: "A \<in> events"
+  shows "prob (space M - A) = 1 - prob A"
+  using finite_measure_compl[OF A] by (simp add: prob_space)
+
+lemma (in prob_space) indep_space: "s \<in> events \<Longrightarrow> indep (space M) s"
+  by (simp add: indep_def prob_space)
+
+lemma (in prob_space) prob_space_increasing: "increasing M prob"
+  by (auto intro!: finite_measure_mono simp: increasing_def)
+
+lemma (in prob_space) prob_zero_union:
+  assumes "s \<in> events" "t \<in> events" "prob t = 0"
+  shows "prob (s \<union> t) = prob s"
+using assms
+proof -
+  have "prob (s \<union> t) \<le> prob s"
+    using finite_measure_subadditive[of s t] assms by auto
+  moreover have "prob (s \<union> t) \<ge> prob s"
+    using assms by (blast intro: finite_measure_mono)
+  ultimately show ?thesis by simp
+qed
+
+lemma (in prob_space) prob_eq_compl:
+  assumes "s \<in> events" "t \<in> events"
+  assumes "prob (space M - s) = prob (space M - t)"
+  shows "prob s = prob t"
+  using assms prob_compl by auto
+
+lemma (in prob_space) prob_one_inter:
+  assumes events:"s \<in> events" "t \<in> events"
+  assumes "prob t = 1"
+  shows "prob (s \<inter> t) = prob s"
+proof -
+  have "prob ((space M - s) \<union> (space M - t)) = prob (space M - s)"
+    using events assms  prob_compl[of "t"] by (auto intro!: prob_zero_union)
+  also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
+    by blast
+  finally show "prob (s \<inter> t) = prob s"
+    using events by (auto intro!: prob_eq_compl[of "s \<inter> t" s])
+qed
+
+lemma (in prob_space) prob_eq_bigunion_image:
+  assumes "range f \<subseteq> events" "range g \<subseteq> events"
+  assumes "disjoint_family f" "disjoint_family g"
+  assumes "\<And> n :: nat. prob (f n) = prob (g n)"
+  shows "(prob (\<Union> i. f i) = prob (\<Union> i. g i))"
+using assms
+proof -
+  have a: "(\<lambda> i. prob (f i)) sums (prob (\<Union> i. f i))"
+    by (rule finite_measure_UNION[OF assms(1,3)])
+  have b: "(\<lambda> i. prob (g i)) sums (prob (\<Union> i. g i))"
+    by (rule finite_measure_UNION[OF assms(2,4)])
+  show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
+qed
+
+lemma (in prob_space) prob_countably_zero:
+  assumes "range c \<subseteq> events"
+  assumes "\<And> i. prob (c i) = 0"
+  shows "prob (\<Union> i :: nat. c i) = 0"
+proof (rule antisym)
+  show "prob (\<Union> i :: nat. c i) \<le> 0"
+    using finite_measure_countably_subadditive[OF assms(1)]
+    by (simp add: assms(2) suminf_zero summable_zero)
+qed simp
+
+lemma (in prob_space) indep_sym:
+   "indep a b \<Longrightarrow> indep b a"
+unfolding indep_def using Int_commute[of a b] by auto
+
+lemma (in prob_space) indep_refl:
+  assumes "a \<in> events"
+  shows "indep a a = (prob a = 0) \<or> (prob a = 1)"
+using assms unfolding indep_def by auto
+
+lemma (in prob_space) prob_equiprobable_finite_unions:
+  assumes "s \<in> events"
+  assumes s_finite: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> events"
+  assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> (prob {x} = prob {y})"
+  shows "prob s = real (card s) * prob {SOME x. x \<in> s}"
+proof (cases "s = {}")
+  case False hence "\<exists> x. x \<in> s" by blast
+  from someI_ex[OF this] assms
+  have prob_some: "\<And> x. x \<in> s \<Longrightarrow> prob {x} = prob {SOME y. y \<in> s}" by blast
+  have "prob s = (\<Sum> x \<in> s. prob {x})"
+    using finite_measure_finite_singleton[OF s_finite] by simp
+  also have "\<dots> = (\<Sum> x \<in> s. prob {SOME y. y \<in> s})" using prob_some by auto
+  also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
+    using setsum_constant assms by (simp add: real_eq_of_nat)
+  finally show ?thesis by simp
+qed simp
+
+lemma (in prob_space) prob_real_sum_image_fn:
+  assumes "e \<in> events"
+  assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events"
+  assumes "finite s"
+  assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
+  assumes upper: "space M \<subseteq> (\<Union> i \<in> s. f i)"
+  shows "prob e = (\<Sum> x \<in> s. prob (e \<inter> f x))"
+proof -
+  have e: "e = (\<Union> i \<in> s. e \<inter> f i)"
+    using `e \<in> events` sets_into_space upper by blast
+  hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp
+  also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))"
+  proof (rule finite_measure_finite_Union)
+    show "finite s" by fact
+    show "\<And>i. i \<in> s \<Longrightarrow> e \<inter> f i \<in> events" by fact
+    show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
+      using disjoint by (auto simp: disjoint_family_on_def)
+  qed
+  finally show ?thesis .
+qed
+
+lemma (in prob_space) prob_space_vimage:
+  assumes S: "sigma_algebra S"
+  assumes T: "T \<in> measure_preserving M S"
+  shows "prob_space S"
+proof -
+  interpret S: measure_space S
+    using S and T by (rule measure_space_vimage)
+  show ?thesis
+  proof
+    from T[THEN measure_preservingD2]
+    have "T -` space S \<inter> space M = space M"
+      by (auto simp: measurable_def)
+    with T[THEN measure_preservingD, of "space S", symmetric]
+    show  "measure S (space S) = 1"
+      using measure_space_1 by simp
+  qed
+qed
+
+lemma (in prob_space) distribution_prob_space:
+  assumes X: "random_variable S X"
+  shows "prob_space (S\<lparr>measure := extreal \<circ> distribution X\<rparr>)" (is "prob_space ?S")
+proof (rule prob_space_vimage)
+  show "X \<in> measure_preserving M ?S"
+    using X
+    unfolding measure_preserving_def distribution_def_raw
+    by (auto simp: finite_measure_eq measurable_sets)
+  show "sigma_algebra ?S" using X by simp
+qed
+
+lemma (in prob_space) AE_distribution:
+  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := extreal \<circ> distribution X\<rparr>. Q x"
+  shows "AE x. Q (X x)"
+proof -
+  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space)
+  obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
+    using assms unfolding X.almost_everywhere_def by auto
+  from X[unfolded measurable_def] N show "AE x. Q (X x)"
+    by (intro AE_I'[where N="X -` N \<inter> space M"])
+       (auto simp: finite_measure_eq distribution_def measurable_sets)
+qed
+
+lemma (in prob_space) distribution_eq_integral:
+  "random_variable S X \<Longrightarrow> A \<in> sets S \<Longrightarrow> distribution X A = expectation (indicator (X -` A \<inter> space M))"
+  using finite_measure_eq[of "X -` A \<inter> space M"]
+  by (auto simp: measurable_sets distribution_def)
+
+lemma (in prob_space) distribution_eq_translated_integral:
+  assumes "random_variable S X" "A \<in> sets S"
+  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := extreal \<circ> distribution X\<rparr>) (indicator A)"
+proof -
+  interpret S: prob_space "S\<lparr>measure := extreal \<circ> distribution X\<rparr>"
+    using assms(1) by (rule distribution_prob_space)
+  show ?thesis
+    using S.positive_integral_indicator(1)[of A] assms by simp
+qed
+
+lemma (in prob_space) finite_expectation1:
+  assumes f: "finite (X`space M)" and rv: "random_variable borel X"
+  shows "expectation X = (\<Sum>r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))" (is "_ = ?r")
+proof (subst integral_on_finite)
+  show "X \<in> borel_measurable M" "finite (X`space M)" using assms by auto
+  show "(\<Sum> r \<in> X ` space M. r * real (\<mu> (X -` {r} \<inter> space M))) = ?r"
+    "\<And>x. \<mu> (X -` {x} \<inter> space M) \<noteq> \<infinity>"
+    using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto
+qed
+
+lemma (in prob_space) finite_expectation:
+  assumes "finite (X`space M)" "random_variable borel X"
+  shows "expectation X = (\<Sum> r \<in> X ` (space M). r * distribution X {r})"
+  using assms unfolding distribution_def using finite_expectation1 by auto
+
+lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0:
+  assumes "{x} \<in> events"
+  assumes "prob {x} = 1"
+  assumes "{y} \<in> events"
+  assumes "y \<noteq> x"
+  shows "prob {y} = 0"
+  using prob_one_inter[of "{y}" "{x}"] assms by auto
+
+lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0"
+  unfolding distribution_def by simp
+
+lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1"
+proof -
+  have "X -` X ` space M \<inter> space M = space M" by auto
+  thus ?thesis unfolding distribution_def by (simp add: prob_space)
+qed
+
+lemma (in prob_space) distribution_one:
+  assumes "random_variable M' X" and "A \<in> sets M'"
+  shows "distribution X A \<le> 1"
+proof -
+  have "distribution X A \<le> \<mu>' (space M)" unfolding distribution_def
+    using assms[unfolded measurable_def] by (auto intro!: finite_measure_mono)
+  thus ?thesis by (simp add: prob_space)
+qed
+
+lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0:
+  assumes X: "random_variable \<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr> X"
+    (is "random_variable ?S X")
+  assumes "distribution X {x} = 1"
+  assumes "y \<noteq> x"
+  shows "distribution X {y} = 0"
+proof cases
+  { fix x have "X -` {x} \<inter> space M \<in> sets M"
+    proof cases
+      assume "x \<in> X`space M" with X show ?thesis
+        by (auto simp: measurable_def image_iff)
+    next
+      assume "x \<notin> X`space M" then have "X -` {x} \<inter> space M = {}" by auto
+      then show ?thesis by auto
+    qed } note single = this
+  have "X -` {x} \<inter> space M - X -` {y} \<inter> space M = X -` {x} \<inter> space M"
+    "X -` {y} \<inter> space M \<inter> (X -` {x} \<inter> space M) = {}"
+    using `y \<noteq> x` by auto
+  with finite_measure_inter_full_set[OF single single, of x y] assms(2)
+  show ?thesis by (auto simp: distribution_def prob_space)
+next
+  assume "{y} \<notin> sets ?S"
+  then have "X -` {y} \<inter> space M = {}" by auto
+  thus "distribution X {y} = 0" unfolding distribution_def by auto
+qed
+
+lemma (in prob_space) joint_distribution_Times_le_fst:
+  assumes X: "random_variable MX X" and Y: "random_variable MY Y"
+    and A: "A \<in> sets MX" and B: "B \<in> sets MY"
+  shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
+  unfolding distribution_def
+proof (intro finite_measure_mono)
+  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
+  show "X -` A \<inter> space M \<in> events"
+    using X A unfolding measurable_def by simp
+  have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
+    (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
+qed
+
+lemma (in prob_space) joint_distribution_commute:
+  "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)"
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
+
+lemma (in prob_space) joint_distribution_Times_le_snd:
+  assumes X: "random_variable MX X" and Y: "random_variable MY Y"
+    and A: "A \<in> sets MX" and B: "B \<in> sets MY"
+  shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B"
+  using assms
+  by (subst joint_distribution_commute)
+     (simp add: swap_product joint_distribution_Times_le_fst)
+
+lemma (in prob_space) random_variable_pairI:
+  assumes "random_variable MX X"
+  assumes "random_variable MY Y"
+  shows "random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))"
+proof
+  interpret MX: sigma_algebra MX using assms by simp
+  interpret MY: sigma_algebra MY using assms by simp
+  interpret P: pair_sigma_algebra MX MY by default
+  show "sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
+  have sa: "sigma_algebra M" by default
+  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
+    unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
+qed
+
+lemma (in prob_space) joint_distribution_commute_singleton:
+  "joint_distribution X Y {(x, y)} = joint_distribution Y X {(y, x)}"
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
+
+lemma (in prob_space) joint_distribution_assoc_singleton:
+  "joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} =
+   joint_distribution (\<lambda>x. (X x, Y x)) Z {((x, y), z)}"
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
+
+locale pair_prob_space = M1: prob_space M1 + M2: prob_space M2 for M1 M2
+
+sublocale pair_prob_space \<subseteq> pair_sigma_finite M1 M2 by default
+
+sublocale pair_prob_space \<subseteq> P: prob_space P
+by default (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure)
+
+lemma countably_additiveI[case_names countably]:
+  assumes "\<And>A. \<lbrakk> range A \<subseteq> sets M ; disjoint_family A ; (\<Union>i. A i) \<in> sets M\<rbrakk> \<Longrightarrow>
+    (\<Sum>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
+  shows "countably_additive M \<mu>"
+  using assms unfolding countably_additive_def by auto
+
+lemma (in prob_space) joint_distribution_prob_space:
+  assumes "random_variable MX X" "random_variable MY Y"
+  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
+  using random_variable_pairI[OF assms] by (rule distribution_prob_space)
+
+section "Probability spaces on finite sets"
+
+locale finite_prob_space = prob_space + finite_measure_space
+
+abbreviation (in prob_space) "finite_random_variable M' X \<equiv> finite_sigma_algebra M' \<and> X \<in> measurable M M'"
+
+lemma (in prob_space) finite_random_variableD:
+  assumes "finite_random_variable M' X" shows "random_variable M' X"
+proof -
+  interpret M': finite_sigma_algebra M' using assms by simp
+  then show "random_variable M' X" using assms by simp default
+qed
+
+lemma (in prob_space) distribution_finite_prob_space:
+  assumes "finite_random_variable MX X"
+  shows "finite_prob_space (MX\<lparr>measure := extreal \<circ> distribution X\<rparr>)"
+proof -
+  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>"
+    using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
+  interpret MX: finite_sigma_algebra MX
+    using assms by auto
+  show ?thesis by default (simp_all add: MX.finite_space)
+qed
+
+lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]:
+  assumes "simple_function M X"
+  shows "finite_random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = x \<rparr> X"
+    (is "finite_random_variable ?X _")
+proof (intro conjI)
+  have [simp]: "finite (X ` space M)" using assms unfolding simple_function_def by simp
+  interpret X: sigma_algebra ?X by (rule sigma_algebra_Pow)
+  show "finite_sigma_algebra ?X"
+    by default auto
+  show "X \<in> measurable M ?X"
+  proof (unfold measurable_def, clarsimp)
+    fix A assume A: "A \<subseteq> X`space M"
+    then have "finite A" by (rule finite_subset) simp
+    then have "X -` (\<Union>a\<in>A. {a}) \<inter> space M \<in> events"
+      unfolding vimage_UN UN_extend_simps
+      apply (rule finite_UN)
+      using A assms unfolding simple_function_def by auto
+    then show "X -` A \<inter> space M \<in> events" by simp
+  qed
+qed
+
+lemma (in prob_space) simple_function_imp_random_variable[simp, intro]:
+  assumes "simple_function M X"
+  shows "random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = ext \<rparr> X"
+  using simple_function_imp_finite_random_variable[OF assms, of ext]
+  by (auto dest!: finite_random_variableD)
+
+lemma (in prob_space) sum_over_space_real_distribution:
+  "simple_function M X \<Longrightarrow> (\<Sum>x\<in>X`space M. distribution X {x}) = 1"
+  unfolding distribution_def prob_space[symmetric]
+  by (subst finite_measure_finite_Union[symmetric])
+     (auto simp add: disjoint_family_on_def simple_function_def
+           intro!: arg_cong[where f=prob])
+
+lemma (in prob_space) finite_random_variable_pairI:
+  assumes "finite_random_variable MX X"
+  assumes "finite_random_variable MY Y"
+  shows "finite_random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))"
+proof
+  interpret MX: finite_sigma_algebra MX using assms by simp
+  interpret MY: finite_sigma_algebra MY using assms by simp
+  interpret P: pair_finite_sigma_algebra MX MY by default
+  show "finite_sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
+  have sa: "sigma_algebra M" by default
+  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
+    unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
+qed
+
+lemma (in prob_space) finite_random_variable_imp_sets:
+  "finite_random_variable MX X \<Longrightarrow> x \<in> space MX \<Longrightarrow> {x} \<in> sets MX"
+  unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp
+
+lemma (in prob_space) finite_random_variable_measurable:
+  assumes X: "finite_random_variable MX X" shows "X -` A \<inter> space M \<in> events"
+proof -
+  interpret X: finite_sigma_algebra MX using X by simp
+  from X have vimage: "\<And>A. A \<subseteq> space MX \<Longrightarrow> X -` A \<inter> space M \<in> events" and
+    "X \<in> space M \<rightarrow> space MX"
+    by (auto simp: measurable_def)
+  then have *: "X -` A \<inter> space M = X -` (A \<inter> space MX) \<inter> space M"
+    by auto
+  show "X -` A \<inter> space M \<in> events"
+    unfolding * by (intro vimage) auto
+qed
+
+lemma (in prob_space) joint_distribution_finite_Times_le_fst:
+  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
+  shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
+  unfolding distribution_def
+proof (intro finite_measure_mono)
+  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
+  show "X -` A \<inter> space M \<in> events"
+    using finite_random_variable_measurable[OF X] .
+  have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
+    (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
+qed
+
+lemma (in prob_space) joint_distribution_finite_Times_le_snd:
+  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
+  shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B"
+  using assms
+  by (subst joint_distribution_commute)
+     (simp add: swap_product joint_distribution_finite_Times_le_fst)
+
+lemma (in prob_space) finite_distribution_order:
+  fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
+  assumes "finite_random_variable MX X" "finite_random_variable MY Y"
+  shows "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
+    and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
+    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
+    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
+    and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
+    and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
+  using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"]
+  using joint_distribution_finite_Times_le_fst[OF assms, of "{x}" "{y}"]
+  by (auto intro: antisym)
+
+lemma (in prob_space) setsum_joint_distribution:
+  assumes X: "finite_random_variable MX X"
+  assumes Y: "random_variable MY Y" "B \<in> sets MY"
+  shows "(\<Sum>a\<in>space MX. joint_distribution X Y ({a} \<times> B)) = distribution Y B"
+  unfolding distribution_def
+proof (subst finite_measure_finite_Union[symmetric])
+  interpret MX: finite_sigma_algebra MX using X by auto
+  show "finite (space MX)" using MX.finite_space .
+  let "?d i" = "(\<lambda>x. (X x, Y x)) -` ({i} \<times> B) \<inter> space M"
+  { fix i assume "i \<in> space MX"
+    moreover have "?d i = (X -` {i} \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
+    ultimately show "?d i \<in> events"
+      using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y
+      using MX.sets_eq_Pow by auto }
+  show "disjoint_family_on ?d (space MX)" by (auto simp: disjoint_family_on_def)
+  show "\<mu>' (\<Union>i\<in>space MX. ?d i) = \<mu>' (Y -` B \<inter> space M)"
+    using X[unfolded measurable_def] by (auto intro!: arg_cong[where f=\<mu>'])
+qed
+
+lemma (in prob_space) setsum_joint_distribution_singleton:
+  assumes X: "finite_random_variable MX X"
+  assumes Y: "finite_random_variable MY Y" "b \<in> space MY"
+  shows "(\<Sum>a\<in>space MX. joint_distribution X Y {(a, b)}) = distribution Y {b}"
+  using setsum_joint_distribution[OF X
+    finite_random_variableD[OF Y(1)]
+    finite_random_variable_imp_sets[OF Y]] by simp
+
+locale pair_finite_prob_space = M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2
+
+sublocale pair_finite_prob_space \<subseteq> pair_prob_space M1 M2 by default
+sublocale pair_finite_prob_space \<subseteq> pair_finite_space M1 M2  by default
+sublocale pair_finite_prob_space \<subseteq> finite_prob_space P by default
+
+lemma (in prob_space) joint_distribution_finite_prob_space:
+  assumes X: "finite_random_variable MX X"
+  assumes Y: "finite_random_variable MY Y"
+  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
+  by (intro distribution_finite_prob_space finite_random_variable_pairI X Y)
+
+lemma finite_prob_space_eq:
+  "finite_prob_space M \<longleftrightarrow> finite_measure_space M \<and> measure M (space M) = 1"
+  unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def
+  by auto
+
+lemma (in prob_space) not_empty: "space M \<noteq> {}"
+  using prob_space empty_measure' by auto
+
+lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1"
+  using measure_space_1 sum_over_space by simp
+
+lemma (in finite_prob_space) joint_distribution_restriction_fst:
+  "joint_distribution X Y A \<le> distribution X (fst ` A)"
+  unfolding distribution_def
+proof (safe intro!: finite_measure_mono)
+  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
+  show "x \<in> X -` fst ` A"
+    by (auto intro!: image_eqI[OF _ *])
+qed (simp_all add: sets_eq_Pow)
+
+lemma (in finite_prob_space) joint_distribution_restriction_snd:
+  "joint_distribution X Y A \<le> distribution Y (snd ` A)"
+  unfolding distribution_def
+proof (safe intro!: finite_measure_mono)
+  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
+  show "x \<in> Y -` snd ` A"
+    by (auto intro!: image_eqI[OF _ *])
+qed (simp_all add: sets_eq_Pow)
+
+lemma (in finite_prob_space) distribution_order:
+  shows "0 \<le> distribution X x'"
+  and "(distribution X x' \<noteq> 0) \<longleftrightarrow> (0 < distribution X x')"
+  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
+  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
+  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
+  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
+  and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
+  and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
+  using
+    joint_distribution_restriction_fst[of X Y "{(x, y)}"]
+    joint_distribution_restriction_snd[of X Y "{(x, y)}"]
+  by (auto intro: antisym)
+
+lemma (in finite_prob_space) distribution_mono:
+  assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
+  shows "distribution X x \<le> distribution Y y"
+  unfolding distribution_def
+  using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono)
+
+lemma (in finite_prob_space) distribution_mono_gt_0:
+  assumes gt_0: "0 < distribution X x"
+  assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
+  shows "0 < distribution Y y"
+  by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
+
+lemma (in finite_prob_space) sum_over_space_distrib:
+  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
+  unfolding distribution_def prob_space[symmetric] using finite_space
+  by (subst finite_measure_finite_Union[symmetric])
+     (auto simp add: disjoint_family_on_def sets_eq_Pow
+           intro!: arg_cong[where f=\<mu>'])
+
+lemma (in finite_prob_space) sum_over_space_real_distribution:
+  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
+  unfolding distribution_def prob_space[symmetric] using finite_space
+  by (subst finite_measure_finite_Union[symmetric])
+     (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob])
+
+lemma (in finite_prob_space) finite_sum_over_space_eq_1:
+  "(\<Sum>x\<in>space M. prob {x}) = 1"
+  using prob_space finite_space
+  by (subst (asm) finite_measure_finite_singleton) auto
+
+lemma (in prob_space) distribution_remove_const:
+  shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
+  and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
+  and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
+  and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
+  and "distribution (\<lambda>x. ()) {()} = 1"
+  by (auto intro!: arg_cong[where f=\<mu>'] simp: distribution_def prob_space[symmetric])
+
+lemma (in finite_prob_space) setsum_distribution_gen:
+  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
+  and "inj_on f (X`space M)"
+  shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}"
+  unfolding distribution_def assms
+  using finite_space assms
+  by (subst finite_measure_finite_Union[symmetric])
+     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
+      intro!: arg_cong[where f=prob])
+
+lemma (in finite_prob_space) setsum_distribution:
+  "(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}"
+  "(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}"
+  "(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}"
+  "(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}"
+  "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
+  by (auto intro!: inj_onI setsum_distribution_gen)
+
+lemma (in finite_prob_space) uniform_prob:
+  assumes "x \<in> space M"
+  assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}"
+  shows "prob {x} = 1 / card (space M)"
+proof -
+  have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}"
+    using assms(2)[OF _ `x \<in> space M`] by blast
+  have "1 = prob (space M)"
+    using prob_space by auto
+  also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
+    using finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
+      sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
+      finite_space unfolding disjoint_family_on_def  prob_space[symmetric]
+    by (auto simp add:setsum_restrict_set)
+  also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
+    using prob_x by auto
+  also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
+  finally have one: "1 = real (card (space M)) * prob {x}"
+    using real_eq_of_nat by auto
+  hence two: "real (card (space M)) \<noteq> 0" by fastsimp 
+  from one have three: "prob {x} \<noteq> 0" by fastsimp
+  thus ?thesis using one two three divide_cancel_right
+    by (auto simp:field_simps)
+qed
+
+lemma (in prob_space) prob_space_subalgebra:
+  assumes "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M"
+    and "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
+  shows "prob_space N"
+proof -
+  interpret N: measure_space N
+    by (rule measure_space_subalgebra[OF assms])
+  show ?thesis
+  proof qed (insert assms(4)[OF N.top], simp add: assms measure_space_1)
+qed
+
+lemma (in prob_space) prob_space_of_restricted_space:
+  assumes "\<mu> A \<noteq> 0" "A \<in> sets M"
+  shows "prob_space (restricted_space A \<lparr>measure := \<lambda>S. \<mu> S / \<mu> A\<rparr>)"
+    (is "prob_space ?P")
+proof -
+  interpret A: measure_space "restricted_space A"
+    using `A \<in> sets M` by (rule restricted_measure_space)
+  interpret A': sigma_algebra ?P
+    by (rule A.sigma_algebra_cong) auto
+  show "prob_space ?P"
+  proof
+    show "measure ?P (space ?P) = 1"
+      using real_measure[OF `A \<in> events`] `\<mu> A \<noteq> 0` by auto
+    show "positive ?P (measure ?P)"
+    proof (simp add: positive_def, safe)
+      show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_extreal_def)
+      fix B assume "B \<in> events"
+      with real_measure[of "A \<inter> B"] real_measure[OF `A \<in> events`] `A \<in> sets M`
+      show "0 \<le> \<mu> (A \<inter> B) / \<mu> A" by (auto simp: Int)
+    qed
+    show "countably_additive ?P (measure ?P)"
+    proof (simp add: countably_additive_def, safe)
+      fix B and F :: "nat \<Rightarrow> 'a set"
+      assume F: "range F \<subseteq> op \<inter> A ` events" "disjoint_family F"
+      { fix i
+        from F have "F i \<in> op \<inter> A ` events" by auto
+        with `A \<in> events` have "F i \<in> events" by auto }
+      moreover then have "range F \<subseteq> events" by auto
+      moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S"
+        by (simp add: mult_commute divide_extreal_def)
+      moreover have "0 \<le> inverse (\<mu> A)"
+        using real_measure[OF `A \<in> events`] by auto
+      ultimately show "(\<Sum>i. \<mu> (F i) / \<mu> A) = \<mu> (\<Union>i. F i) / \<mu> A"
+        using measure_countably_additive[of F] F
+        by (auto simp: suminf_cmult_extreal)
+    qed
+  qed
+qed
+
+lemma finite_prob_spaceI:
+  assumes "finite (space M)" "sets M = Pow(space M)"
+    and "measure M (space M) = 1" "measure M {} = 0" "\<And>A. A \<subseteq> space M \<Longrightarrow> 0 \<le> measure M A"
+    and "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M B"
+  shows "finite_prob_space M"
+  unfolding finite_prob_space_eq
+proof
+  show "finite_measure_space M" using assms
+    by (auto intro!: finite_measure_spaceI)
+  show "measure M (space M) = 1" by fact
+qed
+
+lemma (in finite_prob_space) finite_measure_space:
+  fixes X :: "'a \<Rightarrow> 'x"
+  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X\<rparr>"
+    (is "finite_measure_space ?S")
+proof (rule finite_measure_spaceI, simp_all)
+  show "finite (X ` space M)" using finite_space by simp
+next
+  fix A B :: "'x set" assume "A \<inter> B = {}"
+  then show "distribution X (A \<union> B) = distribution X A + distribution X B"
+    unfolding distribution_def
+    by (subst finite_measure_Union[symmetric])
+       (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow)
+qed
+
+lemma (in finite_prob_space) finite_prob_space_of_images:
+  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X \<rparr>"
+  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_extreal_def)
+
+lemma (in finite_prob_space) finite_product_measure_space:
+  fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
+  assumes "finite s1" "finite s2"
+  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = extreal \<circ> joint_distribution X Y\<rparr>"
+    (is "finite_measure_space ?M")
+proof (rule finite_measure_spaceI, simp_all)
+  show "finite (s1 \<times> s2)"
+    using assms by auto
+next
+  fix A B :: "('x*'y) set" assume "A \<inter> B = {}"
+  then show "joint_distribution X Y (A \<union> B) = joint_distribution X Y A + joint_distribution X Y B"
+    unfolding distribution_def
+    by (subst finite_measure_Union[symmetric])
+       (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow)
+qed
+
+lemma (in finite_prob_space) finite_product_measure_space_of_images:
+  shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
+                                sets = Pow (X ` space M \<times> Y ` space M),
+                                measure = extreal \<circ> joint_distribution X Y \<rparr>"
+  using finite_space by (auto intro!: finite_product_measure_space)
+
+lemma (in finite_prob_space) finite_product_prob_space_of_images:
+  "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M),
+                       measure = extreal \<circ> joint_distribution X Y \<rparr>"
+  (is "finite_prob_space ?S")
+proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_extreal_def)
+  have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
+  thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
+    by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
+qed
+
+section "Conditional Expectation and Probability"
+
+lemma (in prob_space) conditional_expectation_exists:
+  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
+  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
+  and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
+  shows "\<exists>Y\<in>borel_measurable N. (\<forall>x. 0 \<le> Y x) \<and> (\<forall>C\<in>sets N.
+      (\<integral>\<^isup>+x. Y x * indicator C x \<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x \<partial>M))"
+proof -
+  note N(4)[simp]
+  interpret P: prob_space N
+    using prob_space_subalgebra[OF N] .
+
+  let "?f A" = "\<lambda>x. X x * indicator A x"
+  let "?Q A" = "integral\<^isup>P M (?f A)"
+
+  from measure_space_density[OF borel]
+  have Q: "measure_space (N\<lparr> measure := ?Q \<rparr>)"
+    apply (rule measure_space.measure_space_subalgebra[of "M\<lparr> measure := ?Q \<rparr>"])
+    using N by (auto intro!: P.sigma_algebra_cong)
+  then interpret Q: measure_space "N\<lparr> measure := ?Q \<rparr>" .
+
+  have "P.absolutely_continuous ?Q"
+    unfolding P.absolutely_continuous_def
+  proof safe
+    fix A assume "A \<in> sets N" "P.\<mu> A = 0"
+    then have f_borel: "?f A \<in> borel_measurable M" "AE x. x \<notin> A"
+      using borel N by (auto intro!: borel_measurable_indicator AE_not_in)
+    then show "?Q A = 0"
+      by (auto simp add: positive_integral_0_iff_AE)
+  qed
+  from P.Radon_Nikodym[OF Q this]
+  obtain Y where Y: "Y \<in> borel_measurable N" "\<And>x. 0 \<le> Y x"
+    "\<And>A. A \<in> sets N \<Longrightarrow> ?Q A =(\<integral>\<^isup>+x. Y x * indicator A x \<partial>N)"
+    by blast
+  with N(2) show ?thesis
+    by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ _ N(2,3,4,1)])
+qed
+
+definition (in prob_space)
+  "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable N \<and> (\<forall>x. 0 \<le> Y x)
+    \<and> (\<forall>C\<in>sets N. (\<integral>\<^isup>+x. Y x * indicator C x\<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x\<partial>M)))"
+
+abbreviation (in prob_space)
+  "conditional_prob N A \<equiv> conditional_expectation N (indicator A)"
+
+lemma (in prob_space)
+  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
+  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
+  and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
+  shows borel_measurable_conditional_expectation:
+    "conditional_expectation N X \<in> borel_measurable N"
+  and conditional_expectation: "\<And>C. C \<in> sets N \<Longrightarrow>
+      (\<integral>\<^isup>+x. conditional_expectation N X x * indicator C x \<partial>M) =
+      (\<integral>\<^isup>+x. X x * indicator C x \<partial>M)"
+   (is "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C")
+proof -
+  note CE = conditional_expectation_exists[OF assms, unfolded Bex_def]
+  then show "conditional_expectation N X \<in> borel_measurable N"
+    unfolding conditional_expectation_def by (rule someI2_ex) blast
+
+  from CE show "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C"
+    unfolding conditional_expectation_def by (rule someI2_ex) blast
+qed
+
+lemma (in sigma_algebra) factorize_measurable_function_pos:
+  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
+  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
+  assumes Z: "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)"
+  shows "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. max 0 (Z x) = g (Y x)"
+proof -
+  interpret M': sigma_algebra M' by fact
+  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
+  from M'.sigma_algebra_vimage[OF this]
+  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
+
+  from va.borel_measurable_implies_simple_function_sequence'[OF Z] guess f . note f = this
+
+  have "\<forall>i. \<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
+  proof
+    fix i
+    from f(1)[of i] have "finite (f i`space M)" and B_ex:
+      "\<forall>z\<in>(f i)`space M. \<exists>B. B \<in> sets M' \<and> (f i) -` {z} \<inter> space M = Y -` B \<inter> space M"
+      unfolding simple_function_def by auto
+    from B_ex[THEN bchoice] guess B .. note B = this
+
+    let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x"
+
+    show "\<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
+    proof (intro exI[of _ ?g] conjI ballI)
+      show "simple_function M' ?g" using B by auto
+
+      fix x assume "x \<in> space M"
+      then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::extreal)"
+        unfolding indicator_def using B by auto
+      then show "f i x = ?g (Y x)" using `x \<in> space M` f(1)[of i]
+        by (subst va.simple_function_indicator_representation) auto
+    qed
+  qed
+  from choice[OF this] guess g .. note g = this
+
+  show ?thesis
+  proof (intro ballI bexI)
+    show "(\<lambda>x. SUP i. g i x) \<in> borel_measurable M'"
+      using g by (auto intro: M'.borel_measurable_simple_function)
+    fix x assume "x \<in> space M"
+    have "max 0 (Z x) = (SUP i. f i x)" using f by simp
+    also have "\<dots> = (SUP i. g i (Y x))"
+      using g `x \<in> space M` by simp
+    finally show "max 0 (Z x) = (SUP i. g i (Y x))" .
+  qed
+qed
+
+lemma extreal_0_le_iff_le_0[simp]:
+  fixes a :: extreal shows "0 \<le> -a \<longleftrightarrow> a \<le> 0"
+  by (cases rule: extreal2_cases[of a]) auto
+
+lemma (in sigma_algebra) factorize_measurable_function:
+  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
+  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
+  shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
+    \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
+proof safe
+  interpret M': sigma_algebra M' by fact
+  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
+  from M'.sigma_algebra_vimage[OF this]
+  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
+
+  { fix g :: "'c \<Rightarrow> extreal" assume "g \<in> borel_measurable M'"
+    with M'.measurable_vimage_algebra[OF Y]
+    have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+      by (rule measurable_comp)
+    moreover assume "\<forall>x\<in>space M. Z x = g (Y x)"
+    then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow>
+       g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+       by (auto intro!: measurable_cong)
+    ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+      by simp }
+
+  assume Z: "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+  with assms have "(\<lambda>x. - Z x) \<in> borel_measurable M"
+    "(\<lambda>x. - Z x) \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+    by auto
+  from factorize_measurable_function_pos[OF assms(1,2) this] guess n .. note n = this
+  from factorize_measurable_function_pos[OF assms Z] guess p .. note p = this
+  let "?g x" = "p x - n x"
+  show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)"
+  proof (intro bexI ballI)
+    show "?g \<in> borel_measurable M'" using p n by auto
+    fix x assume "x \<in> space M"
+    then have "p (Y x) = max 0 (Z x)" "n (Y x) = max 0 (- Z x)"
+      using p n by auto
+    then show "Z x = ?g (Y x)"
+      by (auto split: split_max)
+  qed
+qed
+
+end
--- a/src/HOL/Probability/Probability_Space.thy	Mon Apr 04 12:39:14 2011 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,935 +0,0 @@
-(*  Title:      HOL/Probability/Probability_Space.thy
-    Author:     Johannes Hölzl, TU München
-    Author:     Armin Heller, TU München
-*)
-
-header {*Probability spaces*}
-
-theory Probability_Space
-imports Lebesgue_Integration Radon_Nikodym Product_Measure
-begin
-
-lemma real_of_extreal_inverse[simp]:
-  fixes X :: extreal
-  shows "real (inverse X) = 1 / real X"
-  by (cases X) (auto simp: inverse_eq_divide)
-
-lemma real_of_extreal_le_0[simp]: "real (X :: extreal) \<le> 0 \<longleftrightarrow> (X \<le> 0 \<or> X = \<infinity>)"
-  by (cases X) auto
-
-lemma abs_real_of_extreal[simp]: "\<bar>real (X :: extreal)\<bar> = real \<bar>X\<bar>"
-  by (cases X) auto
-
-lemma zero_less_real_of_extreal: "0 < real X \<longleftrightarrow> (0 < X \<and> X \<noteq> \<infinity>)"
-  by (cases X) auto
-
-lemma real_of_extreal_le_1: fixes X :: extreal shows "X \<le> 1 \<Longrightarrow> real X \<le> 1"
-  by (cases X) (auto simp: one_extreal_def)
-
-locale prob_space = measure_space +
-  assumes measure_space_1: "measure M (space M) = 1"
-
-sublocale prob_space < finite_measure
-proof
-  from measure_space_1 show "\<mu> (space M) \<noteq> \<infinity>" by simp
-qed
-
-abbreviation (in prob_space) "events \<equiv> sets M"
-abbreviation (in prob_space) "prob \<equiv> \<mu>'"
-abbreviation (in prob_space) "prob_preserving \<equiv> measure_preserving"
-abbreviation (in prob_space) "random_variable M' X \<equiv> sigma_algebra M' \<and> X \<in> measurable M M'"
-abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M"
-
-definition (in prob_space)
-  "indep A B \<longleftrightarrow> A \<in> events \<and> B \<in> events \<and> prob (A \<inter> B) = prob A * prob B"
-
-definition (in prob_space)
-  "indep_families F G \<longleftrightarrow> (\<forall> A \<in> F. \<forall> B \<in> G. indep A B)"
-
-definition (in prob_space)
-  "distribution X A = \<mu>' (X -` A \<inter> space M)"
-
-abbreviation (in prob_space)
-  "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
-
-declare (in finite_measure) positive_measure'[intro, simp]
-
-lemma (in prob_space) distribution_cong:
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
-  shows "distribution X = distribution Y"
-  unfolding distribution_def fun_eq_iff
-  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
-
-lemma (in prob_space) joint_distribution_cong:
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
-  shows "joint_distribution X Y = joint_distribution X' Y'"
-  unfolding distribution_def fun_eq_iff
-  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
-
-lemma (in prob_space) distribution_id[simp]:
-  "N \<in> events \<Longrightarrow> distribution (\<lambda>x. x) N = prob N"
-  by (auto simp: distribution_def intro!: arg_cong[where f=prob])
-
-lemma (in prob_space) prob_space: "prob (space M) = 1"
-  using measure_space_1 unfolding \<mu>'_def by (simp add: one_extreal_def)
-
-lemma (in prob_space) prob_le_1[simp, intro]: "prob A \<le> 1"
-  using bounded_measure[of A] by (simp add: prob_space)
-
-lemma (in prob_space) distribution_positive[simp, intro]:
-  "0 \<le> distribution X A" unfolding distribution_def by auto
-
-lemma (in prob_space) joint_distribution_remove[simp]:
-    "joint_distribution X X {(x, x)} = distribution X {x}"
-  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
-
-lemma (in prob_space) distribution_1:
-  "distribution X A \<le> 1"
-  unfolding distribution_def by simp
-
-lemma (in prob_space) prob_compl:
-  assumes A: "A \<in> events"
-  shows "prob (space M - A) = 1 - prob A"
-  using finite_measure_compl[OF A] by (simp add: prob_space)
-
-lemma (in prob_space) indep_space: "s \<in> events \<Longrightarrow> indep (space M) s"
-  by (simp add: indep_def prob_space)
-
-lemma (in prob_space) prob_space_increasing: "increasing M prob"
-  by (auto intro!: finite_measure_mono simp: increasing_def)
-
-lemma (in prob_space) prob_zero_union:
-  assumes "s \<in> events" "t \<in> events" "prob t = 0"
-  shows "prob (s \<union> t) = prob s"
-using assms
-proof -
-  have "prob (s \<union> t) \<le> prob s"
-    using finite_measure_subadditive[of s t] assms by auto
-  moreover have "prob (s \<union> t) \<ge> prob s"
-    using assms by (blast intro: finite_measure_mono)
-  ultimately show ?thesis by simp
-qed
-
-lemma (in prob_space) prob_eq_compl:
-  assumes "s \<in> events" "t \<in> events"
-  assumes "prob (space M - s) = prob (space M - t)"
-  shows "prob s = prob t"
-  using assms prob_compl by auto
-
-lemma (in prob_space) prob_one_inter:
-  assumes events:"s \<in> events" "t \<in> events"
-  assumes "prob t = 1"
-  shows "prob (s \<inter> t) = prob s"
-proof -
-  have "prob ((space M - s) \<union> (space M - t)) = prob (space M - s)"
-    using events assms  prob_compl[of "t"] by (auto intro!: prob_zero_union)
-  also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
-    by blast
-  finally show "prob (s \<inter> t) = prob s"
-    using events by (auto intro!: prob_eq_compl[of "s \<inter> t" s])
-qed
-
-lemma (in prob_space) prob_eq_bigunion_image:
-  assumes "range f \<subseteq> events" "range g \<subseteq> events"
-  assumes "disjoint_family f" "disjoint_family g"
-  assumes "\<And> n :: nat. prob (f n) = prob (g n)"
-  shows "(prob (\<Union> i. f i) = prob (\<Union> i. g i))"
-using assms
-proof -
-  have a: "(\<lambda> i. prob (f i)) sums (prob (\<Union> i. f i))"
-    by (rule finite_measure_UNION[OF assms(1,3)])
-  have b: "(\<lambda> i. prob (g i)) sums (prob (\<Union> i. g i))"
-    by (rule finite_measure_UNION[OF assms(2,4)])
-  show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
-qed
-
-lemma (in prob_space) prob_countably_zero:
-  assumes "range c \<subseteq> events"
-  assumes "\<And> i. prob (c i) = 0"
-  shows "prob (\<Union> i :: nat. c i) = 0"
-proof (rule antisym)
-  show "prob (\<Union> i :: nat. c i) \<le> 0"
-    using finite_measure_countably_subadditive[OF assms(1)]
-    by (simp add: assms(2) suminf_zero summable_zero)
-qed simp
-
-lemma (in prob_space) indep_sym:
-   "indep a b \<Longrightarrow> indep b a"
-unfolding indep_def using Int_commute[of a b] by auto
-
-lemma (in prob_space) indep_refl:
-  assumes "a \<in> events"
-  shows "indep a a = (prob a = 0) \<or> (prob a = 1)"
-using assms unfolding indep_def by auto
-
-lemma (in prob_space) prob_equiprobable_finite_unions:
-  assumes "s \<in> events"
-  assumes s_finite: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> events"
-  assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> (prob {x} = prob {y})"
-  shows "prob s = real (card s) * prob {SOME x. x \<in> s}"
-proof (cases "s = {}")
-  case False hence "\<exists> x. x \<in> s" by blast
-  from someI_ex[OF this] assms
-  have prob_some: "\<And> x. x \<in> s \<Longrightarrow> prob {x} = prob {SOME y. y \<in> s}" by blast
-  have "prob s = (\<Sum> x \<in> s. prob {x})"
-    using finite_measure_finite_singleton[OF s_finite] by simp
-  also have "\<dots> = (\<Sum> x \<in> s. prob {SOME y. y \<in> s})" using prob_some by auto
-  also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
-    using setsum_constant assms by (simp add: real_eq_of_nat)
-  finally show ?thesis by simp
-qed simp
-
-lemma (in prob_space) prob_real_sum_image_fn:
-  assumes "e \<in> events"
-  assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events"
-  assumes "finite s"
-  assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
-  assumes upper: "space M \<subseteq> (\<Union> i \<in> s. f i)"
-  shows "prob e = (\<Sum> x \<in> s. prob (e \<inter> f x))"
-proof -
-  have e: "e = (\<Union> i \<in> s. e \<inter> f i)"
-    using `e \<in> events` sets_into_space upper by blast
-  hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp
-  also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))"
-  proof (rule finite_measure_finite_Union)
-    show "finite s" by fact
-    show "\<And>i. i \<in> s \<Longrightarrow> e \<inter> f i \<in> events" by fact
-    show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
-      using disjoint by (auto simp: disjoint_family_on_def)
-  qed
-  finally show ?thesis .
-qed
-
-lemma (in prob_space) distribution_prob_space:
-  assumes "random_variable S X"
-  shows "prob_space (S\<lparr>measure := extreal \<circ> distribution X\<rparr>)"
-proof -
-  interpret S: measure_space "S\<lparr>measure := extreal \<circ> distribution X\<rparr>"
-  proof (rule measure_space.measure_space_cong)
-    show "measure_space (S\<lparr> measure := \<lambda>A. \<mu> (X -` A \<inter> space M) \<rparr>)"
-      using assms by (auto intro!: measure_space_vimage simp: measure_preserving_def)
-  qed (insert assms, auto simp add: finite_measure_eq distribution_def measurable_sets)
-  show ?thesis
-  proof (default, simp)
-    have "X -` space S \<inter> space M = space M"
-      using `random_variable S X` by (auto simp: measurable_def)
-    then show "extreal (distribution X (space S)) = 1"
-      by (simp add: distribution_def one_extreal_def prob_space)
-  qed
-qed
-
-lemma (in prob_space) AE_distribution:
-  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := extreal \<circ> distribution X\<rparr>. Q x"
-  shows "AE x. Q (X x)"
-proof -
-  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space)
-  obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
-    using assms unfolding X.almost_everywhere_def by auto
-  from X[unfolded measurable_def] N show "AE x. Q (X x)"
-    by (intro AE_I'[where N="X -` N \<inter> space M"])
-       (auto simp: finite_measure_eq distribution_def measurable_sets)
-qed
-
-lemma (in prob_space) distribution_eq_integral:
-  "random_variable S X \<Longrightarrow> A \<in> sets S \<Longrightarrow> distribution X A = expectation (indicator (X -` A \<inter> space M))"
-  using finite_measure_eq[of "X -` A \<inter> space M"]
-  by (auto simp: measurable_sets distribution_def)
-
-lemma (in prob_space) distribution_eq_translated_integral:
-  assumes "random_variable S X" "A \<in> sets S"
-  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := extreal \<circ> distribution X\<rparr>) (indicator A)"
-proof -
-  interpret S: prob_space "S\<lparr>measure := extreal \<circ> distribution X\<rparr>"
-    using assms(1) by (rule distribution_prob_space)
-  show ?thesis
-    using S.positive_integral_indicator(1)[of A] assms by simp
-qed
-
-lemma (in prob_space) finite_expectation1:
-  assumes f: "finite (X`space M)" and rv: "random_variable borel X"
-  shows "expectation X = (\<Sum>r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))" (is "_ = ?r")
-proof (subst integral_on_finite)
-  show "X \<in> borel_measurable M" "finite (X`space M)" using assms by auto
-  show "(\<Sum> r \<in> X ` space M. r * real (\<mu> (X -` {r} \<inter> space M))) = ?r"
-    "\<And>x. \<mu> (X -` {x} \<inter> space M) \<noteq> \<infinity>"
-    using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto
-qed
-
-lemma (in prob_space) finite_expectation:
-  assumes "finite (X`space M)" "random_variable borel X"
-  shows "expectation X = (\<Sum> r \<in> X ` (space M). r * distribution X {r})"
-  using assms unfolding distribution_def using finite_expectation1 by auto
-
-lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0:
-  assumes "{x} \<in> events"
-  assumes "prob {x} = 1"
-  assumes "{y} \<in> events"
-  assumes "y \<noteq> x"
-  shows "prob {y} = 0"
-  using prob_one_inter[of "{y}" "{x}"] assms by auto
-
-lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0"
-  unfolding distribution_def by simp
-
-lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1"
-proof -
-  have "X -` X ` space M \<inter> space M = space M" by auto
-  thus ?thesis unfolding distribution_def by (simp add: prob_space)
-qed
-
-lemma (in prob_space) distribution_one:
-  assumes "random_variable M' X" and "A \<in> sets M'"
-  shows "distribution X A \<le> 1"
-proof -
-  have "distribution X A \<le> \<mu>' (space M)" unfolding distribution_def
-    using assms[unfolded measurable_def] by (auto intro!: finite_measure_mono)
-  thus ?thesis by (simp add: prob_space)
-qed
-
-lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0:
-  assumes X: "random_variable \<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr> X"
-    (is "random_variable ?S X")
-  assumes "distribution X {x} = 1"
-  assumes "y \<noteq> x"
-  shows "distribution X {y} = 0"
-proof cases
-  { fix x have "X -` {x} \<inter> space M \<in> sets M"
-    proof cases
-      assume "x \<in> X`space M" with X show ?thesis
-        by (auto simp: measurable_def image_iff)
-    next
-      assume "x \<notin> X`space M" then have "X -` {x} \<inter> space M = {}" by auto
-      then show ?thesis by auto
-    qed } note single = this
-  have "X -` {x} \<inter> space M - X -` {y} \<inter> space M = X -` {x} \<inter> space M"
-    "X -` {y} \<inter> space M \<inter> (X -` {x} \<inter> space M) = {}"
-    using `y \<noteq> x` by auto
-  with finite_measure_inter_full_set[OF single single, of x y] assms(2)
-  show ?thesis by (auto simp: distribution_def prob_space)
-next
-  assume "{y} \<notin> sets ?S"
-  then have "X -` {y} \<inter> space M = {}" by auto
-  thus "distribution X {y} = 0" unfolding distribution_def by auto
-qed
-
-lemma (in prob_space) joint_distribution_Times_le_fst:
-  assumes X: "random_variable MX X" and Y: "random_variable MY Y"
-    and A: "A \<in> sets MX" and B: "B \<in> sets MY"
-  shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
-  unfolding distribution_def
-proof (intro finite_measure_mono)
-  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
-  show "X -` A \<inter> space M \<in> events"
-    using X A unfolding measurable_def by simp
-  have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
-    (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
-qed
-
-lemma (in prob_space) joint_distribution_commute:
-  "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)"
-  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
-
-lemma (in prob_space) joint_distribution_Times_le_snd:
-  assumes X: "random_variable MX X" and Y: "random_variable MY Y"
-    and A: "A \<in> sets MX" and B: "B \<in> sets MY"
-  shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B"
-  using assms
-  by (subst joint_distribution_commute)
-     (simp add: swap_product joint_distribution_Times_le_fst)
-
-lemma (in prob_space) random_variable_pairI:
-  assumes "random_variable MX X"
-  assumes "random_variable MY Y"
-  shows "random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))"
-proof
-  interpret MX: sigma_algebra MX using assms by simp
-  interpret MY: sigma_algebra MY using assms by simp
-  interpret P: pair_sigma_algebra MX MY by default
-  show "sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
-  have sa: "sigma_algebra M" by default
-  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
-    unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
-qed
-
-lemma (in prob_space) joint_distribution_commute_singleton:
-  "joint_distribution X Y {(x, y)} = joint_distribution Y X {(y, x)}"
-  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
-
-lemma (in prob_space) joint_distribution_assoc_singleton:
-  "joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} =
-   joint_distribution (\<lambda>x. (X x, Y x)) Z {((x, y), z)}"
-  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
-
-locale pair_prob_space = M1: prob_space M1 + M2: prob_space M2 for M1 M2
-
-sublocale pair_prob_space \<subseteq> pair_sigma_finite M1 M2 by default
-
-sublocale pair_prob_space \<subseteq> P: prob_space P
-by default (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure)
-
-lemma countably_additiveI[case_names countably]:
-  assumes "\<And>A. \<lbrakk> range A \<subseteq> sets M ; disjoint_family A ; (\<Union>i. A i) \<in> sets M\<rbrakk> \<Longrightarrow>
-    (\<Sum>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
-  shows "countably_additive M \<mu>"
-  using assms unfolding countably_additive_def by auto
-
-lemma (in prob_space) joint_distribution_prob_space:
-  assumes "random_variable MX X" "random_variable MY Y"
-  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
-  using random_variable_pairI[OF assms] by (rule distribution_prob_space)
-
-section "Probability spaces on finite sets"
-
-locale finite_prob_space = prob_space + finite_measure_space
-
-abbreviation (in prob_space) "finite_random_variable M' X \<equiv> finite_sigma_algebra M' \<and> X \<in> measurable M M'"
-
-lemma (in prob_space) finite_random_variableD:
-  assumes "finite_random_variable M' X" shows "random_variable M' X"
-proof -
-  interpret M': finite_sigma_algebra M' using assms by simp
-  then show "random_variable M' X" using assms by simp default
-qed
-
-lemma (in prob_space) distribution_finite_prob_space:
-  assumes "finite_random_variable MX X"
-  shows "finite_prob_space (MX\<lparr>measure := extreal \<circ> distribution X\<rparr>)"
-proof -
-  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>"
-    using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
-  interpret MX: finite_sigma_algebra MX
-    using assms by auto
-  show ?thesis by default (simp_all add: MX.finite_space)
-qed
-
-lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]:
-  assumes "simple_function M X"
-  shows "finite_random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = x \<rparr> X"
-    (is "finite_random_variable ?X _")
-proof (intro conjI)
-  have [simp]: "finite (X ` space M)" using assms unfolding simple_function_def by simp
-  interpret X: sigma_algebra ?X by (rule sigma_algebra_Pow)
-  show "finite_sigma_algebra ?X"
-    by default auto
-  show "X \<in> measurable M ?X"
-  proof (unfold measurable_def, clarsimp)
-    fix A assume A: "A \<subseteq> X`space M"
-    then have "finite A" by (rule finite_subset) simp
-    then have "X -` (\<Union>a\<in>A. {a}) \<inter> space M \<in> events"
-      unfolding vimage_UN UN_extend_simps
-      apply (rule finite_UN)
-      using A assms unfolding simple_function_def by auto
-    then show "X -` A \<inter> space M \<in> events" by simp
-  qed
-qed
-
-lemma (in prob_space) simple_function_imp_random_variable[simp, intro]:
-  assumes "simple_function M X"
-  shows "random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = ext \<rparr> X"
-  using simple_function_imp_finite_random_variable[OF assms, of ext]
-  by (auto dest!: finite_random_variableD)
-
-lemma (in prob_space) sum_over_space_real_distribution:
-  "simple_function M X \<Longrightarrow> (\<Sum>x\<in>X`space M. distribution X {x}) = 1"
-  unfolding distribution_def prob_space[symmetric]
-  by (subst finite_measure_finite_Union[symmetric])
-     (auto simp add: disjoint_family_on_def simple_function_def
-           intro!: arg_cong[where f=prob])
-
-lemma (in prob_space) finite_random_variable_pairI:
-  assumes "finite_random_variable MX X"
-  assumes "finite_random_variable MY Y"
-  shows "finite_random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))"
-proof
-  interpret MX: finite_sigma_algebra MX using assms by simp
-  interpret MY: finite_sigma_algebra MY using assms by simp
-  interpret P: pair_finite_sigma_algebra MX MY by default
-  show "finite_sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
-  have sa: "sigma_algebra M" by default
-  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
-    unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
-qed
-
-lemma (in prob_space) finite_random_variable_imp_sets:
-  "finite_random_variable MX X \<Longrightarrow> x \<in> space MX \<Longrightarrow> {x} \<in> sets MX"
-  unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp
-
-lemma (in prob_space) finite_random_variable_measurable:
-  assumes X: "finite_random_variable MX X" shows "X -` A \<inter> space M \<in> events"
-proof -
-  interpret X: finite_sigma_algebra MX using X by simp
-  from X have vimage: "\<And>A. A \<subseteq> space MX \<Longrightarrow> X -` A \<inter> space M \<in> events" and
-    "X \<in> space M \<rightarrow> space MX"
-    by (auto simp: measurable_def)
-  then have *: "X -` A \<inter> space M = X -` (A \<inter> space MX) \<inter> space M"
-    by auto
-  show "X -` A \<inter> space M \<in> events"
-    unfolding * by (intro vimage) auto
-qed
-
-lemma (in prob_space) joint_distribution_finite_Times_le_fst:
-  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
-  shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
-  unfolding distribution_def
-proof (intro finite_measure_mono)
-  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
-  show "X -` A \<inter> space M \<in> events"
-    using finite_random_variable_measurable[OF X] .
-  have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
-    (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
-qed
-
-lemma (in prob_space) joint_distribution_finite_Times_le_snd:
-  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
-  shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B"
-  using assms
-  by (subst joint_distribution_commute)
-     (simp add: swap_product joint_distribution_finite_Times_le_fst)
-
-lemma (in prob_space) finite_distribution_order:
-  fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
-  assumes "finite_random_variable MX X" "finite_random_variable MY Y"
-  shows "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
-    and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
-    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
-    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
-    and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
-    and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
-  using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"]
-  using joint_distribution_finite_Times_le_fst[OF assms, of "{x}" "{y}"]
-  by (auto intro: antisym)
-
-lemma (in prob_space) setsum_joint_distribution:
-  assumes X: "finite_random_variable MX X"
-  assumes Y: "random_variable MY Y" "B \<in> sets MY"
-  shows "(\<Sum>a\<in>space MX. joint_distribution X Y ({a} \<times> B)) = distribution Y B"
-  unfolding distribution_def
-proof (subst finite_measure_finite_Union[symmetric])
-  interpret MX: finite_sigma_algebra MX using X by auto
-  show "finite (space MX)" using MX.finite_space .
-  let "?d i" = "(\<lambda>x. (X x, Y x)) -` ({i} \<times> B) \<inter> space M"
-  { fix i assume "i \<in> space MX"
-    moreover have "?d i = (X -` {i} \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
-    ultimately show "?d i \<in> events"
-      using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y
-      using MX.sets_eq_Pow by auto }
-  show "disjoint_family_on ?d (space MX)" by (auto simp: disjoint_family_on_def)
-  show "\<mu>' (\<Union>i\<in>space MX. ?d i) = \<mu>' (Y -` B \<inter> space M)"
-    using X[unfolded measurable_def] by (auto intro!: arg_cong[where f=\<mu>'])
-qed
-
-lemma (in prob_space) setsum_joint_distribution_singleton:
-  assumes X: "finite_random_variable MX X"
-  assumes Y: "finite_random_variable MY Y" "b \<in> space MY"
-  shows "(\<Sum>a\<in>space MX. joint_distribution X Y {(a, b)}) = distribution Y {b}"
-  using setsum_joint_distribution[OF X
-    finite_random_variableD[OF Y(1)]
-    finite_random_variable_imp_sets[OF Y]] by simp
-
-locale pair_finite_prob_space = M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2
-
-sublocale pair_finite_prob_space \<subseteq> pair_prob_space M1 M2 by default
-sublocale pair_finite_prob_space \<subseteq> pair_finite_space M1 M2  by default
-sublocale pair_finite_prob_space \<subseteq> finite_prob_space P by default
-
-lemma (in prob_space) joint_distribution_finite_prob_space:
-  assumes X: "finite_random_variable MX X"
-  assumes Y: "finite_random_variable MY Y"
-  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
-  by (intro distribution_finite_prob_space finite_random_variable_pairI X Y)
-
-lemma finite_prob_space_eq:
-  "finite_prob_space M \<longleftrightarrow> finite_measure_space M \<and> measure M (space M) = 1"
-  unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def
-  by auto
-
-lemma (in prob_space) not_empty: "space M \<noteq> {}"
-  using prob_space empty_measure' by auto
-
-lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1"
-  using measure_space_1 sum_over_space by simp
-
-lemma (in finite_prob_space) joint_distribution_restriction_fst:
-  "joint_distribution X Y A \<le> distribution X (fst ` A)"
-  unfolding distribution_def
-proof (safe intro!: finite_measure_mono)
-  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
-  show "x \<in> X -` fst ` A"
-    by (auto intro!: image_eqI[OF _ *])
-qed (simp_all add: sets_eq_Pow)
-
-lemma (in finite_prob_space) joint_distribution_restriction_snd:
-  "joint_distribution X Y A \<le> distribution Y (snd ` A)"
-  unfolding distribution_def
-proof (safe intro!: finite_measure_mono)
-  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
-  show "x \<in> Y -` snd ` A"
-    by (auto intro!: image_eqI[OF _ *])
-qed (simp_all add: sets_eq_Pow)
-
-lemma (in finite_prob_space) distribution_order:
-  shows "0 \<le> distribution X x'"
-  and "(distribution X x' \<noteq> 0) \<longleftrightarrow> (0 < distribution X x')"
-  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
-  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
-  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
-  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
-  and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
-  and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
-  using
-    joint_distribution_restriction_fst[of X Y "{(x, y)}"]
-    joint_distribution_restriction_snd[of X Y "{(x, y)}"]
-  by (auto intro: antisym)
-
-lemma (in finite_prob_space) distribution_mono:
-  assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
-  shows "distribution X x \<le> distribution Y y"
-  unfolding distribution_def
-  using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono)
-
-lemma (in finite_prob_space) distribution_mono_gt_0:
-  assumes gt_0: "0 < distribution X x"
-  assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
-  shows "0 < distribution Y y"
-  by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
-
-lemma (in finite_prob_space) sum_over_space_distrib:
-  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
-  unfolding distribution_def prob_space[symmetric] using finite_space
-  by (subst finite_measure_finite_Union[symmetric])
-     (auto simp add: disjoint_family_on_def sets_eq_Pow
-           intro!: arg_cong[where f=\<mu>'])
-
-lemma (in finite_prob_space) sum_over_space_real_distribution:
-  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
-  unfolding distribution_def prob_space[symmetric] using finite_space
-  by (subst finite_measure_finite_Union[symmetric])
-     (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob])
-
-lemma (in finite_prob_space) finite_sum_over_space_eq_1:
-  "(\<Sum>x\<in>space M. prob {x}) = 1"
-  using prob_space finite_space
-  by (subst (asm) finite_measure_finite_singleton) auto
-
-lemma (in prob_space) distribution_remove_const:
-  shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
-  and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
-  and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
-  and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
-  and "distribution (\<lambda>x. ()) {()} = 1"
-  by (auto intro!: arg_cong[where f=\<mu>'] simp: distribution_def prob_space[symmetric])
-
-lemma (in finite_prob_space) setsum_distribution_gen:
-  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
-  and "inj_on f (X`space M)"
-  shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}"
-  unfolding distribution_def assms
-  using finite_space assms
-  by (subst finite_measure_finite_Union[symmetric])
-     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
-      intro!: arg_cong[where f=prob])
-
-lemma (in finite_prob_space) setsum_distribution:
-  "(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}"
-  "(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}"
-  "(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}"
-  "(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}"
-  "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
-  by (auto intro!: inj_onI setsum_distribution_gen)
-
-lemma (in finite_prob_space) uniform_prob:
-  assumes "x \<in> space M"
-  assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}"
-  shows "prob {x} = 1 / card (space M)"
-proof -
-  have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}"
-    using assms(2)[OF _ `x \<in> space M`] by blast
-  have "1 = prob (space M)"
-    using prob_space by auto
-  also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
-    using finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
-      sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
-      finite_space unfolding disjoint_family_on_def  prob_space[symmetric]
-    by (auto simp add:setsum_restrict_set)
-  also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
-    using prob_x by auto
-  also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
-  finally have one: "1 = real (card (space M)) * prob {x}"
-    using real_eq_of_nat by auto
-  hence two: "real (card (space M)) \<noteq> 0" by fastsimp 
-  from one have three: "prob {x} \<noteq> 0" by fastsimp
-  thus ?thesis using one two three divide_cancel_right
-    by (auto simp:field_simps)
-qed
-
-lemma (in prob_space) prob_space_subalgebra:
-  assumes "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M"
-    and "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
-  shows "prob_space N"
-proof -
-  interpret N: measure_space N
-    by (rule measure_space_subalgebra[OF assms])
-  show ?thesis
-  proof qed (insert assms(4)[OF N.top], simp add: assms measure_space_1)
-qed
-
-lemma (in prob_space) prob_space_of_restricted_space:
-  assumes "\<mu> A \<noteq> 0" "A \<in> sets M"
-  shows "prob_space (restricted_space A \<lparr>measure := \<lambda>S. \<mu> S / \<mu> A\<rparr>)"
-    (is "prob_space ?P")
-proof -
-  interpret A: measure_space "restricted_space A"
-    using `A \<in> sets M` by (rule restricted_measure_space)
-  interpret A': sigma_algebra ?P
-    by (rule A.sigma_algebra_cong) auto
-  show "prob_space ?P"
-  proof
-    show "measure ?P (space ?P) = 1"
-      using real_measure[OF `A \<in> events`] `\<mu> A \<noteq> 0` by auto
-    show "positive ?P (measure ?P)"
-    proof (simp add: positive_def, safe)
-      show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_extreal_def)
-      fix B assume "B \<in> events"
-      with real_measure[of "A \<inter> B"] real_measure[OF `A \<in> events`] `A \<in> sets M`
-      show "0 \<le> \<mu> (A \<inter> B) / \<mu> A" by (auto simp: Int)
-    qed
-    show "countably_additive ?P (measure ?P)"
-    proof (simp add: countably_additive_def, safe)
-      fix B and F :: "nat \<Rightarrow> 'a set"
-      assume F: "range F \<subseteq> op \<inter> A ` events" "disjoint_family F"
-      { fix i
-        from F have "F i \<in> op \<inter> A ` events" by auto
-        with `A \<in> events` have "F i \<in> events" by auto }
-      moreover then have "range F \<subseteq> events" by auto
-      moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S"
-        by (simp add: mult_commute divide_extreal_def)
-      moreover have "0 \<le> inverse (\<mu> A)"
-        using real_measure[OF `A \<in> events`] by auto
-      ultimately show "(\<Sum>i. \<mu> (F i) / \<mu> A) = \<mu> (\<Union>i. F i) / \<mu> A"
-        using measure_countably_additive[of F] F
-        by (auto simp: suminf_cmult_extreal)
-    qed
-  qed
-qed
-
-lemma finite_prob_spaceI:
-  assumes "finite (space M)" "sets M = Pow(space M)"
-    and "measure M (space M) = 1" "measure M {} = 0" "\<And>A. A \<subseteq> space M \<Longrightarrow> 0 \<le> measure M A"
-    and "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M B"
-  shows "finite_prob_space M"
-  unfolding finite_prob_space_eq
-proof
-  show "finite_measure_space M" using assms
-    by (auto intro!: finite_measure_spaceI)
-  show "measure M (space M) = 1" by fact
-qed
-
-lemma (in finite_prob_space) finite_measure_space:
-  fixes X :: "'a \<Rightarrow> 'x"
-  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X\<rparr>"
-    (is "finite_measure_space ?S")
-proof (rule finite_measure_spaceI, simp_all)
-  show "finite (X ` space M)" using finite_space by simp
-next
-  fix A B :: "'x set" assume "A \<inter> B = {}"
-  then show "distribution X (A \<union> B) = distribution X A + distribution X B"
-    unfolding distribution_def
-    by (subst finite_measure_Union[symmetric])
-       (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow)
-qed
-
-lemma (in finite_prob_space) finite_prob_space_of_images:
-  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X \<rparr>"
-  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_extreal_def)
-
-lemma (in finite_prob_space) finite_product_measure_space:
-  fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
-  assumes "finite s1" "finite s2"
-  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = extreal \<circ> joint_distribution X Y\<rparr>"
-    (is "finite_measure_space ?M")
-proof (rule finite_measure_spaceI, simp_all)
-  show "finite (s1 \<times> s2)"
-    using assms by auto
-next
-  fix A B :: "('x*'y) set" assume "A \<inter> B = {}"
-  then show "joint_distribution X Y (A \<union> B) = joint_distribution X Y A + joint_distribution X Y B"
-    unfolding distribution_def
-    by (subst finite_measure_Union[symmetric])
-       (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow)
-qed
-
-lemma (in finite_prob_space) finite_product_measure_space_of_images:
-  shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
-                                sets = Pow (X ` space M \<times> Y ` space M),
-                                measure = extreal \<circ> joint_distribution X Y \<rparr>"
-  using finite_space by (auto intro!: finite_product_measure_space)
-
-lemma (in finite_prob_space) finite_product_prob_space_of_images:
-  "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M),
-                       measure = extreal \<circ> joint_distribution X Y \<rparr>"
-  (is "finite_prob_space ?S")
-proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_extreal_def)
-  have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
-  thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
-    by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
-qed
-
-section "Conditional Expectation and Probability"
-
-lemma (in prob_space) conditional_expectation_exists:
-  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
-  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
-  and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
-  shows "\<exists>Y\<in>borel_measurable N. (\<forall>x. 0 \<le> Y x) \<and> (\<forall>C\<in>sets N.
-      (\<integral>\<^isup>+x. Y x * indicator C x \<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x \<partial>M))"
-proof -
-  note N(4)[simp]
-  interpret P: prob_space N
-    using prob_space_subalgebra[OF N] .
-
-  let "?f A" = "\<lambda>x. X x * indicator A x"
-  let "?Q A" = "integral\<^isup>P M (?f A)"
-
-  from measure_space_density[OF borel]
-  have Q: "measure_space (N\<lparr> measure := ?Q \<rparr>)"
-    apply (rule measure_space.measure_space_subalgebra[of "M\<lparr> measure := ?Q \<rparr>"])
-    using N by (auto intro!: P.sigma_algebra_cong)
-  then interpret Q: measure_space "N\<lparr> measure := ?Q \<rparr>" .
-
-  have "P.absolutely_continuous ?Q"
-    unfolding P.absolutely_continuous_def
-  proof safe
-    fix A assume "A \<in> sets N" "P.\<mu> A = 0"
-    then have f_borel: "?f A \<in> borel_measurable M" "AE x. x \<notin> A"
-      using borel N by (auto intro!: borel_measurable_indicator AE_not_in)
-    then show "?Q A = 0"
-      by (auto simp add: positive_integral_0_iff_AE)
-  qed
-  from P.Radon_Nikodym[OF Q this]
-  obtain Y where Y: "Y \<in> borel_measurable N" "\<And>x. 0 \<le> Y x"
-    "\<And>A. A \<in> sets N \<Longrightarrow> ?Q A =(\<integral>\<^isup>+x. Y x * indicator A x \<partial>N)"
-    by blast
-  with N(2) show ?thesis
-    by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ _ N(2,3,4,1)])
-qed
-
-definition (in prob_space)
-  "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable N \<and> (\<forall>x. 0 \<le> Y x)
-    \<and> (\<forall>C\<in>sets N. (\<integral>\<^isup>+x. Y x * indicator C x\<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x\<partial>M)))"
-
-abbreviation (in prob_space)
-  "conditional_prob N A \<equiv> conditional_expectation N (indicator A)"
-
-lemma (in prob_space)
-  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
-  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
-  and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
-  shows borel_measurable_conditional_expectation:
-    "conditional_expectation N X \<in> borel_measurable N"
-  and conditional_expectation: "\<And>C. C \<in> sets N \<Longrightarrow>
-      (\<integral>\<^isup>+x. conditional_expectation N X x * indicator C x \<partial>M) =
-      (\<integral>\<^isup>+x. X x * indicator C x \<partial>M)"
-   (is "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C")
-proof -
-  note CE = conditional_expectation_exists[OF assms, unfolded Bex_def]
-  then show "conditional_expectation N X \<in> borel_measurable N"
-    unfolding conditional_expectation_def by (rule someI2_ex) blast
-
-  from CE show "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C"
-    unfolding conditional_expectation_def by (rule someI2_ex) blast
-qed
-
-lemma (in sigma_algebra) factorize_measurable_function_pos:
-  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
-  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
-  assumes Z: "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)"
-  shows "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. max 0 (Z x) = g (Y x)"
-proof -
-  interpret M': sigma_algebra M' by fact
-  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
-  from M'.sigma_algebra_vimage[OF this]
-  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
-
-  from va.borel_measurable_implies_simple_function_sequence'[OF Z] guess f . note f = this
-
-  have "\<forall>i. \<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
-  proof
-    fix i
-    from f(1)[of i] have "finite (f i`space M)" and B_ex:
-      "\<forall>z\<in>(f i)`space M. \<exists>B. B \<in> sets M' \<and> (f i) -` {z} \<inter> space M = Y -` B \<inter> space M"
-      unfolding simple_function_def by auto
-    from B_ex[THEN bchoice] guess B .. note B = this
-
-    let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x"
-
-    show "\<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
-    proof (intro exI[of _ ?g] conjI ballI)
-      show "simple_function M' ?g" using B by auto
-
-      fix x assume "x \<in> space M"
-      then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::extreal)"
-        unfolding indicator_def using B by auto
-      then show "f i x = ?g (Y x)" using `x \<in> space M` f(1)[of i]
-        by (subst va.simple_function_indicator_representation) auto
-    qed
-  qed
-  from choice[OF this] guess g .. note g = this
-
-  show ?thesis
-  proof (intro ballI bexI)
-    show "(\<lambda>x. SUP i. g i x) \<in> borel_measurable M'"
-      using g by (auto intro: M'.borel_measurable_simple_function)
-    fix x assume "x \<in> space M"
-    have "max 0 (Z x) = (SUP i. f i x)" using f by simp
-    also have "\<dots> = (SUP i. g i (Y x))"
-      using g `x \<in> space M` by simp
-    finally show "max 0 (Z x) = (SUP i. g i (Y x))" .
-  qed
-qed
-
-lemma extreal_0_le_iff_le_0[simp]:
-  fixes a :: extreal shows "0 \<le> -a \<longleftrightarrow> a \<le> 0"
-  by (cases rule: extreal2_cases[of a]) auto
-
-lemma (in sigma_algebra) factorize_measurable_function:
-  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
-  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
-  shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
-    \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
-proof safe
-  interpret M': sigma_algebra M' by fact
-  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
-  from M'.sigma_algebra_vimage[OF this]
-  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
-
-  { fix g :: "'c \<Rightarrow> extreal" assume "g \<in> borel_measurable M'"
-    with M'.measurable_vimage_algebra[OF Y]
-    have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-      by (rule measurable_comp)
-    moreover assume "\<forall>x\<in>space M. Z x = g (Y x)"
-    then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow>
-       g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-       by (auto intro!: measurable_cong)
-    ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-      by simp }
-
-  assume Z: "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-  with assms have "(\<lambda>x. - Z x) \<in> borel_measurable M"
-    "(\<lambda>x. - Z x) \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-    by auto
-  from factorize_measurable_function_pos[OF assms(1,2) this] guess n .. note n = this
-  from factorize_measurable_function_pos[OF assms Z] guess p .. note p = this
-  let "?g x" = "p x - n x"
-  show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)"
-  proof (intro bexI ballI)
-    show "?g \<in> borel_measurable M'" using p n by auto
-    fix x assume "x \<in> space M"
-    then have "p (Y x) = max 0 (Z x)" "n (Y x) = max 0 (- Z x)"
-      using p n by auto
-    then show "Z x = ?g (Y x)"
-      by (auto split: split_max)
-  qed
-qed
-
-end
--- a/src/HOL/Probability/Product_Measure.thy	Mon Apr 04 12:39:14 2011 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1975 +0,0 @@
-(*  Title:      HOL/Probability/Product_Measure.thy
-    Author:     Johannes Hölzl, TU München
-*)
-
-header {*Product measure spaces*}
-
-theory Product_Measure
-imports Lebesgue_Integration
-begin
-
-lemma sigma_sets_subseteq: assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
-proof
-  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
-    by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros)
-qed
-
-lemma times_Int_times: "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)"
-  by auto
-
-lemma Pair_vimage_times[simp]: "\<And>A B x. Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
-  by auto
-
-lemma rev_Pair_vimage_times[simp]: "\<And>A B y. (\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
-  by auto
-
-lemma case_prod_distrib: "f (case x of (x, y) \<Rightarrow> g x y) = (case x of (x, y) \<Rightarrow> f (g x y))"
-  by (cases x) simp
-
-lemma split_beta': "(\<lambda>(x,y). f x y) = (\<lambda>x. f (fst x) (snd x))"
-  by (auto simp: fun_eq_iff)
-
-abbreviation
-  "Pi\<^isub>E A B \<equiv> Pi A B \<inter> extensional A"
-
-syntax
-  "_PiE"  :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3PIE _:_./ _)" 10)
-
-syntax (xsymbols)
-  "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>\<^isub>E _\<in>_./ _)"   10)
-
-syntax (HTML output)
-  "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>\<^isub>E _\<in>_./ _)"   10)
-
-translations
-  "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)"
-
-abbreviation
-  funcset_extensional :: "['a set, 'b set] => ('a => 'b) set"
-    (infixr "->\<^isub>E" 60) where
-  "A ->\<^isub>E B \<equiv> Pi\<^isub>E A (%_. B)"
-
-notation (xsymbols)
-  funcset_extensional  (infixr "\<rightarrow>\<^isub>E" 60)
-
-lemma extensional_empty[simp]: "extensional {} = {\<lambda>x. undefined}"
-  by safe (auto simp add: extensional_def fun_eq_iff)
-
-lemma extensional_insert[intro, simp]:
-  assumes "a \<in> extensional (insert i I)"
-  shows "a(i := b) \<in> extensional (insert i I)"
-  using assms unfolding extensional_def by auto
-
-lemma extensional_Int[simp]:
-  "extensional I \<inter> extensional I' = extensional (I \<inter> I')"
-  unfolding extensional_def by auto
-
-definition
-  "merge I x J y = (\<lambda>i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
-
-lemma merge_apply[simp]:
-  "I \<inter> J = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I x J y i = x i"
-  "I \<inter> J = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I x J y i = y i"
-  "J \<inter> I = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I x J y i = x i"
-  "J \<inter> I = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I x J y i = y i"
-  "i \<notin> I \<Longrightarrow> i \<notin> J \<Longrightarrow> merge I x J y i = undefined"
-  unfolding merge_def by auto
-
-lemma merge_commute:
-  "I \<inter> J = {} \<Longrightarrow> merge I x J y = merge J y I x"
-  by (auto simp: merge_def intro!: ext)
-
-lemma Pi_cancel_merge_range[simp]:
-  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge I A J B) \<longleftrightarrow> x \<in> Pi I A"
-  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge J B I A) \<longleftrightarrow> x \<in> Pi I A"
-  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge I A J B) \<longleftrightarrow> x \<in> Pi I A"
-  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge J B I A) \<longleftrightarrow> x \<in> Pi I A"
-  by (auto simp: Pi_def)
-
-lemma Pi_cancel_merge[simp]:
-  "I \<inter> J = {} \<Longrightarrow> merge I x J y \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
-  "J \<inter> I = {} \<Longrightarrow> merge I x J y \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
-  "I \<inter> J = {} \<Longrightarrow> merge I x J y \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
-  "J \<inter> I = {} \<Longrightarrow> merge I x J y \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
-  by (auto simp: Pi_def)
-
-lemma extensional_merge[simp]: "merge I x J y \<in> extensional (I \<union> J)"
-  by (auto simp: extensional_def)
-
-lemma restrict_Pi_cancel: "restrict x I \<in> Pi I A \<longleftrightarrow> x \<in> Pi I A"
-  by (auto simp: restrict_def Pi_def)
-
-lemma restrict_merge[simp]:
-  "I \<inter> J = {} \<Longrightarrow> restrict (merge I x J y) I = restrict x I"
-  "I \<inter> J = {} \<Longrightarrow> restrict (merge I x J y) J = restrict y J"
-  "J \<inter> I = {} \<Longrightarrow> restrict (merge I x J y) I = restrict x I"
-  "J \<inter> I = {} \<Longrightarrow> restrict (merge I x J y) J = restrict y J"
-  by (auto simp: restrict_def intro!: ext)
-
-lemma extensional_insert_undefined[intro, simp]:
-  assumes "a \<in> extensional (insert i I)"
-  shows "a(i := undefined) \<in> extensional I"
-  using assms unfolding extensional_def by auto
-
-lemma extensional_insert_cancel[intro, simp]:
-  assumes "a \<in> extensional I"
-  shows "a \<in> extensional (insert i I)"
-  using assms unfolding extensional_def by auto
-
-lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I x {i} y = restrict (x(i := y i)) (insert i I)"
-  unfolding merge_def by (auto simp: fun_eq_iff)
-
-lemma Pi_Int: "Pi I E \<inter> Pi I F = (\<Pi> i\<in>I. E i \<inter> F i)"
-  by auto
-
-lemma PiE_Int: "(Pi\<^isub>E I A) \<inter> (Pi\<^isub>E I B) = Pi\<^isub>E I (\<lambda>x. A x \<inter> B x)"
-  by auto
-
-lemma Pi_cancel_fupd_range[simp]: "i \<notin> I \<Longrightarrow> x \<in> Pi I (B(i := b)) \<longleftrightarrow> x \<in> Pi I B"
-  by (auto simp: Pi_def)
-
-lemma Pi_split_insert_domain[simp]: "x \<in> Pi (insert i I) X \<longleftrightarrow> x \<in> Pi I X \<and> x i \<in> X i"
-  by (auto simp: Pi_def)
-
-lemma Pi_split_domain[simp]: "x \<in> Pi (I \<union> J) X \<longleftrightarrow> x \<in> Pi I X \<and> x \<in> Pi J X"
-  by (auto simp: Pi_def)
-
-lemma Pi_cancel_fupd[simp]: "i \<notin> I \<Longrightarrow> x(i := a) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
-  by (auto simp: Pi_def)
-
-lemma restrict_vimage:
-  assumes "I \<inter> J = {}"
-  shows "(\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \<times> Pi\<^isub>E J F) = Pi (I \<union> J) (merge I E J F)"
-  using assms by (auto simp: restrict_Pi_cancel)
-
-lemma merge_vimage:
-  assumes "I \<inter> J = {}"
-  shows "(\<lambda>(x,y). merge I x J y) -` Pi\<^isub>E (I \<union> J) E = Pi I E \<times> Pi J E"
-  using assms by (auto simp: restrict_Pi_cancel)
-
-lemma restrict_fupd[simp]: "i \<notin> I \<Longrightarrow> restrict (f (i := x)) I = restrict f I"
-  by (auto simp: restrict_def intro!: ext)
-
-lemma merge_restrict[simp]:
-  "merge I (restrict x I) J y = merge I x J y"
-  "merge I x J (restrict y J) = merge I x J y"
-  unfolding merge_def by (auto intro!: ext)
-
-lemma merge_x_x_eq_restrict[simp]:
-  "merge I x J x = restrict x (I \<union> J)"
-  unfolding merge_def by (auto intro!: ext)
-
-lemma Pi_fupd_iff: "i \<in> I \<Longrightarrow> f \<in> Pi I (B(i := A)) \<longleftrightarrow> f \<in> Pi (I - {i}) B \<and> f i \<in> A"
-  apply auto
-  apply (drule_tac x=x in Pi_mem)
-  apply (simp_all split: split_if_asm)
-  apply (drule_tac x=i in Pi_mem)
-  apply (auto dest!: Pi_mem)
-  done
-
-lemma Pi_UN:
-  fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set"
-  assumes "finite I" and mono: "\<And>i n m. i \<in> I \<Longrightarrow> n \<le> m \<Longrightarrow> A n i \<subseteq> A m i"
-  shows "(\<Union>n. Pi I (A n)) = (\<Pi> i\<in>I. \<Union>n. A n i)"
-proof (intro set_eqI iffI)
-  fix f assume "f \<in> (\<Pi> i\<in>I. \<Union>n. A n i)"
-  then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" by auto
-  from bchoice[OF this] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto
-  obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k"
-    using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto
-  have "f \<in> Pi I (A k)"
-  proof (intro Pi_I)
-    fix i assume "i \<in> I"
-    from mono[OF this, of "n i" k] k[OF this] n[OF this]
-    show "f i \<in> A k i" by auto
-  qed
-  then show "f \<in> (\<Union>n. Pi I (A n))" by auto
-qed auto
-
-lemma PiE_cong:
-  assumes "\<And>i. i\<in>I \<Longrightarrow> A i = B i"
-  shows "Pi\<^isub>E I A = Pi\<^isub>E I B"
-  using assms by (auto intro!: Pi_cong)
-
-lemma restrict_upd[simp]:
-  "i \<notin> I \<Longrightarrow> (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)"
-  by (auto simp: fun_eq_iff)
-
-lemma Pi_eq_subset:
-  assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
-  assumes eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and "i \<in> I"
-  shows "F i \<subseteq> F' i"
-proof
-  fix x assume "x \<in> F i"
-  with ne have "\<forall>j. \<exists>y. ((j \<in> I \<longrightarrow> y \<in> F j \<and> (i = j \<longrightarrow> x = y)) \<and> (j \<notin> I \<longrightarrow> y = undefined))" by auto
-  from choice[OF this] guess f .. note f = this
-  then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def)
-  then have "f \<in> Pi\<^isub>E I F'" using assms by simp
-  then show "x \<in> F' i" using f `i \<in> I` by auto
-qed
-
-lemma Pi_eq_iff_not_empty:
-  assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
-  shows "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i)"
-proof (intro iffI ballI)
-  fix i assume eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and i: "i \<in> I"
-  show "F i = F' i"
-    using Pi_eq_subset[of I F F', OF ne eq i]
-    using Pi_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i]
-    by auto
-qed auto
-
-lemma Pi_eq_empty_iff:
-  "Pi\<^isub>E I F = {} \<longleftrightarrow> (\<exists>i\<in>I. F i = {})"
-proof
-  assume "Pi\<^isub>E I F = {}"
-  show "\<exists>i\<in>I. F i = {}"
-  proof (rule ccontr)
-    assume "\<not> ?thesis"
-    then have "\<forall>i. \<exists>y. (i \<in> I \<longrightarrow> y \<in> F i) \<and> (i \<notin> I \<longrightarrow> y = undefined)" by auto
-    from choice[OF this] guess f ..
-    then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def)
-    with `Pi\<^isub>E I F = {}` show False by auto
-  qed
-qed auto
-
-lemma Pi_eq_iff:
-  "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i) \<or> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
-proof (intro iffI disjCI)
-  assume eq[simp]: "Pi\<^isub>E I F = Pi\<^isub>E I F'"
-  assume "\<not> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
-  then have "(\<forall>i\<in>I. F i \<noteq> {}) \<and> (\<forall>i\<in>I. F' i \<noteq> {})"
-    using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto
-  with Pi_eq_iff_not_empty[of I F F'] show "\<forall>i\<in>I. F i = F' i" by auto
-next
-  assume "(\<forall>i\<in>I. F i = F' i) \<or> (\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {})"
-  then show "Pi\<^isub>E I F = Pi\<^isub>E I F'"
-    using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto
-qed
-
-section "Binary products"
-
-definition
-  "pair_measure_generator A B =
-    \<lparr> space = space A \<times> space B,
-      sets = {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B},
-      measure = \<lambda>X. \<integral>\<^isup>+x. (\<integral>\<^isup>+y. indicator X (x,y) \<partial>B) \<partial>A \<rparr>"
-
-definition pair_measure (infixr "\<Otimes>\<^isub>M" 80) where
-  "A \<Otimes>\<^isub>M B = sigma (pair_measure_generator A B)"
-
-locale pair_sigma_algebra = M1: sigma_algebra M1 + M2: sigma_algebra M2
-  for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
-
-abbreviation (in pair_sigma_algebra)
-  "E \<equiv> pair_measure_generator M1 M2"
-
-abbreviation (in pair_sigma_algebra)
-  "P \<equiv> M1 \<Otimes>\<^isub>M M2"
-
-lemma sigma_algebra_pair_measure:
-  "sets M1 \<subseteq> Pow (space M1) \<Longrightarrow> sets M2 \<subseteq> Pow (space M2) \<Longrightarrow> sigma_algebra (pair_measure M1 M2)"
-  by (force simp: pair_measure_def pair_measure_generator_def intro!: sigma_algebra_sigma)
-
-sublocale pair_sigma_algebra \<subseteq> sigma_algebra P
-  using M1.space_closed M2.space_closed
-  by (rule sigma_algebra_pair_measure)
-
-lemma pair_measure_generatorI[intro, simp]:
-  "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (pair_measure_generator A B)"
-  by (auto simp add: pair_measure_generator_def)
-
-lemma pair_measureI[intro, simp]:
-  "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^isub>M B)"
-  by (auto simp add: pair_measure_def)
-
-lemma space_pair_measure:
-  "space (A \<Otimes>\<^isub>M B) = space A \<times> space B"
-  by (simp add: pair_measure_def pair_measure_generator_def)
-
-lemma sets_pair_measure_generator:
-  "sets (pair_measure_generator N M) = (\<lambda>(x, y). x \<times> y) ` (sets N \<times> sets M)"
-  unfolding pair_measure_generator_def by auto
-
-lemma pair_measure_generator_sets_into_space:
-  assumes "sets M \<subseteq> Pow (space M)" "sets N \<subseteq> Pow (space N)"
-  shows "sets (pair_measure_generator M N) \<subseteq> Pow (space (pair_measure_generator M N))"
-  using assms by (auto simp: pair_measure_generator_def)
-
-lemma pair_measure_generator_Int_snd:
-  assumes "sets S1 \<subseteq> Pow (space S1)"
-  shows "sets (pair_measure_generator S1 (algebra.restricted_space S2 A)) =
-         sets (algebra.restricted_space (pair_measure_generator S1 S2) (space S1 \<times> A))"
-  (is "?L = ?R")
-  apply (auto simp: pair_measure_generator_def image_iff)
-  using assms
-  apply (rule_tac x="a \<times> xa" in exI)
-  apply force
-  using assms
-  apply (rule_tac x="a" in exI)
-  apply (rule_tac x="b \<inter> A" in exI)
-  apply auto
-  done
-
-lemma (in pair_sigma_algebra)
-  shows measurable_fst[intro!, simp]:
-    "fst \<in> measurable P M1" (is ?fst)
-  and measurable_snd[intro!, simp]:
-    "snd \<in> measurable P M2" (is ?snd)
-proof -
-  { fix X assume "X \<in> sets M1"
-    then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. fst -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2"
-      apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"])
-      using M1.sets_into_space by force+ }
-  moreover
-  { fix X assume "X \<in> sets M2"
-    then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. snd -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2"
-      apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X])
-      using M2.sets_into_space by force+ }
-  ultimately have "?fst \<and> ?snd"
-    by (fastsimp simp: measurable_def sets_sigma space_pair_measure
-                 intro!: sigma_sets.Basic)
-  then show ?fst ?snd by auto
-qed
-
-lemma (in pair_sigma_algebra) measurable_pair_iff:
-  assumes "sigma_algebra M"
-  shows "f \<in> measurable M P \<longleftrightarrow>
-    (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
-proof -
-  interpret M: sigma_algebra M by fact
-  from assms show ?thesis
-  proof (safe intro!: measurable_comp[where b=P])
-    assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2"
-    show "f \<in> measurable M P" unfolding pair_measure_def
-    proof (rule M.measurable_sigma)
-      show "sets (pair_measure_generator M1 M2) \<subseteq> Pow (space E)"
-        unfolding pair_measure_generator_def using M1.sets_into_space M2.sets_into_space by auto
-      show "f \<in> space M \<rightarrow> space E"
-        using f s by (auto simp: mem_Times_iff measurable_def comp_def space_sigma pair_measure_generator_def)
-      fix A assume "A \<in> sets E"
-      then obtain B C where "B \<in> sets M1" "C \<in> sets M2" "A = B \<times> C"
-        unfolding pair_measure_generator_def by auto
-      moreover have "(fst \<circ> f) -` B \<inter> space M \<in> sets M"
-        using f `B \<in> sets M1` unfolding measurable_def by auto
-      moreover have "(snd \<circ> f) -` C \<inter> space M \<in> sets M"
-        using s `C \<in> sets M2` unfolding measurable_def by auto
-      moreover have "f -` A \<inter> space M = ((fst \<circ> f) -` B \<inter> space M) \<inter> ((snd \<circ> f) -` C \<inter> space M)"
-        unfolding `A = B \<times> C` by (auto simp: vimage_Times)
-      ultimately show "f -` A \<inter> space M \<in> sets M" by auto
-    qed
-  qed
-qed
-
-lemma (in pair_sigma_algebra) measurable_pair:
-  assumes "sigma_algebra M"
-  assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2"
-  shows "f \<in> measurable M P"
-  unfolding measurable_pair_iff[OF assms(1)] using assms(2,3) by simp
-
-lemma pair_measure_generatorE:
-  assumes "X \<in> sets (pair_measure_generator M1 M2)"
-  obtains A B where "X = A \<times> B" "A \<in> sets M1" "B \<in> sets M2"
-  using assms unfolding pair_measure_generator_def by auto
-
-lemma (in pair_sigma_algebra) pair_measure_generator_swap:
-  "(\<lambda>X. (\<lambda>(x,y). (y,x)) -` X \<inter> space M2 \<times> space M1) ` sets E = sets (pair_measure_generator M2 M1)"
-proof (safe elim!: pair_measure_generatorE)
-  fix A B assume "A \<in> sets M1" "B \<in> sets M2"
-  moreover then have "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 = B \<times> A"
-    using M1.sets_into_space M2.sets_into_space by auto
-  ultimately show "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 \<in> sets (pair_measure_generator M2 M1)"
-    by (auto intro: pair_measure_generatorI)
-next
-  fix A B assume "A \<in> sets M1" "B \<in> sets M2"
-  then show "B \<times> A \<in> (\<lambda>X. (\<lambda>(x, y). (y, x)) -` X \<inter> space M2 \<times> space M1) ` sets E"
-    using M1.sets_into_space M2.sets_into_space
-    by (auto intro!: image_eqI[where x="A \<times> B"] pair_measure_generatorI)
-qed
-
-lemma (in pair_sigma_algebra) sets_pair_sigma_algebra_swap:
-  assumes Q: "Q \<in> sets P"
-  shows "(\<lambda>(x,y). (y, x)) -` Q \<in> sets (M2 \<Otimes>\<^isub>M M1)" (is "_ \<in> sets ?Q")
-proof -
-  let "?f Q" = "(\<lambda>(x,y). (y, x)) -` Q \<inter> space M2 \<times> space M1"
-  have *: "(\<lambda>(x,y). (y, x)) -` Q = ?f Q"
-    using sets_into_space[OF Q] by (auto simp: space_pair_measure)
-  have "sets (M2 \<Otimes>\<^isub>M M1) = sets (sigma (pair_measure_generator M2 M1))"
-    unfolding pair_measure_def ..
-  also have "\<dots> = sigma_sets (space M2 \<times> space M1) (?f ` sets E)"
-    unfolding sigma_def pair_measure_generator_swap[symmetric]
-    by (simp add: pair_measure_generator_def)
-  also have "\<dots> = ?f ` sigma_sets (space M1 \<times> space M2) (sets E)"
-    using M1.sets_into_space M2.sets_into_space
-    by (intro sigma_sets_vimage) (auto simp: pair_measure_generator_def)
-  also have "\<dots> = ?f ` sets P"
-    unfolding pair_measure_def pair_measure_generator_def sigma_def by simp
-  finally show ?thesis
-    using Q by (subst *) auto
-qed
-
-lemma (in pair_sigma_algebra) pair_sigma_algebra_swap_measurable:
-  shows "(\<lambda>(x,y). (y, x)) \<in> measurable P (M2 \<Otimes>\<^isub>M M1)"
-    (is "?f \<in> measurable ?P ?Q")
-  unfolding measurable_def
-proof (intro CollectI conjI Pi_I ballI)
-  fix x assume "x \<in> space ?P" then show "(case x of (x, y) \<Rightarrow> (y, x)) \<in> space ?Q"
-    unfolding pair_measure_generator_def pair_measure_def by auto
-next
-  fix A assume "A \<in> sets (M2 \<Otimes>\<^isub>M M1)"
-  interpret Q: pair_sigma_algebra M2 M1 by default
-  with Q.sets_pair_sigma_algebra_swap[OF `A \<in> sets (M2 \<Otimes>\<^isub>M M1)`]
-  show "?f -` A \<inter> space ?P \<in> sets ?P" by simp
-qed
-
-lemma (in pair_sigma_algebra) measurable_cut_fst[simp,intro]:
-  assumes "Q \<in> sets P" shows "Pair x -` Q \<in> sets M2"
-proof -
-  let ?Q' = "{Q. Q \<subseteq> space P \<and> Pair x -` Q \<in> sets M2}"
-  let ?Q = "\<lparr> space = space P, sets = ?Q' \<rparr>"
-  interpret Q: sigma_algebra ?Q
-    proof qed (auto simp: vimage_UN vimage_Diff space_pair_measure)
-  have "sets E \<subseteq> sets ?Q"
-    using M1.sets_into_space M2.sets_into_space
-    by (auto simp: pair_measure_generator_def space_pair_measure)
-  then have "sets P \<subseteq> sets ?Q"
-    apply (subst pair_measure_def, intro Q.sets_sigma_subset)
-    by (simp add: pair_measure_def)
-  with assms show ?thesis by auto
-qed
-
-lemma (in pair_sigma_algebra) measurable_cut_snd:
-  assumes Q: "Q \<in> sets P" shows "(\<lambda>x. (x, y)) -` Q \<in> sets M1" (is "?cut Q \<in> sets M1")
-proof -
-  interpret Q: pair_sigma_algebra M2 M1 by default
-  with Q.measurable_cut_fst[OF sets_pair_sigma_algebra_swap[OF Q], of y]
-  show ?thesis by (simp add: vimage_compose[symmetric] comp_def)
-qed
-
-lemma (in pair_sigma_algebra) measurable_pair_image_snd:
-  assumes m: "f \<in> measurable P M" and "x \<in> space M1"
-  shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
-  unfolding measurable_def
-proof (intro CollectI conjI Pi_I ballI)
-  fix y assume "y \<in> space M2" with `f \<in> measurable P M` `x \<in> space M1`
-  show "f (x, y) \<in> space M"
-    unfolding measurable_def pair_measure_generator_def pair_measure_def by auto
-next
-  fix A assume "A \<in> sets M"
-  then have "Pair x -` (f -` A \<inter> space P) \<in> sets M2" (is "?C \<in> _")
-    using `f \<in> measurable P M`
-    by (intro measurable_cut_fst) (auto simp: measurable_def)
-  also have "?C = (\<lambda>y. f (x, y)) -` A \<inter> space M2"
-    using `x \<in> space M1` by (auto simp: pair_measure_generator_def pair_measure_def)
-  finally show "(\<lambda>y. f (x, y)) -` A \<inter> space M2 \<in> sets M2" .
-qed
-
-lemma (in pair_sigma_algebra) measurable_pair_image_fst:
-  assumes m: "f \<in> measurable P M" and "y \<in> space M2"
-  shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"
-proof -
-  interpret Q: pair_sigma_algebra M2 M1 by default
-  from Q.measurable_pair_image_snd[OF measurable_comp `y \<in> space M2`,
-                                      OF Q.pair_sigma_algebra_swap_measurable m]
-  show ?thesis by simp
-qed
-
-lemma (in pair_sigma_algebra) Int_stable_pair_measure_generator: "Int_stable E"
-  unfolding Int_stable_def
-proof (intro ballI)
-  fix A B assume "A \<in> sets E" "B \<in> sets E"
-  then obtain A1 A2 B1 B2 where "A = A1 \<times> A2" "B = B1 \<times> B2"
-    "A1 \<in> sets M1" "A2 \<in> sets M2" "B1 \<in> sets M1" "B2 \<in> sets M2"
-    unfolding pair_measure_generator_def by auto
-  then show "A \<inter> B \<in> sets E"
-    by (auto simp add: times_Int_times pair_measure_generator_def)
-qed
-
-lemma finite_measure_cut_measurable:
-  fixes M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
-  assumes "sigma_finite_measure M1" "finite_measure M2"
-  assumes "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)"
-  shows "(\<lambda>x. measure M2 (Pair x -` Q)) \<in> borel_measurable M1"
-    (is "?s Q \<in> _")
-proof -
-  interpret M1: sigma_finite_measure M1 by fact
-  interpret M2: finite_measure M2 by fact
-  interpret pair_sigma_algebra M1 M2 by default
-  have [intro]: "sigma_algebra M1" by fact
-  have [intro]: "sigma_algebra M2" by fact
-  let ?D = "\<lparr> space = space P, sets = {A\<in>sets P. ?s A \<in> borel_measurable M1}  \<rparr>"
-  note space_pair_measure[simp]
-  interpret dynkin_system ?D
-  proof (intro dynkin_systemI)
-    fix A assume "A \<in> sets ?D" then show "A \<subseteq> space ?D"
-      using sets_into_space by simp
-  next
-    from top show "space ?D \<in> sets ?D"
-      by (auto simp add: if_distrib intro!: M1.measurable_If)
-  next
-    fix A assume "A \<in> sets ?D"
-    with sets_into_space have "\<And>x. measure M2 (Pair x -` (space M1 \<times> space M2 - A)) =
-        (if x \<in> space M1 then measure M2 (space M2) - ?s A x else 0)"
-      by (auto intro!: M2.measure_compl simp: vimage_Diff)
-    with `A \<in> sets ?D` top show "space ?D - A \<in> sets ?D"
-      by (auto intro!: Diff M1.measurable_If M1.borel_measurable_extreal_diff)
-  next
-    fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> sets ?D"
-    moreover then have "\<And>x. measure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
-      by (intro M2.measure_countably_additive[symmetric])
-         (auto simp: disjoint_family_on_def)
-    ultimately show "(\<Union>i. F i) \<in> sets ?D"
-      by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf)
-  qed
-  have "sets P = sets ?D" apply (subst pair_measure_def)
-  proof (intro dynkin_lemma)
-    show "Int_stable E" by (rule Int_stable_pair_measure_generator)
-    from M1.sets_into_space have "\<And>A. A \<in> sets M1 \<Longrightarrow> {x \<in> space M1. x \<in> A} = A"
-      by auto
-    then show "sets E \<subseteq> sets ?D"
-      by (auto simp: pair_measure_generator_def sets_sigma if_distrib
-               intro: sigma_sets.Basic intro!: M1.measurable_If)
-  qed (auto simp: pair_measure_def)
-  with `Q \<in> sets P` have "Q \<in> sets ?D" by simp
-  then show "?s Q \<in> borel_measurable M1" by simp
-qed
-
-subsection {* Binary products of $\sigma$-finite measure spaces *}
-
-locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2
-  for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
-
-sublocale pair_sigma_finite \<subseteq> pair_sigma_algebra M1 M2
-  by default
-
-lemma times_eq_iff: "A \<times> B = C \<times> D \<longleftrightarrow> A = C \<and> B = D \<or> ((A = {} \<or> B = {}) \<and> (C = {} \<or> D = {}))"
-  by auto
-
-lemma (in pair_sigma_finite) measure_cut_measurable_fst:
-  assumes "Q \<in> sets P" shows "(\<lambda>x. measure M2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _")
-proof -
-  have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+
-  have M1: "sigma_finite_measure M1" by default
-  from M2.disjoint_sigma_finite guess F .. note F = this
-  then have F_sets: "\<And>i. F i \<in> sets M2" by auto
-  let "?C x i" = "F i \<inter> Pair x -` Q"
-  { fix i
-    let ?R = "M2.restricted_space (F i)"
-    have [simp]: "space M1 \<times> F i \<inter> space M1 \<times> space M2 = space M1 \<times> F i"
-      using F M2.sets_into_space by auto
-    let ?R2 = "M2.restricted_space (F i)"
-    have "(\<lambda>x. measure ?R2 (Pair x -` (space M1 \<times> space ?R2 \<inter> Q))) \<in> borel_measurable M1"
-    proof (intro finite_measure_cut_measurable[OF M1])
-      show "finite_measure ?R2"
-        using F by (intro M2.restricted_to_finite_measure) auto
-      have "(space M1 \<times> space ?R2) \<inter> Q \<in> (op \<inter> (space M1 \<times> F i)) ` sets P"
-        using `Q \<in> sets P` by (auto simp: image_iff)
-      also have "\<dots> = sigma_sets (space M1 \<times> F i) ((op \<inter> (space M1 \<times> F i)) ` sets E)"
-        unfolding pair_measure_def pair_measure_generator_def sigma_def
-        using `F i \<in> sets M2` M2.sets_into_space
-        by (auto intro!: sigma_sets_Int sigma_sets.Basic)
-      also have "\<dots> \<subseteq> sets (M1 \<Otimes>\<^isub>M ?R2)"
-        using M1.sets_into_space
-        apply (auto simp: times_Int_times pair_measure_def pair_measure_generator_def sigma_def
-                    intro!: sigma_sets_subseteq)
-        apply (rule_tac x="a" in exI)
-        apply (rule_tac x="b \<inter> F i" in exI)
-        by auto
-      finally show "(space M1 \<times> space ?R2) \<inter> Q \<in> sets (M1 \<Otimes>\<^isub>M ?R2)" .
-    qed
-    moreover have "\<And>x. Pair x -` (space M1 \<times> F i \<inter> Q) = ?C x i"
-      using `Q \<in> sets P` sets_into_space by (auto simp: space_pair_measure)
-    ultimately have "(\<lambda>x. measure M2 (?C x i)) \<in> borel_measurable M1"
-      by simp }
-  moreover
-  { fix x
-    have "(\<Sum>i. measure M2 (?C x i)) = measure M2 (\<Union>i. ?C x i)"
-    proof (intro M2.measure_countably_additive)
-      show "range (?C x) \<subseteq> sets M2"
-        using F `Q \<in> sets P` by (auto intro!: M2.Int)
-      have "disjoint_family F" using F by auto
-      show "disjoint_family (?C x)"
-        by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto
-    qed
-    also have "(\<Union>i. ?C x i) = Pair x -` Q"
-      using F sets_into_space `Q \<in> sets P`
-      by (auto simp: space_pair_measure)
-    finally have "measure M2 (Pair x -` Q) = (\<Sum>i. measure M2 (?C x i))"
-      by simp }
-  ultimately show ?thesis using `Q \<in> sets P` F_sets
-    by (auto intro!: M1.borel_measurable_psuminf M2.Int)
-qed
-
-lemma (in pair_sigma_finite) measure_cut_measurable_snd:
-  assumes "Q \<in> sets P" shows "(\<lambda>y. M1.\<mu> ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
-proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
-  note sets_pair_sigma_algebra_swap[OF assms]
-  from Q.measure_cut_measurable_fst[OF this]
-  show ?thesis by (simp add: vimage_compose[symmetric] comp_def)
-qed
-
-lemma (in pair_sigma_algebra) pair_sigma_algebra_measurable:
-  assumes "f \<in> measurable P M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^isub>M M1) M"
-proof -
-  interpret Q: pair_sigma_algebra M2 M1 by default
-  have *: "(\<lambda>(x,y). f (y, x)) = f \<circ> (\<lambda>(x,y). (y, x))" by (simp add: fun_eq_iff)
-  show ?thesis
-    using Q.pair_sigma_algebra_swap_measurable assms
-    unfolding * by (rule measurable_comp)
-qed
-
-lemma (in pair_sigma_finite) pair_measure_alt:
-  assumes "A \<in> sets P"
-  shows "measure (M1 \<Otimes>\<^isub>M M2) A = (\<integral>\<^isup>+ x. measure M2 (Pair x -` A) \<partial>M1)"
-  apply (simp add: pair_measure_def pair_measure_generator_def)
-proof (rule M1.positive_integral_cong)
-  fix x assume "x \<in> space M1"
-  have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: extreal)"
-    unfolding indicator_def by auto
-  show "(\<integral>\<^isup>+ y. indicator A (x, y) \<partial>M2) = measure M2 (Pair x -` A)"
-    unfolding *
-    apply (subst M2.positive_integral_indicator)
-    apply (rule measurable_cut_fst[OF assms])
-    by simp
-qed
-
-lemma (in pair_sigma_finite) pair_measure_times:
-  assumes A: "A \<in> sets M1" and "B \<in> sets M2"
-  shows "measure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = M1.\<mu> A * measure M2 B"
-proof -
-  have "measure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = (\<integral>\<^isup>+ x. measure M2 B * indicator A x \<partial>M1)"
-    using assms by (auto intro!: M1.positive_integral_cong simp: pair_measure_alt)
-  with assms show ?thesis
-    by (simp add: M1.positive_integral_cmult_indicator ac_simps)
-qed
-
-lemma (in measure_space) measure_not_negative[simp,intro]:
-  assumes A: "A \<in> sets M" shows "\<mu> A \<noteq> - \<infinity>"
-  using positive_measure[OF A] by auto
-
-lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
-  "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets E \<and> incseq F \<and> (\<Union>i. F i) = space E \<and>
-    (\<forall>i. measure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
-proof -
-  obtain F1 :: "nat \<Rightarrow> 'a set" and F2 :: "nat \<Rightarrow> 'b set" where
-    F1: "range F1 \<subseteq> sets M1" "incseq F1" "(\<Union>i. F1 i) = space M1" "\<And>i. M1.\<mu> (F1 i) \<noteq> \<infinity>" and
-    F2: "range F2 \<subseteq> sets M2" "incseq F2" "(\<Union>i. F2 i) = space M2" "\<And>i. M2.\<mu> (F2 i) \<noteq> \<infinity>"
-    using M1.sigma_finite_up M2.sigma_finite_up by auto
-  then have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto
-  let ?F = "\<lambda>i. F1 i \<times> F2 i"
-  show ?thesis unfolding space_pair_measure
-  proof (intro exI[of _ ?F] conjI allI)
-    show "range ?F \<subseteq> sets E" using F1 F2
-      by (fastsimp intro!: pair_measure_generatorI)
-  next
-    have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"
-    proof (intro subsetI)
-      fix x assume "x \<in> space M1 \<times> space M2"
-      then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j"
-        by (auto simp: space)
-      then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)"
-        using `incseq F1` `incseq F2` unfolding incseq_def
-        by (force split: split_max)+
-      then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)"
-        by (intro SigmaI) (auto simp add: min_max.sup_commute)
-      then show "x \<in> (\<Union>i. ?F i)" by auto
-    qed
-    then show "(\<Union>i. ?F i) = space E"
-      using space by (auto simp: space pair_measure_generator_def)
-  next
-    fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)"
-      using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto
-  next
-    fix i
-    from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
-    with F1 F2 M1.positive_measure[OF this(1)] M2.positive_measure[OF this(2)]
-    show "measure P (F1 i \<times> F2 i) \<noteq> \<infinity>"
-      by (simp add: pair_measure_times)
-  qed
-qed
-
-sublocale pair_sigma_finite \<subseteq> sigma_finite_measure P
-proof
-  show "positive P (measure P)"
-    unfolding pair_measure_def pair_measure_generator_def sigma_def positive_def
-    by (auto intro: M1.positive_integral_positive M2.positive_integral_positive)
-
-  show "countably_additive P (measure P)"
-    unfolding countably_additive_def
-  proof (intro allI impI)
-    fix F :: "nat \<Rightarrow> ('a \<times> 'b) set"
-    assume F: "range F \<subseteq> sets P" "disjoint_family F"
-    from F have *: "\<And>i. F i \<in> sets P" "(\<Union>i. F i) \<in> sets P" by auto
-    moreover from F have "\<And>i. (\<lambda>x. measure M2 (Pair x -` F i)) \<in> borel_measurable M1"
-      by (intro measure_cut_measurable_fst) auto
-    moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"
-      by (intro disjoint_family_on_bisimulation[OF F(2)]) auto
-    moreover have "\<And>x. x \<in> space M1 \<Longrightarrow> range (\<lambda>i. Pair x -` F i) \<subseteq> sets M2"
-      using F by auto
-    ultimately show "(\<Sum>n. measure P (F n)) = measure P (\<Union>i. F i)"
-      by (simp add: pair_measure_alt vimage_UN M1.positive_integral_suminf[symmetric]
-                    M2.measure_countably_additive
-               cong: M1.positive_integral_cong)
-  qed
-
-  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
-  show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets P \<and> (\<Union>i. F i) = space P \<and> (\<forall>i. measure P (F i) \<noteq> \<infinity>)"
-  proof (rule exI[of _ F], intro conjI)
-    show "range F \<subseteq> sets P" using F by (auto simp: pair_measure_def)
-    show "(\<Union>i. F i) = space P"
-      using F by (auto simp: pair_measure_def pair_measure_generator_def)
-    show "\<forall>i. measure P (F i) \<noteq> \<infinity>" using F by auto
-  qed
-qed
-
-lemma (in pair_sigma_algebra) sets_swap:
-  assumes "A \<in> sets P"
-  shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
-    (is "_ -` A \<inter> space ?Q \<in> sets ?Q")
-proof -
-  have *: "(\<lambda>(x, y). (y, x)) -` A \<inter> space ?Q = (\<lambda>(x, y). (y, x)) -` A"
-    using `A \<in> sets P` sets_into_space by (auto simp: space_pair_measure)
-  show ?thesis
-    unfolding * using assms by (rule sets_pair_sigma_algebra_swap)
-qed
-
-lemma (in pair_sigma_finite) pair_measure_alt2:
-  assumes A: "A \<in> sets P"
-  shows "\<mu> A = (\<integral>\<^isup>+y. M1.\<mu> ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
-    (is "_ = ?\<nu> A")
-proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
-  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
-  have [simp]: "\<And>m. \<lparr> space = space E, sets = sets (sigma E), measure = m \<rparr> = P\<lparr> measure := m \<rparr>"
-    unfolding pair_measure_def by simp
-
-  have "\<mu> A = Q.\<mu> ((\<lambda>(y, x). (x, y)) -` A \<inter> space Q.P)"
-  proof (rule measure_unique_Int_stable_vimage[OF Int_stable_pair_measure_generator])
-    show "measure_space P" "measure_space Q.P" by default
-    show "(\<lambda>(y, x). (x, y)) \<in> measurable Q.P P" by (rule Q.pair_sigma_algebra_swap_measurable)
-    show "sets (sigma E) = sets P" "space E = space P" "A \<in> sets (sigma E)"
-      using assms unfolding pair_measure_def by auto
-    show "range F \<subseteq> sets E" "incseq F" "(\<Union>i. F i) = space E" "\<And>i. \<mu> (F i) \<noteq> \<infinity>"
-      using F `A \<in> sets P` by (auto simp: pair_measure_def)
-    fix X assume "X \<in> sets E"
-    then obtain A B where X[simp]: "X = A \<times> B" and AB: "A \<in> sets M1" "B \<in> sets M2"
-      unfolding pair_measure_def pair_measure_generator_def by auto
-    then have "(\<lambda>(y, x). (x, y)) -` X \<inter> space Q.P = B \<times> A"
-      using M1.sets_into_space M2.sets_into_space by (auto simp: space_pair_measure)
-    then show "\<mu> X = Q.\<mu> ((\<lambda>(y, x). (x, y)) -` X \<inter> space Q.P)"
-      using AB by (simp add: pair_measure_times Q.pair_measure_times ac_simps)
-  qed
-  then show ?thesis
-    using sets_into_space[OF A] Q.pair_measure_alt[OF sets_swap[OF A]]
-    by (auto simp add: Q.pair_measure_alt space_pair_measure
-             intro!: M2.positive_integral_cong arg_cong[where f="M1.\<mu>"])
-qed
-
-lemma pair_sigma_algebra_sigma:
-  assumes 1: "incseq S1" "(\<Union>i. S1 i) = space E1" "range S1 \<subseteq> sets E1" and E1: "sets E1 \<subseteq> Pow (space E1)"
-  assumes 2: "decseq S2" "(\<Union>i. S2 i) = space E2" "range S2 \<subseteq> sets E2" and E2: "sets E2 \<subseteq> Pow (space E2)"
-  shows "sets (sigma (pair_measure_generator (sigma E1) (sigma E2))) = sets (sigma (pair_measure_generator E1 E2))"
-    (is "sets ?S = sets ?E")
-proof -
-  interpret M1: sigma_algebra "sigma E1" using E1 by (rule sigma_algebra_sigma)
-  interpret M2: sigma_algebra "sigma E2" using E2 by (rule sigma_algebra_sigma)
-  have P: "sets (pair_measure_generator E1 E2) \<subseteq> Pow (space E1 \<times> space E2)"
-    using E1 E2 by (auto simp add: pair_measure_generator_def)
-  interpret E: sigma_algebra ?E unfolding pair_measure_generator_def
-    using E1 E2 by (intro sigma_algebra_sigma) auto
-  { fix A assume "A \<in> sets E1"
-    then have "fst -` A \<inter> space ?E = A \<times> (\<Union>i. S2 i)"
-      using E1 2 unfolding pair_measure_generator_def by auto
-    also have "\<dots> = (\<Union>i. A \<times> S2 i)" by auto
-    also have "\<dots> \<in> sets ?E" unfolding pair_measure_generator_def sets_sigma
-      using 2 `A \<in> sets E1`
-      by (intro sigma_sets.Union)
-         (force simp: image_subset_iff intro!: sigma_sets.Basic)
-    finally have "fst -` A \<inter> space ?E \<in> sets ?E" . }
-  moreover
-  { fix B assume "B \<in> sets E2"
-    then have "snd -` B \<inter> space ?E = (\<Union>i. S1 i) \<times> B"
-      using E2 1 unfolding pair_measure_generator_def by auto
-    also have "\<dots> = (\<Union>i. S1 i \<times> B)" by auto
-    also have "\<dots> \<in> sets ?E"
-      using 1 `B \<in> sets E2` unfolding pair_measure_generator_def sets_sigma
-      by (intro sigma_sets.Union)
-         (force simp: image_subset_iff intro!: sigma_sets.Basic)
-    finally have "snd -` B \<inter> space ?E \<in> sets ?E" . }
-  ultimately have proj:
-    "fst \<in> measurable ?E (sigma E1) \<and> snd \<in> measurable ?E (sigma E2)"
-    using E1 E2 by (subst (1 2) E.measurable_iff_sigma)
-                   (auto simp: pair_measure_generator_def sets_sigma)
-  { fix A B assume A: "A \<in> sets (sigma E1)" and B: "B \<in> sets (sigma E2)"
-    with proj have "fst -` A \<inter> space ?E \<in> sets ?E" "snd -` B \<inter> space ?E \<in> sets ?E"
-      unfolding measurable_def by simp_all
-    moreover have "A \<times> B = (fst -` A \<inter> space ?E) \<inter> (snd -` B \<inter> space ?E)"
-      using A B M1.sets_into_space M2.sets_into_space
-      by (auto simp: pair_measure_generator_def)
-    ultimately have "A \<times> B \<in> sets ?E" by auto }
-  then have "sigma_sets (space ?E) (sets (pair_measure_generator (sigma E1) (sigma E2))) \<subseteq> sets ?E"
-    by (intro E.sigma_sets_subset) (auto simp add: pair_measure_generator_def sets_sigma)
-  then have subset: "sets ?S \<subseteq> sets ?E"
-    by (simp add: sets_sigma pair_measure_generator_def)
-  show "sets ?S = sets ?E"
-  proof (intro set_eqI iffI)
-    fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S"
-      unfolding sets_sigma
-    proof induct
-      case (Basic A) then show ?case
-        by (auto simp: pair_measure_generator_def sets_sigma intro: sigma_sets.Basic)
-    qed (auto intro: sigma_sets.intros simp: pair_measure_generator_def)
-  next
-    fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
-  qed
-qed
-
-section "Fubinis theorem"
-
-lemma (in pair_sigma_finite) simple_function_cut:
-  assumes f: "simple_function P f" "\<And>x. 0 \<le> f x"
-  shows "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
-    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
-proof -
-  have f_borel: "f \<in> borel_measurable P"
-    using f(1) by (rule borel_measurable_simple_function)
-  let "?F z" = "f -` {z} \<inter> space P"
-  let "?F' x z" = "Pair x -` ?F z"
-  { fix x assume "x \<in> space M1"
-    have [simp]: "\<And>z y. indicator (?F z) (x, y) = indicator (?F' x z) y"
-      by (auto simp: indicator_def)
-    have "\<And>y. y \<in> space M2 \<Longrightarrow> (x, y) \<in> space P" using `x \<in> space M1`
-      by (simp add: space_pair_measure)
-    moreover have "\<And>x z. ?F' x z \<in> sets M2" using f_borel
-      by (intro borel_measurable_vimage measurable_cut_fst)
-    ultimately have "simple_function M2 (\<lambda> y. f (x, y))"
-      apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _])
-      apply (rule simple_function_indicator_representation[OF f(1)])
-      using `x \<in> space M1` by (auto simp del: space_sigma) }
-  note M2_sf = this
-  { fix x assume x: "x \<in> space M1"
-    then have "(\<integral>\<^isup>+y. f (x, y) \<partial>M2) = (\<Sum>z\<in>f ` space P. z * M2.\<mu> (?F' x z))"
-      unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)]
-      unfolding simple_integral_def
-    proof (safe intro!: setsum_mono_zero_cong_left)
-      from f(1) show "finite (f ` space P)" by (rule simple_functionD)
-    next
-      fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space P"
-        using `x \<in> space M1` by (auto simp: space_pair_measure)
-    next
-      fix x' y assume "(x', y) \<in> space P"
-        "f (x', y) \<notin> (\<lambda>y. f (x, y)) ` space M2"
-      then have *: "?F' x (f (x', y)) = {}"
-        by (force simp: space_pair_measure)
-      show  "f (x', y) * M2.\<mu> (?F' x (f (x', y))) = 0"
-        unfolding * by simp
-    qed (simp add: vimage_compose[symmetric] comp_def
-                   space_pair_measure) }
-  note eq = this
-  moreover have "\<And>z. ?F z \<in> sets P"
-    by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma)
-  moreover then have "\<And>z. (\<lambda>x. M2.\<mu> (?F' x z)) \<in> borel_measurable M1"
-    by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int)
-  moreover have *: "\<And>i x. 0 \<le> M2.\<mu> (Pair x -` (f -` {i} \<inter> space P))"
-    using f(1)[THEN simple_functionD(2)] f(2) by (intro M2.positive_measure measurable_cut_fst)
-  moreover { fix i assume "i \<in> f`space P"
-    with * have "\<And>x. 0 \<le> i * M2.\<mu> (Pair x -` (f -` {i} \<inter> space P))"
-      using f(2) by auto }
-  ultimately
-  show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
-    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f" using f(2)
-    by (auto simp del: vimage_Int cong: measurable_cong
-             intro!: M1.borel_measurable_extreal_setsum setsum_cong
-             simp add: M1.positive_integral_setsum simple_integral_def
-                       M1.positive_integral_cmult
-                       M1.positive_integral_cong[OF eq]
-                       positive_integral_eq_simple_integral[OF f]
-                       pair_measure_alt[symmetric])
-qed
-
-lemma (in pair_sigma_finite) positive_integral_fst_measurable:
-  assumes f: "f \<in> borel_measurable P"
-  shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
-      (is "?C f \<in> borel_measurable M1")
-    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
-proof -
-  from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this
-  then have F_borel: "\<And>i. F i \<in> borel_measurable P"
-    by (auto intro: borel_measurable_simple_function)
-  note sf = simple_function_cut[OF F(1,5)]
-  then have "(\<lambda>x. SUP i. ?C (F i) x) \<in> borel_measurable M1"
-    using F(1) by auto
-  moreover
-  { fix x assume "x \<in> space M1"
-    from F measurable_pair_image_snd[OF F_borel`x \<in> space M1`]
-    have "(\<integral>\<^isup>+y. (SUP i. F i (x, y)) \<partial>M2) = (SUP i. ?C (F i) x)"
-      by (intro M2.positive_integral_monotone_convergence_SUP)
-         (auto simp: incseq_Suc_iff le_fun_def)
-    then have "(SUP i. ?C (F i) x) = ?C f x"
-      unfolding F(4) positive_integral_max_0 by simp }
-  note SUPR_C = this
-  ultimately show "?C f \<in> borel_measurable M1"
-    by (simp cong: measurable_cong)
-  have "(\<integral>\<^isup>+x. (SUP i. F i x) \<partial>P) = (SUP i. integral\<^isup>P P (F i))"
-    using F_borel F
-    by (intro positive_integral_monotone_convergence_SUP) auto
-  also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1)"
-    unfolding sf(2) by simp
-  also have "\<dots> = \<integral>\<^isup>+ x. (SUP i. \<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1" using F sf(1)
-    by (intro M1.positive_integral_monotone_convergence_SUP[symmetric])
-       (auto intro!: M2.positive_integral_mono M2.positive_integral_positive
-                simp: incseq_Suc_iff le_fun_def)
-  also have "\<dots> = \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. (SUP i. F i (x, y)) \<partial>M2) \<partial>M1"
-    using F_borel F(2,5)
-    by (auto intro!: M1.positive_integral_cong M2.positive_integral_monotone_convergence_SUP[symmetric]
-             simp: incseq_Suc_iff le_fun_def measurable_pair_image_snd)
-  finally show "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
-    using F by (simp add: positive_integral_max_0)
-qed
-
-lemma (in pair_sigma_finite) measure_preserving_swap:
-  "(\<lambda>(x,y). (y, x)) \<in> measure_preserving (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
-proof
-  interpret Q: pair_sigma_finite M2 M1 by default
-  show *: "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
-    using pair_sigma_algebra_swap_measurable .
-  fix X assume "X \<in> sets (M2 \<Otimes>\<^isub>M M1)"
-  from measurable_sets[OF * this] this Q.sets_into_space[OF this]
-  show "measure (M1 \<Otimes>\<^isub>M M2) ((\<lambda>(x, y). (y, x)) -` X \<inter> space P) = measure (M2 \<Otimes>\<^isub>M M1) X"
-    by (auto intro!: M1.positive_integral_cong arg_cong[where f="M2.\<mu>"]
-      simp: pair_measure_alt Q.pair_measure_alt2 space_pair_measure)
-qed
-
-lemma (in pair_sigma_finite) positive_integral_product_swap:
-  assumes f: "f \<in> borel_measurable P"
-  shows "(\<integral>\<^isup>+x. f (case x of (x,y)\<Rightarrow>(y,x)) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P P f"
-proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
-  have "sigma_algebra P" by default
-  with f show ?thesis
-    by (subst Q.positive_integral_vimage[OF _ Q.measure_preserving_swap]) auto
-qed
-
-lemma (in pair_sigma_finite) positive_integral_snd_measurable:
-  assumes f: "f \<in> borel_measurable P"
-  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P P f"
-proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
-  note pair_sigma_algebra_measurable[OF f]
-  from Q.positive_integral_fst_measurable[OF this]
-  have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>Q.P)"
-    by simp
-  also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>Q.P) = integral\<^isup>P P f"
-    unfolding positive_integral_product_swap[OF f, symmetric]
-    by (auto intro!: Q.positive_integral_cong)
-  finally show ?thesis .
-qed
-
-lemma (in pair_sigma_finite) Fubini:
-  assumes f: "f \<in> borel_measurable P"
-  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
-  unfolding positive_integral_snd_measurable[OF assms]
-  unfolding positive_integral_fst_measurable[OF assms] ..
-
-lemma (in pair_sigma_finite) AE_pair:
-  assumes "AE x in P. Q x"
-  shows "AE x in M1. (AE y in M2. Q (x, y))"
-proof -
-  obtain N where N: "N \<in> sets P" "\<mu> N = 0" "{x\<in>space P. \<not> Q x} \<subseteq> N"
-    using assms unfolding almost_everywhere_def by auto
-  show ?thesis
-  proof (rule M1.AE_I)
-    from N measure_cut_measurable_fst[OF `N \<in> sets P`]
-    show "M1.\<mu> {x\<in>space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} = 0"
-      by (auto simp: pair_measure_alt M1.positive_integral_0_iff)
-    show "{x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} \<in> sets M1"
-      by (intro M1.borel_measurable_extreal_neq_const measure_cut_measurable_fst N)
-    { fix x assume "x \<in> space M1" "M2.\<mu> (Pair x -` N) = 0"
-      have "M2.almost_everywhere (\<lambda>y. Q (x, y))"
-      proof (rule M2.AE_I)
-        show "M2.\<mu> (Pair x -` N) = 0" by fact
-        show "Pair x -` N \<in> sets M2" by (intro measurable_cut_fst N)
-        show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
-          using N `x \<in> space M1` unfolding space_sigma space_pair_measure by auto
-      qed }
-    then show "{x \<in> space M1. \<not> M2.almost_everywhere (\<lambda>y. Q (x, y))} \<subseteq> {x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0}"
-      by auto
-  qed
-qed
-
-lemma (in pair_sigma_algebra) measurable_product_swap:
-  "f \<in> measurable (M2 \<Otimes>\<^isub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable P M"
-proof -
-  interpret Q: pair_sigma_algebra M2 M1 by default
-  show ?thesis
-    using pair_sigma_algebra_measurable[of "\<lambda>(x,y). f (y, x)"]
-    by (auto intro!: pair_sigma_algebra_measurable Q.pair_sigma_algebra_measurable iffI)
-qed
-
-lemma (in pair_sigma_finite) integrable_product_swap:
-  assumes "integrable P f"
-  shows "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x))"
-proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
-  have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
-  show ?thesis unfolding *
-    using assms unfolding integrable_def
-    apply (subst (1 2) positive_integral_product_swap)
-    using `integrable P f` unfolding integrable_def
-    by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric])
-qed
-
-lemma (in pair_sigma_finite) integrable_product_swap_iff:
-  "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable P f"
-proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
-  from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f]
-  show ?thesis by auto
-qed
-
-lemma (in pair_sigma_finite) integral_product_swap:
-  assumes "integrable P f"
-  shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>L P f"
-proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
-  have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
-  show ?thesis
-    unfolding lebesgue_integral_def *
-    apply (subst (1 2) positive_integral_product_swap)
-    using `integrable P f` unfolding integrable_def
-    by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric])
-qed
-
-lemma (in pair_sigma_finite) integrable_fst_measurable:
-  assumes f: "integrable P f"
-  shows "M1.almost_everywhere (\<lambda>x. integrable M2 (\<lambda> y. f (x, y)))" (is "?AE")
-    and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L P f" (is "?INT")
-proof -
-  let "?pf x" = "extreal (f x)" and "?nf x" = "extreal (- f x)"
-  have
-    borel: "?nf \<in> borel_measurable P""?pf \<in> borel_measurable P" and
-    int: "integral\<^isup>P P ?nf \<noteq> \<infinity>" "integral\<^isup>P P ?pf \<noteq> \<infinity>"
-    using assms by auto
-  have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
-     "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
-    using borel[THEN positive_integral_fst_measurable(1)] int
-    unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all
-  with borel[THEN positive_integral_fst_measurable(1)]
-  have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
-    "AE x in M1. (\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
-    by (auto intro!: M1.positive_integral_PInf_AE )
-  then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
-    "AE x in M1. \<bar>\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
-    by (auto simp: M2.positive_integral_positive)
-  from AE_pos show ?AE using assms
-    by (simp add: measurable_pair_image_snd integrable_def)
-  { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. extreal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
-      using M2.positive_integral_positive
-      by (intro M1.positive_integral_cong_pos) (auto simp: extreal_uminus_le_reorder)
-    then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. extreal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
-  note this[simp]
-  { fix f assume borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable P"
-      and int: "integral\<^isup>P P (\<lambda>x. extreal (f x)) \<noteq> \<infinity>"
-      and AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<noteq> \<infinity>)"
-    have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
-    proof (intro integrable_def[THEN iffD2] conjI)
-      show "?f \<in> borel_measurable M1"
-        using borel by (auto intro!: M1.borel_measurable_real_of_extreal positive_integral_fst_measurable)
-      have "(\<integral>\<^isup>+x. extreal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (f (x, y))  \<partial>M2) \<partial>M1)"
-        using AE M2.positive_integral_positive
-        by (auto intro!: M1.positive_integral_cong_AE simp: extreal_real)
-      then show "(\<integral>\<^isup>+x. extreal (?f x) \<partial>M1) \<noteq> \<infinity>"
-        using positive_integral_fst_measurable[OF borel] int by simp
-      have "(\<integral>\<^isup>+x. extreal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
-        by (intro M1.positive_integral_cong_pos)
-           (simp add: M2.positive_integral_positive real_of_extreal_pos)
-      then show "(\<integral>\<^isup>+x. extreal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
-    qed }
-  with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)]
-  show ?INT
-    unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2]
-      borel[THEN positive_integral_fst_measurable(2), symmetric]
-    using AE[THEN M1.integral_real]
-    by simp
-qed
-
-lemma (in pair_sigma_finite) integrable_snd_measurable:
-  assumes f: "integrable P f"
-  shows "M2.almost_everywhere (\<lambda>y. integrable M1 (\<lambda>x. f (x, y)))" (is "?AE")
-    and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>L P f" (is "?INT")
-proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
-  have Q_int: "integrable Q.P (\<lambda>(x, y). f (y, x))"
-    using f unfolding integrable_product_swap_iff .
-  show ?INT
-    using Q.integrable_fst_measurable(2)[OF Q_int]
-    using integral_product_swap[OF f] by simp
-  show ?AE
-    using Q.integrable_fst_measurable(1)[OF Q_int]
-    by simp
-qed
-
-lemma (in pair_sigma_finite) Fubini_integral:
-  assumes f: "integrable P f"
-  shows "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1)"
-  unfolding integrable_snd_measurable[OF assms]
-  unfolding integrable_fst_measurable[OF assms] ..
-
-section "Finite product spaces"
-
-section "Products"
-
-locale product_sigma_algebra =
-  fixes M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme"
-  assumes sigma_algebras: "\<And>i. sigma_algebra (M i)"
-
-locale finite_product_sigma_algebra = product_sigma_algebra M
-  for M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme" +
-  fixes I :: "'i set"
-  assumes finite_index: "finite I"
-
-definition
-  "product_algebra_generator I M = \<lparr> space = (\<Pi>\<^isub>E i \<in> I. space (M i)),
-    sets = Pi\<^isub>E I ` (\<Pi> i \<in> I. sets (M i)),
-    measure = \<lambda>A. (\<Prod>i\<in>I. measure (M i) ((SOME F. A = Pi\<^isub>E I F) i)) \<rparr>"
-
-definition product_algebra_def:
-  "Pi\<^isub>M I M = sigma (product_algebra_generator I M)
-    \<lparr> measure := (SOME \<mu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<mu> \<rparr>) \<and>
-      (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))))\<rparr>"
-
-syntax
-  "_PiM"  :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
-              ('i => 'a, 'b) measure_space_scheme"  ("(3PIM _:_./ _)" 10)
-
-syntax (xsymbols)
-  "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
-             ('i => 'a, 'b) measure_space_scheme"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"   10)
-
-syntax (HTML output)
-  "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
-             ('i => 'a, 'b) measure_space_scheme"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"   10)
-
-translations
-  "PIM x:I. M" == "CONST Pi\<^isub>M I (%x. M)"
-
-abbreviation (in finite_product_sigma_algebra) "G \<equiv> product_algebra_generator I M"
-abbreviation (in finite_product_sigma_algebra) "P \<equiv> Pi\<^isub>M I M"
-
-sublocale product_sigma_algebra \<subseteq> M: sigma_algebra "M i" for i by (rule sigma_algebras)
-
-lemma sigma_into_space:
-  assumes "sets M \<subseteq> Pow (space M)"
-  shows "sets (sigma M) \<subseteq> Pow (space M)"
-  using sigma_sets_into_sp[OF assms] unfolding sigma_def by auto
-
-lemma (in product_sigma_algebra) product_algebra_generator_into_space:
-  "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))"
-  using M.sets_into_space unfolding product_algebra_generator_def
-  by auto blast
-
-lemma (in product_sigma_algebra) product_algebra_into_space:
-  "sets (Pi\<^isub>M I M) \<subseteq> Pow (space (Pi\<^isub>M I M))"
-  using product_algebra_generator_into_space
-  by (auto intro!: sigma_into_space simp add: product_algebra_def)
-
-lemma (in product_sigma_algebra) sigma_algebra_product_algebra: "sigma_algebra (Pi\<^isub>M I M)"
-  using product_algebra_generator_into_space unfolding product_algebra_def
-  by (rule sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) simp_all
-
-sublocale finite_product_sigma_algebra \<subseteq> sigma_algebra P
-  using sigma_algebra_product_algebra .
-
-lemma product_algebraE:
-  assumes "A \<in> sets (product_algebra_generator I M)"
-  obtains E where "A = Pi\<^isub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
-  using assms unfolding product_algebra_generator_def by auto
-
-lemma product_algebra_generatorI[intro]:
-  assumes "E \<in> (\<Pi> i\<in>I. sets (M i))"
-  shows "Pi\<^isub>E I E \<in> sets (product_algebra_generator I M)"
-  using assms unfolding product_algebra_generator_def by auto
-
-lemma space_product_algebra_generator[simp]:
-  "space (product_algebra_generator I M) = Pi\<^isub>E I (\<lambda>i. space (M i))"
-  unfolding product_algebra_generator_def by simp
-
-lemma space_product_algebra[simp]:
-  "space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E i\<in>I. space (M i))"
-  unfolding product_algebra_def product_algebra_generator_def by simp
-
-lemma sets_product_algebra:
-  "sets (Pi\<^isub>M I M) = sets (sigma (product_algebra_generator I M))"
-  unfolding product_algebra_def sigma_def by simp
-
-lemma product_algebra_generator_sets_into_space:
-  assumes "\<And>i. i\<in>I \<Longrightarrow> sets (M i) \<subseteq> Pow (space (M i))"
-  shows "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))"
-  using assms by (auto simp: product_algebra_generator_def) blast
-
-lemma (in finite_product_sigma_algebra) in_P[simp, intro]:
-  "\<lbrakk> \<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i) \<rbrakk> \<Longrightarrow> Pi\<^isub>E I A \<in> sets P"
-  by (auto simp: sets_product_algebra)
-
-section "Generating set generates also product algebra"
-
-lemma sigma_product_algebra_sigma_eq:
-  assumes "finite I"
-  assumes mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)"
-  assumes union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (E i)"
-  assumes sets_into: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> sets (E i)"
-  and E: "\<And>i. sets (E i) \<subseteq> Pow (space (E i))"
-  shows "sets (\<Pi>\<^isub>M i\<in>I. sigma (E i)) = sets (\<Pi>\<^isub>M i\<in>I. E i)"
-    (is "sets ?S = sets ?E")
-proof cases
-  assume "I = {}" then show ?thesis
-    by (simp add: product_algebra_def product_algebra_generator_def)
-next
-  assume "I \<noteq> {}"
-  interpret E: sigma_algebra "sigma (E i)" for i
-    using E by (rule sigma_algebra_sigma)
-  have into_space[intro]: "\<And>i x A. A \<in> sets (E i) \<Longrightarrow> x i \<in> A \<Longrightarrow> x i \<in> space (E i)"
-    using E by auto
-  interpret G: sigma_algebra ?E
-    unfolding product_algebra_def product_algebra_generator_def using E
-    by (intro sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) (auto dest: Pi_mem)
-  { fix A i assume "i \<in> I" and A: "A \<in> sets (E i)"
-    then have "(\<lambda>x. x i) -` A \<inter> space ?E = (\<Pi>\<^isub>E j\<in>I. if j = i then A else \<Union>n. S j n) \<inter> space ?E"
-      using mono union unfolding incseq_Suc_iff space_product_algebra
-      by (auto dest: Pi_mem)
-    also have "\<dots> = (\<Union>n. (\<Pi>\<^isub>E j\<in>I. if j = i then A else S j n))"
-      unfolding space_product_algebra
-      apply simp
-      apply (subst Pi_UN[OF `finite I`])
-      using mono[THEN incseqD] apply simp
-      apply (simp add: PiE_Int)
-      apply (intro PiE_cong)
-      using A sets_into by (auto intro!: into_space)
-    also have "\<dots> \<in> sets ?E"
-      using sets_into `A \<in> sets (E i)`
-      unfolding sets_product_algebra sets_sigma
-      by (intro sigma_sets.Union)
-         (auto simp: image_subset_iff intro!: sigma_sets.Basic)
-    finally have "(\<lambda>x. x i) -` A \<inter> space ?E \<in> sets ?E" . }
-  then have proj:
-    "\<And>i. i\<in>I \<Longrightarrow> (\<lambda>x. x i) \<in> measurable ?E (sigma (E i))"
-    using E by (subst G.measurable_iff_sigma)
-               (auto simp: sets_product_algebra sets_sigma)
-  { fix A assume A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (sigma (E i))"
-    with proj have basic: "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. x i) -` (A i) \<inter> space ?E \<in> sets ?E"
-      unfolding measurable_def by simp
-    have "Pi\<^isub>E I A = (\<Inter>i\<in>I. (\<lambda>x. x i) -` (A i) \<inter> space ?E)"
-      using A E.sets_into_space `I \<noteq> {}` unfolding product_algebra_def by auto blast
-    then have "Pi\<^isub>E I A \<in> sets ?E"
-      using G.finite_INT[OF `finite I` `I \<noteq> {}` basic, of "\<lambda>i. i"] by simp }
-  then have "sigma_sets (space ?E) (sets (product_algebra_generator I (\<lambda>i. sigma (E i)))) \<subseteq> sets ?E"
-    by (intro G.sigma_sets_subset) (auto simp add: product_algebra_generator_def)
-  then have subset: "sets ?S \<subseteq> sets ?E"
-    by (simp add: sets_sigma sets_product_algebra)
-  show "sets ?S = sets ?E"
-  proof (intro set_eqI iffI)
-    fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S"
-      unfolding sets_sigma sets_product_algebra
-    proof induct
-      case (Basic A) then show ?case
-        by (auto simp: sets_sigma product_algebra_generator_def intro: sigma_sets.Basic)
-    qed (auto intro: sigma_sets.intros simp: product_algebra_generator_def)
-  next
-    fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
-  qed
-qed
-
-lemma product_algebraI[intro]:
-    "E \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> Pi\<^isub>E I E \<in> sets (Pi\<^isub>M I M)"
-  using assms unfolding product_algebra_def by (auto intro: product_algebra_generatorI)
-
-lemma (in product_sigma_algebra) measurable_component_update:
-  assumes "x \<in> space (Pi\<^isub>M I M)" and "i \<notin> I"
-  shows "(\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^isub>M (insert i I) M)" (is "?f \<in> _")
-  unfolding product_algebra_def apply simp
-proof (intro measurable_sigma)
-  let ?G = "product_algebra_generator (insert i I) M"
-  show "sets ?G \<subseteq> Pow (space ?G)" using product_algebra_generator_into_space .
-  show "?f \<in> space (M i) \<rightarrow> space ?G"
-    using M.sets_into_space assms by auto
-  fix A assume "A \<in> sets ?G"
-  from product_algebraE[OF this] guess E . note E = this
-  then have "?f -` A \<inter> space (M i) = E i \<or> ?f -` A \<inter> space (M i) = {}"
-    using M.sets_into_space assms by auto
-  then show "?f -` A \<inter> space (M i) \<in> sets (M i)"
-    using E by (auto intro!: product_algebraI)
-qed
-
-lemma (in product_sigma_algebra) measurable_add_dim:
-  assumes "i \<notin> I"
-  shows "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M)"
-proof -
-  let ?f = "(\<lambda>(f, y). f(i := y))" and ?G = "product_algebra_generator (insert i I) M"
-  interpret Ii: pair_sigma_algebra "Pi\<^isub>M I M" "M i"
-    unfolding pair_sigma_algebra_def
-    by (intro sigma_algebra_product_algebra sigma_algebras conjI)
-  have "?f \<in> measurable Ii.P (sigma ?G)"
-  proof (rule Ii.measurable_sigma)
-    show "sets ?G \<subseteq> Pow (space ?G)"
-      using product_algebra_generator_into_space .
-    show "?f \<in> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<rightarrow> space ?G"
-      by (auto simp: space_pair_measure)
-  next
-    fix A assume "A \<in> sets ?G"
-    then obtain F where "A = Pi\<^isub>E (insert i I) F"
-      and F: "\<And>j. j \<in> I \<Longrightarrow> F j \<in> sets (M j)" "F i \<in> sets (M i)"
-      by (auto elim!: product_algebraE)
-    then have "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = Pi\<^isub>E I F \<times> (F i)"
-      using sets_into_space `i \<notin> I`
-      by (auto simp add: space_pair_measure) blast+
-    then show "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M M i)"
-      using F by (auto intro!: pair_measureI)
-  qed
-  then show ?thesis
-    by (simp add: product_algebra_def)
-qed
-
-lemma (in product_sigma_algebra) measurable_merge:
-  assumes [simp]: "I \<inter> J = {}"
-  shows "(\<lambda>(x, y). merge I x J y) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
-proof -
-  let ?I = "Pi\<^isub>M I M" and ?J = "Pi\<^isub>M J M"
-  interpret P: sigma_algebra "?I \<Otimes>\<^isub>M ?J"
-    by (intro sigma_algebra_pair_measure product_algebra_into_space)
-  let ?f = "\<lambda>(x, y). merge I x J y"
-  let ?G = "product_algebra_generator (I \<union> J) M"
-  have "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (sigma ?G)"
-  proof (rule P.measurable_sigma)
-    fix A assume "A \<in> sets ?G"
-    from product_algebraE[OF this]
-    obtain E where E: "A = Pi\<^isub>E (I \<union> J) E" "E \<in> (\<Pi> i\<in>I \<union> J. sets (M i))" .
-    then have *: "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) = Pi\<^isub>E I E \<times> Pi\<^isub>E J E"
-      using sets_into_space `I \<inter> J = {}`
-      by (auto simp add: space_pair_measure) fast+
-    then show "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) \<in> sets (?I \<Otimes>\<^isub>M ?J)"
-      using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI
-        simp: product_algebra_def)
-  qed (insert product_algebra_generator_into_space, auto simp: space_pair_measure)
-  then show "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (Pi\<^isub>M (I \<union> J) M)"
-    unfolding product_algebra_def[of "I \<union> J"] by simp
-qed
-
-lemma (in product_sigma_algebra) measurable_component_singleton:
-  assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>M I M) (M i)"
-proof (unfold measurable_def, intro CollectI conjI ballI)
-  fix A assume "A \<in> sets (M i)"
-  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E j\<in>I. if i = j then A else space (M j))"
-    using M.sets_into_space `i \<in> I` by (fastsimp dest: Pi_mem split: split_if_asm)
-  then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) \<in> sets (Pi\<^isub>M I M)"
-    using `A \<in> sets (M i)` by (auto intro!: product_algebraI)
-qed (insert `i \<in> I`, auto)
-
-locale product_sigma_finite =
-  fixes M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme"
-  assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)"
-
-locale finite_product_sigma_finite = product_sigma_finite M
-  for M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" +
-  fixes I :: "'i set" assumes finite_index'[intro]: "finite I"
-
-sublocale product_sigma_finite \<subseteq> M: sigma_finite_measure "M i" for i
-  by (rule sigma_finite_measures)
-
-sublocale product_sigma_finite \<subseteq> product_sigma_algebra
-  by default
-
-sublocale finite_product_sigma_finite \<subseteq> finite_product_sigma_algebra
-  by default (fact finite_index')
-
-lemma setprod_extreal_0:
-  fixes f :: "'a \<Rightarrow> extreal"
-  shows "(\<Prod>i\<in>A. f i) = 0 \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. f i = 0))"
-proof cases
-  assume "finite A"
-  then show ?thesis by (induct A) auto
-qed auto
-
-lemma setprod_extreal_pos:
-  fixes f :: "'a \<Rightarrow> extreal" assumes pos: "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i" shows "0 \<le> (\<Prod>i\<in>I. f i)"
-proof cases
-  assume "finite I" from this pos show ?thesis by induct auto
-qed simp
-
-lemma setprod_PInf:
-  assumes "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i"
-  shows "(\<Prod>i\<in>I. f i) = \<infinity> \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = \<infinity>) \<and> (\<forall>i\<in>I. f i \<noteq> 0)"
-proof cases
-  assume "finite I" from this assms show ?thesis
-  proof (induct I)
-    case (insert i I)
-    then have pos: "0 \<le> f i" "0 \<le> setprod f I" by (auto intro!: setprod_extreal_pos)
-    from insert have "(\<Prod>j\<in>insert i I. f j) = \<infinity> \<longleftrightarrow> setprod f I * f i = \<infinity>" by auto
-    also have "\<dots> \<longleftrightarrow> (setprod f I = \<infinity> \<or> f i = \<infinity>) \<and> f i \<noteq> 0 \<and> setprod f I \<noteq> 0"
-      using setprod_extreal_pos[of I f] pos
-      by (cases rule: extreal2_cases[of "f i" "setprod f I"]) auto
-    also have "\<dots> \<longleftrightarrow> finite (insert i I) \<and> (\<exists>j\<in>insert i I. f j = \<infinity>) \<and> (\<forall>j\<in>insert i I. f j \<noteq> 0)"
-      using insert by (auto simp: setprod_extreal_0)
-    finally show ?case .
-  qed simp
-qed simp
-
-lemma setprod_extreal: "(\<Prod>i\<in>A. extreal (f i)) = extreal (setprod f A)"
-proof cases
-  assume "finite A" then show ?thesis
-    by induct (auto simp: one_extreal_def)
-qed (simp add: one_extreal_def)
-
-lemma (in finite_product_sigma_finite) product_algebra_generator_measure:
-  assumes "Pi\<^isub>E I F \<in> sets G"
-  shows "measure G (Pi\<^isub>E I F) = (\<Prod>i\<in>I. M.\<mu> i (F i))"
-proof cases
-  assume ne: "\<forall>i\<in>I. F i \<noteq> {}"
-  have "\<forall>i\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') i = F i"
-    by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
-       (insert ne, auto simp: Pi_eq_iff)
-  then show ?thesis
-    unfolding product_algebra_generator_def by simp
-next
-  assume empty: "\<not> (\<forall>j\<in>I. F j \<noteq> {})"
-  then have "(\<Prod>j\<in>I. M.\<mu> j (F j)) = 0"
-    by (auto simp: setprod_extreal_0 intro!: bexI)
-  moreover
-  have "\<exists>j\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}"
-    by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
-       (insert empty, auto simp: Pi_eq_empty_iff[symmetric])
-  then have "(\<Prod>j\<in>I. M.\<mu> j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0"
-    by (auto simp: setprod_extreal_0 intro!: bexI)
-  ultimately show ?thesis
-    unfolding product_algebra_generator_def by simp
-qed
-
-lemma (in finite_product_sigma_finite) sigma_finite_pairs:
-  "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
-    (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
-    (\<forall>k. \<forall>i\<in>I. \<mu> i (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<and>
-    (\<Union>k. \<Pi>\<^isub>E i\<in>I. F i k) = space G"
-proof -
-  have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. \<mu> i (F k) \<noteq> \<infinity>)"
-    using M.sigma_finite_up by simp
-  from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
-  then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. \<mu> i (F i k) \<noteq> \<infinity>"
-    by auto
-  let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k"
-  note space_product_algebra[simp]
-  show ?thesis
-  proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI)
-    fix i show "range (F i) \<subseteq> sets (M i)" by fact
-  next
-    fix i k show "\<mu> i (F i k) \<noteq> \<infinity>" by fact
-  next
-    fix A assume "A \<in> (\<Union>i. ?F i)" then show "A \<in> space G"
-      using `\<And>i. range (F i) \<subseteq> sets (M i)` M.sets_into_space
-      by (force simp: image_subset_iff)
-  next
-    fix f assume "f \<in> space G"
-    with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F
-    show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def)
-  next
-    fix i show "?F i \<subseteq> ?F (Suc i)"
-      using `\<And>i. incseq (F i)`[THEN incseq_SucD] by auto
-  qed
-qed
-
-lemma sets_pair_cancel_measure[simp]:
-  "sets (M1\<lparr>measure := m1\<rparr> \<Otimes>\<^isub>M M2) = sets (M1 \<Otimes>\<^isub>M M2)"
-  "sets (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m2\<rparr>) = sets (M1 \<Otimes>\<^isub>M M2)"
-  unfolding pair_measure_def pair_measure_generator_def sets_sigma
-  by simp_all
-
-lemma measurable_pair_cancel_measure[simp]:
-  "measurable (M1\<lparr>measure := m1\<rparr> \<Otimes>\<^isub>M M2) M = measurable (M1 \<Otimes>\<^isub>M M2) M"
-  "measurable (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m2\<rparr>) M = measurable (M1 \<Otimes>\<^isub>M M2) M"
-  "measurable M (M1\<lparr>measure := m3\<rparr> \<Otimes>\<^isub>M M2) = measurable M (M1 \<Otimes>\<^isub>M M2)"
-  "measurable M (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m4\<rparr>) = measurable M (M1 \<Otimes>\<^isub>M M2)"
-  unfolding measurable_def by (auto simp add: space_pair_measure)
-
-lemma (in product_sigma_finite) product_measure_exists:
-  assumes "finite I"
-  shows "\<exists>\<nu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<nu> \<rparr>) \<and>
-    (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i)))"
-using `finite I` proof induct
-  case empty
-  interpret finite_product_sigma_finite M "{}" by default simp
-  let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> extreal"
-  show ?case
-  proof (intro exI conjI ballI)
-    have "sigma_algebra (sigma G \<lparr>measure := ?\<nu>\<rparr>)"
-      by (rule sigma_algebra_cong) (simp_all add: product_algebra_def)
-    then have "measure_space (sigma G\<lparr>measure := ?\<nu>\<rparr>)"
-      by (rule finite_additivity_sufficient)
-         (simp_all add: positive_def additive_def sets_sigma
-                        product_algebra_generator_def image_constant)
-    then show "sigma_finite_measure (sigma G\<lparr>measure := ?\<nu>\<rparr>)"
-      by (auto intro!: exI[of _ "\<lambda>i. {\<lambda>_. undefined}"]
-               simp: sigma_finite_measure_def sigma_finite_measure_axioms_def
-                     product_algebra_generator_def)
-  qed auto
-next
-  case (insert i I)
-  interpret finite_product_sigma_finite M I by default fact
-  have "finite (insert i I)" using `finite I` by auto
-  interpret I': finite_product_sigma_finite M "insert i I" by default fact
-  from insert obtain \<nu> where
-    prod: "\<And>A. A \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" and
-    "sigma_finite_measure (sigma G\<lparr> measure := \<nu> \<rparr>)" by auto
-  then interpret I: sigma_finite_measure "P\<lparr> measure := \<nu>\<rparr>" unfolding product_algebra_def by simp
-  interpret P: pair_sigma_finite "P\<lparr> measure := \<nu>\<rparr>" "M i" ..
-  let ?h = "(\<lambda>(f, y). f(i := y))"
-  let ?\<nu> = "\<lambda>A. P.\<mu> (?h -` A \<inter> space P.P)"
-  have I': "sigma_algebra (I'.P\<lparr> measure := ?\<nu> \<rparr>)"
-    by (rule I'.sigma_algebra_cong) simp_all
-  interpret I'': measure_space "I'.P\<lparr> measure := ?\<nu> \<rparr>"
-    using measurable_add_dim[OF `i \<notin> I`]
-    by (intro P.measure_space_vimage[OF I']) (auto simp add: measure_preserving_def)
-  show ?case
-  proof (intro exI[of _ ?\<nu>] conjI ballI)
-    let "?m A" = "measure (Pi\<^isub>M I M\<lparr>measure := \<nu>\<rparr> \<Otimes>\<^isub>M M i) (?h -` A \<inter> space P.P)"
-    { fix A assume A: "A \<in> (\<Pi> i\<in>insert i I. sets (M i))"
-      then have *: "?h -` Pi\<^isub>E (insert i I) A \<inter> space P.P = Pi\<^isub>E I A \<times> A i"
-        using `i \<notin> I` M.sets_into_space by (auto simp: space_pair_measure space_product_algebra) blast
-      show "?m (Pi\<^isub>E (insert i I) A) = (\<Prod>i\<in>insert i I. M.\<mu> i (A i))"
-        unfolding * using A
-        apply (subst P.pair_measure_times)
-        using A apply fastsimp
-        using A apply fastsimp
-        using `i \<notin> I` `finite I` prod[of A] A by (auto simp: ac_simps) }
-    note product = this
-    have *: "sigma I'.G\<lparr> measure := ?\<nu> \<rparr> = I'.P\<lparr> measure := ?\<nu> \<rparr>"
-      by (simp add: product_algebra_def)
-    show "sigma_finite_measure (sigma I'.G\<lparr> measure := ?\<nu> \<rparr>)"
-    proof (unfold *, default, simp)
-      from I'.sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
-      then have F: "\<And>j. j \<in> insert i I \<Longrightarrow> range (F j) \<subseteq> sets (M j)"
-        "incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k)"
-        "(\<Union>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) = space I'.G"
-        "\<And>k. \<And>j. j \<in> insert i I \<Longrightarrow> \<mu> j (F j k) \<noteq> \<infinity>"
-        by blast+
-      let "?F k" = "\<Pi>\<^isub>E j \<in> insert i I. F j k"
-      show "\<exists>F::nat \<Rightarrow> ('i \<Rightarrow> 'a) set. range F \<subseteq> sets I'.P \<and>
-          (\<Union>i. F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i)) \<and> (\<forall>i. ?m (F i) \<noteq> \<infinity>)"
-      proof (intro exI[of _ ?F] conjI allI)
-        show "range ?F \<subseteq> sets I'.P" using F(1) by auto
-      next
-        from F(3) show "(\<Union>i. ?F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i))" by simp
-      next
-        fix j
-        have "\<And>k. k \<in> insert i I \<Longrightarrow> 0 \<le> measure (M k) (F k j)"
-          using F(1) by auto
-        with F `finite I` setprod_PInf[of "insert i I", OF this] show "?\<nu> (?F j) \<noteq> \<infinity>"
-          by (subst product) auto
-      qed
-    qed
-  qed
-qed
-
-sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure P
-  unfolding product_algebra_def
-  using product_measure_exists[OF finite_index]
-  by (rule someI2_ex) auto
-
-lemma (in finite_product_sigma_finite) measure_times:
-  assumes "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)"
-  shows "\<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
-  unfolding product_algebra_def
-  using product_measure_exists[OF finite_index]
-  proof (rule someI2_ex, elim conjE)
-    fix \<nu> assume *: "\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
-    have "Pi\<^isub>E I A = Pi\<^isub>E I (\<lambda>i\<in>I. A i)" by (auto dest: Pi_mem)
-    then have "\<nu> (Pi\<^isub>E I A) = \<nu> (Pi\<^isub>E I (\<lambda>i\<in>I. A i))" by simp
-    also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i ((\<lambda>i\<in>I. A i) i))" using assms * by auto
-    finally show "measure (sigma G\<lparr>measure := \<nu>\<rparr>) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
-      by (simp add: sigma_def)
-  qed
-
-lemma (in product_sigma_finite) product_measure_empty[simp]:
-  "measure (Pi\<^isub>M {} M) {\<lambda>x. undefined} = 1"
-proof -
-  interpret finite_product_sigma_finite M "{}" by default auto
-  from measure_times[of "\<lambda>x. {}"] show ?thesis by simp
-qed
-
-lemma (in finite_product_sigma_algebra) P_empty:
-  assumes "I = {}"
-  shows "space P = {\<lambda>k. undefined}" "sets P = { {}, {\<lambda>k. undefined} }"
-  unfolding product_algebra_def product_algebra_generator_def `I = {}`
-  by (simp_all add: sigma_def image_constant)
-
-lemma (in product_sigma_finite) positive_integral_empty:
-  assumes pos: "0 \<le> f (\<lambda>k. undefined)"
-  shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)"
-proof -
-  interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI)
-  have "\<And>A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1"
-    using assms by (subst measure_times) auto
-  then show ?thesis
-    unfolding positive_integral_def simple_function_def simple_integral_def_raw
-  proof (simp add: P_empty, intro antisym)
-    show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined))"
-      by (intro le_SUPI) (auto simp: le_fun_def split: split_max)
-    show "(SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)" using pos
-      by (intro SUP_leI) (auto simp: le_fun_def simp: max_def split: split_if_asm)
-  qed
-qed
-
-lemma (in product_sigma_finite) measure_fold:
-  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
-  assumes A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
-  shows "measure (Pi\<^isub>M (I \<union> J) M) A =
-    measure (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) ((\<lambda>(x,y). merge I x J y) -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M))"
-proof -
-  interpret I: finite_product_sigma_finite M I by default fact
-  interpret J: finite_product_sigma_finite M J by default fact
-  have "finite (I \<union> J)" using fin by auto
-  interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
-  interpret P: pair_sigma_finite I.P J.P by default
-  let ?g = "\<lambda>(x,y). merge I x J y"
-  let "?X S" = "?g -` S \<inter> space P.P"
-  from IJ.sigma_finite_pairs obtain F where
-    F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)"
-       "incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k)"
-       "(\<Union>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) = space IJ.G"
-       "\<And>k. \<forall>i\<in>I\<union>J. \<mu> i (F i k) \<noteq> \<infinity>"
-    by auto
-  let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k"
-  show "IJ.\<mu> A = P.\<mu> (?X A)"
-  proof (rule measure_unique_Int_stable_vimage)
-    show "measure_space IJ.P" "measure_space P.P" by default
-    show "sets (sigma IJ.G) = sets IJ.P" "space IJ.G = space IJ.P" "A \<in> sets (sigma IJ.G)"
-      using A unfolding product_algebra_def by auto
-  next
-    show "Int_stable IJ.G"
-      by (simp add: PiE_Int Int_stable_def product_algebra_def
-                    product_algebra_generator_def)
-          auto
-    show "range ?F \<subseteq> sets IJ.G" using F
-      by (simp add: image_subset_iff product_algebra_def
-                    product_algebra_generator_def)
-    show "incseq ?F" "(\<Union>i. ?F i) = space IJ.G " by fact+
-  next
-    fix k
-    have "\<And>j. j \<in> I \<union> J \<Longrightarrow> 0 \<le> measure (M j) (F j k)"
-      using F(1) by auto
-    with F `finite I` setprod_PInf[of "I \<union> J", OF this]
-    show "IJ.\<mu> (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto
-  next
-    fix A assume "A \<in> sets IJ.G"
-    then obtain F where A: "A = Pi\<^isub>E (I \<union> J) F"
-      and F: "\<And>i. i \<in> I \<union> J \<Longrightarrow> F i \<in> sets (M i)"
-      by (auto simp: product_algebra_generator_def)
-    then have X: "?X A = (Pi\<^isub>E I F \<times> Pi\<^isub>E J F)"
-      using sets_into_space by (auto simp: space_pair_measure) blast+
-    then have "P.\<mu> (?X A) = (\<Prod>i\<in>I. \<mu> i (F i)) * (\<Prod>i\<in>J. \<mu> i (F i))"
-      using `finite J` `finite I` F
-      by (simp add: P.pair_measure_times I.measure_times J.measure_times)
-    also have "\<dots> = (\<Prod>i\<in>I \<union> J. \<mu> i (F i))"
-      using `finite J` `finite I` `I \<inter> J = {}`  by (simp add: setprod_Un_one)
-    also have "\<dots> = IJ.\<mu> A"
-      using `finite J` `finite I` F unfolding A
-      by (intro IJ.measure_times[symmetric]) auto
-    finally show "IJ.\<mu> A = P.\<mu> (?X A)" by (rule sym)
-  qed (rule measurable_merge[OF IJ])
-qed
-
-lemma (in product_sigma_finite) measure_preserving_merge:
-  assumes IJ: "I \<inter> J = {}" and "finite I" "finite J"
-  shows "(\<lambda>(x,y). merge I x J y) \<in> measure_preserving (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
-  by (intro measure_preservingI measurable_merge[OF IJ] measure_fold[symmetric] assms)
-
-lemma (in product_sigma_finite) product_positive_integral_fold:
-  assumes IJ[simp]: "I \<inter> J = {}" "finite I" "finite J"
-  and f: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
-  shows "integral\<^isup>P (Pi\<^isub>M (I \<union> J) M) f =
-    (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (merge I x J y) \<partial>(Pi\<^isub>M J M)) \<partial>(Pi\<^isub>M I M))"
-proof -
-  interpret I: finite_product_sigma_finite M I by default fact
-  interpret J: finite_product_sigma_finite M J by default fact
-  interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
-  interpret IJ: finite_product_sigma_finite M "I \<union> J" by default simp
-  have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable P.P"
-    using measurable_comp[OF measurable_merge[OF IJ(1)] f] by (simp add: comp_def)
-  show ?thesis
-    unfolding P.positive_integral_fst_measurable[OF P_borel, simplified]
-  proof (rule P.positive_integral_vimage)
-    show "sigma_algebra IJ.P" by default
-    show "(\<lambda>(x, y). merge I x J y) \<in> measure_preserving P.P IJ.P"
-      using IJ by (rule measure_preserving_merge)
-    show "f \<in> borel_measurable IJ.P" using f by simp
-  qed
-qed
-
-lemma (in product_sigma_finite) measure_preserving_component_singelton:
-  "(\<lambda>x. x i) \<in> measure_preserving (Pi\<^isub>M {i} M) (M i)"
-proof (intro measure_preservingI measurable_component_singleton)
-  interpret I: finite_product_sigma_finite M "{i}" by default simp
-  fix A let ?P = "(\<lambda>x. x i) -` A \<inter> space I.P"
-  assume A: "A \<in> sets (M i)"
-  then have *: "?P = {i} \<rightarrow>\<^isub>E A" using sets_into_space by auto
-  show "I.\<mu> ?P = M.\<mu> i A" unfolding *
-    using A I.measure_times[of "\<lambda>_. A"] by auto
-qed auto
-
-lemma (in product_sigma_finite) product_positive_integral_singleton:
-  assumes f: "f \<in> borel_measurable (M i)"
-  shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\<lambda>x. f (x i)) = integral\<^isup>P (M i) f"
-proof -
-  interpret I: finite_product_sigma_finite M "{i}" by default simp
-  show ?thesis
-  proof (rule I.positive_integral_vimage[symmetric])
-    show "sigma_algebra (M i)" by (rule sigma_algebras)
-    show "(\<lambda>x. x i) \<in> measure_preserving (Pi\<^isub>M {i} M) (M i)"
-      by (rule measure_preserving_component_singelton)
-    show "f \<in> borel_measurable (M i)" by fact
-  qed
-qed
-
-lemma (in product_sigma_finite) product_positive_integral_insert:
-  assumes [simp]: "finite I" "i \<notin> I"
-    and f: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)"
-  shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^isub>M I M))"
-proof -
-  interpret I: finite_product_sigma_finite M I by default auto
-  interpret i: finite_product_sigma_finite M "{i}" by default auto
-  interpret P: pair_sigma_algebra I.P i.P ..
-  have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I"
-    using f by auto
-  show ?thesis
-    unfolding product_positive_integral_fold[OF IJ, unfolded insert, simplified, OF f]
-  proof (rule I.positive_integral_cong, subst product_positive_integral_singleton)
-    fix x assume x: "x \<in> space I.P"
-    let "?f y" = "f (restrict (x(i := y)) (insert i I))"
-    have f'_eq: "\<And>y. ?f y = f (x(i := y))"
-      using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def)
-    show "?f \<in> borel_measurable (M i)" unfolding f'_eq
-      using measurable_comp[OF measurable_component_update f] x `i \<notin> I`
-      by (simp add: comp_def)
-    show "integral\<^isup>P (M i) ?f = \<integral>\<^isup>+ y. f (x(i:=y)) \<partial>M i"
-      unfolding f'_eq by simp
-  qed
-qed
-
-lemma (in product_sigma_finite) product_positive_integral_setprod:
-  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> extreal"
-  assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
-  and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x"
-  shows "(\<integral>\<^isup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>P (M i) (f i))"
-using assms proof induct
-  case empty
-  interpret finite_product_sigma_finite M "{}" by default auto
-  then show ?case by simp
-next
-  case (insert i I)
-  note `finite I`[intro, simp]
-  interpret I: finite_product_sigma_finite M I by default auto
-  have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
-    using insert by (auto intro!: setprod_cong)
-  have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^isub>M J M)"
-    using sets_into_space insert
-    by (intro sigma_algebra.borel_measurable_extreal_setprod sigma_algebra_product_algebra
-              measurable_comp[OF measurable_component_singleton, unfolded comp_def])
-       auto
-  then show ?case
-    apply (simp add: product_positive_integral_insert[OF insert(1,2) prod])
-    apply (simp add: insert * pos borel setprod_extreal_pos M.positive_integral_multc)
-    apply (subst I.positive_integral_cmult)
-    apply (auto simp add: pos borel insert setprod_extreal_pos positive_integral_positive)
-    done
-qed
-
-lemma (in product_sigma_finite) product_integral_singleton:
-  assumes f: "f \<in> borel_measurable (M i)"
-  shows "(\<integral>x. f (x i) \<partial>Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f"
-proof -
-  interpret I: finite_product_sigma_finite M "{i}" by default simp
-  have *: "(\<lambda>x. extreal (f x)) \<in> borel_measurable (M i)"
-    "(\<lambda>x. extreal (- f x)) \<in> borel_measurable (M i)"
-    using assms by auto
-  show ?thesis
-    unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] ..
-qed
-
-lemma (in product_sigma_finite) product_integral_fold:
-  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
-  and f: "integrable (Pi\<^isub>M (I \<union> J) M) f"
-  shows "integral\<^isup>L (Pi\<^isub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I x J y) \<partial>Pi\<^isub>M J M) \<partial>Pi\<^isub>M I M)"
-proof -
-  interpret I: finite_product_sigma_finite M I by default fact
-  interpret J: finite_product_sigma_finite M J by default fact
-  have "finite (I \<union> J)" using fin by auto
-  interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
-  interpret P: pair_sigma_finite I.P J.P by default
-  let ?M = "\<lambda>(x, y). merge I x J y"
-  let ?f = "\<lambda>x. f (?M x)"
-  show ?thesis
-  proof (subst P.integrable_fst_measurable(2)[of ?f, simplified])
-    have 1: "sigma_algebra IJ.P" by default
-    have 2: "?M \<in> measure_preserving P.P IJ.P" using measure_preserving_merge[OF assms(1,2,3)] .
-    have 3: "integrable (Pi\<^isub>M (I \<union> J) M) f" by fact
-    then have 4: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
-      by (simp add: integrable_def)
-    show "integrable P.P ?f"
-      by (rule P.integrable_vimage[where f=f, OF 1 2 3])
-    show "integral\<^isup>L IJ.P f = integral\<^isup>L P.P ?f"
-      by (rule P.integral_vimage[where f=f, OF 1 2 4])
-  qed
-qed
-
-lemma (in product_sigma_finite) product_integral_insert:
-  assumes [simp]: "finite I" "i \<notin> I"
-    and f: "integrable (Pi\<^isub>M (insert i I) M) f"
-  shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^isub>M I M)"
-proof -
-  interpret I: finite_product_sigma_finite M I by default auto
-  interpret I': finite_product_sigma_finite M "insert i I" by default auto
-  interpret i: finite_product_sigma_finite M "{i}" by default auto
-  interpret P: pair_sigma_finite I.P i.P ..
-  have IJ: "I \<inter> {i} = {}" by auto
-  show ?thesis
-    unfolding product_integral_fold[OF IJ, simplified, OF f]
-  proof (rule I.integral_cong, subst product_integral_singleton)
-    fix x assume x: "x \<in> space I.P"
-    let "?f y" = "f (restrict (x(i := y)) (insert i I))"
-    have f'_eq: "\<And>y. ?f y = f (x(i := y))"
-      using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def)
-    have f: "f \<in> borel_measurable I'.P" using f unfolding integrable_def by auto
-    show "?f \<in> borel_measurable (M i)"
-      unfolding measurable_cong[OF f'_eq]
-      using measurable_comp[OF measurable_component_update f] x `i \<notin> I`
-      by (simp add: comp_def)
-    show "integral\<^isup>L (M i) ?f = integral\<^isup>L (M i) (\<lambda>y. f (x(i := y)))"
-      unfolding f'_eq by simp
-  qed
-qed
-
-lemma (in product_sigma_finite) product_integrable_setprod:
-  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
-  assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
-  shows "integrable (Pi\<^isub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f")
-proof -
-  interpret finite_product_sigma_finite M I by default fact
-  have f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
-    using integrable unfolding integrable_def by auto
-  then have borel: "?f \<in> borel_measurable P"
-    using measurable_comp[OF measurable_component_singleton f]
-    by (auto intro!: borel_measurable_setprod simp: comp_def)
-  moreover have "integrable (Pi\<^isub>M I M) (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)"
-  proof (unfold integrable_def, intro conjI)
-    show "(\<lambda>x. abs (?f x)) \<in> borel_measurable P"
-      using borel by auto
-    have "(\<integral>\<^isup>+x. extreal (abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. extreal (abs (f i (x i)))) \<partial>P)"
-      by (simp add: setprod_extreal abs_setprod)
-    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. extreal (abs (f i x)) \<partial>M i))"
-      using f by (subst product_positive_integral_setprod) auto
-    also have "\<dots> < \<infinity>"
-      using integrable[THEN M.integrable_abs]
-      by (simp add: setprod_PInf integrable_def M.positive_integral_positive)
-    finally show "(\<integral>\<^isup>+x. extreal (abs (?f x)) \<partial>P) \<noteq> \<infinity>" by auto
-    have "(\<integral>\<^isup>+x. extreal (- abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. 0 \<partial>P)"
-      by (intro positive_integral_cong_pos) auto
-    then show "(\<integral>\<^isup>+x. extreal (- abs (?f x)) \<partial>P) \<noteq> \<infinity>" by simp
-  qed
-  ultimately show ?thesis
-    by (rule integrable_abs_iff[THEN iffD1])
-qed
-
-lemma (in product_sigma_finite) product_integral_setprod:
-  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
-  assumes "finite I" "I \<noteq> {}" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
-  shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>L (M i) (f i))"
-using assms proof (induct rule: finite_ne_induct)
-  case (singleton i)
-  then show ?case by (simp add: product_integral_singleton integrable_def)
-next
-  case (insert i I)
-  then have iI: "finite (insert i I)" by auto
-  then have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow>
-    integrable (Pi\<^isub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))"
-    by (intro product_integrable_setprod insert(5)) (auto intro: finite_subset)
-  interpret I: finite_product_sigma_finite M I by default fact
-  have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
-    using `i \<notin> I` by (auto intro!: setprod_cong)
-  show ?case
-    unfolding product_integral_insert[OF insert(1,3) prod[OF subset_refl]]
-    by (simp add: * insert integral_multc I.integral_cmult[OF prod] subset_insertI)
-qed
-
-section "Products on finite spaces"
-
-lemma sigma_sets_pair_measure_generator_finite:
-  assumes "finite A" and "finite B"
-  shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<in> Pow A \<and> b \<in> Pow B} = Pow (A \<times> B)"
-  (is "sigma_sets ?prod ?sets = _")
-proof safe
-  have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
-  fix x assume subset: "x \<subseteq> A \<times> B"
-  hence "finite x" using fin by (rule finite_subset)
-  from this subset show "x \<in> sigma_sets ?prod ?sets"
-  proof (induct x)
-    case empty show ?case by (rule sigma_sets.Empty)
-  next
-    case (insert a x)
-    hence "{a} \<in> sigma_sets ?prod ?sets"
-      by (auto simp: pair_measure_generator_def intro!: sigma_sets.Basic)
-    moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
-    ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)
-  qed
-next
-  fix x a b
-  assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x"
-  from sigma_sets_into_sp[OF _ this(1)] this(2)
-  show "a \<in> A" and "b \<in> B" by auto
-qed
-
-locale pair_finite_sigma_algebra = M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2
-  for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
-
-sublocale pair_finite_sigma_algebra \<subseteq> pair_sigma_algebra by default
-
-lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra:
-  shows "P = \<lparr> space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2), \<dots> = algebra.more P \<rparr>"
-proof -
-  show ?thesis
-    using sigma_sets_pair_measure_generator_finite[OF M1.finite_space M2.finite_space]
-    by (intro algebra.equality) (simp_all add: pair_measure_def pair_measure_generator_def sigma_def)
-qed
-
-sublocale pair_finite_sigma_algebra \<subseteq> finite_sigma_algebra P
-  apply default
-  using M1.finite_space M2.finite_space
-  apply (subst finite_pair_sigma_algebra) apply simp
-  apply (subst (1 2) finite_pair_sigma_algebra) apply simp
-  done
-
-locale pair_finite_space = M1: finite_measure_space M1 + M2: finite_measure_space M2
-  for M1 M2
-
-sublocale pair_finite_space \<subseteq> pair_finite_sigma_algebra
-  by default
-
-sublocale pair_finite_space \<subseteq> pair_sigma_finite
-  by default
-
-lemma (in pair_finite_space) pair_measure_Pair[simp]:
-  assumes "a \<in> space M1" "b \<in> space M2"
-  shows "\<mu> {(a, b)} = M1.\<mu> {a} * M2.\<mu> {b}"
-proof -
-  have "\<mu> ({a}\<times>{b}) = M1.\<mu> {a} * M2.\<mu> {b}"
-    using M1.sets_eq_Pow M2.sets_eq_Pow assms
-    by (subst pair_measure_times) auto
-  then show ?thesis by simp
-qed
-
-lemma (in pair_finite_space) pair_measure_singleton[simp]:
-  assumes "x \<in> space M1 \<times> space M2"
-  shows "\<mu> {x} = M1.\<mu> {fst x} * M2.\<mu> {snd x}"
-  using pair_measure_Pair assms by (cases x) auto
-
-sublocale pair_finite_space \<subseteq> finite_measure_space P
-  by default (auto simp: space_pair_measure)
-
-end
\ No newline at end of file
--- a/src/HOL/Probability/Sigma_Algebra.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Probability/Sigma_Algebra.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -9,7 +9,7 @@
 
 theory Sigma_Algebra
 imports
-  Main
+  Complex_Main
   "~~/src/HOL/Library/Countable"
   "~~/src/HOL/Library/FuncSet"
   "~~/src/HOL/Library/Indicator_Function"
@@ -195,13 +195,17 @@
   ultimately show ?thesis by metis
 qed
 
+lemma ring_of_sets_Pow:
+ "ring_of_sets \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
+  by default auto
+
 lemma algebra_Pow:
-     "algebra \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
-  by (auto simp add: algebra_iff_Un)
+  "algebra \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
+  by default auto
 
 lemma sigma_algebra_Pow:
      "sigma_algebra \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
-  by (auto simp add: sigma_algebra_def sigma_algebra_axioms_def algebra_Pow)
+  by default auto
 
 lemma sigma_algebra_iff:
      "sigma_algebra M \<longleftrightarrow>
@@ -714,6 +718,18 @@
   "range A \<subseteq> sets M \<Longrightarrow> range (disjointed A) \<subseteq> sets M"
   using range_disjointed_sets .
 
+lemma disjointed_0[simp]: "disjointed A 0 = A 0"
+  by (simp add: disjointed_def)
+
+lemma incseq_Un:
+  "incseq A \<Longrightarrow> (\<Union>i\<le>n. A i) = A n"
+  unfolding incseq_def by auto
+
+lemma disjointed_incseq:
+  "incseq A \<Longrightarrow> disjointed A (Suc n) = A (Suc n) - A n"
+  using incseq_Un[of A]
+  by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost)
+
 lemma sigma_algebra_disjoint_iff:
      "sigma_algebra M \<longleftrightarrow>
       algebra M &
--- a/src/HOL/Quickcheck.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Quickcheck.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -164,7 +164,7 @@
    in Predicate.Seq (%u. Predicate.Insert x (iter (nrandom - 1) sz seed')))"
 unfolding iter_def iter'.simps[of _ nrandom] ..
 
-types 'a randompred = "Random.seed \<Rightarrow> ('a Predicate.pred \<times> Random.seed)"
+type_synonym 'a randompred = "Random.seed \<Rightarrow> ('a Predicate.pred \<times> Random.seed)"
 
 definition empty :: "'a randompred"
   where "empty = Pair (bot_class.bot)"
--- a/src/HOL/Quickcheck_Exhaustive.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Quickcheck_Exhaustive.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -355,7 +355,22 @@
 
 end
 
+subsection {* Bounded universal quantifiers *}
 
+class bounded_forall =
+  fixes bounded_forall :: "('a \<Rightarrow> bool) \<Rightarrow> code_numeral \<Rightarrow> bool"
+
+
+instantiation nat :: bounded_forall
+begin
+
+fun bounded_forall_nat :: "(nat => bool) => code_numeral => bool"
+where
+  "bounded_forall P d = ((P 0) \<and> (if d > 0 then bounded_forall (%n. P (Suc n)) (d - 1) else True))"
+
+instance ..
+
+end
 
 subsection {* Defining combinators for any first-order data type *}
 
--- a/src/HOL/Random.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Random.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -25,7 +25,7 @@
 
 subsection {* Random seeds *}
 
-types seed = "code_numeral \<times> code_numeral"
+type_synonym seed = "code_numeral \<times> code_numeral"
 
 primrec "next" :: "seed \<Rightarrow> code_numeral \<times> seed" where
   "next (v, w) = (let
--- a/src/HOL/Random_Sequence.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Random_Sequence.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -5,7 +5,7 @@
 imports Quickcheck DSequence
 begin
 
-types 'a random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a DSequence.dseq \<times> Random.seed)"
+type_synonym 'a random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a DSequence.dseq \<times> Random.seed)"
 
 definition empty :: "'a random_dseq"
 where
--- a/src/HOL/Set.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Set.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -8,7 +8,7 @@
 
 subsection {* Sets as predicates *}
 
-types 'a set = "'a \<Rightarrow> bool"
+type_synonym 'a set = "'a \<Rightarrow> bool"
 
 definition Collect :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set" where -- "comprehension"
   "Collect P = P"
--- a/src/HOL/String.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/String.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -74,7 +74,7 @@
 
 subsection {* Strings *}
 
-types string = "char list"
+type_synonym string = "char list"
 
 syntax
   "_String" :: "xstr => string"    ("_")
--- a/src/HOL/Tools/Predicate_Compile/code_prolog.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Predicate_Compile/code_prolog.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -39,7 +39,7 @@
   val write_program : logic_program -> string
   val run : (Time.time * prolog_system) -> logic_program -> (string * prol_term list) -> string list -> int option -> prol_term list list
   
-  val quickcheck : Proof.context -> term * term list -> int -> term list option * Quickcheck.report option
+  val quickcheck : Proof.context -> (term * term list) list -> int list -> term list option * Quickcheck.report option
 
   val trace : bool Unsynchronized.ref
   
@@ -1009,7 +1009,7 @@
 
 (* FIXME: a small clone of Predicate_Compile_Quickcheck - maybe refactor out commons *)
 
-fun quickcheck ctxt (t, []) size =
+fun quickcheck ctxt [(t, [])] [_, size] =
   let
     val t' = list_abs_free (Term.add_frees t [], t)
     val options = code_options_of (ProofContext.theory_of ctxt)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -22,7 +22,7 @@
   val tracing : bool Unsynchronized.ref;
   val quiet : bool Unsynchronized.ref;
   val quickcheck_compile_term : Predicate_Compile_Aux.compilation -> bool -> bool -> int ->
-    Proof.context -> term * term list -> int -> term list option * Quickcheck.report option;
+    Proof.context -> (term * term list) list -> int list -> term list option * Quickcheck.report option;
 (*  val test_term : Proof.context -> bool -> int -> int -> int -> int -> term -> *)
   val nrandom : int Unsynchronized.ref;
   val debug : bool Unsynchronized.ref;
@@ -215,7 +215,7 @@
   let val ({cpu, ...}, result) = Timing.timing e ()
   in (result, (description, Time.toMilliseconds cpu)) end
 
-fun compile_term compilation options ctxt (t, eval_terms) =
+fun compile_term compilation options ctxt [(t, eval_terms)] =
   let
     val t' = list_abs_free (Term.add_frees t [], t)
     val thy = Theory.copy (ProofContext.theory_of ctxt)
@@ -353,7 +353,7 @@
   let
     val c = compile_term compilation options ctxt t
   in
-    fn size => (try_upto (!quiet) (c size (!nrandom)) depth, NONE)
+    fn [card, size] => (try_upto (!quiet) (c size (!nrandom)) depth, NONE)
   end
 
 fun quickcheck_compile_term compilation function_flattening fail_safe_function_flattening depth =
--- a/src/HOL/Tools/Quickcheck/exhaustive_generators.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Quickcheck/exhaustive_generators.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -7,15 +7,16 @@
 signature EXHAUSTIVE_GENERATORS =
 sig
   val compile_generator_expr:
-    Proof.context -> term * term list -> int -> term list option * Quickcheck.report option
-  val compile_generator_exprs:
-    Proof.context -> term list -> (int -> term list option) list
-  val put_counterexample: (unit -> int -> term list option)
+    Proof.context -> (term * term list) list -> int list -> term list option * Quickcheck.report option
+  val compile_generator_exprs: Proof.context -> term list -> (int -> term list option) list
+  val compile_validator_exprs: Proof.context -> term list -> (int -> bool) list
+  val put_counterexample: (unit -> int -> int -> term list option)
     -> Proof.context -> Proof.context
   val put_counterexample_batch: (unit -> (int -> term list option) list)
     -> Proof.context -> Proof.context
-  val smart_quantifier : bool Config.T;
-  val quickcheck_pretty : bool Config.T;
+  val put_validator_batch: (unit -> (int -> bool) list) -> Proof.context -> Proof.context
+  val smart_quantifier : bool Config.T
+  val quickcheck_pretty : bool Config.T
   val setup: theory -> theory
 end;
 
@@ -239,8 +240,6 @@
       @{term "Quickcheck_Exhaustive.catch_match :: term list option => term list option => term list option"} $
         (@{term "If :: bool => term list option => term list option => term list option"}
         $ cond $ then_t $ else_t) $ none_t;
-    fun strip_imp (Const(@{const_name HOL.implies},_) $ A $ B) = apfst (cons A) (strip_imp B)
-        | strip_imp A = ([], A)
     fun lookup v = the (AList.lookup (op =) (names ~~ (frees ~~ term_vars)) v)
     fun mk_naive_test_term t =
       fold_rev mk_exhaustive_closure (frees ~~ term_vars) (mk_safe_if (t, none_t, return)) 
@@ -256,7 +255,7 @@
       end
     fun mk_smart_test_term t =
       let
-        val (assms, concl) = strip_imp t
+        val (assms, concl) = Quickcheck_Common.strip_imp t
       in
         mk_smart_test_term' concl [] assms
       end
@@ -264,11 +263,50 @@
       if Config.get ctxt smart_quantifier then mk_smart_test_term else mk_naive_test_term 
   in lambda depth (mk_test_term t) end
 
+val mk_parametric_generator_expr =
+  Quickcheck_Common.gen_mk_parametric_generator_expr 
+    ((mk_generator_expr, absdummy (@{typ "code_numeral"}, @{term "None :: term list option"})),
+      @{typ "code_numeral => term list option"})
+
+fun mk_validator_expr ctxt t =
+  let
+    fun bounded_forallT T = (T --> @{typ bool}) --> @{typ code_numeral} --> @{typ bool}
+    val thy = ProofContext.theory_of ctxt
+    val ctxt' = Variable.auto_fixes t ctxt
+    val ([depth_name], ctxt'') = Variable.variant_fixes ["depth"] ctxt'
+    val depth = Free (depth_name, @{typ code_numeral})
+    fun mk_bounded_forall (s, T) t =
+      Const (@{const_name "Quickcheck_Exhaustive.bounded_forall_class.bounded_forall"}, bounded_forallT T)
+        $ lambda (Free (s, T)) t $ depth
+    fun mk_implies (cond, then_t) =
+      @{term "If :: bool => bool => bool => bool"} $ cond $ then_t $ @{term False}
+    fun mk_naive_test_term t = fold_rev mk_bounded_forall (Term.add_frees t []) t 
+    fun mk_smart_test_term' concl bound_frees assms =
+      let
+        fun vars_of t = subtract (op =) bound_frees (Term.add_frees t [])
+        val (vars, check) =
+          case assms of [] => (vars_of concl, concl)
+            | assm :: assms => (vars_of assm, mk_implies (assm,
+                mk_smart_test_term' concl (union (op =) (vars_of assm) bound_frees) assms))
+      in
+        fold_rev mk_bounded_forall vars check
+      end
+    fun mk_smart_test_term t =
+      let
+        val (assms, concl) = Quickcheck_Common.strip_imp t
+      in
+        mk_smart_test_term' concl [] assms
+      end
+    val mk_test_term =
+      if Config.get ctxt smart_quantifier then mk_smart_test_term else mk_naive_test_term 
+  in lambda depth (mk_test_term t) end
+
+
 (** generator compiliation **)
 
 structure Counterexample = Proof_Data
 (
-  type T = unit -> int -> term list option
+  type T = unit -> int -> int -> term list option
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Counterexample"
 );
@@ -282,21 +320,31 @@
 );
 val put_counterexample_batch = Counterexample_Batch.put;
 
-val target = "Quickcheck";
+structure Validator_Batch = Proof_Data
+(
+  type T = unit -> (int -> bool) list
+  (* FIXME avoid user error with non-user text *)
+  fun init _ () = error "Counterexample"
+);
+val put_validator_batch = Validator_Batch.put;
 
-fun compile_generator_expr ctxt (t, eval_terms) =
+
+val target = "Quickcheck";
+  
+fun compile_generator_expr ctxt ts =
   let
     val thy = ProofContext.theory_of ctxt
-    val t' = mk_generator_expr ctxt (t, eval_terms);
+    val t' = mk_parametric_generator_expr ctxt ts;
     val compile = Code_Runtime.dynamic_value_strict
       (Counterexample.get, put_counterexample, "Exhaustive_Generators.put_counterexample")
-      thy (SOME target) (fn proc => fn g => g #> (Option.map o map) proc) t' [];
+      thy (SOME target) (fn proc => fn g =>
+        fn card => fn size => g card size |> (Option.map o map) proc) t' []
   in
-    fn size => rpair NONE (compile size |> 
+    fn [card, size] => rpair NONE (compile card size |> 
       (if Config.get ctxt quickcheck_pretty then
         Option.map (map Quickcheck_Common.post_process_term) else I))
   end;
-
+ 
 fun compile_generator_exprs ctxt ts =
   let
     val thy = ProofContext.theory_of ctxt
@@ -305,13 +353,22 @@
       (Counterexample_Batch.get, put_counterexample_batch,
         "Exhaustive_Generators.put_counterexample_batch")
       thy (SOME target) (fn proc => map (fn g => g #> (Option.map o map) proc))
-      (HOLogic.mk_list @{typ "code_numeral => term list option"} ts') [];
+      (HOLogic.mk_list @{typ "code_numeral => term list option"} ts') []
   in
     map (fn compile => fn size => compile size
       |> Option.map (map Quickcheck_Common.post_process_term)) compiles
   end;
 
-  
+fun compile_validator_exprs ctxt ts =
+  let
+    val thy = ProofContext.theory_of ctxt
+    val ts' = map (mk_validator_expr ctxt) ts;
+  in
+    Code_Runtime.dynamic_value_strict
+      (Validator_Batch.get, put_validator_batch, "Exhaustive_Generators.put_validator_batch")
+      thy (SOME target) (K I) (HOLogic.mk_list @{typ "code_numeral => bool"} ts') []
+  end;
+
 (** setup **)
 
 val setup =
@@ -320,6 +377,7 @@
   #> setup_smart_quantifier
   #> setup_quickcheck_pretty
   #> Context.theory_map (Quickcheck.add_generator ("exhaustive", compile_generator_expr))
-  #> Context.theory_map (Quickcheck.add_batch_generator ("exhaustive", compile_generator_exprs));
+  #> Context.theory_map (Quickcheck.add_batch_generator ("exhaustive", compile_generator_exprs))
+  #> Context.theory_map (Quickcheck.add_batch_validator ("exhaustive", compile_validator_exprs));
 
 end;
--- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -7,7 +7,7 @@
 signature NARROWING_GENERATORS =
 sig
   val compile_generator_expr:
-    Proof.context -> term * term list -> int -> term list option * Quickcheck.report option
+    Proof.context -> (term * term list) list -> int list -> term list option * Quickcheck.report option
   val put_counterexample: (unit -> term list option) -> Proof.context -> Proof.context
   val finite_functions : bool Config.T
   val setup: theory -> theory
@@ -212,7 +212,7 @@
     list_abs (names ~~ Ts', t'')
   end
 
-fun compile_generator_expr ctxt (t, eval_terms) size =
+fun compile_generator_expr ctxt [(t, eval_terms)] [_, size] =
   let
     val thy = ProofContext.theory_of ctxt
     val t' = list_abs_free (Term.add_frees t [], t)
--- a/src/HOL/Tools/Quickcheck/quickcheck_common.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Quickcheck/quickcheck_common.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -6,18 +6,26 @@
 
 signature QUICKCHECK_COMMON =
 sig
+  val strip_imp : term -> (term list * term)
   val perhaps_constrain: theory -> (typ * sort) list -> (string * sort) list
     -> (string * sort -> string * sort) option
   val ensure_sort_datatype:
     sort * (Datatype.config -> Datatype.descr -> (string * sort) list -> string list -> string ->
       string list * string list -> typ list * typ list -> theory -> theory)
     -> Datatype.config -> string list -> theory -> theory
+  val gen_mk_parametric_generator_expr :
+   (((Proof.context -> term * term list -> term) * term) * typ) -> Proof.context -> (term * term list) list -> term
   val post_process_term : term -> term
 end;
 
 structure Quickcheck_Common : QUICKCHECK_COMMON =
 struct
 
+(* HOLogic's term functions *)
+
+fun strip_imp (Const(@{const_name HOL.implies},_) $ A $ B) = apfst (cons A) (strip_imp B)
+  | strip_imp A = ([], A)
+
 (** ensuring sort constraints **)
 
 fun perhaps_constrain thy insts raw_vs =
@@ -53,6 +61,18 @@
       | NONE => thy
   end;
   
+(** generic parametric compilation **)
+
+fun gen_mk_parametric_generator_expr ((mk_generator_expr, out_of_bounds), T) ctxt ts =
+  let
+    val if_t = Const (@{const_name "If"}, @{typ bool} --> T --> T --> T)
+    fun mk_if (index, (t, eval_terms)) else_t =
+      if_t $ (HOLogic.eq_const @{typ code_numeral} $ Bound 0 $ HOLogic.mk_number @{typ code_numeral} index) $
+        (mk_generator_expr ctxt (t, eval_terms)) $ else_t
+  in
+    absdummy (@{typ "code_numeral"}, fold_rev mk_if (1 upto (length ts) ~~ ts) out_of_bounds)
+  end
+
 (** post-processing of function terms **)
 
 fun dest_fun_upd (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) = (t0, (t1, t2))
--- a/src/HOL/Tools/Quickcheck/random_generators.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Quickcheck/random_generators.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -11,10 +11,10 @@
     -> (seed -> ('b * (unit -> term)) * seed) -> (seed -> seed * seed)
     -> seed -> (('a -> 'b) * (unit -> term)) * seed
   val compile_generator_expr:
-    Proof.context -> term * term list -> int -> term list option * Quickcheck.report option
-  val put_counterexample: (unit -> int -> seed -> term list option * seed)
+    Proof.context -> (term * term list) list -> int list -> term list option * Quickcheck.report option
+  val put_counterexample: (unit -> int -> int -> seed -> term list option * seed)
     -> Proof.context -> Proof.context
-  val put_counterexample_report: (unit -> int -> seed -> (term list option * (bool list * bool)) * seed)
+  val put_counterexample_report: (unit -> int -> int -> seed -> (term list option * (bool list * bool)) * seed)
     -> Proof.context -> Proof.context
   val setup: theory -> theory
 end;
@@ -272,7 +272,7 @@
 
 structure Counterexample = Proof_Data
 (
-  type T = unit -> int -> int * int -> term list option * (int * int)
+  type T = unit -> int -> int -> int * int -> term list option * (int * int)
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Counterexample"
 );
@@ -280,7 +280,7 @@
 
 structure Counterexample_Report = Proof_Data
 (
-  type T = unit -> int -> seed -> (term list option * (bool list * bool)) * seed
+  type T = unit -> int -> int -> seed -> (term list option * (bool list * bool)) * seed
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Counterexample_Report"
 );
@@ -288,8 +288,11 @@
 
 val target = "Quickcheck";
 
-fun mk_generator_expr thy prop Ts =
-  let
+fun mk_generator_expr ctxt (t, eval_terms) =
+  let  
+    val thy = ProofContext.theory_of ctxt
+    val prop = list_abs_free (Term.add_frees t [], t)
+    val Ts = (map snd o fst o strip_abs) prop
     val bound_max = length Ts - 1;
     val bounds = map_index (fn (i, ty) =>
       (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts;
@@ -311,16 +314,17 @@
       (Sign.mk_const thy (@{const_name Quickcheck.random}, [T]) $ Bound i);
   in Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check)) end;
 
-fun mk_reporting_generator_expr thy prop Ts =
+fun mk_reporting_generator_expr ctxt (t, eval_terms) =
   let
-    val bound_max = length Ts - 1;
+    val thy = ProofContext.theory_of ctxt
+    val prop = list_abs_free (Term.add_frees t [], t)
+    val Ts = (map snd o fst o strip_abs) prop
+    val bound_max = length Ts - 1
     val bounds = map_index (fn (i, ty) =>
       (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts;
-    fun strip_imp (Const(@{const_name HOL.implies},_) $ A $ B) = apfst (cons A) (strip_imp B)
-      | strip_imp A = ([], A)
     val prop' = betapplys (prop, map (fn (i, _, _, _) => Bound i) bounds);
     val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds)
-    val (assms, concl) = strip_imp prop'
+    val (assms, concl) = Quickcheck_Common.strip_imp prop'
     val return =
       @{term "Pair :: term list option * (bool list * bool) => Random.seed => (term list option * (bool list * bool)) * Random.seed"};
     fun mk_assms_report i =
@@ -354,6 +358,18 @@
     Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check))
   end
 
+val mk_parametric_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr 
+  ((mk_generator_expr, 
+    absdummy (@{typ code_numeral}, @{term "Pair None :: Random.seed => term list option * Random.seed"})),
+    @{typ "code_numeral => Random.seed => term list option * Random.seed"})
+
+val mk_parametric_reporting_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr 
+  ((mk_reporting_generator_expr,
+    absdummy (@{typ code_numeral},
+      @{term "Pair (None, ([], False)) :: Random.seed => (term list option * (bool list * bool)) * Random.seed"})),
+    @{typ "code_numeral => Random.seed => (term list option * (bool list * bool)) * Random.seed"})
+    
+    
 (* single quickcheck report *)
 
 datatype single_report = Run of bool list * bool | MatchExc
@@ -375,52 +391,50 @@
 val empty_report = Quickcheck.Report { iterations = 0, raised_match_errors = 0,
   satisfied_assms = [], positive_concl_tests = 0 }
     
-fun compile_generator_expr ctxt (t, eval_terms) =
+fun compile_generator_expr ctxt ts =
   let
-    val t' = list_abs_free (Term.add_frees t [], t)
-    val Ts = (map snd o fst o strip_abs) t';
     val thy = ProofContext.theory_of ctxt
     val iterations = Config.get ctxt Quickcheck.iterations
   in
     if Config.get ctxt Quickcheck.report then
       let
-        val t'' = mk_reporting_generator_expr thy t' Ts;
+        val t' = mk_parametric_reporting_generator_expr ctxt ts;
         val compile = Code_Runtime.dynamic_value_strict
           (Counterexample_Report.get, put_counterexample_report, "Random_Generators.put_counterexample_report")
-          thy (SOME target) (fn proc => fn g => fn s => g s #>> (apfst o Option.map o map) proc) t'' [];
-        val single_tester = compile #> Random_Engine.run
-        fun iterate_and_collect size 0 report = (NONE, report)
-          | iterate_and_collect size j report =
+          thy (SOME target) (fn proc => fn g => fn c => fn s => g c s #>> (apfst o Option.map o map) proc) t' [];
+        fun single_tester c s = compile c s |> Random_Engine.run
+        fun iterate_and_collect (card, size) 0 report = (NONE, report)
+          | iterate_and_collect (card, size) j report =
             let
-              val (test_result, single_report) = apsnd Run (single_tester size) handle Match => 
+              val (test_result, single_report) = apsnd Run (single_tester card size) handle Match => 
                 (if Config.get ctxt Quickcheck.quiet then ()
                  else warning "Exception Match raised during quickcheck"; (NONE, MatchExc))
               val report = collect_single_report single_report report
             in
-              case test_result of NONE => iterate_and_collect size (j - 1) report
+              case test_result of NONE => iterate_and_collect (card, size) (j - 1) report
                 | SOME q => (SOME q, report)
             end
       in
-        fn size => apsnd SOME (iterate_and_collect size iterations empty_report)
+        fn [card, size] => apsnd SOME (iterate_and_collect (card, size) iterations empty_report)
       end
     else
       let
-        val t'' = mk_generator_expr thy t' Ts;
+        val t' = mk_parametric_generator_expr ctxt ts;
         val compile = Code_Runtime.dynamic_value_strict
           (Counterexample.get, put_counterexample, "Random_Generators.put_counterexample")
-          thy (SOME target) (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) t'' [];
-        val single_tester = compile #> Random_Engine.run
-        fun iterate size 0 = NONE
-          | iterate size j =
+          thy (SOME target) (fn proc => fn g => fn c => fn s => g c s #>> (Option.map o map) proc) t' [];
+        fun single_tester c s = compile c s |> Random_Engine.run
+        fun iterate (card, size) 0 = NONE
+          | iterate (card, size) j =
             let
-              val result = single_tester size handle Match => 
+              val result = single_tester card size handle Match => 
                 (if Config.get ctxt Quickcheck.quiet then ()
                  else warning "Exception Match raised during quickcheck"; NONE)
             in
-              case result of NONE => iterate size (j - 1) | SOME q => SOME q
+              case result of NONE => iterate (card, size) (j - 1) | SOME q => SOME q
             end
       in
-        fn size => (rpair NONE (iterate size iterations))
+        fn [card, size] => (rpair NONE (iterate (card, size) iterations))
       end
   end;
 
--- a/src/HOL/Tools/SMT/smt_config.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/SMT/smt_config.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -33,6 +33,7 @@
   val trace_used_facts: bool Config.T
   val monomorph_limit: int Config.T
   val monomorph_instances: int Config.T
+  val monomorph_full: bool Config.T
   val infer_triggers: bool Config.T
   val drop_bad_facts: bool Config.T
   val filter_only_facts: bool Config.T
@@ -181,6 +182,10 @@
 val (monomorph_instances, setup_monomorph_instances) =
   Attrib.config_int monomorph_instancesN (K 500)
 
+val monomorph_fullN = "smt_monomorph_full"
+val (monomorph_full, setup_monomorph_full) =
+  Attrib.config_bool monomorph_fullN (K true)
+
 val infer_triggersN = "smt_infer_triggers"
 val (infer_triggers, setup_infer_triggers) =
   Attrib.config_bool infer_triggersN (K false)
@@ -207,6 +212,7 @@
   setup_verbose #>
   setup_monomorph_limit #>
   setup_monomorph_instances #>
+  setup_monomorph_full #>
   setup_infer_triggers #>
   setup_drop_bad_facts #>
   setup_filter_only_facts #>
--- a/src/HOL/Tools/SMT/smt_monomorph.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/SMT/smt_monomorph.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -28,8 +28,9 @@
 
 signature SMT_MONOMORPH =
 sig
-  val monomorph: (int * thm) list -> Proof.context ->
-    (int * thm) list * Proof.context
+  val typ_has_tvars: typ -> bool
+  val monomorph: ('a * thm) list -> Proof.context ->
+    ('a * thm) list * Proof.context
 end
 
 structure SMT_Monomorph: SMT_MONOMORPH =
@@ -159,31 +160,38 @@
 
   in most_specific [] end
 
-fun instantiate (i, thm) substs (ithms, ctxt) =
+fun instantiate full (i, thm) substs (ithms, ctxt) =
   let
+    val thy = ProofContext.theory_of ctxt
+
     val (vs, Ss) = split_list (Term.add_tvars (Thm.prop_of thm) [])
     val (Tenv, ctxt') =
       ctxt
       |> Variable.invent_types Ss
       |>> map2 (fn v => fn (n, S) => (v, (S, TFree (n, S)))) vs
 
-    val thy = ProofContext.theory_of ctxt'
+    exception PARTIAL_INST of unit
+
+    fun update_subst vT subst =
+      if full then Vartab.update vT subst
+      else raise PARTIAL_INST ()
 
     fun replace (v, (_, T)) (U as TVar (u, _)) = if u = v then T else U
       | replace _ T = T
 
     fun complete (vT as (v, _)) subst =
       subst
-      |> not (Vartab.defined subst v) ? Vartab.update vT
+      |> not (Vartab.defined subst v) ? update_subst vT
       |> Vartab.map (K (apsnd (Term.map_atyps (replace vT))))
 
     fun cert (ix, (S, T)) = pairself (Thm.ctyp_of thy) (TVar (ix, S), T)
 
     fun inst subst =
       let val cTs = Vartab.fold (cons o cert) (fold complete Tenv subst) []
-      in (i, Thm.instantiate (cTs, []) thm) end
+      in SOME (i, Thm.instantiate (cTs, []) thm) end
+      handle PARTIAL_INST () => NONE
 
-  in (map inst substs @ ithms, ctxt') end
+  in (map_filter inst substs @ ithms, if full then ctxt' else ctxt) end
 
 
 
@@ -203,12 +211,13 @@
 
     val limit = Config.get ctxt SMT_Config.monomorph_limit
     val instances = Config.get ctxt SMT_Config.monomorph_instances
+    val full = Config.get ctxt SMT_Config.monomorph_full
   in
     scss
     |> search_substitutions ctxt limit instances Symtab.empty grounds
     |> map (filter_most_specific (ProofContext.theory_of ctxt))
     |> rpair (monos, ctxt)
-    |-> fold2 instantiate polys
+    |-> fold2 (instantiate full) polys
   end
 
 fun monomorph irules ctxt =
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -596,76 +596,44 @@
 end
 
 
-(* c = SOME x. P x |- (EX x. P x) = P c
-   c = SOME x. ~ P x |- ~(ALL x. P x) = ~ P c *)
+(* |- (EX x. P x) = P c     |- ~(ALL x. P x) = ~ P c *)
 local
-  val elim_ex = @{lemma "EX x. P == P" by simp}
-  val elim_all = @{lemma "~ (ALL x. P) == ~P" by simp}
-  val sk_ex = @{lemma "c == SOME x. P x ==> EX x. P x == P c"
-    by simp (intro eq_reflection some_eq_ex[symmetric])}
-  val sk_all = @{lemma "c == SOME x. ~ P x ==> ~(ALL x. P x) == ~ P c"
-    by (simp only: not_all) (intro eq_reflection some_eq_ex[symmetric])}
-  val sk_ex_rule = ((sk_ex, I), elim_ex)
-  and sk_all_rule = ((sk_all, Thm.dest_arg), elim_all)
+  val forall =
+    SMT_Utils.mk_const_pat @{theory} @{const_name all}
+      (SMT_Utils.destT1 o SMT_Utils.destT1)
+  fun mk_forall cv ct =
+    Thm.capply (SMT_Utils.instT' cv forall) (Thm.cabs cv ct)
 
-  fun dest f sk_rule = 
-    Thm.dest_comb (f (Thm.dest_arg (Thm.dest_arg (Thm.cprop_of sk_rule))))
-  fun type_of f sk_rule = Thm.ctyp_of_term (snd (dest f sk_rule))
-  fun pair2 (a, b) (c, d) = [(a, c), (b, d)]
-  fun inst_sk (sk_rule, f) p c =
-    Thm.instantiate ([(type_of f sk_rule, Thm.ctyp_of_term c)], []) sk_rule
-    |> (fn sk' => Thm.instantiate ([], (pair2 (dest f sk') (p, c))) sk')
-    |> Conv.fconv_rule (Thm.beta_conversion true)
-
-  fun kind (Const (@{const_name Ex}, _) $ _) = (sk_ex_rule, I, I)
-    | kind (@{const Not} $ (Const (@{const_name All}, _) $ _)) =
-        (sk_all_rule, Thm.dest_arg, Z3_Proof_Literals.negate)
-    | kind t = raise TERM ("skolemize", [t])
-
-  fun dest_abs_type (Abs (_, T, _)) = T
-    | dest_abs_type t = raise TERM ("dest_abs_type", [t])
-
-  fun bodies_of thy lhs rhs =
-    let
-      val (rule, dest, make) = kind (Thm.term_of lhs)
+  fun get_vars f mk pred ctxt t =
+    Term.fold_aterms f t []
+    |> map_filter (fn v =>
+         if pred v then SOME (SMT_Utils.certify ctxt (mk v)) else NONE)
 
-      fun dest_body idx cbs ct =
-        let
-          val cb = Thm.dest_arg (dest ct)
-          val T = dest_abs_type (Thm.term_of cb)
-          val cv = Thm.cterm_of thy (Var (("x", idx), T))
-          val cu = make (Drule.beta_conv cb cv)
-          val cbs' = (cv, cb) :: cbs
-        in
-          (snd (Thm.first_order_match (cu, rhs)), rev cbs')
-          handle Pattern.MATCH => dest_body (idx+1) cbs' cu
-        end
-    in (rule, dest_body 1 [] lhs) end
-
-  fun transitive f thm = Thm.transitive thm (f (Thm.rhs_of thm))
+  fun close vars f ct ctxt =
+    let
+      val frees_of = get_vars Term.add_frees Free (member (op =) vars o fst)
+      val vs = frees_of ctxt (Thm.term_of ct)
+      val (thm, ctxt') = f (fold_rev mk_forall vs ct) ctxt
+      val vars_of = get_vars Term.add_vars Var (K true) ctxt'
+    in (Thm.instantiate ([], vars_of (Thm.prop_of thm) ~~ vs) thm, ctxt') end
 
-  fun sk_step (rule, elim) (cv, mct, cb) ((is, thm), ctxt) =
-    (case mct of
-      SOME ct =>
-        ctxt
-        |> Z3_Proof_Tools.make_hyp_def
-             (inst_sk rule (Thm.instantiate_cterm ([], is) cb) ct)
-        |>> pair ((cv, ct) :: is) o Thm.transitive thm
-    | NONE => ((is, transitive (Conv.rewr_conv elim) thm), ctxt))
+  val sk_rules = @{lemma
+    "(EX x. P x) = P (SOME x. P x)"   "(~(ALL x. P x)) = (~P (SOME x. ~P x))"
+    by (metis someI_ex)+}
 in
-fun skolemize ct ctxt =
-  let
-    val (lhs, rhs) = Thm.dest_binop (Thm.dest_arg ct)
-    val (rule, (ctab, cbs)) = bodies_of (ProofContext.theory_of ctxt) lhs rhs
-    fun lookup_var (cv, cb) = (cv, AList.lookup (op aconvc) ctab cv, cb)
-  in
-    (([], Thm.reflexive lhs), ctxt)
-    |> fold (sk_step rule) (map lookup_var cbs)
-    |>> MetaEq o snd
-  end
+
+fun skolemize vars =
+  apfst Thm oo close vars (yield_singleton Assumption.add_assumes)
+
+fun discharge_sk_tac i st = (
+  Tactic.rtac @{thm trans}
+  THEN' Tactic.resolve_tac sk_rules
+  THEN' (Tactic.rtac @{thm refl} ORELSE' discharge_sk_tac)) i st
+
 end
 
 
+
 (** theory proof rules **)
 
 (* theory lemmas: linear arithmetic, arrays *)
@@ -816,7 +784,7 @@
     | (Z3_Proof_Parser.Elim_Unused_Vars, _) => (elim_unused_vars cx ct, cxp)
     | (Z3_Proof_Parser.Dest_Eq_Res, _) => (dest_eq_res cx ct, cxp)
     | (Z3_Proof_Parser.Quant_Inst, _) => (quant_inst ct, cxp)
-    | (Z3_Proof_Parser.Skolemize, _) => skolemize ct cx ||> rpair ptab
+    | (Z3_Proof_Parser.Skolemize, _) => skolemize vars ct cx ||> rpair ptab
 
       (* theory rules *)
     | (Z3_Proof_Parser.Th_Lemma _, _) =>  (* FIXME: use arguments *)
@@ -850,21 +818,23 @@
         |> tap (check_after idx r ps prop)
     in (thm, (ctxt', Inttab.update (idx, thm) ptab')) end
 
-  val disch_rules = [@{thm allI}, @{thm refl}, @{thm reflexive},
-    Z3_Proof_Literals.true_thm]
-  fun all_disch_rules rules = map (pair false) (disch_rules @ rules)
+  fun make_discharge_rules rules = rules @ [@{thm allI}, @{thm refl},
+    @{thm reflexive}, Z3_Proof_Literals.true_thm]
+
+  fun discharge_tac rules =
+    Tactic.resolve_tac rules ORELSE' SOLVED' discharge_sk_tac
 
-  fun disch_assm rules thm =
-    if Thm.nprems_of thm = 0 then Drule.flexflex_unique thm
+  fun discharge_assms rules thm =
+    if Thm.nprems_of thm = 0 then Goal.norm_result thm
     else
-      (case Seq.pull (Thm.biresolution false rules 1 thm) of
-        SOME (thm', _) => disch_assm rules thm'
+      (case Seq.pull (discharge_tac rules 1 thm) of
+        SOME (thm', _) => discharge_assms rules thm'
       | NONE => raise THM ("failed to discharge premise", 1, [thm]))
 
   fun discharge rules outer_ctxt (p, (inner_ctxt, _)) =
     thm_of p
     |> singleton (ProofContext.export inner_ctxt outer_ctxt)
-    |> disch_assm rules
+    |> discharge_assms (make_discharge_rules rules)
 in
 
 fun reconstruct outer_ctxt recon output =
@@ -882,7 +852,7 @@
     else
       (Thm @{thm TrueI}, cxp)
       |> fold (prove simpset vars) steps 
-      |> discharge (all_disch_rules rules) outer_ctxt
+      |> discharge rules outer_ctxt
       |> pair []
   end
 
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -88,6 +88,8 @@
   #> raw_explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation
   #> fst
 
+val unprefix_fact_number = space_implode "_" o tl o space_explode "_"
+
 fun repair_conjecture_shape_and_fact_names output conjecture_shape fact_names =
   if String.isSubstring set_ClauseFormulaRelationN output then
     let
@@ -100,7 +102,8 @@
         |> map (fn s => find_index (curry (op =) s) seq + 1)
       fun names_for_number j =
         j |> AList.lookup (op =) name_map |> these
-          |> map_filter (try (unprefix fact_prefix)) |> map unascii_of
+          |> map_filter (try (unascii_of o unprefix_fact_number
+                              o unprefix fact_prefix))
           |> map (fn name =>
                      (name, name |> find_first_in_list_vector fact_names |> the)
                      handle Option.Option =>
@@ -145,16 +148,19 @@
       "\nTo minimize the number of lemmas, try this: " ^
       Markup.markup Markup.sendback command ^ "."
 
-val vampire_fact_prefix = "f" (* grrr... *)
+val vampire_step_prefix = "f" (* grrr... *)
 
 fun resolve_fact fact_names ((_, SOME s)) =
-    (case strip_prefix_and_unascii fact_prefix s of
-       SOME s' => (case find_first_in_list_vector fact_names s' of
-                     SOME x => [(s', x)]
-                   | NONE => [])
+    (case try (unprefix fact_prefix) s of
+       SOME s' =>
+       let val s' = s' |> unprefix_fact_number |> unascii_of in
+         case find_first_in_list_vector fact_names s' of
+           SOME x => [(s', x)]
+         | NONE => []
+       end
      | NONE => [])
   | resolve_fact fact_names (num, NONE) =
-    case Int.fromString (perhaps (try (unprefix vampire_fact_prefix)) num) of
+    case Int.fromString (perhaps (try (unprefix vampire_step_prefix)) num) of
       SOME j =>
       if j > 0 andalso j <= Vector.length fact_names then
         Vector.sub (fact_names, j - 1)
@@ -233,7 +239,7 @@
 
 val raw_prefix = "X"
 val assum_prefix = "A"
-val fact_prefix = "F"
+val have_prefix = "F"
 
 fun resolve_conjecture conjecture_shape (num, s_opt) =
   let
@@ -847,7 +853,7 @@
               (l, subst, next_fact)
             else
               let
-                val l' = (prefix_for_depth depth fact_prefix, next_fact)
+                val l' = (prefix_for_depth depth have_prefix, next_fact)
               in (l', (l, l') :: subst, next_fact + 1) end
           val relabel_facts =
             apfst (maps (the_list o AList.lookup (op =) subst))
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -442,9 +442,13 @@
                 (atp_type_literals_for_types type_sys ctypes_sorts))
            (formula_for_combformula ctxt type_sys combformula)
 
-fun problem_line_for_fact ctxt prefix type_sys (formula as {name, kind, ...}) =
-  Fof (prefix ^ ascii_of name, kind, formula_for_fact ctxt type_sys formula,
-       NONE)
+(* Each fact is given a unique fact number to avoid name clashes (e.g., because
+   of monomorphization). The TPTP explicitly forbids name clashes, and some of
+   the remote provers might care. *)
+fun problem_line_for_fact ctxt prefix type_sys
+                          (j, formula as {name, kind, ...}) =
+  Fof (prefix ^ string_of_int j ^ "_" ^ ascii_of name, kind,
+       formula_for_fact ctxt type_sys formula, NONE)
 
 fun problem_line_for_class_rel_clause (ClassRelClause {name, subclass,
                                                        superclass, ...}) =
@@ -626,7 +630,8 @@
     (* Reordering these might or might not confuse the proof reconstruction
        code or the SPASS Flotter hack. *)
     val problem =
-      [(factsN, map (problem_line_for_fact ctxt fact_prefix type_sys) facts),
+      [(factsN, map (problem_line_for_fact ctxt fact_prefix type_sys)
+                    (0 upto length facts - 1 ~~ facts)),
        (class_relsN, map problem_line_for_class_rel_clause class_rel_clauses),
        (aritiesN, map problem_line_for_arity_clause arity_clauses),
        (helpersN, []),
@@ -638,7 +643,8 @@
     val helper_lines =
       get_helper_facts ctxt explicit_forall type_sys
                        (maps (map (#3 o dest_Fof) o snd) problem)
-      |>> map (problem_line_for_fact ctxt helper_prefix type_sys
+      |>> map (pair 0
+               #> problem_line_for_fact ctxt helper_prefix type_sys
                #> repair_problem_line thy explicit_forall type_sys const_tab)
       |> op @
     val (problem, pool) =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -79,10 +79,12 @@
    ("verbose", "false"),
    ("overlord", "false"),
    ("blocking", "false"),
+   ("relevance_thresholds", "0.45 0.85"),
+   ("max_relevant", "smart"),
+   ("monomorphize", "false"),
+   ("monomorphize_limit", "4"),
    ("type_sys", "smart"),
    ("explicit_apply", "false"),
-   ("relevance_thresholds", "0.45 0.85"),
-   ("max_relevant", "smart"),
    ("isar_proof", "false"),
    ("isar_shrink_factor", "1")]
 
@@ -95,6 +97,7 @@
    ("quiet", "verbose"),
    ("no_overlord", "overlord"),
    ("non_blocking", "blocking"),
+   ("dont_monomorphize", "monomorphize"),
    ("partial_types", "full_types"),
    ("implicit_apply", "explicit_apply"),
    ("no_isar_proof", "isar_proof")]
@@ -233,6 +236,10 @@
     val blocking = auto orelse debug orelse lookup_bool "blocking"
     val provers = lookup_string "provers" |> space_explode " "
                   |> auto ? single o hd
+    val relevance_thresholds = lookup_real_pair "relevance_thresholds"
+    val max_relevant = lookup_int_option "max_relevant"
+    val monomorphize = lookup_bool "monomorphize"
+    val monomorphize_limit = lookup_int "monomorphize_limit"
     val type_sys =
       case (lookup_string "type_sys", lookup_bool "full_types") of
         ("tags", full_types) => Tags full_types
@@ -245,18 +252,18 @@
         else
           error ("Unknown type system: " ^ quote type_sys ^ ".")
     val explicit_apply = lookup_bool "explicit_apply"
-    val relevance_thresholds = lookup_real_pair "relevance_thresholds"
-    val max_relevant = lookup_int_option "max_relevant"
     val isar_proof = lookup_bool "isar_proof"
     val isar_shrink_factor = Int.max (1, lookup_int "isar_shrink_factor")
     val timeout = (if auto then NONE else lookup_time "timeout") |> the_timeout
     val expect = lookup_string "expect"
   in
     {debug = debug, verbose = verbose, overlord = overlord, blocking = blocking,
-     provers = provers, type_sys = type_sys, explicit_apply = explicit_apply,
-     relevance_thresholds = relevance_thresholds, max_relevant = max_relevant,
-     isar_proof = isar_proof, isar_shrink_factor = isar_shrink_factor,
-     timeout = timeout, expect = expect}
+     provers = provers, relevance_thresholds = relevance_thresholds,
+     max_relevant = max_relevant, monomorphize = monomorphize,
+     monomorphize_limit = monomorphize_limit, type_sys = type_sys,
+     explicit_apply = explicit_apply, isar_proof = isar_proof,
+     isar_shrink_factor = isar_shrink_factor, timeout = timeout,
+     expect = expect}
   end
 
 fun get_params auto ctxt = extract_params auto (default_raw_params ctxt)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -42,8 +42,8 @@
 
 fun print silent f = if silent then () else Output.urgent_message (f ())
 
-fun test_facts ({debug, verbose, overlord, provers, type_sys, isar_proof,
-                 isar_shrink_factor, ...} : params)
+fun test_facts ({debug, verbose, overlord, provers, monomorphize_limit,
+                 type_sys, isar_proof, isar_shrink_factor, ...} : params)
         explicit_apply_opt silent (prover : prover) timeout i n state facts =
   let
     val thy = Proof.theory_of state
@@ -65,6 +65,7 @@
       {debug = debug, verbose = verbose, overlord = overlord, blocking = true,
        provers = provers, type_sys = type_sys, explicit_apply = explicit_apply,
        relevance_thresholds = (1.01, 1.01), max_relevant = NONE,
+       monomorphize = false, monomorphize_limit = monomorphize_limit,
        isar_proof = isar_proof, isar_shrink_factor = isar_shrink_factor,
        timeout = timeout, expect = ""}
     val facts =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -21,10 +21,12 @@
      overlord: bool,
      blocking: bool,
      provers: string list,
+     relevance_thresholds: real * real,
+     max_relevant: int option,
+     monomorphize: bool,
+     monomorphize_limit: int,
      type_sys: type_system,
      explicit_apply: bool,
-     relevance_thresholds: real * real,
-     max_relevant: int option,
      isar_proof: bool,
      isar_shrink_factor: int,
      timeout: Time.time,
@@ -66,7 +68,6 @@
   val smt_iter_fact_frac : real Unsynchronized.ref
   val smt_iter_time_frac : real Unsynchronized.ref
   val smt_iter_min_msecs : int Unsynchronized.ref
-  val smt_monomorphize_limit : int Unsynchronized.ref
 
   val das_Tool : string
   val select_smt_solver : string -> Proof.context -> Proof.context
@@ -229,10 +230,12 @@
    overlord: bool,
    blocking: bool,
    provers: string list,
+   relevance_thresholds: real * real,
+   max_relevant: int option,
+   monomorphize: bool,
+   monomorphize_limit: int,
    type_sys: type_system,
    explicit_apply: bool,
-   relevance_thresholds: real * real,
-   max_relevant: int option,
    isar_proof: bool,
    isar_shrink_factor: int,
    timeout: Time.time,
@@ -333,13 +336,36 @@
         ({exec, required_execs, arguments, has_incomplete_mode, proof_delims,
           known_failures, explicit_forall, use_conjecture_for_hypotheses, ...}
           : atp_config)
-        ({debug, verbose, overlord, type_sys, explicit_apply, isar_proof,
-          isar_shrink_factor, timeout, ...} : params)
+        ({debug, verbose, overlord, monomorphize, monomorphize_limit, type_sys,
+          explicit_apply, isar_proof, isar_shrink_factor, timeout, ...}
+         : params)
         minimize_command ({state, goal, subgoal, facts, ...} : prover_problem) =
   let
+    val thy = Proof.theory_of state
     val ctxt = Proof.context_of state
     val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal
-    val facts = facts |> map (atp_translated_fact ctxt)
+    fun monomorphize_facts facts =
+      let
+        val repair_context =
+          Config.put SMT_Config.verbose debug
+          #> Config.put SMT_Config.monomorph_full false
+          #> Config.put SMT_Config.monomorph_limit monomorphize_limit
+        val facts = facts |> map untranslated_fact
+        (* pseudo-theorem involving the same constants as the subgoal *)
+        val subgoal_th =
+          Logic.list_implies (hyp_ts, concl_t) |> Skip_Proof.make_thm thy
+        val indexed_facts =
+          (~1, subgoal_th) :: (0 upto length facts - 1 ~~ map snd facts)
+        val mono_facts =
+          SMT_Monomorph.monomorph indexed_facts (repair_context ctxt) |> fst
+      in
+        mono_facts
+        |> sort (int_ord o pairself fst)
+        |> filter_out (curry (op =) ~1 o fst)
+        |> map (Untranslated_Fact o apfst (fst o nth facts))
+      end
+    val facts = facts |> monomorphize ? monomorphize_facts
+                      |> map (atp_translated_fact ctxt)
     val (dest_dir, problem_prefix) =
       if overlord then overlord_file_location_for_prover name
       else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix)
@@ -513,9 +539,9 @@
 val smt_iter_fact_frac = Unsynchronized.ref 0.5
 val smt_iter_time_frac = Unsynchronized.ref 0.5
 val smt_iter_min_msecs = Unsynchronized.ref 5000
-val smt_monomorphize_limit = Unsynchronized.ref 4
 
-fun smt_filter_loop name ({debug, verbose, overlord, timeout, ...} : params)
+fun smt_filter_loop name ({debug, verbose, overlord, monomorphize_limit,
+                           timeout, ...} : params)
                     state i smt_filter =
   let
     val ctxt = Proof.context_of state
@@ -529,7 +555,8 @@
           else
             I)
       #> Config.put SMT_Config.infer_triggers (!smt_triggers)
-      #> Config.put SMT_Config.monomorph_limit (!smt_monomorphize_limit)
+      #> Config.put SMT_Config.monomorph_full false
+      #> Config.put SMT_Config.monomorph_limit monomorphize_limit
     val state = state |> Proof.map_context repair_context
 
     fun iter timeout iter_num outcome0 time_so_far facts =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -164,9 +164,9 @@
 (* FUDGE *)
 val auto_max_relevant_divisor = 2
 
-fun run_sledgehammer (params as {debug, blocking, provers, type_sys,
-                                 relevance_thresholds, max_relevant, timeout,
-                                 ...})
+fun run_sledgehammer (params as {debug, blocking, provers, monomorphize,
+                                 type_sys, relevance_thresholds, max_relevant,
+                                 timeout, ...})
                      auto i (relevance_override as {only, ...}) minimize_command
                      state =
   if null provers then
@@ -246,7 +246,10 @@
         else
           launch_provers state
               (get_facts "ATP" no_dangerous_types atp_relevance_fudge o K atps)
-              (ATP_Translated_Fact oo K (translate_atp_fact ctxt false o fst))
+              (if monomorphize then
+                 K (Untranslated_Fact o fst)
+               else
+                 ATP_Translated_Fact oo K (translate_atp_fact ctxt false o fst))
               (K (K NONE)) atps
       fun launch_smts accum =
         if null smts then
--- a/src/HOL/Tools/inductive_codegen.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/inductive_codegen.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -9,7 +9,7 @@
   val add : string option -> int option -> attribute
   val test_fn : (int * int * int -> term list option) Unsynchronized.ref  (* FIXME *)
   val test_term:
-    Proof.context -> term * term list -> int -> term list option * Quickcheck.report option
+    Proof.context -> (term * term list) list -> int list -> term list option * Quickcheck.report option
   val setup : theory -> theory
   val quickcheck_setup : theory -> theory
 end;
@@ -808,7 +808,7 @@
 val (random_values, setup_random_values) = Attrib.config_int "ind_quickcheck_random" (K 5);
 val (size_offset, setup_size_offset) = Attrib.config_int "ind_quickcheck_size_offset" (K 0);
 
-fun test_term ctxt (t, []) =
+fun test_term ctxt [(t, [])] =
   let
     val t' = list_abs_free (Term.add_frees t [], t)
     val thy = ProofContext.theory_of ctxt;
@@ -852,11 +852,12 @@
     val bound = Config.get ctxt depth_bound;
     val start = Config.get ctxt depth_start;
     val offset = Config.get ctxt size_offset;
-    fun test k = (deepen bound (fn i =>
+    fun test [k] = (deepen bound (fn i =>
       (Output.urgent_message ("Search depth: " ^ string_of_int i);
        test_fn' (i, values, k+offset))) start, NONE);
   in test end
-  | test_term ctxt (t, _) = error "Option eval is not supported by tester SML_inductive";
+  | test_term ctxt [(t, _)] = error "Option eval is not supported by tester SML_inductive"
+  | test_term ctxt _ = error "Compilation of multiple instances is not supported by tester SML_inductive";
 
 val quickcheck_setup =
   setup_depth_bound #>
--- a/src/HOL/Tools/list_to_set_comprehension.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/list_to_set_comprehension.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -8,7 +8,7 @@
 signature LIST_TO_SET_COMPREHENSION =
 sig
   val simproc : simpset -> cterm -> thm option
-end;
+end
 
 structure List_to_Set_Comprehension : LIST_TO_SET_COMPREHENSION =
 struct
@@ -16,66 +16,69 @@
 (* conversion *)
 
 fun all_exists_conv cv ctxt ct =
-  case Thm.term_of ct of
-    Const(@{const_name HOL.Ex}, _) $ Abs(_, _, _) =>
+  (case Thm.term_of ct of
+    Const (@{const_name HOL.Ex}, _) $ Abs _ =>
       Conv.arg_conv (Conv.abs_conv (all_exists_conv cv o #2) ctxt) ct
-  | _ => cv ctxt ct
+  | _ => cv ctxt ct)
 
 fun all_but_last_exists_conv cv ctxt ct =
-  case Thm.term_of ct of
+  (case Thm.term_of ct of
     Const (@{const_name HOL.Ex}, _) $ Abs (_, _, Const (@{const_name HOL.Ex}, _) $ _) =>
       Conv.arg_conv (Conv.abs_conv (all_but_last_exists_conv cv o #2) ctxt) ct
-  | _ => cv ctxt ct
+  | _ => cv ctxt ct)
 
 fun Collect_conv cv ctxt ct =
   (case Thm.term_of ct of
     Const (@{const_name Set.Collect}, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv cv ctxt) ct
-  | _ => raise CTERM ("Collect_conv", [ct]));
+  | _ => raise CTERM ("Collect_conv", [ct]))
 
 fun Trueprop_conv cv ct =
   (case Thm.term_of ct of
     Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct
-  | _ => raise CTERM ("Trueprop_conv", [ct]));
+  | _ => raise CTERM ("Trueprop_conv", [ct]))
 
 fun eq_conv cv1 cv2 ct =
   (case Thm.term_of ct of
     Const (@{const_name HOL.eq}, _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv1) cv2 ct
-  | _ => raise CTERM ("eq_conv", [ct]));
+  | _ => raise CTERM ("eq_conv", [ct]))
 
 fun conj_conv cv1 cv2 ct =
   (case Thm.term_of ct of
     Const (@{const_name HOL.conj}, _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv1) cv2 ct
-  | _ => raise CTERM ("conj_conv", [ct]));
+  | _ => raise CTERM ("conj_conv", [ct]))
 
 fun rewr_conv' th = Conv.rewr_conv (mk_meta_eq th)
-  
-fun conjunct_assoc_conv ct = Conv.try_conv
-    ((rewr_conv' @{thm conj_assoc}) then_conv conj_conv Conv.all_conv conjunct_assoc_conv) ct
-  
-fun right_hand_set_comprehension_conv conv ctxt = Trueprop_conv (eq_conv Conv.all_conv
-  (Collect_conv (all_exists_conv conv o #2) ctxt))
+
+fun conjunct_assoc_conv ct =
+  Conv.try_conv
+    (rewr_conv' @{thm conj_assoc} then_conv conj_conv Conv.all_conv conjunct_assoc_conv) ct
+
+fun right_hand_set_comprehension_conv conv ctxt =
+  Trueprop_conv (eq_conv Conv.all_conv
+    (Collect_conv (all_exists_conv conv o #2) ctxt))
+
 
 (* term abstraction of list comprehension patterns *)
- 
+
 datatype termlets = If | Case of (typ * int)
 
 fun simproc ss redex =
   let
     val ctxt = Simplifier.the_context ss
-    val thy = ProofContext.theory_of ctxt 
+    val thy = ProofContext.theory_of ctxt
     val set_Nil_I = @{thm trans} OF [@{thm set.simps(1)}, @{thm empty_def}]
     val set_singleton = @{lemma "set [a] = {x. x = a}" by simp}
     val inst_Collect_mem_eq = @{lemma "set A = {x. x : set A}" by simp}
-    val del_refl_eq = @{lemma "(t = t & P) == P" by simp} 
+    val del_refl_eq = @{lemma "(t = t & P) == P" by simp}
     fun mk_set T = Const (@{const_name List.set}, HOLogic.listT T --> HOLogic.mk_setT T)
     fun dest_set (Const (@{const_name List.set}, _) $ xs) = xs
     fun dest_singleton_list (Const (@{const_name List.Cons}, _)
-      $ t $ (Const (@{const_name List.Nil}, _))) = t
+          $ t $ (Const (@{const_name List.Nil}, _))) = t
       | dest_singleton_list t = raise TERM ("dest_singleton_list", [t])
     (* We check that one case returns a singleton list and all other cases
-       return [], and return the index of the one singleton list case *) 
+       return [], and return the index of the one singleton list case *)
     fun possible_index_of_singleton_case cases =
-      let  
+      let
         fun check (i, case_t) s =
           (case strip_abs_body case_t of
             (Const (@{const_name List.Nil}, _)) => s
@@ -88,112 +91,132 @@
       let
         val (case_const, args) = strip_comb case_term
       in
-        case try dest_Const case_const of
-          SOME (c, T) => (case Datatype_Data.info_of_case thy c of
-            SOME _ => (case possible_index_of_singleton_case (fst (split_last args)) of
-              SOME i => 
-                let
-                  val (Ts, _) = strip_type T
-                  val T' = List.last Ts
-                in SOME (List.last args, T', i, nth args i) end
+        (case try dest_Const case_const of
+          SOME (c, T) =>
+            (case Datatype_Data.info_of_case thy c of
+              SOME _ =>
+                (case possible_index_of_singleton_case (fst (split_last args)) of
+                  SOME i =>
+                    let
+                      val (Ts, _) = strip_type T
+                      val T' = List.last Ts
+                    in SOME (List.last args, T', i, nth args i) end
+                | NONE => NONE)
             | NONE => NONE)
-          | NONE => NONE)
-        | NONE => NONE
+        | NONE => NONE)
       end
     (* returns condition continuing term option *)
     fun dest_if (Const (@{const_name If}, _) $ cond $ then_t $ Const (@{const_name Nil}, _)) =
-        SOME (cond, then_t)
+          SOME (cond, then_t)
       | dest_if _ = NONE
-    fun tac _ [] =
-      rtac set_singleton 1 ORELSE rtac inst_Collect_mem_eq 1
-    | tac ctxt (If :: cont) =
-      Splitter.split_tac [@{thm split_if}] 1
-      THEN rtac @{thm conjI} 1
-      THEN rtac @{thm impI} 1
-      THEN Subgoal.FOCUS (fn {prems, context, ...} =>
-        CONVERSION (right_hand_set_comprehension_conv (K
-          (conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv
-           then_conv rewr_conv' @{thm simp_thms(22)})) context) 1) ctxt 1
-      THEN tac ctxt cont
-      THEN rtac @{thm impI} 1
-      THEN Subgoal.FOCUS (fn {prems, context, ...} =>
-          CONVERSION (right_hand_set_comprehension_conv (K
-            (conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv
-             then_conv rewr_conv' @{thm simp_thms(24)})) context) 1) ctxt 1
-      THEN rtac set_Nil_I 1
-    | tac ctxt (Case (T, i) :: cont) =
-      let
-        val info = Datatype.the_info thy (fst (dest_Type T))
-      in
-        (* do case distinction *)
-        Splitter.split_tac [#split info] 1
-        THEN EVERY (map_index (fn (i', case_rewrite) =>
-          (if i' < length (#case_rewrites info) - 1 then rtac @{thm conjI} 1 else all_tac)
-          THEN REPEAT_DETERM (rtac @{thm allI} 1)
+    fun tac _ [] = rtac set_singleton 1 ORELSE rtac inst_Collect_mem_eq 1
+      | tac ctxt (If :: cont) =
+          Splitter.split_tac [@{thm split_if}] 1
+          THEN rtac @{thm conjI} 1
+          THEN rtac @{thm impI} 1
+          THEN Subgoal.FOCUS (fn {prems, context, ...} =>
+            CONVERSION (right_hand_set_comprehension_conv (K
+              (conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv
+               then_conv
+               rewr_conv' @{lemma "(True & P) = P" by simp})) context) 1) ctxt 1
+          THEN tac ctxt cont
           THEN rtac @{thm impI} 1
-          THEN (if i' = i then
-            (* continue recursively *)
-            Subgoal.FOCUS (fn {prems, context, ...} =>
-              CONVERSION (Thm.eta_conversion then_conv right_hand_set_comprehension_conv (K
-                  ((conj_conv 
-                    (eq_conv Conv.all_conv (rewr_conv' (List.last prems))
-                    then_conv (Conv.try_conv (Conv.rewrs_conv (map mk_meta_eq (#inject info))))) Conv.all_conv)
-                    then_conv (Conv.try_conv (Conv.rewr_conv del_refl_eq))
-                    then_conv conjunct_assoc_conv)) context
-                then_conv (Trueprop_conv (eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt) =>
-                  Conv.repeat_conv (all_but_last_exists_conv (K (rewr_conv' @{thm simp_thms(39)})) ctxt)) context)))) 1) ctxt 1
-            THEN tac ctxt cont
-          else
-            Subgoal.FOCUS (fn {prems, context, ...} =>
-              CONVERSION ((right_hand_set_comprehension_conv (K
-                (conj_conv
-                  ((eq_conv Conv.all_conv
-                    (rewr_conv' (List.last prems)))
-                     then_conv (Conv.rewrs_conv (map (fn th => th RS @{thm Eq_FalseI}) (#distinct info)))) Conv.all_conv
-                  then_conv (rewr_conv' @{thm simp_thms(24)}))) context)
-               then_conv (Trueprop_conv (eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt) =>
-                   Conv.repeat_conv (Conv.bottom_conv (K (rewr_conv' @{thm simp_thms(36)})) ctxt)) context)))) 1) ctxt 1
-            THEN rtac set_Nil_I 1)) (#case_rewrites info))
-      end
+          THEN Subgoal.FOCUS (fn {prems, context, ...} =>
+              CONVERSION (right_hand_set_comprehension_conv (K
+                (conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv
+                 then_conv rewr_conv' @{lemma "(False & P) = False" by simp})) context) 1) ctxt 1
+          THEN rtac set_Nil_I 1
+      | tac ctxt (Case (T, i) :: cont) =
+          let
+            val info = Datatype.the_info thy (fst (dest_Type T))
+          in
+            (* do case distinction *)
+            Splitter.split_tac [#split info] 1
+            THEN EVERY (map_index (fn (i', case_rewrite) =>
+              (if i' < length (#case_rewrites info) - 1 then rtac @{thm conjI} 1 else all_tac)
+              THEN REPEAT_DETERM (rtac @{thm allI} 1)
+              THEN rtac @{thm impI} 1
+              THEN (if i' = i then
+                (* continue recursively *)
+                Subgoal.FOCUS (fn {prems, context, ...} =>
+                  CONVERSION (Thm.eta_conversion then_conv right_hand_set_comprehension_conv (K
+                      ((conj_conv
+                        (eq_conv Conv.all_conv (rewr_conv' (List.last prems)) then_conv
+                          (Conv.try_conv (Conv.rewrs_conv (map mk_meta_eq (#inject info)))))
+                        Conv.all_conv)
+                        then_conv (Conv.try_conv (Conv.rewr_conv del_refl_eq))
+                        then_conv conjunct_assoc_conv)) context
+                    then_conv (Trueprop_conv (eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt) =>
+                      Conv.repeat_conv
+                        (all_but_last_exists_conv
+                          (K (rewr_conv'
+                            @{lemma "(EX x. x = t & P x) = P t" by simp})) ctxt)) context)))) 1) ctxt 1
+                THEN tac ctxt cont
+              else
+                Subgoal.FOCUS (fn {prems, context, ...} =>
+                  CONVERSION
+                    (right_hand_set_comprehension_conv (K
+                      (conj_conv
+                        ((eq_conv Conv.all_conv
+                          (rewr_conv' (List.last prems))) then_conv
+                          (Conv.rewrs_conv (map (fn th => th RS @{thm Eq_FalseI}) (#distinct info))))
+                        Conv.all_conv then_conv
+                        (rewr_conv' @{lemma "(False & P) = False" by simp}))) context then_conv
+                      Trueprop_conv
+                        (eq_conv Conv.all_conv
+                          (Collect_conv (fn (_, ctxt) =>
+                            Conv.repeat_conv
+                              (Conv.bottom_conv
+                                (K (rewr_conv'
+                                  @{lemma "(EX x. P) = P" by simp})) ctxt)) context))) 1) ctxt 1
+                THEN rtac set_Nil_I 1)) (#case_rewrites info))
+          end
     fun make_inner_eqs bound_vs Tis eqs t =
-      case dest_case t of
+      (case dest_case t of
         SOME (x, T, i, cont) =>
           let
             val (vs, body) = strip_abs (Pattern.eta_long (map snd bound_vs) cont)
             val x' = incr_boundvars (length vs) x
             val eqs' = map (incr_boundvars (length vs)) eqs
             val (constr_name, _) = nth (the (Datatype_Data.get_constrs thy (fst (dest_Type T)))) i
-            val constr_t = list_comb (Const (constr_name, map snd vs ---> T), map Bound (((length vs) - 1) downto 0))
+            val constr_t =
+              list_comb
+                (Const (constr_name, map snd vs ---> T), map Bound (((length vs) - 1) downto 0))
             val constr_eq = Const (@{const_name HOL.eq}, T --> T --> @{typ bool}) $ constr_t $ x'
           in
             make_inner_eqs (rev vs @ bound_vs) (Case (T, i) :: Tis) (constr_eq :: eqs') body
           end
       | NONE =>
-        case dest_if t of
-          SOME (condition, cont) => make_inner_eqs bound_vs (If :: Tis) (condition :: eqs) cont
-        | NONE =>
-          if eqs = [] then NONE (* no rewriting, nothing to be done *)
-          else
-            let
-              val Type (@{type_name List.list}, [rT]) = fastype_of1 (map snd bound_vs, t)
-              val pat_eq =
-                case try dest_singleton_list t of
-                  SOME t' => Const (@{const_name HOL.eq}, rT --> rT --> @{typ bool})
-                    $ Bound (length bound_vs) $ t'
-                | NONE => Const (@{const_name Set.member}, rT --> HOLogic.mk_setT rT --> @{typ bool})
-                  $ Bound (length bound_vs) $ (mk_set rT $ t)
-              val reverse_bounds = curry subst_bounds
-                ((map Bound ((length bound_vs - 1) downto 0)) @ [Bound (length bound_vs)])
-              val eqs' = map reverse_bounds eqs
-              val pat_eq' = reverse_bounds pat_eq
-              val inner_t = fold (fn (v, T) => fn t => HOLogic.exists_const T $ absdummy (T, t))
-                (rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq')
-              val lhs = term_of redex
-              val rhs = HOLogic.mk_Collect ("x", rT, inner_t)
-              val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
-            in
-              SOME ((Goal.prove ctxt [] [] rewrite_rule_t (fn {context, ...} => tac context (rev Tis))) RS @{thm eq_reflection})
-            end
+          (case dest_if t of
+            SOME (condition, cont) => make_inner_eqs bound_vs (If :: Tis) (condition :: eqs) cont
+          | NONE =>
+            if eqs = [] then NONE (* no rewriting, nothing to be done *)
+            else
+              let
+                val Type (@{type_name List.list}, [rT]) = fastype_of1 (map snd bound_vs, t)
+                val pat_eq =
+                  (case try dest_singleton_list t of
+                    SOME t' =>
+                      Const (@{const_name HOL.eq}, rT --> rT --> @{typ bool}) $
+                        Bound (length bound_vs) $ t'
+                  | NONE =>
+                      Const (@{const_name Set.member}, rT --> HOLogic.mk_setT rT --> @{typ bool}) $
+                        Bound (length bound_vs) $ (mk_set rT $ t))
+                val reverse_bounds = curry subst_bounds
+                  ((map Bound ((length bound_vs - 1) downto 0)) @ [Bound (length bound_vs)])
+                val eqs' = map reverse_bounds eqs
+                val pat_eq' = reverse_bounds pat_eq
+                val inner_t =
+                  fold (fn (v, T) => fn t => HOLogic.exists_const T $ absdummy (T, t))
+                    (rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq')
+                val lhs = term_of redex
+                val rhs = HOLogic.mk_Collect ("x", rT, inner_t)
+                val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
+              in
+                SOME
+                  ((Goal.prove ctxt [] [] rewrite_rule_t
+                    (fn {context, ...} => tac context (rev Tis))) RS @{thm eq_reflection})
+              end))
   in
     make_inner_eqs [] [] [] (dest_set (term_of redex))
   end
--- a/src/HOL/Tools/try.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/Tools/try.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -8,8 +8,8 @@
 sig
   val auto : bool Unsynchronized.ref
   val invoke_try :
-    Time.time option -> string list * string list * string list -> Proof.state
-    -> bool
+    Time.time option -> string list * string list * string list * string list
+    -> Proof.state -> bool
   val setup : theory -> theory
 end;
 
@@ -61,13 +61,13 @@
   | add_attr_text (_, []) s = s
   | add_attr_text (SOME x, fs) s =
     s ^ " " ^ (if x = "" then "" else x ^ ": ") ^ space_implode " " fs
-fun attrs_text (sx, ix, ex) (ss, is, es) =
-  "" |> fold add_attr_text [(sx, ss), (ix, is), (ex, es)]
+fun attrs_text (sx, ix, ex, dx) (ss, is, es, ds) =
+  "" |> fold add_attr_text [(sx, ss), (ix, is), (ex, es), (dx, ds)]
 
 fun do_named_method (name, ((all_goals, run_if_auto), attrs)) auto timeout_opt
-                    triple st =
+                    quad st =
   if not auto orelse run_if_auto then
-    let val attrs = attrs_text attrs triple in
+    let val attrs = attrs_text attrs quad in
       do_generic timeout_opt
                  (name ^ (if all_goals andalso
                              nprems_of (#goal (Proof.goal st)) > 1 then
@@ -81,13 +81,13 @@
   else
     NONE
 
-val full_attrs = (SOME "simp", SOME "intro", SOME "elim")
-val clas_attrs = (NONE, SOME "intro", SOME "elim")
-val simp_attrs = (SOME "add", NONE, NONE)
-val metis_attrs = (SOME "", SOME "", SOME "")
-val no_attrs = (NONE, NONE, NONE)
+val full_attrs = (SOME "simp", SOME "intro", SOME "elim", SOME "dest")
+val clas_attrs = (NONE, SOME "intro", SOME "elim", SOME "dest")
+val simp_attrs = (SOME "add", NONE, NONE, NONE)
+val metis_attrs = (SOME "", SOME "", SOME "", SOME "")
+val no_attrs = (NONE, NONE, NONE, NONE)
 
-(* name * ((all_goals, run_if_auto), (simp, intro, elim) *)
+(* name * ((all_goals, run_if_auto), (simp, intro, elim, dest) *)
 val named_methods =
   [("simp", ((false, true), simp_attrs)),
    ("auto", ((true, true), full_attrs)),
@@ -102,11 +102,11 @@
 
 fun time_string (s, ms) = s ^ ": " ^ string_of_int ms ^ " ms"
 
-fun do_try auto timeout_opt triple st =
+fun do_try auto timeout_opt quad st =
   let
     val st = st |> Proof.map_context (Config.put Metis_Tactics.verbose false)
   in
-    case do_methods |> Par_List.map (fn f => f auto timeout_opt triple st)
+    case do_methods |> Par_List.map (fn f => f auto timeout_opt quad st)
                     |> map_filter I |> sort (int_ord o pairself snd) of
       [] => (if auto then () else writeln "No proof found."; (false, st))
     | xs as (s, _) :: _ =>
@@ -137,11 +137,12 @@
 
 val tryN = "try"
 
-fun try_trans triple =
-  Toplevel.keep (K () o do_try false (SOME default_timeout) triple
+fun try_trans quad =
+  Toplevel.keep (K () o do_try false (SOME default_timeout) quad
                  o Toplevel.proof_of)
 
-fun merge_attrs (s1, i1, e1) (s2, i2, e2) = (s1 @ s2, i1 @ i2, e1 @ e2)
+fun merge_attrs (s1, i1, e1, d1) (s2, i2, e2, d2) =
+  (s1 @ s2, i1 @ i2, e1 @ e2, d1 @ d2)
 
 fun string_of_xthm (xref, args) =
   Facts.string_of_ref xref ^
@@ -153,23 +154,25 @@
                             (Parse_Spec.xthm >> string_of_xthm))
 val parse_attr =
      Args.$$$ "simp" |-- Args.colon |-- parse_fact_refs
-     >> (fn ss => (ss, [], []))
+     >> (fn ss => (ss, [], [], []))
   || Args.$$$ "intro" |-- Args.colon |-- parse_fact_refs
-     >> (fn is => ([], is, []))
+     >> (fn is => ([], is, [], []))
   || Args.$$$ "elim" |-- Args.colon |-- parse_fact_refs
-     >> (fn es => ([], [], es))
+     >> (fn es => ([], [], es, []))
+  || Args.$$$ "dest" |-- Args.colon |-- parse_fact_refs
+     >> (fn ds => ([], [], [], ds))
 fun parse_attrs x =
     (Args.parens parse_attrs
   || Scan.repeat parse_attr
-     >> (fn triple => fold merge_attrs triple ([], [], []))) x
+     >> (fn quad => fold merge_attrs quad ([], [], [], []))) x
 
-val parse_try_command = Scan.optional parse_attrs ([], [], []) #>> try_trans
+val parse_try_command = Scan.optional parse_attrs ([], [], [], []) #>> try_trans
 
 val _ =
   Outer_Syntax.improper_command tryN
       "try a combination of proof methods" Keyword.diag parse_try_command
 
-val auto_try = do_try true NONE ([], [], [])
+val auto_try = do_try true NONE ([], [], [], [])
 
 val setup = Auto_Tools.register_tool (auto, auto_try)
 
--- a/src/HOL/ex/Quickcheck_Examples.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/ex/Quickcheck_Examples.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -50,9 +50,14 @@
   oops
 
 theorem "rev xs = xs"
-  quickcheck[random, expect = counterexample]
-  quickcheck[exhaustive, expect = counterexample]
-  oops
+  quickcheck[tester = random, finite_types = true, report = false, expect = counterexample]
+  quickcheck[tester = random, finite_types = false, report = false, expect = counterexample]
+  quickcheck[tester = random, finite_types = true, report = true, expect = counterexample]
+  quickcheck[tester = random, finite_types = false, report = true, expect = counterexample]
+  quickcheck[tester = exhaustive, finite_types = true, expect = counterexample]
+  quickcheck[tester = exhaustive, finite_types = false, expect = counterexample]
+oops
+
 
 text {* An example involving functions inside other data structures *}
 
--- a/src/HOL/ex/Quickcheck_Narrowing_Examples.thy	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/ex/Quickcheck_Narrowing_Examples.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -32,15 +32,15 @@
 lemma "rev xs = xs"
   quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample]
 oops
-
+(*
 lemma "rev xs = xs"
   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
 oops
-
+*)
 subsection {* Simple examples with functions *}
 
 declare [[quickcheck_finite_functions]]
-
+(*
 lemma "map f xs = map g xs"
   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
 oops
@@ -57,7 +57,7 @@
 lemma "map f xs = F f xs"
   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
 oops
-
+*)
 lemma "map f xs = F f xs"
   quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
 oops
@@ -163,7 +163,7 @@
 
 lemma is_ord_l_bal:
  "\<lbrakk> is_ord(MKT (x :: nat) l r h); height l = height r + 2 \<rbrakk> \<Longrightarrow> is_ord(l_bal(x,l,r))"
-quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, timeout = 100, expect = counterexample]
+quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, timeout = 1000, expect = counterexample]
 oops
 
 
--- a/src/HOL/ex/ROOT.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/HOL/ex/ROOT.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -64,6 +64,7 @@
   "HarmonicSeries",
   "Refute_Examples",
   "Quickcheck_Examples",
+  "SML_Quickcheck_Examples",
   "Quickcheck_Lattice_Examples",
   "Landau",
   "Execute_Choice",
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/SML_Quickcheck_Examples.thy	Mon Apr 04 12:40:00 2011 +0100
@@ -0,0 +1,142 @@
+(*  Title:      HOL/ex/SML_Quickcheck_Examples.thy
+    Author:     Stefan Berghofer, Lukas Bulwahn
+    Copyright   2011 TU Muenchen
+*)
+
+theory SML_Quickcheck_Examples
+imports "~~/src/HOL/Library/SML_Quickcheck"
+begin
+
+text {*
+This is a regression test for the 'quickcheck' counterexample generator based on the
+SML code generator.
+
+The counterexample generator has been superseded by counterexample generators based on
+the generic code generator.
+*}
+
+declare [[quickcheck_finite_types = false]]
+
+subsection {* Lists *}
+
+theorem "map g (map f xs) = map (g o f) xs"
+  quickcheck[SML, expect = no_counterexample]
+  oops
+
+theorem "map g (map f xs) = map (f o g) xs"
+  quickcheck[SML, expect = counterexample]
+  oops
+
+theorem "rev (xs @ ys) = rev ys @ rev xs"
+  quickcheck[SML, expect = no_counterexample]
+  oops
+
+theorem "rev (xs @ ys) = rev xs @ rev ys"
+  quickcheck[SML, expect = counterexample]
+  oops
+
+theorem "rev (rev xs) = xs"
+  quickcheck[SML, expect = no_counterexample]
+  oops
+
+theorem "rev xs = xs"
+  quickcheck[tester = SML, expect = counterexample]
+oops
+
+
+text {* An example involving functions inside other data structures *}
+
+primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where
+  "app [] x = x"
+  | "app (f # fs) x = app fs (f x)"
+
+lemma "app (fs @ gs) x = app gs (app fs x)"
+  quickcheck[SML, expect = no_counterexample]
+  by (induct fs arbitrary: x) simp_all
+
+lemma "app (fs @ gs) x = app fs (app gs x)"
+  quickcheck[SML, expect = counterexample]
+  oops
+
+primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
+  "occurs a [] = 0"
+  | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)"
+
+primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+  "del1 a [] = []"
+  | "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))"
+
+text {* A lemma, you'd think to be true from our experience with delAll *}
+lemma "Suc (occurs a (del1 a xs)) = occurs a xs"
+  -- {* Wrong. Precondition needed.*}
+  quickcheck[SML, expect = counterexample]
+  oops
+
+lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
+  quickcheck[SML, expect = counterexample]
+    -- {* Also wrong.*}
+  oops
+
+lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
+  quickcheck[SML, expect = no_counterexample]
+  by (induct xs) auto
+
+primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+  "replace a b [] = []"
+  | "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs)) 
+                            else (x#(replace a b xs)))"
+
+lemma "occurs a xs = occurs b (replace a b xs)"
+  quickcheck[SML, expect = counterexample]
+  -- {* Wrong. Precondition needed.*}
+  oops
+
+lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)"
+  quickcheck[SML, expect = no_counterexample]
+  by (induct xs) simp_all
+
+
+subsection {* Trees *}
+
+datatype 'a tree = Twig |  Leaf 'a | Branch "'a tree" "'a tree"
+
+primrec leaves :: "'a tree \<Rightarrow> 'a list" where
+  "leaves Twig = []"
+  | "leaves (Leaf a) = [a]"
+  | "leaves (Branch l r) = (leaves l) @ (leaves r)"
+
+primrec plant :: "'a list \<Rightarrow> 'a tree" where
+  "plant [] = Twig "
+  | "plant (x#xs) = Branch (Leaf x) (plant xs)"
+
+primrec mirror :: "'a tree \<Rightarrow> 'a tree" where
+  "mirror (Twig) = Twig "
+  | "mirror (Leaf a) = Leaf a "
+  | "mirror (Branch l r) = Branch (mirror r) (mirror l)"
+
+theorem "plant (rev (leaves xt)) = mirror xt"
+  quickcheck[SML, expect = counterexample]
+    --{* Wrong! *} 
+  oops
+
+theorem "plant((leaves xt) @ (leaves yt)) = Branch xt yt"
+  quickcheck[SML, expect = counterexample]
+    --{* Wrong! *} 
+  oops
+
+datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree"
+
+primrec inOrder :: "'a ntree \<Rightarrow> 'a list" where
+  "inOrder (Tip a)= [a]"
+  | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)"
+
+primrec root :: "'a ntree \<Rightarrow> 'a" where
+  "root (Tip a) = a"
+  | "root (Node f x y) = f"
+
+theorem "hd (inOrder xt) = root xt"
+  quickcheck[SML, expect = counterexample]
+  --{* Wrong! *} 
+  oops
+
+end
\ No newline at end of file
--- a/src/Pure/Isar/proof_context.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/Pure/Isar/proof_context.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -670,9 +670,11 @@
   get_const = fn a => ((true, #1 (Term.dest_Const (read_const_proper ctxt false a)))
     handle ERROR _ => (false, Consts.intern (consts_of ctxt) a)),
   get_free = intern_skolem ctxt (Variable.def_type ctxt false),
-  markup_const = Name_Space.markup_entry (const_space ctxt),
-  markup_free = fn x => Markup.name x Markup.free,
-  markup_var = fn xi => Markup.name (Term.string_of_vname xi) Markup.var};
+  markup_const = fn x => [Name_Space.markup_entry (const_space ctxt) x],
+  markup_free = fn x =>
+    [if can Name.dest_skolem x then Markup.skolem else Markup.free] @
+    (if not (Variable.is_body ctxt) orelse Variable.is_fixed ctxt x then [] else [Markup.hilite]),
+  markup_var = fn xi => [Markup.name (Term.string_of_vname xi) Markup.var]};
 
 val decode_term = Syntax.decode_term o term_context;
 
--- a/src/Pure/Syntax/syn_trans.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/Pure/Syntax/syn_trans.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -36,6 +36,7 @@
   val non_typed_tr': (term list -> term) -> bool -> typ -> term list -> term
   val non_typed_tr'': ('a -> term list -> term) -> 'a -> bool -> typ -> term list -> term
   val constrainAbsC: string
+  val abs_tr: term list -> term
   val pure_trfuns:
     (string * (Ast.ast list -> Ast.ast)) list *
     (string * (term list -> term)) list *
@@ -97,7 +98,6 @@
 fun idtypdummy_ast_tr (*"_idtypdummy"*) [ty] =
       Ast.Appl [Ast.Constant "_constrain", Ast.Constant "_idtdummy", ty]
   | idtypdummy_ast_tr (*"_idtypdummy"*) asts = raise Ast.AST ("idtyp_ast_tr", asts);
-
 fun lambda_ast_tr (*"_lambda"*) [pats, body] =
       Ast.fold_ast_p "_abs" (Ast.unfold_ast "_pttrns" pats, body)
   | lambda_ast_tr (*"_lambda"*) asts = raise Ast.AST ("lambda_ast_tr", asts);
--- a/src/Pure/Syntax/type_ext.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/Pure/Syntax/type_ext.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -130,15 +130,15 @@
 (* decode_term -- transform parse tree into raw term *)
 
 fun markup_bound def id =
-  Markup.properties [(if def then Markup.defN else Markup.refN, id)] Markup.bound;
+  [Markup.properties [(if def then Markup.defN else Markup.refN, id)] Markup.bound];
 
 type term_context =
  {get_sort: (indexname * sort) list -> indexname -> sort,
   get_const: string -> bool * string,
   get_free: string -> string option,
-  markup_const: string -> Markup.T,
-  markup_free: string -> Markup.T,
-  markup_var: indexname -> Markup.T};
+  markup_const: string -> Markup.T list,
+  markup_free: string -> Markup.T list,
+  markup_var: indexname -> Markup.T list};
 
 fun decode_term (term_context: term_context) tm =
   let
@@ -148,8 +148,8 @@
     val reports = Unsynchronized.ref ([]: (Position.T * Markup.T) list);
     fun report [] _ _ = ()
       | report ps markup x =
-          let val m = markup x
-          in Unsynchronized.change reports (fold (fn p => cons (p, m)) ps) end;
+          let val ms = markup x
+          in Unsynchronized.change reports (fold (fn p => fold (fn m => cons (p, m)) ms) ps) end;
 
     fun decode ps qs bs (Const ("_constrain", _) $ t $ typ) =
           (case decode_position typ of
--- a/src/Pure/System/session.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/Pure/System/session.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -103,10 +103,13 @@
         let
           val start = Timing.start ();
           val _ = use root;
+          val timing = Timing.result start;
+          val factor = Time.toReal (#cpu timing) / Time.toReal (#elapsed timing)
+            |> Real.fmt (StringCvt.FIX (SOME 2));
           val _ =
             Output.raw_stderr ("Timing " ^ item ^ " (" ^
               string_of_int (Multithreading.max_threads_value ()) ^ " threads, " ^
-              Timing.message (Timing.result start) ^ ")\n");
+              Timing.message timing ^ ", factor " ^ factor ^ ")\n");
         in () end
       else use root;
       finish ()))
--- a/src/Pure/type_infer.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/Pure/type_infer.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -93,7 +93,7 @@
 
 (* prepare_term *)
 
-fun prepare_term const_type tm (vparams, params, idx) =
+fun prepare_term ctxt const_type tm (vparams, params, idx) =
   let
     fun add_vparm xi (ps_idx as (ps, idx)) =
       if not (Vartab.defined ps xi) then
@@ -118,9 +118,12 @@
 
     fun prepare (Const ("_type_constraint_", T) $ t) ps_idx =
           let
-            val (T', ps_idx') = prepare_typ T ps_idx;
+            fun err () =
+              error ("Malformed internal type constraint: " ^ Syntax.string_of_typ ctxt T);
+            val A = (case T of Type ("fun", [A, B]) => if A = B then A else err () | _ => err ());
+            val (A', ps_idx') = prepare_typ A ps_idx;
             val (t', ps_idx'') = prepare t ps_idx';
-          in (Const ("_type_constraint_", T') $ t', ps_idx'') end
+          in (Const ("_type_constraint_", A' --> A') $ t', ps_idx'') end
       | prepare (Const (c, T)) (ps, idx) =
           (case const_type c of
             SOME U =>
@@ -344,8 +347,8 @@
 
     val ts = burrow_types (Syntax.check_typs ctxt) raw_ts;
     val (ts', (_, _, idx)) =
-      fold_map (prepare_term const_type o constrain_vars) ts
-      (Vartab.empty, Vartab.empty, 0);
+      fold_map (prepare_term ctxt const_type o constrain_vars) ts
+        (Vartab.empty, Vartab.empty, 0);
   in (idx, ts') end;
 
 fun infer_types ctxt const_type var_type raw_ts =
--- a/src/Tools/WWW_Find/IsaMakefile	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/Tools/WWW_Find/IsaMakefile	Mon Apr 04 12:40:00 2011 +0100
@@ -9,6 +9,7 @@
 images:
 test: Pure-WWW_Find
 all: images test
+smlnj: all
 
 
 ## global settings
--- a/src/Tools/jEdit/dist-template/etc/isabelle-jedit.css	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/Tools/jEdit/dist-template/etc/isabelle-jedit.css	Mon Apr 04 12:40:00 2011 +0100
@@ -9,7 +9,7 @@
 
 .report { display: none; }
 
-.hilite { background-color: #FFFACD; }
+.hilite { background-color: #FFCC66; }
 
 .keyword { font-weight: bold; color: #009966; }
 .operator { font-weight: bold; }
--- a/src/Tools/jEdit/src/jedit/isabelle_markup.scala	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/Tools/jEdit/src/jedit/isabelle_markup.scala	Mon Apr 04 12:40:00 2011 +0100
@@ -26,6 +26,7 @@
   val warning_color = new Color(255, 140, 0)
   val error_color = new Color(178, 34, 34)
   val bad_color = new Color(255, 106, 106, 100)
+  val hilite_color = new Color(255, 204, 102, 100)
 
   class Icon(val priority: Int, val icon: javax.swing.Icon)
   {
@@ -100,6 +101,7 @@
   val background1: Markup_Tree.Select[Color] =
   {
     case Text.Info(_, XML.Elem(Markup(Markup.BAD, _), _)) => bad_color
+    case Text.Info(_, XML.Elem(Markup(Markup.HILITE, _), _)) => hilite_color
   }
 
   val background2: Markup_Tree.Select[Color] =
@@ -154,7 +156,7 @@
       Markup.TFREE -> NULL,
       Markup.FREE -> MARKUP,
       Markup.TVAR -> NULL,
-      Markup.SKOLEM -> NULL,
+      Markup.SKOLEM -> COMMENT2,
       Markup.BOUND -> LABEL,
       Markup.VAR -> NULL,
       Markup.NUM -> DIGIT,
--- a/src/Tools/quickcheck.ML	Mon Apr 04 12:39:14 2011 +0100
+++ b/src/Tools/quickcheck.ML	Mon Apr 04 12:40:00 2011 +0100
@@ -29,11 +29,14 @@
       satisfied_assms : int list, positive_concl_tests : int }
   (* registering generators *)
   val add_generator:
-    string * (Proof.context -> term * term list -> int -> term list option * report option)
+    string * (Proof.context -> (term * term list) list -> int list -> term list option * report option)
       -> Context.generic -> Context.generic
   val add_batch_generator:
     string * (Proof.context -> term list -> (int -> term list option) list)
       -> Context.generic -> Context.generic
+  val add_batch_validator:
+    string * (Proof.context -> term list -> (int -> bool) list)
+      -> Context.generic -> Context.generic
   (* quickcheck's result *)
   datatype result =
     Result of
@@ -50,7 +53,9 @@
       -> result list
   val quickcheck: (string * string list) list -> int -> Proof.state -> (string * term) list option
   (* testing a batch of terms *)
-  val test_terms: Proof.context -> term list -> (string * term) list option list option
+  val test_terms:
+    Proof.context -> term list -> (string * term) list option list option * (string * int) list
+  val validate_terms: Proof.context -> term list -> bool list option * (string * int) list
 end;
 
 structure Quickcheck : QUICKCHECK =
@@ -110,12 +115,14 @@
     timings = #timings r, reports = cons (size, report) (#reports r)}
   | cons_report _ NONE result = result
 
-fun add_timing timing result_ref = (result_ref := cons_timing timing (!result_ref))
+fun add_timing timing result_ref =
+  Unsynchronized.change result_ref (cons_timing timing)
 
-fun add_report size report result_ref = (result_ref := cons_report size report (!result_ref))
+fun add_report size report result_ref =
+  Unsynchronized.change result_ref (cons_report size report)
 
 fun add_response names eval_terms response result_ref =
-  (result_ref := set_reponse names eval_terms response (!result_ref))
+  Unsynchronized.change result_ref (set_reponse names eval_terms response)
 
 (* expectation *)
 
@@ -161,14 +168,17 @@
 structure Data = Generic_Data
 (
   type T =
-    ((string * (Proof.context -> term * term list -> int -> term list option * report option)) list
-      * (string * (Proof.context -> term list -> (int -> term list option) list)) list)
+    ((string * (Proof.context -> (term * term list) list -> int list -> term list option * report option)) list
+      * ((string * (Proof.context -> term list -> (int -> term list option) list)) list
+      * ((string * (Proof.context -> term list -> (int -> bool) list)) list)))
       * test_params;
-  val empty = (([], []), Test_Params {default_type = [], expect = No_Expectation});
+  val empty = (([], ([], [])), Test_Params {default_type = [], expect = No_Expectation});
   val extend = I;
-  fun merge (((generators1, batch_generators1), params1), ((generators2, batch_generators2), params2)) : T =
+  fun merge (((generators1, (batch_generators1, batch_validators1)), params1),
+    ((generators2, (batch_generators2, batch_validators2)), params2)) : T =
     ((AList.merge (op =) (K true) (generators1, generators2),
-    AList.merge (op =) (K true) (batch_generators1, batch_generators2)),
+     (AList.merge (op =) (K true) (batch_generators1, batch_generators2),
+      AList.merge (op =) (K true) (batch_validators1, batch_validators2))),
       merge_test_params (params1, params2));
 );
 
@@ -182,7 +192,9 @@
 
 val add_generator = Data.map o apfst o apfst o AList.update (op =);
 
-val add_batch_generator = Data.map o apfst o apsnd o AList.update (op =);
+val add_batch_generator = Data.map o apfst o apsnd o apfst o AList.update (op =);
+
+val add_batch_validator = Data.map o apfst o apsnd o apsnd o AList.update (op =);
 
 (* generating tests *)
 
@@ -204,10 +216,13 @@
       end
   end
 
-val mk_tester = gen_mk_tester (fn ctxt =>
-  AList.lookup (op =) ((fst o fst o Data.get o Context.Proof) ctxt))
-val mk_batch_tester = gen_mk_tester (fn ctxt => AList.lookup (op =) ((snd o fst o Data.get o Context.Proof) ctxt))
-  
+val mk_tester =
+  gen_mk_tester (fn ctxt => AList.lookup (op =) ((fst o fst o Data.get o Context.Proof) ctxt))
+val mk_batch_tester =
+  gen_mk_tester (fn ctxt => AList.lookup (op =) ((fst o snd o fst o Data.get o Context.Proof) ctxt))
+val mk_batch_validator =
+  gen_mk_tester (fn ctxt => AList.lookup (op =) ((snd o snd o fst o Data.get o Context.Proof) ctxt))
+
 (* testing propositions *)
 
 fun check_test_term t =
@@ -239,9 +254,6 @@
     val current_result = Unsynchronized.ref empty_result 
     fun excipit () =
       "Quickcheck ran out of time while testing at size " ^ string_of_int (!current_size)
-    val (test_fun, comp_time) = cpu_time "quickcheck compilation"
-      (fn () => mk_tester ctxt (t, eval_terms));
-    val _ = add_timing comp_time current_result
     fun with_size test_fun k =
       if k > Config.get ctxt size then
         NONE
@@ -250,40 +262,66 @@
           val _ = message ("Test data size: " ^ string_of_int k)
           val _ = current_size := k
           val ((result, report), timing) =
-            cpu_time ("size " ^ string_of_int k) (fn () => test_fun (k - 1))
+            cpu_time ("size " ^ string_of_int k) (fn () => test_fun [1, k - 1])
           val _ = add_timing timing current_result
           val _ = add_report k report current_result
         in
           case result of NONE => with_size test_fun (k + 1) | SOME q => SOME q
         end;
   in
-    case test_fun of NONE => !current_result
-      | SOME test_fun =>
-        limit ctxt (limit_time, is_interactive) (fn () =>
-          let
-            val (response, exec_time) =
-              cpu_time "quickcheck execution" (fn () => with_size test_fun 1)
-            val _ = add_response names eval_terms response current_result
-            val _ = add_timing exec_time current_result
-          in
-            !current_result
-          end) (fn () => (message (excipit ()); !current_result)) ()
+    limit ctxt (limit_time, is_interactive) (fn () =>
+      let
+        val (test_fun, comp_time) = cpu_time "quickcheck compilation"
+          (fn () => mk_tester ctxt [(t, eval_terms)]);
+        val _ = add_timing comp_time current_result
+      in
+        case test_fun of NONE => !current_result
+          | SOME test_fun =>
+            let
+              val (response, exec_time) =
+                cpu_time "quickcheck execution" (fn () => with_size test_fun 1)
+              val _ = add_response names eval_terms response current_result
+              val _ = add_timing exec_time current_result
+            in
+              !current_result
+            end
+       end)
+     (fn () => (message (excipit ()); !current_result)) ()
   end;
 
+fun validate_terms ctxt ts =
+  let
+    val _ = map check_test_term ts
+    val size = Config.get ctxt size
+    val (test_funs, comp_time) = cpu_time "quickcheck batch compilation"
+      (fn () => mk_batch_validator ctxt ts) 
+    fun with_size tester k =
+      if k > size then true
+      else if tester k then with_size tester (k + 1) else false
+    val (results, exec_time) = cpu_time "quickcheck batch execution" (fn () =>
+        Option.map (map (fn test_fun => TimeLimit.timeLimit (seconds (Config.get ctxt timeout))
+              (fn () => with_size test_fun 1) ()
+             handle TimeLimit.TimeOut => true)) test_funs)
+  in
+    (results, [comp_time, exec_time])
+  end
+  
 fun test_terms ctxt ts =
   let
     val _ = map check_test_term ts
+    val size = Config.get ctxt size
     val namess = map (fn t => Term.add_free_names t []) ts
-    val test_funs = mk_batch_tester ctxt ts
+    val (test_funs, comp_time) = cpu_time "quickcheck batch compilation" (fn () => mk_batch_tester ctxt ts) 
     fun with_size tester k =
-      if k > Config.get ctxt size then NONE
+      if k > size then NONE
       else case tester k of SOME ts => SOME ts | NONE => with_size tester (k + 1)
-    val results =
-      Option.map (map (fn test_fun => TimeLimit.timeLimit (seconds (Config.get ctxt timeout))
-            (fn () => with_size test_fun 1)  ()
-           handle TimeLimit.TimeOut => NONE)) test_funs
+    val (results, exec_time) = cpu_time "quickcheck batch execution" (fn () =>
+        Option.map (map (fn test_fun => TimeLimit.timeLimit (seconds (Config.get ctxt timeout))
+              (fn () => with_size test_fun 1) ()
+             handle TimeLimit.TimeOut => NONE)) test_funs)
   in
-    Option.map (map2 (fn names => Option.map (fn ts => names ~~ ts)) namess) results
+    (Option.map (map2 (fn names => Option.map (fn ts => names ~~ ts)) namess) results,
+      [comp_time, exec_time])
   end
 
 (* FIXME: this function shows that many assumptions are made upon the generation *)
@@ -293,19 +331,16 @@
   let
     val thy = ProofContext.theory_of ctxt
     fun message s = if Config.get ctxt quiet then () else Output.urgent_message s
-    val (ts, eval_terms) = split_list ts
-    val _ = map check_test_term ts
-    val names = Term.add_free_names (hd ts) []
-    val Ts = map snd (Term.add_frees (hd ts) [])
-    val current_result = Unsynchronized.ref empty_result 
-    val (test_funs, comp_time) = cpu_time "quickcheck compilation"
-      (fn () => map (mk_tester ctxt) (ts ~~ eval_terms))
-    val _ = add_timing comp_time current_result
-    fun test_card_size (card, size) =
+    val (ts', eval_terms) = split_list ts
+    val _ = map check_test_term ts'
+    val names = Term.add_free_names (hd ts') []
+    val Ts = map snd (Term.add_frees (hd ts') [])
+    val current_result = Unsynchronized.ref empty_result
+    fun test_card_size test_fun (card, size) =
       (* FIXME: why decrement size by one? *)
       let
         val (ts, timing) = cpu_time ("size " ^ string_of_int size ^ " and card " ^ string_of_int card)
-          (fn () => fst (the (nth test_funs (card - 1)) (size - 1)))
+          (fn () => fst (test_fun [card - 1, size - 1]))
         val _ = add_timing timing current_result
       in
         Option.map (pair card) ts
@@ -318,19 +353,23 @@
         (* size does matter *)
         map_product pair (1 upto (length ts)) (1 upto (Config.get ctxt size))
         |> sort (fn ((c1, s1), (c2, s2)) => int_ord ((c1 + s1), (c2 + s2)))
-    in
-      if (forall is_none test_funs) then !current_result
-      else if (forall is_some test_funs) then
-        limit ctxt (limit_time, is_interactive) (fn () =>
+  in
+    limit ctxt (limit_time, is_interactive) (fn () =>
+      let
+        val (test_fun, comp_time) = cpu_time "quickcheck compilation" (fn () => mk_tester ctxt ts)
+        val _ = add_timing comp_time current_result
+      in
+        case test_fun of
+          NONE => !current_result
+        | SOME test_fun =>
           let
-            val _ = case get_first test_card_size enumeration_card_size of
+            val _ = case get_first (test_card_size test_fun) enumeration_card_size of
               SOME (card, ts) => add_response names (nth eval_terms (card - 1)) (SOME ts) current_result
             | NONE => ()
-          in !current_result end)
-          (fn () => (message "Quickcheck ran out of time"; !current_result)) ()
-      else
-        error "Unexpectedly, testers of some cardinalities are executable but others are not"
-    end
+          in !current_result end
+      end)
+      (fn () => (message "Quickcheck ran out of time"; !current_result)) ()
+  end
 
 fun get_finite_types ctxt =
   fst (chop (Config.get ctxt finite_type_size)