More mathematical symbols for ZF examples
authorpaulson
Tue, 06 Mar 2012 17:01:37 +0000
changeset 46823 57bf0cecb366
parent 46822 95f1e700b712
child 46824 1257c80988cd
child 46841 49b91b716cbe
More mathematical symbols for ZF examples
src/ZF/Constructible/AC_in_L.thy
src/ZF/Constructible/DPow_absolute.thy
src/ZF/Constructible/Datatype_absolute.thy
src/ZF/Constructible/Formula.thy
src/ZF/Constructible/Internalize.thy
src/ZF/Constructible/L_axioms.thy
src/ZF/Constructible/Normal.thy
src/ZF/Constructible/Rank.thy
src/ZF/Constructible/Rank_Separation.thy
src/ZF/Constructible/Rec_Separation.thy
src/ZF/Constructible/Reflection.thy
src/ZF/Constructible/Relative.thy
src/ZF/Constructible/Satisfies_absolute.thy
src/ZF/Constructible/Separation.thy
src/ZF/Constructible/WF_absolute.thy
src/ZF/Constructible/WFrec.thy
src/ZF/Constructible/Wellorderings.thy
src/ZF/Resid/Confluence.thy
src/ZF/Resid/Redex.thy
src/ZF/Resid/Reduction.thy
src/ZF/Resid/Residuals.thy
src/ZF/Resid/Substitution.thy
src/ZF/UNITY/AllocBase.thy
src/ZF/UNITY/AllocImpl.thy
src/ZF/UNITY/ClientImpl.thy
src/ZF/UNITY/Comp.thy
src/ZF/UNITY/Constrains.thy
src/ZF/UNITY/Distributor.thy
src/ZF/UNITY/FP.thy
src/ZF/UNITY/Follows.thy
src/ZF/UNITY/GenPrefix.thy
src/ZF/UNITY/Guar.thy
src/ZF/UNITY/Increasing.thy
src/ZF/UNITY/Merge.thy
src/ZF/UNITY/Monotonicity.thy
src/ZF/UNITY/MultisetSum.thy
src/ZF/UNITY/Mutex.thy
src/ZF/UNITY/State.thy
src/ZF/UNITY/SubstAx.thy
src/ZF/UNITY/UNITY.thy
src/ZF/UNITY/Union.thy
src/ZF/UNITY/WFair.thy
--- a/src/ZF/Constructible/AC_in_L.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/AC_in_L.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -86,7 +86,7 @@
  apply (simp (no_asm_use))
 apply clarify
 apply (simp (no_asm_use))
-apply (subgoal_tac "\<forall>l2 \<in> list(A). length(l2) = x --> Cons(a,l2) \<in> B", blast)
+apply (subgoal_tac "\<forall>l2 \<in> list(A). length(l2) = x \<longrightarrow> Cons(a,l2) \<in> B", blast)
 apply (erule_tac a=a in wf_on_induct, assumption)
 apply (rule ballI)
 apply (rule impI)
@@ -163,7 +163,7 @@
 
 lemma (in Nat_Times_Nat) fn_iff:
     "[|x \<in> nat; y \<in> nat; u \<in> nat; v \<in> nat|]
-     ==> (fn`<x,y> = fn`<u,v>) <-> (x=u & y=v)"
+     ==> (fn`<x,y> = fn`<u,v>) \<longleftrightarrow> (x=u & y=v)"
 by (blast dest: inj_apply_equality [OF fn_inj])
 
 lemma (in Nat_Times_Nat) enum_type [TC,simp]:
@@ -171,7 +171,7 @@
 by (induct_tac p, simp_all)
 
 lemma (in Nat_Times_Nat) enum_inject [rule_format]:
-    "p \<in> formula ==> \<forall>q\<in>formula. enum(fn,p) = enum(fn,q) --> p=q"
+    "p \<in> formula ==> \<forall>q\<in>formula. enum(fn,p) = enum(fn,q) \<longrightarrow> p=q"
 apply (induct_tac p, simp_all)
    apply (rule ballI)
    apply (erule formula.cases)
@@ -382,7 +382,7 @@
 apply (simp add: ltI, clarify)
 apply (rename_tac u v)
 apply (rule_tac i="lrank(u)" and j="lrank(v)" in Ord_linear_lt, simp_all) 
-apply (drule_tac x="succ(lrank(u) Un lrank(v))" in ospec)
+apply (drule_tac x="succ(lrank(u) \<union> lrank(v))" in ospec)
  apply (simp add: ltI)
 apply (drule_tac x=u in spec, simp)
 apply (drule_tac x=v in spec, simp)
--- a/src/ZF/Constructible/DPow_absolute.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/DPow_absolute.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -40,13 +40,13 @@
   assumes MH_iff_sats: 
       "!!a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10. 
         [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A; a8\<in>A; a9\<in>A; a10\<in>A|]
-        ==> MH(a2, a1, a0) <-> 
+        ==> MH(a2, a1, a0) \<longleftrightarrow> 
             sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,
                           Cons(a4,Cons(a5,Cons(a6,Cons(a7,
                                   Cons(a8,Cons(a9,Cons(a10,env))))))))))))"
   shows 
       "[|x \<in> nat; z \<in> nat; env \<in> list(A)|]
-       ==> sats(A, formula_rec_fm(p,x,z), env) <-> 
+       ==> sats(A, formula_rec_fm(p,x,z), env) \<longleftrightarrow> 
            is_formula_rec(##A, MH, nth(x,env), nth(z,env))"
 by (simp add: formula_rec_fm_def sats_is_transrec_fm is_formula_rec_def 
               MH_iff_sats [THEN iff_sym])
@@ -55,14 +55,14 @@
   assumes MH_iff_sats: 
       "!!a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10. 
         [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A; a8\<in>A; a9\<in>A; a10\<in>A|]
-        ==> MH(a2, a1, a0) <-> 
+        ==> MH(a2, a1, a0) \<longleftrightarrow> 
             sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,
                           Cons(a4,Cons(a5,Cons(a6,Cons(a7,
                                   Cons(a8,Cons(a9,Cons(a10,env))))))))))))"
   shows
   "[|nth(i,env) = x; nth(k,env) = z; 
       i \<in> nat; k \<in> nat; env \<in> list(A)|]
-   ==> is_formula_rec(##A, MH, x, z) <-> sats(A, formula_rec_fm(p,i,k), env)" 
+   ==> is_formula_rec(##A, MH, x, z) \<longleftrightarrow> sats(A, formula_rec_fm(p,i,k), env)" 
 by (simp add: sats_formula_rec_fm [OF MH_iff_sats])
 
 theorem formula_rec_reflection:
@@ -90,7 +90,7 @@
 
 lemma sats_satisfies_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, satisfies_fm(x,y,z), env) <->
+    ==> sats(A, satisfies_fm(x,y,z), env) \<longleftrightarrow>
         is_satisfies(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: satisfies_fm_def is_satisfies_def sats_satisfies_MH_fm
               sats_formula_rec_fm)
@@ -98,7 +98,7 @@
 lemma satisfies_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_satisfies(##A, x, y, z) <-> sats(A, satisfies_fm(i,j,k), env)"
+       ==> is_satisfies(##A, x, y, z) \<longleftrightarrow> sats(A, satisfies_fm(i,j,k), env)"
 by (simp add: sats_satisfies_fm)
 
 theorem satisfies_reflection:
@@ -124,12 +124,12 @@
   is_DPow_sats :: "[i=>o,i,i,i,i] => o" where
    "is_DPow_sats(M,A,env,p,x) ==
       \<forall>n1[M]. \<forall>e[M]. \<forall>sp[M]. 
-             is_satisfies(M,A,p,sp) --> is_Cons(M,x,env,e) --> 
-             fun_apply(M, sp, e, n1) --> number1(M, n1)"
+             is_satisfies(M,A,p,sp) \<longrightarrow> is_Cons(M,x,env,e) \<longrightarrow> 
+             fun_apply(M, sp, e, n1) \<longrightarrow> number1(M, n1)"
 
 lemma (in M_satisfies) DPow_sats_abs:
     "[| M(A); env \<in> list(A); p \<in> formula; M(x) |]
-    ==> is_DPow_sats(M,A,env,p,x) <-> sats(A, p, Cons(x,env))"
+    ==> is_DPow_sats(M,A,env,p,x) \<longleftrightarrow> sats(A, p, Cons(x,env))"
 apply (subgoal_tac "M(env)") 
  apply (simp add: is_DPow_sats_def satisfies_closed satisfies_abs) 
 apply (blast dest: transM) 
@@ -146,8 +146,8 @@
 
 (* is_DPow_sats(M,A,env,p,x) ==
       \<forall>n1[M]. \<forall>e[M]. \<forall>sp[M]. 
-             is_satisfies(M,A,p,sp) --> is_Cons(M,x,env,e) --> 
-             fun_apply(M, sp, e, n1) --> number1(M, n1) *)
+             is_satisfies(M,A,p,sp) \<longrightarrow> is_Cons(M,x,env,e) \<longrightarrow> 
+             fun_apply(M, sp, e, n1) \<longrightarrow> number1(M, n1) *)
 
 definition
   DPow_sats_fm :: "[i,i,i,i]=>i" where
@@ -164,14 +164,14 @@
 
 lemma sats_DPow_sats_fm [simp]:
    "[| u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, DPow_sats_fm(u,x,y,z), env) <->
+    ==> sats(A, DPow_sats_fm(u,x,y,z), env) \<longleftrightarrow>
         is_DPow_sats(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: DPow_sats_fm_def is_DPow_sats_def)
 
 lemma DPow_sats_iff_sats:
   "[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
       u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-   ==> is_DPow_sats(##A,nu,nx,ny,nz) <->
+   ==> is_DPow_sats(##A,nu,nx,ny,nz) \<longleftrightarrow>
        sats(A, DPow_sats_fm(u,x,y,z), env)"
 by simp
 
@@ -223,13 +223,13 @@
 definition 
   is_DPow' :: "[i=>o,i,i] => o" where
     "is_DPow'(M,A,Z) == 
-       \<forall>X[M]. X \<in> Z <-> 
+       \<forall>X[M]. X \<in> Z \<longleftrightarrow> 
          subset(M,X,A) & 
            (\<exists>env[M]. \<exists>p[M]. mem_formula(M,p) & mem_list(M,A,env) &
                     is_Collect(M, A, is_DPow_sats(M,A,env,p), X))"
 
 lemma (in M_DPow) DPow'_abs:
-    "[|M(A); M(Z)|] ==> is_DPow'(M,A,Z) <-> Z = DPow'(A)"
+    "[|M(A); M(Z)|] ==> is_DPow'(M,A,Z) \<longleftrightarrow> Z = DPow'(A)"
 apply (rule iffI)
  prefer 2 apply (simp add: is_DPow'_def DPow'_def Collect_DPow_sats_abs) 
 apply (rule M_equalityI) 
@@ -310,7 +310,7 @@
 enclosed within a single quantifier.*}
 
 (* is_Collect :: "[i=>o,i,i=>o,i] => o"
-    "is_Collect(M,A,P,z) == \<forall>x[M]. x \<in> z <-> x \<in> A & P(x)" *)
+    "is_Collect(M,A,P,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A & P(x)" *)
 
 definition
   Collect_fm :: "[i, i, i]=>i" where
@@ -325,20 +325,20 @@
 
 lemma sats_Collect_fm:
   assumes is_P_iff_sats: 
-      "!!a. a \<in> A ==> is_P(a) <-> sats(A, p, Cons(a, env))"
+      "!!a. a \<in> A ==> is_P(a) \<longleftrightarrow> sats(A, p, Cons(a, env))"
   shows 
       "[|x \<in> nat; y \<in> nat; env \<in> list(A)|]
-       ==> sats(A, Collect_fm(x,p,y), env) <->
+       ==> sats(A, Collect_fm(x,p,y), env) \<longleftrightarrow>
            is_Collect(##A, nth(x,env), is_P, nth(y,env))"
 by (simp add: Collect_fm_def is_Collect_def is_P_iff_sats [THEN iff_sym])
 
 lemma Collect_iff_sats:
   assumes is_P_iff_sats: 
-      "!!a. a \<in> A ==> is_P(a) <-> sats(A, p, Cons(a, env))"
+      "!!a. a \<in> A ==> is_P(a) \<longleftrightarrow> sats(A, p, Cons(a, env))"
   shows 
   "[| nth(i,env) = x; nth(j,env) = y;
       i \<in> nat; j \<in> nat; env \<in> list(A)|]
-   ==> is_Collect(##A, x, is_P, y) <-> sats(A, Collect_fm(i,p,j), env)"
+   ==> is_Collect(##A, x, is_P, y) \<longleftrightarrow> sats(A, Collect_fm(i,p,j), env)"
 by (simp add: sats_Collect_fm [OF is_P_iff_sats])
 
 
@@ -361,7 +361,7 @@
  and not the usual 1, 0!  It is enclosed within two quantifiers.*}
 
 (*  is_Replace :: "[i=>o,i,[i,i]=>o,i] => o"
-    "is_Replace(M,A,P,z) == \<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & P(x,u))" *)
+    "is_Replace(M,A,P,z) == \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,u))" *)
 
 definition
   Replace_fm :: "[i, i, i]=>i" where
@@ -377,21 +377,21 @@
 lemma sats_Replace_fm:
   assumes is_P_iff_sats: 
       "!!a b. [|a \<in> A; b \<in> A|] 
-              ==> is_P(a,b) <-> sats(A, p, Cons(a,Cons(b,env)))"
+              ==> is_P(a,b) \<longleftrightarrow> sats(A, p, Cons(a,Cons(b,env)))"
   shows 
       "[|x \<in> nat; y \<in> nat; env \<in> list(A)|]
-       ==> sats(A, Replace_fm(x,p,y), env) <->
+       ==> sats(A, Replace_fm(x,p,y), env) \<longleftrightarrow>
            is_Replace(##A, nth(x,env), is_P, nth(y,env))"
 by (simp add: Replace_fm_def is_Replace_def is_P_iff_sats [THEN iff_sym])
 
 lemma Replace_iff_sats:
   assumes is_P_iff_sats: 
       "!!a b. [|a \<in> A; b \<in> A|] 
-              ==> is_P(a,b) <-> sats(A, p, Cons(a,Cons(b,env)))"
+              ==> is_P(a,b) \<longleftrightarrow> sats(A, p, Cons(a,Cons(b,env)))"
   shows 
   "[| nth(i,env) = x; nth(j,env) = y;
       i \<in> nat; j \<in> nat; env \<in> list(A)|]
-   ==> is_Replace(##A, x, is_P, y) <-> sats(A, Replace_fm(i,p,j), env)"
+   ==> is_Replace(##A, x, is_P, y) \<longleftrightarrow> sats(A, Replace_fm(i,p,j), env)"
 by (simp add: sats_Replace_fm [OF is_P_iff_sats])
 
 
@@ -412,7 +412,7 @@
 subsubsection{*The Operator @{term is_DPow'}, Internalized*}
 
 (*  "is_DPow'(M,A,Z) == 
-       \<forall>X[M]. X \<in> Z <-> 
+       \<forall>X[M]. X \<in> Z \<longleftrightarrow> 
          subset(M,X,A) & 
            (\<exists>env[M]. \<exists>p[M]. mem_formula(M,p) & mem_list(M,A,env) &
                     is_Collect(M, A, is_DPow_sats(M,A,env,p), X))" *)
@@ -435,14 +435,14 @@
 
 lemma sats_DPow'_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, DPow'_fm(x,y), env) <->
+    ==> sats(A, DPow'_fm(x,y), env) \<longleftrightarrow>
         is_DPow'(##A, nth(x,env), nth(y,env))"
 by (simp add: DPow'_fm_def is_DPow'_def sats_subset_fm' sats_Collect_fm)
 
 lemma DPow'_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; 
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> is_DPow'(##A, x, y) <-> sats(A, DPow'_fm(i,j), env)"
+       ==> is_DPow'(##A, x, y) \<longleftrightarrow> sats(A, DPow'_fm(i,j), env)"
 by (simp add: sats_DPow'_fm)
 
 theorem DPow'_reflection:
@@ -463,7 +463,7 @@
 
 lemma (in M_DPow) transrec_body_abs:
      "[|M(x); M(g); M(z)|]
-    ==> transrec_body(M,g,x,y,z) <-> y \<in> x & z = DPow'(g`y)"
+    ==> transrec_body(M,g,x,y,z) \<longleftrightarrow> y \<in> x & z = DPow'(g`y)"
 by (simp add: transrec_body_def DPow'_abs transM [of _ x])
 
 locale M_Lset = M_DPow +
@@ -490,7 +490,7 @@
 
 lemma (in M_Lset) RepFun_DPow_abs:
      "[|M(x); M(f); M(r) |]
-      ==> is_Replace(M, x, \<lambda>y z. transrec_body(M,f,x,y,z), r) <->
+      ==> is_Replace(M, x, \<lambda>y z. transrec_body(M,f,x,y,z), r) \<longleftrightarrow>
           r =  {DPow'(f`y). y\<in>x}"
 apply (simp add: transrec_body_abs RepFun_def) 
 apply (rule iff_trans) 
@@ -517,7 +517,7 @@
 
 lemma (in M_Lset) Lset_abs:
   "[|Ord(i);  M(i);  M(z)|] 
-   ==> is_Lset(M,i,z) <-> z = Lset(i)"
+   ==> is_Lset(M,i,z) \<longleftrightarrow> z = Lset(i)"
 apply (simp add: is_Lset_def Lset_eq_transrec_DPow') 
 apply (rule transrec_abs)  
 apply (simp_all add: transrec_rep' relation2_def RepFun_DPow_apply_closed)
--- a/src/ZF/Constructible/Datatype_absolute.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Datatype_absolute.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -15,15 +15,15 @@
 
 definition
   contin :: "(i=>i) => o" where
-   "contin(h) == (\<forall>A. directed(A) --> h(\<Union>A) = (\<Union>X\<in>A. h(X)))"
+   "contin(h) == (\<forall>A. directed(A) \<longrightarrow> h(\<Union>A) = (\<Union>X\<in>A. h(X)))"
 
-lemma bnd_mono_iterates_subset: "[|bnd_mono(D, h); n \<in> nat|] ==> h^n (0) <= D"
+lemma bnd_mono_iterates_subset: "[|bnd_mono(D, h); n \<in> nat|] ==> h^n (0) \<subseteq> D"
 apply (induct_tac n) 
  apply (simp_all add: bnd_mono_def, blast) 
 done
 
 lemma bnd_mono_increasing [rule_format]:
-     "[|i \<in> nat; j \<in> nat; bnd_mono(D,h)|] ==> i \<le> j --> h^i(0) \<subseteq> h^j(0)"
+     "[|i \<in> nat; j \<in> nat; bnd_mono(D,h)|] ==> i \<le> j \<longrightarrow> h^i(0) \<subseteq> h^j(0)"
 apply (rule_tac m=i and n=j in diff_induct, simp_all)
 apply (blast del: subsetI
              intro: bnd_mono_iterates_subset bnd_monoD2 [of concl: h]) 
@@ -56,14 +56,14 @@
 done
 
 lemma lfp_subset_Union:
-     "[|bnd_mono(D, h); contin(h)|] ==> lfp(D,h) <= (\<Union>n\<in>nat. h^n(0))"
+     "[|bnd_mono(D, h); contin(h)|] ==> lfp(D,h) \<subseteq> (\<Union>n\<in>nat. h^n(0))"
 apply (rule lfp_lowerbound) 
  apply (simp add: contin_iterates_eq) 
 apply (simp add: contin_def bnd_mono_iterates_subset UN_subset_iff) 
 done
 
 lemma Union_subset_lfp:
-     "bnd_mono(D,h) ==> (\<Union>n\<in>nat. h^n(0)) <= lfp(D,h)"
+     "bnd_mono(D,h) ==> (\<Union>n\<in>nat. h^n(0)) \<subseteq> lfp(D,h)"
 apply (simp add: UN_subset_iff)
 apply (rule ballI)  
 apply (induct_tac n, simp_all) 
@@ -128,12 +128,12 @@
 definition
   iterates_replacement :: "[i=>o, [i,i]=>o, i] => o" where
    "iterates_replacement(M,isF,v) ==
-      \<forall>n[M]. n\<in>nat --> 
+      \<forall>n[M]. n\<in>nat \<longrightarrow> 
          wfrec_replacement(M, iterates_MH(M,isF,v), Memrel(succ(n)))"
 
 lemma (in M_basic) iterates_MH_abs:
   "[| relation1(M,isF,F); M(n); M(g); M(z) |] 
-   ==> iterates_MH(M,isF,v,n,g,z) <-> z = nat_case(v, \<lambda>m. F(g`m), n)"
+   ==> iterates_MH(M,isF,v,n,g,z) \<longleftrightarrow> z = nat_case(v, \<lambda>m. F(g`m), n)"
 by (simp add: nat_case_abs [of _ "\<lambda>m. F(g ` m)"]
               relation1_def iterates_MH_def)  
 
@@ -146,7 +146,7 @@
 theorem (in M_trancl) iterates_abs:
   "[| iterates_replacement(M,isF,v); relation1(M,isF,F);
       n \<in> nat; M(v); M(z); \<forall>x[M]. M(F(x)) |] 
-   ==> is_iterates(M,isF,v,n,z) <-> z = iterates(F,n,v)" 
+   ==> is_iterates(M,isF,v,n,z) \<longleftrightarrow> z = iterates(F,n,v)" 
 apply (frule iterates_imp_wfrec_replacement, assumption+)
 apply (simp add: wf_Memrel trans_Memrel relation_Memrel nat_into_M
                  is_iterates_def relation2_def iterates_MH_abs 
@@ -215,7 +215,7 @@
          number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)"
 
 lemma (in M_basic) list_functor_abs [simp]: 
-     "[| M(A); M(X); M(Z) |] ==> is_list_functor(M,A,X,Z) <-> (Z = {0} + A*X)"
+     "[| M(A); M(X); M(Z) |] ==> is_list_functor(M,A,X,Z) \<longleftrightarrow> (Z = {0} + A*X)"
 by (simp add: is_list_functor_def singleton_0 nat_into_M)
 
 
@@ -272,7 +272,7 @@
 
 lemma (in M_basic) formula_functor_abs [simp]: 
      "[| M(X); M(Z) |] 
-      ==> is_formula_functor(M,X,Z) <-> 
+      ==> is_formula_functor(M,X,Z) \<longleftrightarrow> 
           Z = ((nat*nat) + (nat*nat)) + (X*X + X)"
 by (simp add: is_formula_functor_def) 
 
@@ -287,7 +287,7 @@
 by (simp add: list_N_def Nil_def)
 
 lemma Cons_in_list_N [simp]:
-     "Cons(a,l) \<in> list_N(A,succ(n)) <-> a\<in>A & l \<in> list_N(A,n)"
+     "Cons(a,l) \<in> list_N(A,succ(n)) \<longleftrightarrow> a\<in>A & l \<in> list_N(A,n)"
 by (simp add: list_N_def Cons_def) 
 
 text{*These two aren't simprules because they reveal the underlying
@@ -310,7 +310,7 @@
 done
 
 lemma list_imp_list_N [rule_format]:
-     "l \<in> list(A) ==> \<forall>n\<in>nat. length(l) < n --> l \<in> list_N(A, n)"
+     "l \<in> list(A) ==> \<forall>n\<in>nat. length(l) < n \<longrightarrow> l \<in> list_N(A, n)"
 apply (induct_tac l)
 apply (force elim: natE)+
 done
@@ -354,7 +354,7 @@
 
 definition
   is_list :: "[i=>o,i,i] => o" where
-    "is_list(M,A,Z) == \<forall>l[M]. l \<in> Z <-> mem_list(M,A,l)"
+    "is_list(M,A,Z) == \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_list(M,A,l)"
 
 subsubsection{*Towards Absoluteness of @{term formula_rec}*}
 
@@ -374,19 +374,19 @@
     "formula_N(n) == (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0)"
 
 lemma Member_in_formula_N [simp]:
-     "Member(x,y) \<in> formula_N(succ(n)) <-> x \<in> nat & y \<in> nat"
+     "Member(x,y) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> nat & y \<in> nat"
 by (simp add: formula_N_def Member_def) 
 
 lemma Equal_in_formula_N [simp]:
-     "Equal(x,y) \<in> formula_N(succ(n)) <-> x \<in> nat & y \<in> nat"
+     "Equal(x,y) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> nat & y \<in> nat"
 by (simp add: formula_N_def Equal_def) 
 
 lemma Nand_in_formula_N [simp]:
-     "Nand(x,y) \<in> formula_N(succ(n)) <-> x \<in> formula_N(n) & y \<in> formula_N(n)"
+     "Nand(x,y) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> formula_N(n) & y \<in> formula_N(n)"
 by (simp add: formula_N_def Nand_def) 
 
 lemma Forall_in_formula_N [simp]:
-     "Forall(x) \<in> formula_N(succ(n)) <-> x \<in> formula_N(n)"
+     "Forall(x) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> formula_N(n)"
 by (simp add: formula_N_def Forall_def) 
 
 text{*These two aren't simprules because they reveal the underlying
@@ -413,7 +413,7 @@
 done
 
 lemma formula_imp_formula_N [rule_format]:
-     "p \<in> formula ==> \<forall>n\<in>nat. depth(p) < n --> p \<in> formula_N(n)"
+     "p \<in> formula ==> \<forall>n\<in>nat. depth(p) < n \<longrightarrow> p \<in> formula_N(n)"
 apply (induct_tac p)
 apply (simp_all add: succ_Un_distrib Un_least_lt_iff) 
 apply (force elim: natE)+
@@ -432,7 +432,7 @@
 
 text{*This result and the next are unused.*}
 lemma formula_N_mono [rule_format]:
-  "[| m \<in> nat; n \<in> nat |] ==> m\<le>n --> formula_N(m) \<subseteq> formula_N(n)"
+  "[| m \<in> nat; n \<in> nat |] ==> m\<le>n \<longrightarrow> formula_N(m) \<subseteq> formula_N(n)"
 apply (rule_tac m = m and n = n in diff_induct)
 apply (simp_all add: formula_N_0 formula_N_succ, blast) 
 done
@@ -459,7 +459,7 @@
 
 definition
   is_formula :: "[i=>o,i] => o" where
-    "is_formula(M,Z) == \<forall>p[M]. p \<in> Z <-> mem_formula(M,p)"
+    "is_formula(M,Z) == \<forall>p[M]. p \<in> Z \<longleftrightarrow> mem_formula(M,p)"
 
 locale M_datatypes = M_trancl +
  assumes list_replacement1:
@@ -498,7 +498,7 @@
 
 lemma (in M_datatypes) list_N_abs [simp]:
      "[|M(A); n\<in>nat; M(Z)|]
-      ==> is_list_N(M,A,n,Z) <-> Z = list_N(A,n)"
+      ==> is_list_N(M,A,n,Z) \<longleftrightarrow> Z = list_N(A,n)"
 apply (insert list_replacement1)
 apply (simp add: is_list_N_def list_N_def relation1_def nat_into_M
                  iterates_abs [of "is_list_functor(M,A)" _ "\<lambda>X. {0} + A*X"])
@@ -512,14 +512,14 @@
 done
 
 lemma (in M_datatypes) mem_list_abs [simp]:
-     "M(A) ==> mem_list(M,A,l) <-> l \<in> list(A)"
+     "M(A) ==> mem_list(M,A,l) \<longleftrightarrow> l \<in> list(A)"
 apply (insert list_replacement1)
 apply (simp add: mem_list_def list_N_def relation1_def list_eq_Union
                  iterates_closed [of "is_list_functor(M,A)"])
 done
 
 lemma (in M_datatypes) list_abs [simp]:
-     "[|M(A); M(Z)|] ==> is_list(M,A,Z) <-> Z = list(A)"
+     "[|M(A); M(Z)|] ==> is_list(M,A,Z) \<longleftrightarrow> Z = list(A)"
 apply (simp add: is_list_def, safe)
 apply (rule M_equalityI, simp_all)
 done
@@ -546,7 +546,7 @@
 
 lemma (in M_datatypes) formula_N_abs [simp]:
      "[|n\<in>nat; M(Z)|]
-      ==> is_formula_N(M,n,Z) <-> Z = formula_N(n)"
+      ==> is_formula_N(M,n,Z) \<longleftrightarrow> Z = formula_N(n)"
 apply (insert formula_replacement1)
 apply (simp add: is_formula_N_def formula_N_def relation1_def nat_into_M
                  iterates_abs [of "is_formula_functor(M)" _
@@ -561,14 +561,14 @@
 done
 
 lemma (in M_datatypes) mem_formula_abs [simp]:
-     "mem_formula(M,l) <-> l \<in> formula"
+     "mem_formula(M,l) \<longleftrightarrow> l \<in> formula"
 apply (insert formula_replacement1)
 apply (simp add: mem_formula_def relation1_def formula_eq_Union formula_N_def
                  iterates_closed [of "is_formula_functor(M)"])
 done
 
 lemma (in M_datatypes) formula_abs [simp]:
-     "[|M(Z)|] ==> is_formula(M,Z) <-> Z = formula"
+     "[|M(Z)|] ==> is_formula(M,Z) \<longleftrightarrow> Z = formula"
 apply (simp add: is_formula_def, safe)
 apply (rule M_equalityI, simp_all)
 done
@@ -599,7 +599,7 @@
 
 definition
   is_eclose :: "[i=>o,i,i] => o" where
-    "is_eclose(M,A,Z) == \<forall>u[M]. u \<in> Z <-> mem_eclose(M,A,u)"
+    "is_eclose(M,A,Z) == \<forall>u[M]. u \<in> Z \<longleftrightarrow> mem_eclose(M,A,u)"
 
 
 locale M_eclose = M_datatypes +
@@ -625,21 +625,21 @@
                iterates_closed [of "big_union(M)"])
 
 lemma (in M_eclose) is_eclose_n_abs [simp]:
-     "[|M(A); n\<in>nat; M(Z)|] ==> is_eclose_n(M,A,n,Z) <-> Z = Union^n (A)"
+     "[|M(A); n\<in>nat; M(Z)|] ==> is_eclose_n(M,A,n,Z) \<longleftrightarrow> Z = Union^n (A)"
 apply (insert eclose_replacement1)
 apply (simp add: is_eclose_n_def relation1_def nat_into_M
                  iterates_abs [of "big_union(M)" _ "Union"])
 done
 
 lemma (in M_eclose) mem_eclose_abs [simp]:
-     "M(A) ==> mem_eclose(M,A,l) <-> l \<in> eclose(A)"
+     "M(A) ==> mem_eclose(M,A,l) \<longleftrightarrow> l \<in> eclose(A)"
 apply (insert eclose_replacement1)
 apply (simp add: mem_eclose_def relation1_def eclose_eq_Union
                  iterates_closed [of "big_union(M)"])
 done
 
 lemma (in M_eclose) eclose_abs [simp]:
-     "[|M(A); M(Z)|] ==> is_eclose(M,A,Z) <-> Z = eclose(A)"
+     "[|M(A); M(Z)|] ==> is_eclose(M,A,Z) \<longleftrightarrow> Z = eclose(A)"
 apply (simp add: is_eclose_def, safe)
 apply (rule M_equalityI, simp_all)
 done
@@ -669,8 +669,8 @@
 theorem (in M_eclose) transrec_abs:
   "[|transrec_replacement(M,MH,i);  relation2(M,MH,H);
      Ord(i);  M(i);  M(z);
-     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|]
-   ==> is_transrec(M,MH,i,z) <-> z = transrec(i,H)"
+     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]
+   ==> is_transrec(M,MH,i,z) \<longleftrightarrow> z = transrec(i,H)"
 by (simp add: trans_wfrec_abs transrec_replacement_def is_transrec_def
        transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel)
 
@@ -678,7 +678,7 @@
 theorem (in M_eclose) transrec_closed:
      "[|transrec_replacement(M,MH,i);  relation2(M,MH,H);
         Ord(i);  M(i);
-        \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|]
+        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]
       ==> M(transrec(i,H))"
 by (simp add: trans_wfrec_closed transrec_replacement_def is_transrec_def
         transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel)
@@ -706,7 +706,7 @@
 
 
 lemma (in M_datatypes) length_abs [simp]:
-     "[|M(A); l \<in> list(A); n \<in> nat|] ==> is_length(M,A,l,n) <-> n = length(l)"
+     "[|M(A); l \<in> list(A); n \<in> nat|] ==> is_length(M,A,l,n) \<longleftrightarrow> n = length(l)"
 apply (subgoal_tac "M(l) & M(n)")
  prefer 2 apply (blast dest: transM)
 apply (simp add: is_length_def)
@@ -752,7 +752,7 @@
 
 lemma (in M_datatypes) nth_abs [simp]:
      "[|M(A); n \<in> nat; l \<in> list(A); M(Z)|]
-      ==> is_nth(M,n,l,Z) <-> Z = nth(n,l)"
+      ==> is_nth(M,n,l,Z) \<longleftrightarrow> Z = nth(n,l)"
 apply (subgoal_tac "M(l)")
  prefer 2 apply (blast intro: transM)
 apply (simp add: is_nth_def nth_eq_hd_iterates_tl nat_into_M
@@ -770,11 +770,11 @@
         \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)"
 
 lemma (in M_trivial) Member_abs [simp]:
-     "[|M(x); M(y); M(Z)|] ==> is_Member(M,x,y,Z) <-> (Z = Member(x,y))"
+     "[|M(x); M(y); M(Z)|] ==> is_Member(M,x,y,Z) \<longleftrightarrow> (Z = Member(x,y))"
 by (simp add: is_Member_def Member_def)
 
 lemma (in M_trivial) Member_in_M_iff [iff]:
-     "M(Member(x,y)) <-> M(x) & M(y)"
+     "M(Member(x,y)) \<longleftrightarrow> M(x) & M(y)"
 by (simp add: Member_def)
 
 definition
@@ -784,10 +784,10 @@
         \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)"
 
 lemma (in M_trivial) Equal_abs [simp]:
-     "[|M(x); M(y); M(Z)|] ==> is_Equal(M,x,y,Z) <-> (Z = Equal(x,y))"
+     "[|M(x); M(y); M(Z)|] ==> is_Equal(M,x,y,Z) \<longleftrightarrow> (Z = Equal(x,y))"
 by (simp add: is_Equal_def Equal_def)
 
-lemma (in M_trivial) Equal_in_M_iff [iff]: "M(Equal(x,y)) <-> M(x) & M(y)"
+lemma (in M_trivial) Equal_in_M_iff [iff]: "M(Equal(x,y)) \<longleftrightarrow> M(x) & M(y)"
 by (simp add: Equal_def)
 
 definition
@@ -797,10 +797,10 @@
         \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)"
 
 lemma (in M_trivial) Nand_abs [simp]:
-     "[|M(x); M(y); M(Z)|] ==> is_Nand(M,x,y,Z) <-> (Z = Nand(x,y))"
+     "[|M(x); M(y); M(Z)|] ==> is_Nand(M,x,y,Z) \<longleftrightarrow> (Z = Nand(x,y))"
 by (simp add: is_Nand_def Nand_def)
 
-lemma (in M_trivial) Nand_in_M_iff [iff]: "M(Nand(x,y)) <-> M(x) & M(y)"
+lemma (in M_trivial) Nand_in_M_iff [iff]: "M(Nand(x,y)) \<longleftrightarrow> M(x) & M(y)"
 by (simp add: Nand_def)
 
 definition
@@ -809,10 +809,10 @@
     "is_Forall(M,p,Z) == \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)"
 
 lemma (in M_trivial) Forall_abs [simp]:
-     "[|M(x); M(Z)|] ==> is_Forall(M,x,Z) <-> (Z = Forall(x))"
+     "[|M(x); M(Z)|] ==> is_Forall(M,x,Z) \<longleftrightarrow> (Z = Forall(x))"
 by (simp add: is_Forall_def Forall_def)
 
-lemma (in M_trivial) Forall_in_M_iff [iff]: "M(Forall(x)) <-> M(x)"
+lemma (in M_trivial) Forall_in_M_iff [iff]: "M(Forall(x)) \<longleftrightarrow> M(x)"
 by (simp add: Forall_def)
 
 
@@ -862,7 +862,7 @@
 
 
 lemma (in M_datatypes) depth_abs [simp]:
-     "[|p \<in> formula; n \<in> nat|] ==> is_depth(M,p,n) <-> n = depth(p)"
+     "[|p \<in> formula; n \<in> nat|] ==> is_depth(M,p,n) \<longleftrightarrow> n = depth(p)"
 apply (subgoal_tac "M(p) & M(n)")
  prefer 2 apply (blast dest: transM)
 apply (simp add: is_depth_def)
@@ -883,19 +883,19 @@
     "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" where
   --{*no constraint on non-formulas*}
   "is_formula_case(M, is_a, is_b, is_c, is_d, p, z) ==
-      (\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) --> finite_ordinal(M,y) -->
-                      is_Member(M,x,y,p) --> is_a(x,y,z)) &
-      (\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) --> finite_ordinal(M,y) -->
-                      is_Equal(M,x,y,p) --> is_b(x,y,z)) &
-      (\<forall>x[M]. \<forall>y[M]. mem_formula(M,x) --> mem_formula(M,y) -->
-                     is_Nand(M,x,y,p) --> is_c(x,y,z)) &
-      (\<forall>x[M]. mem_formula(M,x) --> is_Forall(M,x,p) --> is_d(x,z))"
+      (\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) \<longrightarrow> finite_ordinal(M,y) \<longrightarrow>
+                      is_Member(M,x,y,p) \<longrightarrow> is_a(x,y,z)) &
+      (\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) \<longrightarrow> finite_ordinal(M,y) \<longrightarrow>
+                      is_Equal(M,x,y,p) \<longrightarrow> is_b(x,y,z)) &
+      (\<forall>x[M]. \<forall>y[M]. mem_formula(M,x) \<longrightarrow> mem_formula(M,y) \<longrightarrow>
+                     is_Nand(M,x,y,p) \<longrightarrow> is_c(x,y,z)) &
+      (\<forall>x[M]. mem_formula(M,x) \<longrightarrow> is_Forall(M,x,p) \<longrightarrow> is_d(x,z))"
 
 lemma (in M_datatypes) formula_case_abs [simp]:
      "[| Relation2(M,nat,nat,is_a,a); Relation2(M,nat,nat,is_b,b);
          Relation2(M,formula,formula,is_c,c); Relation1(M,formula,is_d,d);
          p \<in> formula; M(z) |]
-      ==> is_formula_case(M,is_a,is_b,is_c,is_d,p,z) <->
+      ==> is_formula_case(M,is_a,is_b,is_c,is_d,p,z) \<longleftrightarrow>
           z = formula_case(a,b,c,d,p)"
 apply (simp add: formula_into_M is_formula_case_def)
 apply (erule formula.cases)
@@ -904,10 +904,10 @@
 
 lemma (in M_datatypes) formula_case_closed [intro,simp]:
   "[|p \<in> formula;
-     \<forall>x[M]. \<forall>y[M]. x\<in>nat --> y\<in>nat --> M(a(x,y));
-     \<forall>x[M]. \<forall>y[M]. x\<in>nat --> y\<in>nat --> M(b(x,y));
-     \<forall>x[M]. \<forall>y[M]. x\<in>formula --> y\<in>formula --> M(c(x,y));
-     \<forall>x[M]. x\<in>formula --> M(d(x))|] ==> M(formula_case(a,b,c,d,p))"
+     \<forall>x[M]. \<forall>y[M]. x\<in>nat \<longrightarrow> y\<in>nat \<longrightarrow> M(a(x,y));
+     \<forall>x[M]. \<forall>y[M]. x\<in>nat \<longrightarrow> y\<in>nat \<longrightarrow> M(b(x,y));
+     \<forall>x[M]. \<forall>y[M]. x\<in>formula \<longrightarrow> y\<in>formula \<longrightarrow> M(c(x,y));
+     \<forall>x[M]. x\<in>formula \<longrightarrow> M(d(x))|] ==> M(formula_case(a,b,c,d,p))"
 by (erule formula.cases, simp_all)
 
 
@@ -946,7 +946,7 @@
   fixes a and is_a and b and is_b and c and is_c and d and is_d and MH
   defines
       "MH(u::i,f,z) ==
-        \<forall>fml[M]. is_formula(M,fml) -->
+        \<forall>fml[M]. is_formula(M,fml) \<longrightarrow>
              is_lambda
          (M, fml, is_formula_case (M, is_a, is_b, is_c(f), is_d(f)), z)"
 
@@ -1003,7 +1003,7 @@
 
 theorem (in Formula_Rec) formula_rec_abs:
   "[| p \<in> formula; M(z)|]
-   ==> is_formula_rec(M,MH,p,z) <-> z = formula_rec(a,b,c,d,p)"
+   ==> is_formula_rec(M,MH,p,z) \<longleftrightarrow> z = formula_rec(a,b,c,d,p)"
 by (simp add: is_formula_rec_def formula_rec_eq transM [OF _ formula_closed]
               transrec_abs [OF fr_replace MH_rel2] depth_type
               fr_transrec_closed formula_rec_lam_closed eq_commute)
--- a/src/ZF/Constructible/Formula.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Formula.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -45,38 +45,38 @@
   "Exists(p) == Neg(Forall(Neg(p)))";
 
 lemma Neg_type [TC]: "p \<in> formula ==> Neg(p) \<in> formula"
-by (simp add: Neg_def) 
+by (simp add: Neg_def)
 
 lemma And_type [TC]: "[| p \<in> formula; q \<in> formula |] ==> And(p,q) \<in> formula"
-by (simp add: And_def) 
+by (simp add: And_def)
 
 lemma Or_type [TC]: "[| p \<in> formula; q \<in> formula |] ==> Or(p,q) \<in> formula"
-by (simp add: Or_def) 
+by (simp add: Or_def)
 
 lemma Implies_type [TC]:
      "[| p \<in> formula; q \<in> formula |] ==> Implies(p,q) \<in> formula"
-by (simp add: Implies_def) 
+by (simp add: Implies_def)
 
 lemma Iff_type [TC]:
      "[| p \<in> formula; q \<in> formula |] ==> Iff(p,q) \<in> formula"
-by (simp add: Iff_def) 
+by (simp add: Iff_def)
 
 lemma Exists_type [TC]: "p \<in> formula ==> Exists(p) \<in> formula"
-by (simp add: Exists_def) 
+by (simp add: Exists_def)
 
 
 consts   satisfies :: "[i,i]=>i"
 primrec (*explicit lambda is required because the environment varies*)
-  "satisfies(A,Member(x,y)) = 
+  "satisfies(A,Member(x,y)) =
       (\<lambda>env \<in> list(A). bool_of_o (nth(x,env) \<in> nth(y,env)))"
 
-  "satisfies(A,Equal(x,y)) = 
+  "satisfies(A,Equal(x,y)) =
       (\<lambda>env \<in> list(A). bool_of_o (nth(x,env) = nth(y,env)))"
 
   "satisfies(A,Nand(p,q)) =
       (\<lambda>env \<in> list(A). not ((satisfies(A,p)`env) and (satisfies(A,q)`env)))"
 
-  "satisfies(A,Forall(p)) = 
+  "satisfies(A,Forall(p)) =
       (\<lambda>env \<in> list(A). bool_of_o (\<forall>x\<in>A. satisfies(A,p) ` (Cons(x,env)) = 1))"
 
 
@@ -88,57 +88,57 @@
   "sats(A,p,env) == satisfies(A,p)`env = 1"
 
 lemma [simp]:
-  "env \<in> list(A) 
-   ==> sats(A, Member(x,y), env) <-> nth(x,env) \<in> nth(y,env)"
+  "env \<in> list(A)
+   ==> sats(A, Member(x,y), env) \<longleftrightarrow> nth(x,env) \<in> nth(y,env)"
 by simp
 
 lemma [simp]:
-  "env \<in> list(A) 
-   ==> sats(A, Equal(x,y), env) <-> nth(x,env) = nth(y,env)"
+  "env \<in> list(A)
+   ==> sats(A, Equal(x,y), env) \<longleftrightarrow> nth(x,env) = nth(y,env)"
 by simp
 
 lemma sats_Nand_iff [simp]:
-  "env \<in> list(A) 
-   ==> (sats(A, Nand(p,q), env)) <-> ~ (sats(A,p,env) & sats(A,q,env))" 
-by (simp add: Bool.and_def Bool.not_def cond_def) 
+  "env \<in> list(A)
+   ==> (sats(A, Nand(p,q), env)) \<longleftrightarrow> ~ (sats(A,p,env) & sats(A,q,env))"
+by (simp add: Bool.and_def Bool.not_def cond_def)
 
 lemma sats_Forall_iff [simp]:
-  "env \<in> list(A) 
-   ==> sats(A, Forall(p), env) <-> (\<forall>x\<in>A. sats(A, p, Cons(x,env)))"
+  "env \<in> list(A)
+   ==> sats(A, Forall(p), env) \<longleftrightarrow> (\<forall>x\<in>A. sats(A, p, Cons(x,env)))"
 by simp
 
-declare satisfies.simps [simp del]; 
+declare satisfies.simps [simp del];
 
 subsection{*Dividing line between primitive and derived connectives*}
 
 lemma sats_Neg_iff [simp]:
-  "env \<in> list(A) 
-   ==> sats(A, Neg(p), env) <-> ~ sats(A,p,env)"
-by (simp add: Neg_def) 
+  "env \<in> list(A)
+   ==> sats(A, Neg(p), env) \<longleftrightarrow> ~ sats(A,p,env)"
+by (simp add: Neg_def)
 
 lemma sats_And_iff [simp]:
-  "env \<in> list(A) 
-   ==> (sats(A, And(p,q), env)) <-> sats(A,p,env) & sats(A,q,env)"
-by (simp add: And_def) 
+  "env \<in> list(A)
+   ==> (sats(A, And(p,q), env)) \<longleftrightarrow> sats(A,p,env) & sats(A,q,env)"
+by (simp add: And_def)
 
 lemma sats_Or_iff [simp]:
-  "env \<in> list(A) 
-   ==> (sats(A, Or(p,q), env)) <-> sats(A,p,env) | sats(A,q,env)"
+  "env \<in> list(A)
+   ==> (sats(A, Or(p,q), env)) \<longleftrightarrow> sats(A,p,env) | sats(A,q,env)"
 by (simp add: Or_def)
 
 lemma sats_Implies_iff [simp]:
-  "env \<in> list(A) 
-   ==> (sats(A, Implies(p,q), env)) <-> (sats(A,p,env) --> sats(A,q,env))"
-by (simp add: Implies_def, blast) 
+  "env \<in> list(A)
+   ==> (sats(A, Implies(p,q), env)) \<longleftrightarrow> (sats(A,p,env) \<longrightarrow> sats(A,q,env))"
+by (simp add: Implies_def, blast)
 
 lemma sats_Iff_iff [simp]:
-  "env \<in> list(A) 
-   ==> (sats(A, Iff(p,q), env)) <-> (sats(A,p,env) <-> sats(A,q,env))"
-by (simp add: Iff_def, blast) 
+  "env \<in> list(A)
+   ==> (sats(A, Iff(p,q), env)) \<longleftrightarrow> (sats(A,p,env) \<longleftrightarrow> sats(A,q,env))"
+by (simp add: Iff_def, blast)
 
 lemma sats_Exists_iff [simp]:
-  "env \<in> list(A) 
-   ==> sats(A, Exists(p), env) <-> (\<exists>x\<in>A. sats(A, p, Cons(x,env)))"
+  "env \<in> list(A)
+   ==> sats(A, Exists(p), env) \<longleftrightarrow> (\<exists>x\<in>A. sats(A, p, Cons(x,env)))"
 by (simp add: Exists_def)
 
 
@@ -146,50 +146,50 @@
 
 lemma mem_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; env \<in> list(A)|]
-       ==> (x\<in>y) <-> sats(A, Member(i,j), env)" 
+       ==> (x\<in>y) \<longleftrightarrow> sats(A, Member(i,j), env)"
 by (simp add: satisfies.simps)
 
 lemma equal_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; env \<in> list(A)|]
-       ==> (x=y) <-> sats(A, Equal(i,j), env)" 
+       ==> (x=y) \<longleftrightarrow> sats(A, Equal(i,j), env)"
 by (simp add: satisfies.simps)
 
 lemma not_iff_sats:
-      "[| P <-> sats(A,p,env); env \<in> list(A)|]
-       ==> (~P) <-> sats(A, Neg(p), env)"
+      "[| P \<longleftrightarrow> sats(A,p,env); env \<in> list(A)|]
+       ==> (~P) \<longleftrightarrow> sats(A, Neg(p), env)"
 by simp
 
 lemma conj_iff_sats:
-      "[| P <-> sats(A,p,env); Q <-> sats(A,q,env); env \<in> list(A)|]
-       ==> (P & Q) <-> sats(A, And(p,q), env)"
+      "[| P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)|]
+       ==> (P & Q) \<longleftrightarrow> sats(A, And(p,q), env)"
 by (simp add: sats_And_iff)
 
 lemma disj_iff_sats:
-      "[| P <-> sats(A,p,env); Q <-> sats(A,q,env); env \<in> list(A)|]
-       ==> (P | Q) <-> sats(A, Or(p,q), env)"
+      "[| P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)|]
+       ==> (P | Q) \<longleftrightarrow> sats(A, Or(p,q), env)"
 by (simp add: sats_Or_iff)
 
 lemma iff_iff_sats:
-      "[| P <-> sats(A,p,env); Q <-> sats(A,q,env); env \<in> list(A)|]
-       ==> (P <-> Q) <-> sats(A, Iff(p,q), env)"
-by (simp add: sats_Forall_iff) 
+      "[| P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)|]
+       ==> (P \<longleftrightarrow> Q) \<longleftrightarrow> sats(A, Iff(p,q), env)"
+by (simp add: sats_Forall_iff)
 
 lemma imp_iff_sats:
-      "[| P <-> sats(A,p,env); Q <-> sats(A,q,env); env \<in> list(A)|]
-       ==> (P --> Q) <-> sats(A, Implies(p,q), env)"
-by (simp add: sats_Forall_iff) 
+      "[| P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)|]
+       ==> (P \<longrightarrow> Q) \<longleftrightarrow> sats(A, Implies(p,q), env)"
+by (simp add: sats_Forall_iff)
 
 lemma ball_iff_sats:
-      "[| !!x. x\<in>A ==> P(x) <-> sats(A, p, Cons(x, env)); env \<in> list(A)|]
-       ==> (\<forall>x\<in>A. P(x)) <-> sats(A, Forall(p), env)"
-by (simp add: sats_Forall_iff) 
+      "[| !!x. x\<in>A ==> P(x) \<longleftrightarrow> sats(A, p, Cons(x, env)); env \<in> list(A)|]
+       ==> (\<forall>x\<in>A. P(x)) \<longleftrightarrow> sats(A, Forall(p), env)"
+by (simp add: sats_Forall_iff)
 
 lemma bex_iff_sats:
-      "[| !!x. x\<in>A ==> P(x) <-> sats(A, p, Cons(x, env)); env \<in> list(A)|]
-       ==> (\<exists>x\<in>A. P(x)) <-> sats(A, Exists(p), env)"
-by (simp add: sats_Exists_iff) 
+      "[| !!x. x\<in>A ==> P(x) \<longleftrightarrow> sats(A, p, Cons(x, env)); env \<in> list(A)|]
+       ==> (\<exists>x\<in>A. P(x)) \<longleftrightarrow> sats(A, Exists(p), env)"
+by (simp add: sats_Exists_iff)
 
-lemmas FOL_iff_sats = 
+lemmas FOL_iff_sats =
         mem_iff_sats equal_iff_sats not_iff_sats conj_iff_sats
         disj_iff_sats imp_iff_sats iff_iff_sats imp_iff_sats ball_iff_sats
         bex_iff_sats
@@ -209,43 +209,43 @@
 
 
 lemma arity_type [TC]: "p \<in> formula ==> arity(p) \<in> nat"
-by (induct_tac p, simp_all) 
+by (induct_tac p, simp_all)
 
 lemma arity_Neg [simp]: "arity(Neg(p)) = arity(p)"
-by (simp add: Neg_def) 
+by (simp add: Neg_def)
 
 lemma arity_And [simp]: "arity(And(p,q)) = arity(p) \<union> arity(q)"
-by (simp add: And_def) 
+by (simp add: And_def)
 
 lemma arity_Or [simp]: "arity(Or(p,q)) = arity(p) \<union> arity(q)"
-by (simp add: Or_def) 
+by (simp add: Or_def)
 
 lemma arity_Implies [simp]: "arity(Implies(p,q)) = arity(p) \<union> arity(q)"
-by (simp add: Implies_def) 
+by (simp add: Implies_def)
 
 lemma arity_Iff [simp]: "arity(Iff(p,q)) = arity(p) \<union> arity(q)"
 by (simp add: Iff_def, blast)
 
 lemma arity_Exists [simp]: "arity(Exists(p)) = Arith.pred(arity(p))"
-by (simp add: Exists_def) 
+by (simp add: Exists_def)
 
 
 lemma arity_sats_iff [rule_format]:
   "[| p \<in> formula; extra \<in> list(A) |]
-   ==> \<forall>env \<in> list(A). 
-           arity(p) \<le> length(env) --> 
-           sats(A, p, env @ extra) <-> sats(A, p, env)"
+   ==> \<forall>env \<in> list(A).
+           arity(p) \<le> length(env) \<longrightarrow>
+           sats(A, p, env @ extra) \<longleftrightarrow> sats(A, p, env)"
 apply (induct_tac p)
 apply (simp_all add: Arith.pred_def nth_append Un_least_lt_iff nat_imp_quasinat
-                split: split_nat_case, auto) 
+                split: split_nat_case, auto)
 done
 
 lemma arity_sats1_iff:
-  "[| arity(p) \<le> succ(length(env)); p \<in> formula; x \<in> A; env \<in> list(A); 
+  "[| arity(p) \<le> succ(length(env)); p \<in> formula; x \<in> A; env \<in> list(A);
       extra \<in> list(A) |]
-   ==> sats(A, p, Cons(x, env @ extra)) <-> sats(A, p, Cons(x, env))"
+   ==> sats(A, p, Cons(x, env @ extra)) \<longleftrightarrow> sats(A, p, Cons(x, env))"
 apply (insert arity_sats_iff [of p extra A "Cons(x,env)"])
-apply simp 
+apply simp
 done
 
 
@@ -259,30 +259,30 @@
 by (simp add: incr_var_def)
 
 lemma incr_var_le: "nq\<le>x ==> incr_var(x,nq) = succ(x)"
-apply (simp add: incr_var_def) 
-apply (blast dest: lt_trans1) 
+apply (simp add: incr_var_def)
+apply (blast dest: lt_trans1)
 done
 
 consts   incr_bv :: "i=>i"
 primrec
-  "incr_bv(Member(x,y)) = 
+  "incr_bv(Member(x,y)) =
       (\<lambda>nq \<in> nat. Member (incr_var(x,nq), incr_var(y,nq)))"
 
-  "incr_bv(Equal(x,y)) = 
+  "incr_bv(Equal(x,y)) =
       (\<lambda>nq \<in> nat. Equal (incr_var(x,nq), incr_var(y,nq)))"
 
   "incr_bv(Nand(p,q)) =
       (\<lambda>nq \<in> nat. Nand (incr_bv(p)`nq, incr_bv(q)`nq))"
 
-  "incr_bv(Forall(p)) = 
+  "incr_bv(Forall(p)) =
       (\<lambda>nq \<in> nat. Forall (incr_bv(p) ` succ(nq)))"
 
 
 lemma [TC]: "x \<in> nat ==> incr_var(x,nq) \<in> nat"
-by (simp add: incr_var_def) 
+by (simp add: incr_var_def)
 
 lemma incr_bv_type [TC]: "p \<in> formula ==> incr_bv(p) \<in> nat -> formula"
-by (induct_tac p, simp_all) 
+by (induct_tac p, simp_all)
 
 text{*Obviously, @{term DPow} is closed under complements and finite
 intersections and unions.  Needs an inductive lemma to allow two lists of
@@ -290,8 +290,8 @@
 
 lemma sats_incr_bv_iff [rule_format]:
   "[| p \<in> formula; env \<in> list(A); x \<in> A |]
-   ==> \<forall>bvs \<in> list(A). 
-           sats(A, incr_bv(p) ` length(bvs), bvs @ Cons(x,env)) <-> 
+   ==> \<forall>bvs \<in> list(A).
+           sats(A, incr_bv(p) ` length(bvs), bvs @ Cons(x,env)) \<longleftrightarrow>
            sats(A, p, bvs@env)"
 apply (induct_tac p)
 apply (simp_all add: incr_var_def nth_append succ_lt_iff length_type)
@@ -305,35 +305,35 @@
       ==> succ(x) \<union> incr_var(y,nq) = succ(x \<union> y)"
 apply (simp add: incr_var_def Ord_Un_if, auto)
   apply (blast intro: leI)
- apply (simp add: not_lt_iff_le)  
- apply (blast intro: le_anti_sym) 
-apply (blast dest: lt_trans2) 
+ apply (simp add: not_lt_iff_le)
+ apply (blast intro: le_anti_sym)
+apply (blast dest: lt_trans2)
 done
 
 lemma incr_And_lemma:
      "y < x ==> y \<union> succ(x) = succ(x \<union> y)"
-apply (simp add: Ord_Un_if lt_Ord lt_Ord2 succ_lt_iff) 
-apply (blast dest: lt_asym) 
+apply (simp add: Ord_Un_if lt_Ord lt_Ord2 succ_lt_iff)
+apply (blast dest: lt_asym)
 done
 
 lemma arity_incr_bv_lemma [rule_format]:
-  "p \<in> formula 
-   ==> \<forall>n \<in> nat. arity (incr_bv(p) ` n) = 
+  "p \<in> formula
+   ==> \<forall>n \<in> nat. arity (incr_bv(p) ` n) =
                  (if n < arity(p) then succ(arity(p)) else arity(p))"
-apply (induct_tac p) 
+apply (induct_tac p)
 apply (simp_all add: imp_disj not_lt_iff_le Un_least_lt_iff lt_Un_iff le_Un_iff
                      succ_Un_distrib [symmetric] incr_var_lt incr_var_le
                      Un_commute incr_var_lemma Arith.pred_def nat_imp_quasinat
-            split: split_nat_case) 
+            split: split_nat_case)
  txt{*the Forall case reduces to linear arithmetic*}
  prefer 2
- apply clarify 
- apply (blast dest: lt_trans1) 
+ apply clarify
+ apply (blast dest: lt_trans1)
 txt{*left with the And case*}
 apply safe
- apply (blast intro: incr_And_lemma lt_trans1) 
+ apply (blast intro: incr_And_lemma lt_trans1)
 apply (subst incr_And_lemma)
- apply (blast intro: lt_trans1) 
+ apply (blast intro: lt_trans1)
 apply (simp add: Un_commute)
 done
 
@@ -346,26 +346,26 @@
 
 
 lemma incr_bv1_type [TC]: "p \<in> formula ==> incr_bv1(p) \<in> formula"
-by (simp add: incr_bv1_def) 
+by (simp add: incr_bv1_def)
 
 (*For renaming all but the bound variable at level 0*)
 lemma sats_incr_bv1_iff:
   "[| p \<in> formula; env \<in> list(A); x \<in> A; y \<in> A |]
-   ==> sats(A, incr_bv1(p), Cons(x, Cons(y, env))) <-> 
+   ==> sats(A, incr_bv1(p), Cons(x, Cons(y, env))) \<longleftrightarrow>
        sats(A, p, Cons(x,env))"
 apply (insert sats_incr_bv_iff [of p env A y "Cons(x,Nil)"])
-apply (simp add: incr_bv1_def) 
+apply (simp add: incr_bv1_def)
 done
 
 lemma formula_add_params1 [rule_format]:
   "[| p \<in> formula; n \<in> nat; x \<in> A |]
-   ==> \<forall>bvs \<in> list(A). \<forall>env \<in> list(A). 
-          length(bvs) = n --> 
-          sats(A, iterates(incr_bv1, n, p), Cons(x, bvs@env)) <-> 
+   ==> \<forall>bvs \<in> list(A). \<forall>env \<in> list(A).
+          length(bvs) = n \<longrightarrow>
+          sats(A, iterates(incr_bv1, n, p), Cons(x, bvs@env)) \<longleftrightarrow>
           sats(A, p, Cons(x,env))"
-apply (induct_tac n, simp, clarify) 
+apply (induct_tac n, simp, clarify)
 apply (erule list.cases)
-apply (simp_all add: sats_incr_bv1_iff) 
+apply (simp_all add: sats_incr_bv1_iff)
 done
 
 
@@ -374,17 +374,17 @@
    ==> arity(incr_bv1(p)) =
         (if 1 < arity(p) then succ(arity(p)) else arity(p))"
 apply (insert arity_incr_bv_lemma [of p 1])
-apply (simp add: incr_bv1_def) 
+apply (simp add: incr_bv1_def)
 done
 
 lemma arity_iterates_incr_bv1_eq:
   "[| p \<in> formula; n \<in> nat |]
    ==> arity(incr_bv1^n(p)) =
          (if 1 < arity(p) then n #+ arity(p) else arity(p))"
-apply (induct_tac n) 
+apply (induct_tac n)
 apply (simp_all add: arity_incr_bv1_eq)
 apply (simp add: not_lt_iff_le)
-apply (blast intro: le_trans add_le_self2 arity_type) 
+apply (blast intro: le_trans add_le_self2 arity_type)
 done
 
 
@@ -394,112 +394,112 @@
 text{*The definable powerset operation: Kunen's definition VI 1.1, page 165.*}
 definition
   DPow :: "i => i" where
-  "DPow(A) == {X \<in> Pow(A). 
-               \<exists>env \<in> list(A). \<exists>p \<in> formula. 
-                 arity(p) \<le> succ(length(env)) & 
+  "DPow(A) == {X \<in> Pow(A).
+               \<exists>env \<in> list(A). \<exists>p \<in> formula.
+                 arity(p) \<le> succ(length(env)) &
                  X = {x\<in>A. sats(A, p, Cons(x,env))}}"
 
 lemma DPowI:
   "[|env \<in> list(A);  p \<in> formula;  arity(p) \<le> succ(length(env))|]
    ==> {x\<in>A. sats(A, p, Cons(x,env))} \<in> DPow(A)"
-by (simp add: DPow_def, blast) 
+by (simp add: DPow_def, blast)
 
 text{*With this rule we can specify @{term p} later.*}
 lemma DPowI2 [rule_format]:
-  "[|\<forall>x\<in>A. P(x) <-> sats(A, p, Cons(x,env));
+  "[|\<forall>x\<in>A. P(x) \<longleftrightarrow> sats(A, p, Cons(x,env));
      env \<in> list(A);  p \<in> formula;  arity(p) \<le> succ(length(env))|]
    ==> {x\<in>A. P(x)} \<in> DPow(A)"
-by (simp add: DPow_def, blast) 
+by (simp add: DPow_def, blast)
 
 lemma DPowD:
-  "X \<in> DPow(A) 
-   ==> X <= A &
-       (\<exists>env \<in> list(A). 
-        \<exists>p \<in> formula. arity(p) \<le> succ(length(env)) & 
+  "X \<in> DPow(A)
+   ==> X \<subseteq> A &
+       (\<exists>env \<in> list(A).
+        \<exists>p \<in> formula. arity(p) \<le> succ(length(env)) &
                       X = {x\<in>A. sats(A, p, Cons(x,env))})"
-by (simp add: DPow_def) 
+by (simp add: DPow_def)
 
 lemmas DPow_imp_subset = DPowD [THEN conjunct1]
 
 (*Kunen's Lemma VI 1.2*)
-lemma "[| p \<in> formula; env \<in> list(A); arity(p) \<le> succ(length(env)) |] 
+lemma "[| p \<in> formula; env \<in> list(A); arity(p) \<le> succ(length(env)) |]
        ==> {x\<in>A. sats(A, p, Cons(x,env))} \<in> DPow(A)"
 by (blast intro: DPowI)
 
-lemma DPow_subset_Pow: "DPow(A) <= Pow(A)"
+lemma DPow_subset_Pow: "DPow(A) \<subseteq> Pow(A)"
 by (simp add: DPow_def, blast)
 
 lemma empty_in_DPow: "0 \<in> DPow(A)"
 apply (simp add: DPow_def)
-apply (rule_tac x=Nil in bexI) 
- apply (rule_tac x="Neg(Equal(0,0))" in bexI) 
-  apply (auto simp add: Un_least_lt_iff) 
+apply (rule_tac x=Nil in bexI)
+ apply (rule_tac x="Neg(Equal(0,0))" in bexI)
+  apply (auto simp add: Un_least_lt_iff)
 done
 
 lemma Compl_in_DPow: "X \<in> DPow(A) ==> (A-X) \<in> DPow(A)"
-apply (simp add: DPow_def, clarify, auto) 
-apply (rule bexI) 
- apply (rule_tac x="Neg(p)" in bexI) 
-  apply auto 
+apply (simp add: DPow_def, clarify, auto)
+apply (rule bexI)
+ apply (rule_tac x="Neg(p)" in bexI)
+  apply auto
 done
 
-lemma Int_in_DPow: "[| X \<in> DPow(A); Y \<in> DPow(A) |] ==> X Int Y \<in> DPow(A)"
-apply (simp add: DPow_def, auto) 
-apply (rename_tac envp p envq q) 
-apply (rule_tac x="envp@envq" in bexI) 
+lemma Int_in_DPow: "[| X \<in> DPow(A); Y \<in> DPow(A) |] ==> X \<inter> Y \<in> DPow(A)"
+apply (simp add: DPow_def, auto)
+apply (rename_tac envp p envq q)
+apply (rule_tac x="envp@envq" in bexI)
  apply (rule_tac x="And(p, iterates(incr_bv1,length(envp),q))" in bexI)
   apply typecheck
-apply (rule conjI) 
+apply (rule conjI)
 (*finally check the arity!*)
  apply (simp add: arity_iterates_incr_bv1_eq length_app Un_least_lt_iff)
- apply (force intro: add_le_self le_trans) 
-apply (simp add: arity_sats1_iff formula_add_params1, blast) 
+ apply (force intro: add_le_self le_trans)
+apply (simp add: arity_sats1_iff formula_add_params1, blast)
 done
 
-lemma Un_in_DPow: "[| X \<in> DPow(A); Y \<in> DPow(A) |] ==> X Un Y \<in> DPow(A)"
-apply (subgoal_tac "X Un Y = A - ((A-X) Int (A-Y))") 
-apply (simp add: Int_in_DPow Compl_in_DPow) 
-apply (simp add: DPow_def, blast) 
+lemma Un_in_DPow: "[| X \<in> DPow(A); Y \<in> DPow(A) |] ==> X \<union> Y \<in> DPow(A)"
+apply (subgoal_tac "X \<union> Y = A - ((A-X) \<inter> (A-Y))")
+apply (simp add: Int_in_DPow Compl_in_DPow)
+apply (simp add: DPow_def, blast)
 done
 
 lemma singleton_in_DPow: "a \<in> A ==> {a} \<in> DPow(A)"
 apply (simp add: DPow_def)
-apply (rule_tac x="Cons(a,Nil)" in bexI) 
- apply (rule_tac x="Equal(0,1)" in bexI) 
+apply (rule_tac x="Cons(a,Nil)" in bexI)
+ apply (rule_tac x="Equal(0,1)" in bexI)
   apply typecheck
-apply (force simp add: succ_Un_distrib [symmetric])  
+apply (force simp add: succ_Un_distrib [symmetric])
 done
 
 lemma cons_in_DPow: "[| a \<in> A; X \<in> DPow(A) |] ==> cons(a,X) \<in> DPow(A)"
-apply (rule cons_eq [THEN subst]) 
-apply (blast intro: singleton_in_DPow Un_in_DPow) 
+apply (rule cons_eq [THEN subst])
+apply (blast intro: singleton_in_DPow Un_in_DPow)
 done
 
 (*Part of Lemma 1.3*)
 lemma Fin_into_DPow: "X \<in> Fin(A) ==> X \<in> DPow(A)"
-apply (erule Fin.induct) 
- apply (rule empty_in_DPow) 
-apply (blast intro: cons_in_DPow) 
+apply (erule Fin.induct)
+ apply (rule empty_in_DPow)
+apply (blast intro: cons_in_DPow)
 done
 
 text{*@{term DPow} is not monotonic.  For example, let @{term A} be some
 non-constructible set of natural numbers, and let @{term B} be @{term nat}.
-Then @{term "A<=B"} and obviously @{term "A \<in> DPow(A)"} but @{term "A ~:
+Then @{term "A<=B"} and obviously @{term "A \<in> DPow(A)"} but @{term "A \<notin>
 DPow(B)"}.*}
 
-(*This may be true but the proof looks difficult, requiring relativization 
-lemma DPow_insert: "DPow (cons(a,A)) = DPow(A) Un {cons(a,X) . X: DPow(A)}"
+(*This may be true but the proof looks difficult, requiring relativization
+lemma DPow_insert: "DPow (cons(a,A)) = DPow(A) \<union> {cons(a,X) . X: DPow(A)}"
 apply (rule equalityI, safe)
 oops
 *)
 
-lemma Finite_Pow_subset_Pow: "Finite(A) ==> Pow(A) <= DPow(A)" 
+lemma Finite_Pow_subset_Pow: "Finite(A) ==> Pow(A) \<subseteq> DPow(A)"
 by (blast intro: Fin_into_DPow Finite_into_Fin Fin_subset)
 
 lemma Finite_DPow_eq_Pow: "Finite(A) ==> DPow(A) = Pow(A)"
-apply (rule equalityI) 
-apply (rule DPow_subset_Pow) 
-apply (erule Finite_Pow_subset_Pow) 
+apply (rule equalityI)
+apply (rule DPow_subset_Pow)
+apply (erule Finite_Pow_subset_Pow)
 done
 
 
@@ -520,18 +520,18 @@
   "subset_fm(x,y) == Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))"
 
 lemma subset_type [TC]: "[| x \<in> nat; y \<in> nat |] ==> subset_fm(x,y) \<in> formula"
-by (simp add: subset_fm_def) 
+by (simp add: subset_fm_def)
 
 lemma arity_subset_fm [simp]:
      "[| x \<in> nat; y \<in> nat |] ==> arity(subset_fm(x,y)) = succ(x) \<union> succ(y)"
-by (simp add: subset_fm_def succ_Un_distrib [symmetric]) 
+by (simp add: subset_fm_def succ_Un_distrib [symmetric])
 
 lemma sats_subset_fm [simp]:
    "[|x < length(env); y \<in> nat; env \<in> list(A); Transset(A)|]
-    ==> sats(A, subset_fm(x,y), env) <-> nth(x,env) \<subseteq> nth(y,env)"
-apply (frule lt_length_in_nat, assumption)  
-apply (simp add: subset_fm_def Transset_def) 
-apply (blast intro: nth_type) 
+    ==> sats(A, subset_fm(x,y), env) \<longleftrightarrow> nth(x,env) \<subseteq> nth(y,env)"
+apply (frule lt_length_in_nat, assumption)
+apply (simp add: subset_fm_def Transset_def)
+apply (blast intro: nth_type)
 done
 
 subsubsection{*Transitive sets*}
@@ -541,50 +541,50 @@
   "transset_fm(x) == Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))"
 
 lemma transset_type [TC]: "x \<in> nat ==> transset_fm(x) \<in> formula"
-by (simp add: transset_fm_def) 
+by (simp add: transset_fm_def)
 
 lemma arity_transset_fm [simp]:
      "x \<in> nat ==> arity(transset_fm(x)) = succ(x)"
-by (simp add: transset_fm_def succ_Un_distrib [symmetric]) 
+by (simp add: transset_fm_def succ_Un_distrib [symmetric])
 
 lemma sats_transset_fm [simp]:
    "[|x < length(env); env \<in> list(A); Transset(A)|]
-    ==> sats(A, transset_fm(x), env) <-> Transset(nth(x,env))"
-apply (frule lt_nat_in_nat, erule length_type) 
-apply (simp add: transset_fm_def Transset_def) 
-apply (blast intro: nth_type) 
+    ==> sats(A, transset_fm(x), env) \<longleftrightarrow> Transset(nth(x,env))"
+apply (frule lt_nat_in_nat, erule length_type)
+apply (simp add: transset_fm_def Transset_def)
+apply (blast intro: nth_type)
 done
 
 subsubsection{*Ordinals*}
 
 definition
   ordinal_fm :: "i=>i" where
-  "ordinal_fm(x) == 
+  "ordinal_fm(x) ==
     And(transset_fm(x), Forall(Implies(Member(0,succ(x)), transset_fm(0))))"
 
 lemma ordinal_type [TC]: "x \<in> nat ==> ordinal_fm(x) \<in> formula"
-by (simp add: ordinal_fm_def) 
+by (simp add: ordinal_fm_def)
 
 lemma arity_ordinal_fm [simp]:
      "x \<in> nat ==> arity(ordinal_fm(x)) = succ(x)"
-by (simp add: ordinal_fm_def succ_Un_distrib [symmetric]) 
+by (simp add: ordinal_fm_def succ_Un_distrib [symmetric])
 
 lemma sats_ordinal_fm:
    "[|x < length(env); env \<in> list(A); Transset(A)|]
-    ==> sats(A, ordinal_fm(x), env) <-> Ord(nth(x,env))"
-apply (frule lt_nat_in_nat, erule length_type) 
+    ==> sats(A, ordinal_fm(x), env) \<longleftrightarrow> Ord(nth(x,env))"
+apply (frule lt_nat_in_nat, erule length_type)
 apply (simp add: ordinal_fm_def Ord_def Transset_def)
-apply (blast intro: nth_type) 
+apply (blast intro: nth_type)
 done
 
 text{*The subset consisting of the ordinals is definable.  Essential lemma for
 @{text Ord_in_Lset}.  This result is the objective of the present subsection.*}
 theorem Ords_in_DPow: "Transset(A) ==> {x \<in> A. Ord(x)} \<in> DPow(A)"
-apply (simp add: DPow_def Collect_subset) 
-apply (rule_tac x=Nil in bexI) 
- apply (rule_tac x="ordinal_fm(0)" in bexI) 
+apply (simp add: DPow_def Collect_subset)
+apply (rule_tac x=Nil in bexI)
+ apply (rule_tac x="ordinal_fm(0)" in bexI)
 apply (simp_all add: sats_ordinal_fm)
-done 
+done
 
 
 subsection{* Constant Lset: Levels of the Constructible Universe *}
@@ -596,37 +596,37 @@
 definition
   L :: "i=>o" where --{*Kunen's definition VI 1.5, page 167*}
   "L(x) == \<exists>i. Ord(i) & x \<in> Lset(i)"
-  
+
 text{*NOT SUITABLE FOR REWRITING -- RECURSIVE!*}
-lemma Lset: "Lset(i) = (UN j:i. DPow(Lset(j)))"
+lemma Lset: "Lset(i) = (\<Union>j\<in>i. DPow(Lset(j)))"
 by (subst Lset_def [THEN def_transrec], simp)
 
 lemma LsetI: "[|y\<in>x; A \<in> DPow(Lset(y))|] ==> A \<in> Lset(x)";
 by (subst Lset, blast)
 
 lemma LsetD: "A \<in> Lset(x) ==> \<exists>y\<in>x. A \<in> DPow(Lset(y))";
-apply (insert Lset [of x]) 
-apply (blast intro: elim: equalityE) 
+apply (insert Lset [of x])
+apply (blast intro: elim: equalityE)
 done
 
 subsubsection{* Transitivity *}
 
 lemma elem_subset_in_DPow: "[|X \<in> A; X \<subseteq> A|] ==> X \<in> DPow(A)"
 apply (simp add: Transset_def DPow_def)
-apply (rule_tac x="[X]" in bexI) 
- apply (rule_tac x="Member(0,1)" in bexI) 
-  apply (auto simp add: Un_least_lt_iff) 
+apply (rule_tac x="[X]" in bexI)
+ apply (rule_tac x="Member(0,1)" in bexI)
+  apply (auto simp add: Un_least_lt_iff)
 done
 
-lemma Transset_subset_DPow: "Transset(A) ==> A <= DPow(A)"
-apply clarify  
+lemma Transset_subset_DPow: "Transset(A) ==> A \<subseteq> DPow(A)"
+apply clarify
 apply (simp add: Transset_def)
-apply (blast intro: elem_subset_in_DPow) 
+apply (blast intro: elem_subset_in_DPow)
 done
 
 lemma Transset_DPow: "Transset(A) ==> Transset(DPow(A))"
-apply (simp add: Transset_def) 
-apply (blast intro: elem_subset_in_DPow dest: DPowD) 
+apply (simp add: Transset_def)
+apply (blast intro: elem_subset_in_DPow dest: DPowD)
 done
 
 text{*Kunen's VI 1.6 (a)*}
@@ -637,61 +637,61 @@
 done
 
 lemma mem_Lset_imp_subset_Lset: "a \<in> Lset(i) ==> a \<subseteq> Lset(i)"
-apply (insert Transset_Lset) 
-apply (simp add: Transset_def) 
+apply (insert Transset_Lset)
+apply (simp add: Transset_def)
 done
 
 subsubsection{* Monotonicity *}
 
 text{*Kunen's VI 1.6 (b)*}
 lemma Lset_mono [rule_format]:
-     "ALL j. i<=j --> Lset(i) <= Lset(j)"
+     "\<forall>j. i<=j \<longrightarrow> Lset(i) \<subseteq> Lset(j)"
 proof (induct i rule: eps_induct, intro allI impI)
   fix x j
   assume "\<forall>y\<in>x. \<forall>j. y \<subseteq> j \<longrightarrow> Lset(y) \<subseteq> Lset(j)"
      and "x \<subseteq> j"
   thus "Lset(x) \<subseteq> Lset(j)"
-    by (force simp add: Lset [of x] Lset [of j]) 
+    by (force simp add: Lset [of x] Lset [of j])
 qed
 
 text{*This version lets us remove the premise @{term "Ord(i)"} sometimes.*}
 lemma Lset_mono_mem [rule_format]:
-     "ALL j. i:j --> Lset(i) <= Lset(j)"
+     "\<forall>j. i:j \<longrightarrow> Lset(i) \<subseteq> Lset(j)"
 proof (induct i rule: eps_induct, intro allI impI)
   fix x j
   assume "\<forall>y\<in>x. \<forall>j. y \<in> j \<longrightarrow> Lset(y) \<subseteq> Lset(j)"
      and "x \<in> j"
   thus "Lset(x) \<subseteq> Lset(j)"
-    by (force simp add: Lset [of j] 
-              intro!: bexI intro: elem_subset_in_DPow dest: LsetD DPowD) 
+    by (force simp add: Lset [of j]
+              intro!: bexI intro: elem_subset_in_DPow dest: LsetD DPowD)
 qed
 
 
 text{*Useful with Reflection to bump up the ordinal*}
 lemma subset_Lset_ltD: "[|A \<subseteq> Lset(i); i < j|] ==> A \<subseteq> Lset(j)"
-by (blast dest: ltD [THEN Lset_mono_mem]) 
+by (blast dest: ltD [THEN Lset_mono_mem])
 
 subsubsection{* 0, successor and limit equations for Lset *}
 
 lemma Lset_0 [simp]: "Lset(0) = 0"
 by (subst Lset, blast)
 
-lemma Lset_succ_subset1: "DPow(Lset(i)) <= Lset(succ(i))"
+lemma Lset_succ_subset1: "DPow(Lset(i)) \<subseteq> Lset(succ(i))"
 by (subst Lset, rule succI1 [THEN RepFunI, THEN Union_upper])
 
-lemma Lset_succ_subset2: "Lset(succ(i)) <= DPow(Lset(i))"
+lemma Lset_succ_subset2: "Lset(succ(i)) \<subseteq> DPow(Lset(i))"
 apply (subst Lset, rule UN_least)
-apply (erule succE) 
- apply blast 
+apply (erule succE)
+ apply blast
 apply clarify
 apply (rule elem_subset_in_DPow)
  apply (subst Lset)
- apply blast 
-apply (blast intro: dest: DPowD Lset_mono_mem) 
+ apply blast
+apply (blast intro: dest: DPowD Lset_mono_mem)
 done
 
 lemma Lset_succ: "Lset(succ(i)) = DPow(Lset(i))"
-by (intro equalityI Lset_succ_subset1 Lset_succ_subset2) 
+by (intro equalityI Lset_succ_subset1 Lset_succ_subset2)
 
 lemma Lset_Union [simp]: "Lset(\<Union>(X)) = (\<Union>y\<in>X. Lset(y))"
 apply (subst Lset)
@@ -722,7 +722,7 @@
 apply (rule classical)
 apply (rule Limit_Lset_eq [THEN equalityD1, THEN subsetD, THEN UN_E])
   prefer 2 apply assumption
- apply blast 
+ apply blast
 apply (blast intro: ltI  Limit_is_Ord)
 done
 
@@ -734,7 +734,7 @@
 lemma notin_Lset: "x \<notin> Lset(x)"
 apply (rule_tac a=x in eps_induct)
 apply (subst Lset)
-apply (blast dest: DPowD)  
+apply (blast dest: DPowD)
 done
 
 
@@ -743,20 +743,20 @@
 lemma Ords_of_Lset_eq: "Ord(i) ==> {x\<in>Lset(i). Ord(x)} = i"
 apply (erule trans_induct3)
   apply (simp_all add: Lset_succ Limit_Lset_eq Limit_Union_eq)
-txt{*The successor case remains.*} 
+txt{*The successor case remains.*}
 apply (rule equalityI)
 txt{*First inclusion*}
- apply clarify  
- apply (erule Ord_linear_lt, assumption) 
-   apply (blast dest: DPow_imp_subset ltD notE [OF notin_Lset]) 
-  apply blast 
+ apply clarify
+ apply (erule Ord_linear_lt, assumption)
+   apply (blast dest: DPow_imp_subset ltD notE [OF notin_Lset])
+  apply blast
  apply (blast dest: ltD)
 txt{*Opposite inclusion, @{term "succ(x) \<subseteq> DPow(Lset(x)) \<inter> ON"}*}
 apply auto
 txt{*Key case: *}
-  apply (erule subst, rule Ords_in_DPow [OF Transset_Lset]) 
- apply (blast intro: elem_subset_in_DPow dest: OrdmemD elim: equalityE) 
-apply (blast intro: Ord_in_Ord) 
+  apply (erule subst, rule Ords_in_DPow [OF Transset_Lset])
+ apply (blast intro: elem_subset_in_DPow dest: OrdmemD elim: equalityE)
+apply (blast intro: Ord_in_Ord)
 done
 
 
@@ -765,8 +765,8 @@
 
 lemma Ord_in_Lset: "Ord(i) ==> i \<in> Lset(succ(i))"
 apply (simp add: Lset_succ)
-apply (subst Ords_of_Lset_eq [symmetric], assumption, 
-       rule Ords_in_DPow [OF Transset_Lset]) 
+apply (subst Ords_of_Lset_eq [symmetric], assumption,
+       rule Ords_in_DPow [OF Transset_Lset])
 done
 
 lemma Ord_in_L: "Ord(i) ==> L(i)"
@@ -775,34 +775,34 @@
 subsubsection{* Unions *}
 
 lemma Union_in_Lset:
-     "X \<in> Lset(i) ==> Union(X) \<in> Lset(succ(i))"
+     "X \<in> Lset(i) ==> \<Union>(X) \<in> Lset(succ(i))"
 apply (insert Transset_Lset)
 apply (rule LsetI [OF succI1])
-apply (simp add: Transset_def DPow_def) 
+apply (simp add: Transset_def DPow_def)
 apply (intro conjI, blast)
 txt{*Now to create the formula @{term "\<exists>y. y \<in> X \<and> x \<in> y"} *}
-apply (rule_tac x="Cons(X,Nil)" in bexI) 
- apply (rule_tac x="Exists(And(Member(0,2), Member(1,0)))" in bexI) 
+apply (rule_tac x="Cons(X,Nil)" in bexI)
+ apply (rule_tac x="Exists(And(Member(0,2), Member(1,0)))" in bexI)
   apply typecheck
-apply (simp add: succ_Un_distrib [symmetric], blast) 
+apply (simp add: succ_Un_distrib [symmetric], blast)
 done
 
-theorem Union_in_L: "L(X) ==> L(Union(X))"
-by (simp add: L_def, blast dest: Union_in_Lset) 
+theorem Union_in_L: "L(X) ==> L(\<Union>(X))"
+by (simp add: L_def, blast dest: Union_in_Lset)
 
 subsubsection{* Finite sets and ordered pairs *}
 
 lemma singleton_in_Lset: "a: Lset(i) ==> {a} \<in> Lset(succ(i))"
-by (simp add: Lset_succ singleton_in_DPow) 
+by (simp add: Lset_succ singleton_in_DPow)
 
 lemma doubleton_in_Lset:
      "[| a: Lset(i);  b: Lset(i) |] ==> {a,b} \<in> Lset(succ(i))"
-by (simp add: Lset_succ empty_in_DPow cons_in_DPow) 
+by (simp add: Lset_succ empty_in_DPow cons_in_DPow)
 
 lemma Pair_in_Lset:
     "[| a: Lset(i);  b: Lset(i); Ord(i) |] ==> <a,b> \<in> Lset(succ(succ(i)))"
 apply (unfold Pair_def)
-apply (blast intro: doubleton_in_Lset) 
+apply (blast intro: doubleton_in_Lset)
 done
 
 lemmas Lset_UnI1 = Un_upper1 [THEN Lset_mono [THEN subsetD]]
@@ -818,9 +818,9 @@
 done
 
 theorem doubleton_in_L: "[| L(a); L(b) |] ==> L({a, b})"
-apply (simp add: L_def, clarify) 
-apply (drule Ord2_imp_greater_Limit, assumption) 
-apply (blast intro: lt_LsetI doubleton_in_LLimit Limit_is_Ord) 
+apply (simp add: L_def, clarify)
+apply (drule Ord2_imp_greater_Limit, assumption)
+apply (blast intro: lt_LsetI doubleton_in_LLimit Limit_is_Ord)
 done
 
 lemma Pair_in_LLimit:
@@ -828,7 +828,7 @@
 txt{*Infer that a, b occur at ordinals x,xa < i.*}
 apply (erule Limit_LsetE, assumption)
 apply (erule Limit_LsetE, assumption)
-txt{*Infer that succ(succ(x Un xa)) < i *}
+txt{*Infer that @{term"succ(succ(x \<union> xa)) < i"} *}
 apply (blast intro: lt_Ord lt_LsetI [OF Pair_in_Lset]
                     Lset_UnI1 Lset_UnI2 Limit_has_succ Un_least_lt)
 done
@@ -849,38 +849,38 @@
 lemma Ord_lrank [simp]: "Ord(lrank(a))"
 by (simp add: lrank_def)
 
-lemma Lset_lrank_lt [rule_format]: "Ord(i) ==> x \<in> Lset(i) --> lrank(x) < i"
+lemma Lset_lrank_lt [rule_format]: "Ord(i) ==> x \<in> Lset(i) \<longrightarrow> lrank(x) < i"
 apply (erule trans_induct3)
-  apply simp   
- apply (simp only: lrank_def) 
- apply (blast intro: Least_le) 
-apply (simp_all add: Limit_Lset_eq) 
-apply (blast intro: ltI Limit_is_Ord lt_trans) 
+  apply simp
+ apply (simp only: lrank_def)
+ apply (blast intro: Least_le)
+apply (simp_all add: Limit_Lset_eq)
+apply (blast intro: ltI Limit_is_Ord lt_trans)
 done
 
 text{*Kunen's VI 1.8.  The proof is much harder than the text would
 suggest.  For a start, it needs the previous lemma, which is proved by
 induction.*}
-lemma Lset_iff_lrank_lt: "Ord(i) ==> x \<in> Lset(i) <-> L(x) & lrank(x) < i"
-apply (simp add: L_def, auto) 
- apply (blast intro: Lset_lrank_lt) 
- apply (unfold lrank_def) 
-apply (drule succI1 [THEN Lset_mono_mem, THEN subsetD]) 
-apply (drule_tac P="\<lambda>i. x \<in> Lset(succ(i))" in LeastI, assumption) 
-apply (blast intro!: le_imp_subset Lset_mono [THEN subsetD]) 
+lemma Lset_iff_lrank_lt: "Ord(i) ==> x \<in> Lset(i) \<longleftrightarrow> L(x) & lrank(x) < i"
+apply (simp add: L_def, auto)
+ apply (blast intro: Lset_lrank_lt)
+ apply (unfold lrank_def)
+apply (drule succI1 [THEN Lset_mono_mem, THEN subsetD])
+apply (drule_tac P="\<lambda>i. x \<in> Lset(succ(i))" in LeastI, assumption)
+apply (blast intro!: le_imp_subset Lset_mono [THEN subsetD])
 done
 
-lemma Lset_succ_lrank_iff [simp]: "x \<in> Lset(succ(lrank(x))) <-> L(x)"
+lemma Lset_succ_lrank_iff [simp]: "x \<in> Lset(succ(lrank(x))) \<longleftrightarrow> L(x)"
 by (simp add: Lset_iff_lrank_lt)
 
 text{*Kunen's VI 1.9 (a)*}
 lemma lrank_of_Ord: "Ord(i) ==> lrank(i) = i"
-apply (unfold lrank_def) 
-apply (rule Least_equality) 
-  apply (erule Ord_in_Lset) 
+apply (unfold lrank_def)
+apply (rule Least_equality)
+  apply (erule Ord_in_Lset)
  apply assumption
-apply (insert notin_Lset [of i]) 
-apply (blast intro!: le_imp_subset Lset_mono [THEN subsetD]) 
+apply (insert notin_Lset [of i])
+apply (blast intro!: le_imp_subset Lset_mono [THEN subsetD])
 done
 
 
@@ -889,41 +889,41 @@
 
 text{*Kunen's VI 1.10 *}
 lemma Lset_in_Lset_succ: "Lset(i) \<in> Lset(succ(i))";
-apply (simp add: Lset_succ DPow_def) 
-apply (rule_tac x=Nil in bexI) 
- apply (rule_tac x="Equal(0,0)" in bexI) 
-apply auto 
+apply (simp add: Lset_succ DPow_def)
+apply (rule_tac x=Nil in bexI)
+ apply (rule_tac x="Equal(0,0)" in bexI)
+apply auto
 done
 
 lemma lrank_Lset: "Ord(i) ==> lrank(Lset(i)) = i"
-apply (unfold lrank_def) 
-apply (rule Least_equality) 
-  apply (rule Lset_in_Lset_succ) 
+apply (unfold lrank_def)
+apply (rule Least_equality)
+  apply (rule Lset_in_Lset_succ)
  apply assumption
-apply clarify 
-apply (subgoal_tac "Lset(succ(ia)) <= Lset(i)")
- apply (blast dest: mem_irrefl) 
-apply (blast intro!: le_imp_subset Lset_mono) 
+apply clarify
+apply (subgoal_tac "Lset(succ(ia)) \<subseteq> Lset(i)")
+ apply (blast dest: mem_irrefl)
+apply (blast intro!: le_imp_subset Lset_mono)
 done
 
 text{*Kunen's VI 1.11 *}
-lemma Lset_subset_Vset: "Ord(i) ==> Lset(i) <= Vset(i)";
+lemma Lset_subset_Vset: "Ord(i) ==> Lset(i) \<subseteq> Vset(i)";
 apply (erule trans_induct)
-apply (subst Lset) 
-apply (subst Vset) 
-apply (rule UN_mono [OF subset_refl]) 
-apply (rule subset_trans [OF DPow_subset_Pow]) 
-apply (rule Pow_mono, blast) 
+apply (subst Lset)
+apply (subst Vset)
+apply (rule UN_mono [OF subset_refl])
+apply (rule subset_trans [OF DPow_subset_Pow])
+apply (rule Pow_mono, blast)
 done
 
 text{*Kunen's VI 1.12 *}
 lemma Lset_subset_Vset': "i \<in> nat ==> Lset(i) = Vset(i)";
 apply (erule nat_induct)
- apply (simp add: Vfrom_0) 
-apply (simp add: Lset_succ Vset_succ Finite_Vset Finite_DPow_eq_Pow) 
+ apply (simp add: Vfrom_0)
+apply (simp add: Lset_succ Vset_succ Finite_Vset Finite_DPow_eq_Pow)
 done
 
-text{*Every set of constructible sets is included in some @{term Lset}*} 
+text{*Every set of constructible sets is included in some @{term Lset}*}
 lemma subset_Lset:
      "(\<forall>x\<in>A. L(x)) ==> \<exists>i. Ord(i) & A \<subseteq> Lset(i)"
 by (rule_tac x = "\<Union>x\<in>A. succ(lrank(x))" in exI, force)
@@ -932,32 +932,32 @@
      "[|\<forall>x\<in>A. L(x);
         !!i. [|Ord(i); A \<subseteq> Lset(i)|] ==> P|]
       ==> P"
-by (blast dest: subset_Lset) 
+by (blast dest: subset_Lset)
 
 subsubsection{*For L to satisfy the Powerset axiom *}
 
 lemma LPow_env_typing:
-    "[| y \<in> Lset(i); Ord(i); y \<subseteq> X |] 
+    "[| y \<in> Lset(i); Ord(i); y \<subseteq> X |]
      ==> \<exists>z \<in> Pow(X). y \<in> Lset(succ(lrank(z)))"
-by (auto intro: L_I iff: Lset_succ_lrank_iff) 
+by (auto intro: L_I iff: Lset_succ_lrank_iff)
 
 lemma LPow_in_Lset:
      "[|X \<in> Lset(i); Ord(i)|] ==> \<exists>j. Ord(j) & {y \<in> Pow(X). L(y)} \<in> Lset(j)"
 apply (rule_tac x="succ(\<Union>y \<in> Pow(X). succ(lrank(y)))" in exI)
-apply simp 
+apply simp
 apply (rule LsetI [OF succI1])
-apply (simp add: DPow_def) 
-apply (intro conjI, clarify) 
- apply (rule_tac a=x in UN_I, simp+)  
+apply (simp add: DPow_def)
+apply (intro conjI, clarify)
+ apply (rule_tac a=x in UN_I, simp+)
 txt{*Now to create the formula @{term "y \<subseteq> X"} *}
-apply (rule_tac x="Cons(X,Nil)" in bexI) 
- apply (rule_tac x="subset_fm(0,1)" in bexI) 
+apply (rule_tac x="Cons(X,Nil)" in bexI)
+ apply (rule_tac x="subset_fm(0,1)" in bexI)
   apply typecheck
- apply (rule conjI) 
-apply (simp add: succ_Un_distrib [symmetric]) 
-apply (rule equality_iffI) 
+ apply (rule conjI)
+apply (simp add: succ_Un_distrib [symmetric])
+apply (rule equality_iffI)
 apply (simp add: Transset_UN [OF Transset_Lset] LPow_env_typing)
-apply (auto intro: L_I iff: Lset_succ_lrank_iff) 
+apply (auto intro: L_I iff: Lset_succ_lrank_iff)
 done
 
 theorem LPow_in_L: "L(X) ==> L({y \<in> Pow(X). L(y)})"
@@ -971,7 +971,7 @@
 
 lemma sats_app_0_iff [rule_format]:
   "[| p \<in> formula; 0 \<in> A |]
-   ==> \<forall>env \<in> list(A). sats(A,p, env@[0]) <-> sats(A,p,env)"
+   ==> \<forall>env \<in> list(A). sats(A,p, env@[0]) \<longleftrightarrow> sats(A,p,env)"
 apply (induct_tac p)
 apply (simp_all del: app_Cons add: app_Cons [symmetric]
                 add: nth_zero_eq_0 nth_append not_lt_iff_le nth_eq_0)
@@ -979,17 +979,17 @@
 
 lemma sats_app_zeroes_iff:
   "[| p \<in> formula; 0 \<in> A; env \<in> list(A); n \<in> nat |]
-   ==> sats(A,p,env @ repeat(0,n)) <-> sats(A,p,env)"
-apply (induct_tac n, simp) 
+   ==> sats(A,p,env @ repeat(0,n)) \<longleftrightarrow> sats(A,p,env)"
+apply (induct_tac n, simp)
 apply (simp del: repeat.simps
-            add: repeat_succ_app sats_app_0_iff app_assoc [symmetric]) 
+            add: repeat_succ_app sats_app_0_iff app_assoc [symmetric])
 done
 
 lemma exists_bigger_env:
   "[| p \<in> formula; 0 \<in> A; env \<in> list(A) |]
-   ==> \<exists>env' \<in> list(A). arity(p) \<le> succ(length(env')) & 
-              (\<forall>a\<in>A. sats(A,p,Cons(a,env')) <-> sats(A,p,Cons(a,env)))"
-apply (rule_tac x="env @ repeat(0,arity(p))" in bexI) 
+   ==> \<exists>env' \<in> list(A). arity(p) \<le> succ(length(env')) &
+              (\<forall>a\<in>A. sats(A,p,Cons(a,env')) \<longleftrightarrow> sats(A,p,Cons(a,env)))"
+apply (rule_tac x="env @ repeat(0,arity(p))" in bexI)
 apply (simp del: app_Cons add: app_Cons [symmetric]
             add: length_repeat sats_app_zeroes_iff, typecheck)
 done
@@ -998,28 +998,28 @@
 text{*A simpler version of @{term DPow}: no arity check!*}
 definition
   DPow' :: "i => i" where
-  "DPow'(A) == {X \<in> Pow(A). 
-                \<exists>env \<in> list(A). \<exists>p \<in> formula. 
+  "DPow'(A) == {X \<in> Pow(A).
+                \<exists>env \<in> list(A). \<exists>p \<in> formula.
                     X = {x\<in>A. sats(A, p, Cons(x,env))}}"
 
-lemma DPow_subset_DPow': "DPow(A) <= DPow'(A)";
+lemma DPow_subset_DPow': "DPow(A) \<subseteq> DPow'(A)";
 by (simp add: DPow_def DPow'_def, blast)
 
 lemma DPow'_0: "DPow'(0) = {0}"
 by (auto simp add: DPow'_def)
 
 lemma DPow'_subset_DPow: "0 \<in> A ==> DPow'(A) \<subseteq> DPow(A)"
-apply (auto simp add: DPow'_def DPow_def) 
-apply (frule exists_bigger_env, assumption+, force)  
+apply (auto simp add: DPow'_def DPow_def)
+apply (frule exists_bigger_env, assumption+, force)
 done
 
 lemma DPow_eq_DPow': "Transset(A) ==> DPow(A) = DPow'(A)"
-apply (drule Transset_0_disj) 
-apply (erule disjE) 
- apply (simp add: DPow'_0 Finite_DPow_eq_Pow) 
+apply (drule Transset_0_disj)
+apply (erule disjE)
+ apply (simp add: DPow'_0 Finite_DPow_eq_Pow)
 apply (rule equalityI)
- apply (rule DPow_subset_DPow') 
-apply (erule DPow'_subset_DPow) 
+ apply (rule DPow_subset_DPow')
+apply (erule DPow'_subset_DPow)
 done
 
 text{*And thus we can relativize @{term Lset} without bothering with
@@ -1028,15 +1028,15 @@
 apply (rule_tac a=i in eps_induct)
 apply (subst Lset)
 apply (subst transrec)
-apply (simp only: DPow_eq_DPow' [OF Transset_Lset], simp) 
+apply (simp only: DPow_eq_DPow' [OF Transset_Lset], simp)
 done
 
 text{*With this rule we can specify @{term p} later and don't worry about
       arities at all!*}
 lemma DPow_LsetI [rule_format]:
-  "[|\<forall>x\<in>Lset(i). P(x) <-> sats(Lset(i), p, Cons(x,env));
+  "[|\<forall>x\<in>Lset(i). P(x) \<longleftrightarrow> sats(Lset(i), p, Cons(x,env));
      env \<in> list(Lset(i));  p \<in> formula|]
    ==> {x\<in>Lset(i). P(x)} \<in> DPow(Lset(i))"
-by (simp add: DPow_eq_DPow' [OF Transset_Lset] DPow'_def, blast) 
+by (simp add: DPow_eq_DPow' [OF Transset_Lset] DPow'_def, blast)
 
 end
--- a/src/ZF/Constructible/Internalize.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Internalize.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -19,13 +19,13 @@
 
 lemma sats_Inl_fm [simp]:
    "[| x \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, Inl_fm(x,z), env) <-> is_Inl(##A, nth(x,env), nth(z,env))"
+    ==> sats(A, Inl_fm(x,z), env) \<longleftrightarrow> is_Inl(##A, nth(x,env), nth(z,env))"
 by (simp add: Inl_fm_def is_Inl_def)
 
 lemma Inl_iff_sats:
       "[| nth(i,env) = x; nth(k,env) = z;
           i \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_Inl(##A, x, z) <-> sats(A, Inl_fm(i,k), env)"
+       ==> is_Inl(##A, x, z) \<longleftrightarrow> sats(A, Inl_fm(i,k), env)"
 by simp
 
 theorem Inl_reflection:
@@ -49,13 +49,13 @@
 
 lemma sats_Inr_fm [simp]:
    "[| x \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, Inr_fm(x,z), env) <-> is_Inr(##A, nth(x,env), nth(z,env))"
+    ==> sats(A, Inr_fm(x,z), env) \<longleftrightarrow> is_Inr(##A, nth(x,env), nth(z,env))"
 by (simp add: Inr_fm_def is_Inr_def)
 
 lemma Inr_iff_sats:
       "[| nth(i,env) = x; nth(k,env) = z;
           i \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_Inr(##A, x, z) <-> sats(A, Inr_fm(i,k), env)"
+       ==> is_Inr(##A, x, z) \<longleftrightarrow> sats(A, Inr_fm(i,k), env)"
 by simp
 
 theorem Inr_reflection:
@@ -79,12 +79,12 @@
 
 lemma sats_Nil_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, Nil_fm(x), env) <-> is_Nil(##A, nth(x,env))"
+    ==> sats(A, Nil_fm(x), env) \<longleftrightarrow> is_Nil(##A, nth(x,env))"
 by (simp add: Nil_fm_def is_Nil_def)
 
 lemma Nil_iff_sats:
       "[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|]
-       ==> is_Nil(##A, x) <-> sats(A, Nil_fm(i), env)"
+       ==> is_Nil(##A, x) \<longleftrightarrow> sats(A, Nil_fm(i), env)"
 by simp
 
 theorem Nil_reflection:
@@ -110,14 +110,14 @@
 
 lemma sats_Cons_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, Cons_fm(x,y,z), env) <->
+    ==> sats(A, Cons_fm(x,y,z), env) \<longleftrightarrow>
        is_Cons(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: Cons_fm_def is_Cons_def)
 
 lemma Cons_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==>is_Cons(##A, x, y, z) <-> sats(A, Cons_fm(i,j,k), env)"
+       ==>is_Cons(##A, x, y, z) \<longleftrightarrow> sats(A, Cons_fm(i,j,k), env)"
 by simp
 
 theorem Cons_reflection:
@@ -141,12 +141,12 @@
 
 lemma sats_quasilist_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, quasilist_fm(x), env) <-> is_quasilist(##A, nth(x,env))"
+    ==> sats(A, quasilist_fm(x), env) \<longleftrightarrow> is_quasilist(##A, nth(x,env))"
 by (simp add: quasilist_fm_def is_quasilist_def)
 
 lemma quasilist_iff_sats:
       "[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|]
-       ==> is_quasilist(##A, x) <-> sats(A, quasilist_fm(i), env)"
+       ==> is_quasilist(##A, x) \<longleftrightarrow> sats(A, quasilist_fm(i), env)"
 by simp
 
 theorem quasilist_reflection:
@@ -163,7 +163,7 @@
 subsubsection{*The Formula @{term is_hd}, Internalized*}
 
 (*   "is_hd(M,xs,H) == 
-       (is_Nil(M,xs) --> empty(M,H)) &
+       (is_Nil(M,xs) \<longrightarrow> empty(M,H)) &
        (\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | H=x) &
        (is_quasilist(M,xs) | empty(M,H))" *)
 definition
@@ -179,13 +179,13 @@
 
 lemma sats_hd_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, hd_fm(x,y), env) <-> is_hd(##A, nth(x,env), nth(y,env))"
+    ==> sats(A, hd_fm(x,y), env) \<longleftrightarrow> is_hd(##A, nth(x,env), nth(y,env))"
 by (simp add: hd_fm_def is_hd_def)
 
 lemma hd_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> is_hd(##A, x, y) <-> sats(A, hd_fm(i,j), env)"
+       ==> is_hd(##A, x, y) \<longleftrightarrow> sats(A, hd_fm(i,j), env)"
 by simp
 
 theorem hd_reflection:
@@ -200,7 +200,7 @@
 subsubsection{*The Formula @{term is_tl}, Internalized*}
 
 (*     "is_tl(M,xs,T) ==
-       (is_Nil(M,xs) --> T=xs) &
+       (is_Nil(M,xs) \<longrightarrow> T=xs) &
        (\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | T=l) &
        (is_quasilist(M,xs) | empty(M,T))" *)
 definition
@@ -216,13 +216,13 @@
 
 lemma sats_tl_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, tl_fm(x,y), env) <-> is_tl(##A, nth(x,env), nth(y,env))"
+    ==> sats(A, tl_fm(x,y), env) \<longleftrightarrow> is_tl(##A, nth(x,env), nth(y,env))"
 by (simp add: tl_fm_def is_tl_def)
 
 lemma tl_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> is_tl(##A, x, y) <-> sats(A, tl_fm(i,j), env)"
+       ==> is_tl(##A, x, y) \<longleftrightarrow> sats(A, tl_fm(i,j), env)"
 by simp
 
 theorem tl_reflection:
@@ -251,16 +251,16 @@
 by (simp add: bool_of_o_fm_def)
 
 lemma sats_bool_of_o_fm:
-  assumes p_iff_sats: "P <-> sats(A, p, env)"
+  assumes p_iff_sats: "P \<longleftrightarrow> sats(A, p, env)"
   shows 
       "[|z \<in> nat; env \<in> list(A)|]
-       ==> sats(A, bool_of_o_fm(p,z), env) <->
+       ==> sats(A, bool_of_o_fm(p,z), env) \<longleftrightarrow>
            is_bool_of_o(##A, P, nth(z,env))"
 by (simp add: bool_of_o_fm_def is_bool_of_o_def p_iff_sats [THEN iff_sym])
 
 lemma is_bool_of_o_iff_sats:
-  "[| P <-> sats(A, p, env); nth(k,env) = z; k \<in> nat; env \<in> list(A)|]
-   ==> is_bool_of_o(##A, P, z) <-> sats(A, bool_of_o_fm(p,k), env)"
+  "[| P \<longleftrightarrow> sats(A, p, env); nth(k,env) = z; k \<in> nat; env \<in> list(A)|]
+   ==> is_bool_of_o(##A, P, z) \<longleftrightarrow> sats(A, bool_of_o_fm(p,k), env)"
 by (simp add: sats_bool_of_o_fm)
 
 theorem bool_of_o_reflection:
@@ -281,7 +281,7 @@
 
 (* is_lambda :: "[i=>o, i, [i,i]=>o, i] => o"
     "is_lambda(M, A, is_b, z) == 
-       \<forall>p[M]. p \<in> z <->
+       \<forall>p[M]. p \<in> z \<longleftrightarrow>
         (\<exists>u[M]. \<exists>v[M]. u\<in>A & pair(M,u,v,p) & is_b(u,v))" *)
 definition
   lambda_fm :: "[i, i, i]=>i" where
@@ -302,10 +302,10 @@
   assumes is_b_iff_sats: 
       "!!a0 a1 a2. 
         [|a0\<in>A; a1\<in>A; a2\<in>A|] 
-        ==> is_b(a1, a0) <-> sats(A, p, Cons(a0,Cons(a1,Cons(a2,env))))"
+        ==> is_b(a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,env))))"
   shows 
       "[|x \<in> nat; y \<in> nat; env \<in> list(A)|]
-       ==> sats(A, lambda_fm(p,x,y), env) <-> 
+       ==> sats(A, lambda_fm(p,x,y), env) \<longleftrightarrow> 
            is_lambda(##A, nth(x,env), is_b, nth(y,env))"
 by (simp add: lambda_fm_def is_lambda_def is_b_iff_sats [THEN iff_sym]) 
 
@@ -335,14 +335,14 @@
 
 lemma sats_Member_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, Member_fm(x,y,z), env) <->
+    ==> sats(A, Member_fm(x,y,z), env) \<longleftrightarrow>
         is_Member(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: Member_fm_def is_Member_def)
 
 lemma Member_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_Member(##A, x, y, z) <-> sats(A, Member_fm(i,j,k), env)"
+       ==> is_Member(##A, x, y, z) \<longleftrightarrow> sats(A, Member_fm(i,j,k), env)"
 by (simp add: sats_Member_fm)
 
 theorem Member_reflection:
@@ -368,14 +368,14 @@
 
 lemma sats_Equal_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, Equal_fm(x,y,z), env) <->
+    ==> sats(A, Equal_fm(x,y,z), env) \<longleftrightarrow>
         is_Equal(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: Equal_fm_def is_Equal_def)
 
 lemma Equal_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_Equal(##A, x, y, z) <-> sats(A, Equal_fm(i,j,k), env)"
+       ==> is_Equal(##A, x, y, z) \<longleftrightarrow> sats(A, Equal_fm(i,j,k), env)"
 by (simp add: sats_Equal_fm)
 
 theorem Equal_reflection:
@@ -401,14 +401,14 @@
 
 lemma sats_Nand_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, Nand_fm(x,y,z), env) <->
+    ==> sats(A, Nand_fm(x,y,z), env) \<longleftrightarrow>
         is_Nand(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: Nand_fm_def is_Nand_def)
 
 lemma Nand_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_Nand(##A, x, y, z) <-> sats(A, Nand_fm(i,j,k), env)"
+       ==> is_Nand(##A, x, y, z) \<longleftrightarrow> sats(A, Nand_fm(i,j,k), env)"
 by (simp add: sats_Nand_fm)
 
 theorem Nand_reflection:
@@ -432,14 +432,14 @@
 
 lemma sats_Forall_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, Forall_fm(x,y), env) <->
+    ==> sats(A, Forall_fm(x,y), env) \<longleftrightarrow>
         is_Forall(##A, nth(x,env), nth(y,env))"
 by (simp add: Forall_fm_def is_Forall_def)
 
 lemma Forall_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; 
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> is_Forall(##A, x, y) <-> sats(A, Forall_fm(i,j), env)"
+       ==> is_Forall(##A, x, y) \<longleftrightarrow> sats(A, Forall_fm(i,j), env)"
 by (simp add: sats_Forall_fm)
 
 theorem Forall_reflection:
@@ -466,14 +466,14 @@
 
 lemma sats_and_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, and_fm(x,y,z), env) <->
+    ==> sats(A, and_fm(x,y,z), env) \<longleftrightarrow>
         is_and(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: and_fm_def is_and_def)
 
 lemma is_and_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_and(##A, x, y, z) <-> sats(A, and_fm(i,j,k), env)"
+       ==> is_and(##A, x, y, z) \<longleftrightarrow> sats(A, and_fm(i,j,k), env)"
 by simp
 
 theorem is_and_reflection:
@@ -501,14 +501,14 @@
 
 lemma sats_or_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, or_fm(x,y,z), env) <->
+    ==> sats(A, or_fm(x,y,z), env) \<longleftrightarrow>
         is_or(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: or_fm_def is_or_def)
 
 lemma is_or_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_or(##A, x, y, z) <-> sats(A, or_fm(i,j,k), env)"
+       ==> is_or(##A, x, y, z) \<longleftrightarrow> sats(A, or_fm(i,j,k), env)"
 by simp
 
 theorem is_or_reflection:
@@ -536,13 +536,13 @@
 
 lemma sats_is_not_fm [simp]:
    "[| x \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, not_fm(x,z), env) <-> is_not(##A, nth(x,env), nth(z,env))"
+    ==> sats(A, not_fm(x,z), env) \<longleftrightarrow> is_not(##A, nth(x,env), nth(z,env))"
 by (simp add: not_fm_def is_not_def)
 
 lemma is_not_iff_sats:
       "[| nth(i,env) = x; nth(k,env) = z;
           i \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_not(##A, x, z) <-> sats(A, not_fm(i,k), env)"
+       ==> is_not(##A, x, z) \<longleftrightarrow> sats(A, not_fm(i,k), env)"
 by simp
 
 theorem is_not_reflection:
@@ -565,8 +565,8 @@
 
 text{*Alternative definition, minimizing nesting of quantifiers around MH*}
 lemma M_is_recfun_iff:
-   "M_is_recfun(M,MH,r,a,f) <->
-    (\<forall>z[M]. z \<in> f <-> 
+   "M_is_recfun(M,MH,r,a,f) \<longleftrightarrow>
+    (\<forall>z[M]. z \<in> f \<longleftrightarrow> 
      (\<exists>x[M]. \<exists>f_r_sx[M]. \<exists>y[M]. 
              MH(x, f_r_sx, y) & pair(M,x,y,z) &
              (\<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. 
@@ -580,7 +580,7 @@
 
 (* M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
    "M_is_recfun(M,MH,r,a,f) ==
-     \<forall>z[M]. z \<in> f <->
+     \<forall>z[M]. z \<in> f \<longleftrightarrow>
                2      1           0
 new def     (\<exists>x[M]. \<exists>f_r_sx[M]. \<exists>y[M]. 
              MH(x, f_r_sx, y) & pair(M,x,y,z) &
@@ -614,10 +614,10 @@
   assumes MH_iff_sats: 
       "!!a0 a1 a2 a3. 
         [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A|] 
-        ==> MH(a2, a1, a0) <-> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))"
+        ==> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))"
   shows 
       "[|x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-       ==> sats(A, is_recfun_fm(p,x,y,z), env) <->
+       ==> sats(A, is_recfun_fm(p,x,y,z), env) \<longleftrightarrow>
            M_is_recfun(##A, MH, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: is_recfun_fm_def M_is_recfun_iff MH_iff_sats [THEN iff_sym])
 
@@ -625,11 +625,11 @@
   assumes MH_iff_sats: 
       "!!a0 a1 a2 a3. 
         [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A|] 
-        ==> MH(a2, a1, a0) <-> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))"
+        ==> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))"
   shows
   "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
       i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-   ==> M_is_recfun(##A, MH, x, y, z) <-> sats(A, is_recfun_fm(p,i,j,k), env)"
+   ==> M_is_recfun(##A, MH, x, y, z) \<longleftrightarrow> sats(A, is_recfun_fm(p,i,j,k), env)"
 by (simp add: sats_is_recfun_fm [OF MH_iff_sats]) 
 
 text{*The additional variable in the premise, namely @{term f'}, is essential.
@@ -676,10 +676,10 @@
   assumes MH_iff_sats: 
       "!!a0 a1 a2 a3 a4. 
         [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A|] 
-        ==> MH(a2, a1, a0) <-> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))"
+        ==> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))"
   shows 
       "[|x \<in> nat; y < length(env); z < length(env); env \<in> list(A)|]
-       ==> sats(A, is_wfrec_fm(p,x,y,z), env) <-> 
+       ==> sats(A, is_wfrec_fm(p,x,y,z), env) \<longleftrightarrow> 
            is_wfrec(##A, MH, nth(x,env), nth(y,env), nth(z,env))"
 apply (frule_tac x=z in lt_length_in_nat, assumption)  
 apply (frule lt_length_in_nat, assumption)  
@@ -691,11 +691,11 @@
   assumes MH_iff_sats: 
       "!!a0 a1 a2 a3 a4. 
         [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A|] 
-        ==> MH(a2, a1, a0) <-> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))"
+        ==> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))"
   shows
   "[|nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
       i \<in> nat; j < length(env); k < length(env); env \<in> list(A)|]
-   ==> is_wfrec(##A, MH, x, y, z) <-> sats(A, is_wfrec_fm(p,i,j,k), env)" 
+   ==> is_wfrec(##A, MH, x, y, z) \<longleftrightarrow> sats(A, is_wfrec_fm(p,i,j,k), env)" 
 by (simp add: sats_is_wfrec_fm [OF MH_iff_sats])
 
 theorem is_wfrec_reflection:
@@ -716,7 +716,7 @@
 definition
   cartprod_fm :: "[i,i,i]=>i" where
 (* "cartprod(M,A,B,z) ==
-        \<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))" *)
+        \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))" *)
     "cartprod_fm(A,B,z) ==
        Forall(Iff(Member(0,succ(z)),
                   Exists(And(Member(0,succ(succ(A))),
@@ -729,14 +729,14 @@
 
 lemma sats_cartprod_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, cartprod_fm(x,y,z), env) <->
+    ==> sats(A, cartprod_fm(x,y,z), env) \<longleftrightarrow>
         cartprod(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: cartprod_fm_def cartprod_def)
 
 lemma cartprod_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> cartprod(##A, x, y, z) <-> sats(A, cartprod_fm(i,j,k), env)"
+       ==> cartprod(##A, x, y, z) \<longleftrightarrow> sats(A, cartprod_fm(i,j,k), env)"
 by (simp add: sats_cartprod_fm)
 
 theorem cartprod_reflection:
@@ -769,14 +769,14 @@
 
 lemma sats_sum_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, sum_fm(x,y,z), env) <->
+    ==> sats(A, sum_fm(x,y,z), env) \<longleftrightarrow>
         is_sum(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: sum_fm_def is_sum_def)
 
 lemma sum_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_sum(##A, x, y, z) <-> sats(A, sum_fm(i,j,k), env)"
+       ==> is_sum(##A, x, y, z) \<longleftrightarrow> sats(A, sum_fm(i,j,k), env)"
 by simp
 
 theorem sum_reflection:
@@ -800,13 +800,13 @@
 
 lemma sats_quasinat_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, quasinat_fm(x), env) <-> is_quasinat(##A, nth(x,env))"
+    ==> sats(A, quasinat_fm(x), env) \<longleftrightarrow> is_quasinat(##A, nth(x,env))"
 by (simp add: quasinat_fm_def is_quasinat_def)
 
 lemma quasinat_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; env \<in> list(A)|]
-       ==> is_quasinat(##A, x) <-> sats(A, quasinat_fm(i), env)"
+       ==> is_quasinat(##A, x) \<longleftrightarrow> sats(A, quasinat_fm(i), env)"
 by simp
 
 theorem quasinat_reflection:
@@ -824,8 +824,8 @@
 
 (* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o"
     "is_nat_case(M, a, is_b, k, z) ==
-       (empty(M,k) --> z=a) &
-       (\<forall>m[M]. successor(M,m,k) --> is_b(m,z)) &
+       (empty(M,k) \<longrightarrow> z=a) &
+       (\<forall>m[M]. successor(M,m,k) \<longrightarrow> is_b(m,z)) &
        (is_quasinat(M,k) | empty(M,z))" *)
 text{*The formula @{term is_b} has free variables 1 and 0.*}
 definition
@@ -844,22 +844,22 @@
 
 lemma sats_is_nat_case_fm:
   assumes is_b_iff_sats: 
-      "!!a. a \<in> A ==> is_b(a,nth(z, env)) <-> 
+      "!!a. a \<in> A ==> is_b(a,nth(z, env)) \<longleftrightarrow> 
                       sats(A, p, Cons(nth(z,env), Cons(a, env)))"
   shows 
       "[|x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A)|]
-       ==> sats(A, is_nat_case_fm(x,p,y,z), env) <->
+       ==> sats(A, is_nat_case_fm(x,p,y,z), env) \<longleftrightarrow>
            is_nat_case(##A, nth(x,env), is_b, nth(y,env), nth(z,env))"
 apply (frule lt_length_in_nat, assumption)
 apply (simp add: is_nat_case_fm_def is_nat_case_def is_b_iff_sats [THEN iff_sym])
 done
 
 lemma is_nat_case_iff_sats:
-  "[| (!!a. a \<in> A ==> is_b(a,z) <->
+  "[| (!!a. a \<in> A ==> is_b(a,z) \<longleftrightarrow>
                       sats(A, p, Cons(z, Cons(a,env))));
       nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
       i \<in> nat; j \<in> nat; k < length(env); env \<in> list(A)|]
-   ==> is_nat_case(##A, x, is_b, y, z) <-> sats(A, is_nat_case_fm(i,p,j,k), env)"
+   ==> is_nat_case(##A, x, is_b, y, z) \<longleftrightarrow> sats(A, is_nat_case_fm(i,p,j,k), env)"
 by (simp add: sats_is_nat_case_fm [of A is_b])
 
 
@@ -901,11 +901,11 @@
 lemma sats_iterates_MH_fm:
   assumes is_F_iff_sats:
       "!!a b c d. [| a \<in> A; b \<in> A; c \<in> A; d \<in> A|]
-              ==> is_F(a,b) <->
+              ==> is_F(a,b) \<longleftrightarrow>
                   sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d,env)))))"
   shows 
       "[|v \<in> nat; x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A)|]
-       ==> sats(A, iterates_MH_fm(p,v,x,y,z), env) <->
+       ==> sats(A, iterates_MH_fm(p,v,x,y,z), env) \<longleftrightarrow>
            iterates_MH(##A, is_F, nth(v,env), nth(x,env), nth(y,env), nth(z,env))"
 apply (frule lt_length_in_nat, assumption)  
 apply (simp add: iterates_MH_fm_def iterates_MH_def sats_is_nat_case_fm 
@@ -917,12 +917,12 @@
 lemma iterates_MH_iff_sats:
   assumes is_F_iff_sats:
       "!!a b c d. [| a \<in> A; b \<in> A; c \<in> A; d \<in> A|]
-              ==> is_F(a,b) <->
+              ==> is_F(a,b) \<longleftrightarrow>
                   sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d,env)))))"
   shows 
   "[| nth(i',env) = v; nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
       i' \<in> nat; i \<in> nat; j \<in> nat; k < length(env); env \<in> list(A)|]
-   ==> iterates_MH(##A, is_F, v, x, y, z) <->
+   ==> iterates_MH(##A, is_F, v, x, y, z) \<longleftrightarrow>
        sats(A, iterates_MH_fm(p,i',i,j,k), env)"
 by (simp add: sats_iterates_MH_fm [OF is_F_iff_sats]) 
 
@@ -973,12 +973,12 @@
       "!!a b c d e f g h i j k. 
               [| a \<in> A; b \<in> A; c \<in> A; d \<in> A; e \<in> A; f \<in> A; 
                  g \<in> A; h \<in> A; i \<in> A; j \<in> A; k \<in> A|]
-              ==> is_F(a,b) <->
+              ==> is_F(a,b) \<longleftrightarrow>
                   sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d, Cons(e, Cons(f, 
                       Cons(g, Cons(h, Cons(i, Cons(j, Cons(k, env))))))))))))"
   shows 
       "[|x \<in> nat; y < length(env); z < length(env); env \<in> list(A)|]
-       ==> sats(A, is_iterates_fm(p,x,y,z), env) <->
+       ==> sats(A, is_iterates_fm(p,x,y,z), env) \<longleftrightarrow>
            is_iterates(##A, is_F, nth(x,env), nth(y,env), nth(z,env))"
 apply (frule_tac x=z in lt_length_in_nat, assumption)  
 apply (frule lt_length_in_nat, assumption)  
@@ -992,13 +992,13 @@
       "!!a b c d e f g h i j k. 
               [| a \<in> A; b \<in> A; c \<in> A; d \<in> A; e \<in> A; f \<in> A; 
                  g \<in> A; h \<in> A; i \<in> A; j \<in> A; k \<in> A|]
-              ==> is_F(a,b) <->
+              ==> is_F(a,b) \<longleftrightarrow>
                   sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d, Cons(e, Cons(f, 
                       Cons(g, Cons(h, Cons(i, Cons(j, Cons(k, env))))))))))))"
   shows 
   "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
       i \<in> nat; j < length(env); k < length(env); env \<in> list(A)|]
-   ==> is_iterates(##A, is_F, x, y, z) <->
+   ==> is_iterates(##A, is_F, x, y, z) \<longleftrightarrow>
        sats(A, is_iterates_fm(p,i,j,k), env)"
 by (simp add: sats_is_iterates_fm [OF is_F_iff_sats]) 
 
@@ -1031,7 +1031,7 @@
 
 lemma sats_eclose_n_fm [simp]:
    "[| x \<in> nat; y < length(env); z < length(env); env \<in> list(A)|]
-    ==> sats(A, eclose_n_fm(x,y,z), env) <->
+    ==> sats(A, eclose_n_fm(x,y,z), env) \<longleftrightarrow>
         is_eclose_n(##A, nth(x,env), nth(y,env), nth(z,env))"
 apply (frule_tac x=z in lt_length_in_nat, assumption)  
 apply (frule_tac x=y in lt_length_in_nat, assumption)  
@@ -1042,7 +1042,7 @@
 lemma eclose_n_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j < length(env); k < length(env); env \<in> list(A)|]
-       ==> is_eclose_n(##A, x, y, z) <-> sats(A, eclose_n_fm(i,j,k), env)"
+       ==> is_eclose_n(##A, x, y, z) \<longleftrightarrow> sats(A, eclose_n_fm(i,j,k), env)"
 by (simp add: sats_eclose_n_fm)
 
 theorem eclose_n_reflection:
@@ -1071,13 +1071,13 @@
 
 lemma sats_mem_eclose_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, mem_eclose_fm(x,y), env) <-> mem_eclose(##A, nth(x,env), nth(y,env))"
+    ==> sats(A, mem_eclose_fm(x,y), env) \<longleftrightarrow> mem_eclose(##A, nth(x,env), nth(y,env))"
 by (simp add: mem_eclose_fm_def mem_eclose_def)
 
 lemma mem_eclose_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> mem_eclose(##A, x, y) <-> sats(A, mem_eclose_fm(i,j), env)"
+       ==> mem_eclose(##A, x, y) \<longleftrightarrow> sats(A, mem_eclose_fm(i,j), env)"
 by simp
 
 theorem mem_eclose_reflection:
@@ -1090,7 +1090,7 @@
 
 subsubsection{*The Predicate ``Is @{term "eclose(A)"}''*}
 
-(* is_eclose(M,A,Z) == \<forall>l[M]. l \<in> Z <-> mem_eclose(M,A,l) *)
+(* is_eclose(M,A,Z) == \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_eclose(M,A,l) *)
 definition
   is_eclose_fm :: "[i,i]=>i" where
     "is_eclose_fm(A,Z) ==
@@ -1102,13 +1102,13 @@
 
 lemma sats_is_eclose_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, is_eclose_fm(x,y), env) <-> is_eclose(##A, nth(x,env), nth(y,env))"
+    ==> sats(A, is_eclose_fm(x,y), env) \<longleftrightarrow> is_eclose(##A, nth(x,env), nth(y,env))"
 by (simp add: is_eclose_fm_def is_eclose_def)
 
 lemma is_eclose_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> is_eclose(##A, x, y) <-> sats(A, is_eclose_fm(i,j), env)"
+       ==> is_eclose(##A, x, y) \<longleftrightarrow> sats(A, is_eclose_fm(i,j), env)"
 by simp
 
 theorem is_eclose_reflection:
@@ -1137,14 +1137,14 @@
 
 lemma sats_list_functor_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, list_functor_fm(x,y,z), env) <->
+    ==> sats(A, list_functor_fm(x,y,z), env) \<longleftrightarrow>
         is_list_functor(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: list_functor_fm_def is_list_functor_def)
 
 lemma list_functor_iff_sats:
   "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
       i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-   ==> is_list_functor(##A, x, y, z) <-> sats(A, list_functor_fm(i,j,k), env)"
+   ==> is_list_functor(##A, x, y, z) \<longleftrightarrow> sats(A, list_functor_fm(i,j,k), env)"
 by simp
 
 theorem list_functor_reflection:
@@ -1175,7 +1175,7 @@
 
 lemma sats_list_N_fm [simp]:
    "[| x \<in> nat; y < length(env); z < length(env); env \<in> list(A)|]
-    ==> sats(A, list_N_fm(x,y,z), env) <->
+    ==> sats(A, list_N_fm(x,y,z), env) \<longleftrightarrow>
         is_list_N(##A, nth(x,env), nth(y,env), nth(z,env))"
 apply (frule_tac x=z in lt_length_in_nat, assumption)  
 apply (frule_tac x=y in lt_length_in_nat, assumption)  
@@ -1185,7 +1185,7 @@
 lemma list_N_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j < length(env); k < length(env); env \<in> list(A)|]
-       ==> is_list_N(##A, x, y, z) <-> sats(A, list_N_fm(i,j,k), env)"
+       ==> is_list_N(##A, x, y, z) \<longleftrightarrow> sats(A, list_N_fm(i,j,k), env)"
 by (simp add: sats_list_N_fm)
 
 theorem list_N_reflection:
@@ -1216,13 +1216,13 @@
 
 lemma sats_mem_list_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, mem_list_fm(x,y), env) <-> mem_list(##A, nth(x,env), nth(y,env))"
+    ==> sats(A, mem_list_fm(x,y), env) \<longleftrightarrow> mem_list(##A, nth(x,env), nth(y,env))"
 by (simp add: mem_list_fm_def mem_list_def)
 
 lemma mem_list_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> mem_list(##A, x, y) <-> sats(A, mem_list_fm(i,j), env)"
+       ==> mem_list(##A, x, y) \<longleftrightarrow> sats(A, mem_list_fm(i,j), env)"
 by simp
 
 theorem mem_list_reflection:
@@ -1235,7 +1235,7 @@
 
 subsubsection{*The Predicate ``Is @{term "list(A)"}''*}
 
-(* is_list(M,A,Z) == \<forall>l[M]. l \<in> Z <-> mem_list(M,A,l) *)
+(* is_list(M,A,Z) == \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_list(M,A,l) *)
 definition
   is_list_fm :: "[i,i]=>i" where
     "is_list_fm(A,Z) ==
@@ -1247,13 +1247,13 @@
 
 lemma sats_is_list_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, is_list_fm(x,y), env) <-> is_list(##A, nth(x,env), nth(y,env))"
+    ==> sats(A, is_list_fm(x,y), env) \<longleftrightarrow> is_list(##A, nth(x,env), nth(y,env))"
 by (simp add: is_list_fm_def is_list_def)
 
 lemma is_list_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> is_list(##A, x, y) <-> sats(A, is_list_fm(i,j), env)"
+       ==> is_list(##A, x, y) \<longleftrightarrow> sats(A, is_list_fm(i,j), env)"
 by simp
 
 theorem is_list_reflection:
@@ -1288,14 +1288,14 @@
 
 lemma sats_formula_functor_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, formula_functor_fm(x,y), env) <->
+    ==> sats(A, formula_functor_fm(x,y), env) \<longleftrightarrow>
         is_formula_functor(##A, nth(x,env), nth(y,env))"
 by (simp add: formula_functor_fm_def is_formula_functor_def)
 
 lemma formula_functor_iff_sats:
   "[| nth(i,env) = x; nth(j,env) = y;
       i \<in> nat; j \<in> nat; env \<in> list(A)|]
-   ==> is_formula_functor(##A, x, y) <-> sats(A, formula_functor_fm(i,j), env)"
+   ==> is_formula_functor(##A, x, y) \<longleftrightarrow> sats(A, formula_functor_fm(i,j), env)"
 by simp
 
 theorem formula_functor_reflection:
@@ -1325,7 +1325,7 @@
 
 lemma sats_formula_N_fm [simp]:
    "[| x < length(env); y < length(env); env \<in> list(A)|]
-    ==> sats(A, formula_N_fm(x,y), env) <->
+    ==> sats(A, formula_N_fm(x,y), env) \<longleftrightarrow>
         is_formula_N(##A, nth(x,env), nth(y,env))"
 apply (frule_tac x=y in lt_length_in_nat, assumption)  
 apply (frule lt_length_in_nat, assumption)  
@@ -1335,7 +1335,7 @@
 lemma formula_N_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; 
           i < length(env); j < length(env); env \<in> list(A)|]
-       ==> is_formula_N(##A, x, y) <-> sats(A, formula_N_fm(i,j), env)"
+       ==> is_formula_N(##A, x, y) \<longleftrightarrow> sats(A, formula_N_fm(i,j), env)"
 by (simp add: sats_formula_N_fm)
 
 theorem formula_N_reflection:
@@ -1366,12 +1366,12 @@
 
 lemma sats_mem_formula_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, mem_formula_fm(x), env) <-> mem_formula(##A, nth(x,env))"
+    ==> sats(A, mem_formula_fm(x), env) \<longleftrightarrow> mem_formula(##A, nth(x,env))"
 by (simp add: mem_formula_fm_def mem_formula_def)
 
 lemma mem_formula_iff_sats:
       "[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|]
-       ==> mem_formula(##A, x) <-> sats(A, mem_formula_fm(i), env)"
+       ==> mem_formula(##A, x) \<longleftrightarrow> sats(A, mem_formula_fm(i), env)"
 by simp
 
 theorem mem_formula_reflection:
@@ -1385,7 +1385,7 @@
 
 subsubsection{*The Predicate ``Is @{term "formula"}''*}
 
-(* is_formula(M,Z) == \<forall>p[M]. p \<in> Z <-> mem_formula(M,p) *)
+(* is_formula(M,Z) == \<forall>p[M]. p \<in> Z \<longleftrightarrow> mem_formula(M,p) *)
 definition
   is_formula_fm :: "i=>i" where
     "is_formula_fm(Z) == Forall(Iff(Member(0,succ(Z)), mem_formula_fm(0)))"
@@ -1396,12 +1396,12 @@
 
 lemma sats_is_formula_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, is_formula_fm(x), env) <-> is_formula(##A, nth(x,env))"
+    ==> sats(A, is_formula_fm(x), env) \<longleftrightarrow> is_formula(##A, nth(x,env))"
 by (simp add: is_formula_fm_def is_formula_def)
 
 lemma is_formula_iff_sats:
       "[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|]
-       ==> is_formula(##A, x) <-> sats(A, is_formula_fm(i), env)"
+       ==> is_formula(##A, x) \<longleftrightarrow> sats(A, is_formula_fm(i), env)"
 by simp
 
 theorem is_formula_reflection:
@@ -1443,12 +1443,12 @@
   assumes MH_iff_sats: 
       "!!a0 a1 a2 a3 a4 a5 a6 a7. 
         [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A|] 
-        ==> MH(a2, a1, a0) <-> 
+        ==> MH(a2, a1, a0) \<longleftrightarrow> 
             sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,
                           Cons(a4,Cons(a5,Cons(a6,Cons(a7,env)))))))))"
   shows 
       "[|x < length(env); z < length(env); env \<in> list(A)|]
-       ==> sats(A, is_transrec_fm(p,x,z), env) <-> 
+       ==> sats(A, is_transrec_fm(p,x,z), env) \<longleftrightarrow> 
            is_transrec(##A, MH, nth(x,env), nth(z,env))"
 apply (frule_tac x=z in lt_length_in_nat, assumption)  
 apply (frule_tac x=x in lt_length_in_nat, assumption)  
@@ -1460,13 +1460,13 @@
   assumes MH_iff_sats: 
       "!!a0 a1 a2 a3 a4 a5 a6 a7. 
         [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A|] 
-        ==> MH(a2, a1, a0) <-> 
+        ==> MH(a2, a1, a0) \<longleftrightarrow> 
             sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,
                           Cons(a4,Cons(a5,Cons(a6,Cons(a7,env)))))))))"
   shows
   "[|nth(i,env) = x; nth(k,env) = z; 
       i < length(env); k < length(env); env \<in> list(A)|]
-   ==> is_transrec(##A, MH, x, z) <-> sats(A, is_transrec_fm(p,i,k), env)" 
+   ==> is_transrec(##A, MH, x, z) \<longleftrightarrow> sats(A, is_transrec_fm(p,i,k), env)" 
 by (simp add: sats_is_transrec_fm [OF MH_iff_sats])
 
 theorem is_transrec_reflection:
--- a/src/ZF/Constructible/L_axioms.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/L_axioms.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -26,7 +26,7 @@
 
 theorem Union_ax: "Union_ax(L)"
 apply (simp add: Union_ax_def big_union_def, clarify)
-apply (rule_tac x="Union(x)" in rexI)
+apply (rule_tac x="\<Union>(x)" in rexI)
 apply (simp_all add: Union_in_L, auto)
 apply (blast intro: transL)
 done
@@ -116,7 +116,7 @@
 
 definition
   L_F0 :: "[i=>o,i] => i" where
-    "L_F0(P,y) == \<mu> b. (\<exists>z. L(z) \<and> P(<y,z>)) --> (\<exists>z\<in>Lset(b). P(<y,z>))"
+    "L_F0(P,y) == \<mu> b. (\<exists>z. L(z) \<and> P(<y,z>)) \<longrightarrow> (\<exists>z\<in>Lset(b). P(<y,z>))"
 
 definition
   L_FF :: "[i=>o,i] => i" where
@@ -132,7 +132,7 @@
 definition
   L_Reflects :: "[i=>o,[i,i]=>o] => prop"  ("(3REFLECTS/ [_,/ _])") where
     "REFLECTS[P,Q] == (??Cl. Closed_Unbounded(Cl) &
-                           (\<forall>a. Cl(a) --> (\<forall>x \<in> Lset(a). P(x) <-> Q(a,x))))"
+                           (\<forall>a. Cl(a) \<longrightarrow> (\<forall>x \<in> Lset(a). P(x) \<longleftrightarrow> Q(a,x))))"
 
 
 theorem Triv_reflection:
@@ -169,7 +169,7 @@
 
 theorem Imp_reflection:
      "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |]
-      ==> REFLECTS[\<lambda>x. P(x) --> P'(x), \<lambda>a x. Q(a,x) --> Q'(a,x)]"
+      ==> REFLECTS[\<lambda>x. P(x) \<longrightarrow> P'(x), \<lambda>a x. Q(a,x) \<longrightarrow> Q'(a,x)]"
 apply (unfold L_Reflects_def)
 apply (elim meta_exE)
 apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI)
@@ -178,7 +178,7 @@
 
 theorem Iff_reflection:
      "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |]
-      ==> REFLECTS[\<lambda>x. P(x) <-> P'(x), \<lambda>a x. Q(a,x) <-> Q'(a,x)]"
+      ==> REFLECTS[\<lambda>x. P(x) \<longleftrightarrow> P'(x), \<lambda>a x. Q(a,x) \<longleftrightarrow> Q'(a,x)]"
 apply (unfold L_Reflects_def)
 apply (elim meta_exE)
 apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI)
@@ -202,7 +202,7 @@
 
 theorem All_reflection:
      "REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
-      ==> REFLECTS[\<lambda>x. \<forall>z. L(z) --> P(x,z), \<lambda>a x. \<forall>z\<in>Lset(a). Q(a,x,z)]"
+      ==> REFLECTS[\<lambda>x. \<forall>z. L(z) \<longrightarrow> P(x,z), \<lambda>a x. \<forall>z\<in>Lset(a). Q(a,x,z)]"
 apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def)
 apply (elim meta_exE)
 apply (rule meta_exI)
@@ -247,7 +247,7 @@
 
 lemma ReflectsD:
      "[|REFLECTS[P,Q]; Ord(i)|]
-      ==> \<exists>j. i<j & (\<forall>x \<in> Lset(j). P(x) <-> Q(j,x))"
+      ==> \<exists>j. i<j & (\<forall>x \<in> Lset(j). P(x) \<longleftrightarrow> Q(j,x))"
 apply (unfold L_Reflects_def Closed_Unbounded_def)
 apply (elim meta_exE, clarify)
 apply (blast dest!: UnboundedD)
@@ -255,7 +255,7 @@
 
 lemma ReflectsE:
      "[| REFLECTS[P,Q]; Ord(i);
-         !!j. [|i<j;  \<forall>x \<in> Lset(j). P(x) <-> Q(j,x)|] ==> R |]
+         !!j. [|i<j;  \<forall>x \<in> Lset(j). P(x) \<longleftrightarrow> Q(j,x)|] ==> R |]
       ==> R"
 by (drule ReflectsD, assumption, blast)
 
@@ -301,13 +301,13 @@
 
 lemma sats_empty_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, empty_fm(x), env) <-> empty(##A, nth(x,env))"
+    ==> sats(A, empty_fm(x), env) \<longleftrightarrow> empty(##A, nth(x,env))"
 by (simp add: empty_fm_def empty_def)
 
 lemma empty_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; env \<in> list(A)|]
-       ==> empty(##A, x) <-> sats(A, empty_fm(i), env)"
+       ==> empty(##A, x) \<longleftrightarrow> sats(A, empty_fm(i), env)"
 by simp
 
 theorem empty_reflection:
@@ -320,7 +320,7 @@
 text{*Not used.  But maybe useful?*}
 lemma Transset_sats_empty_fm_eq_0:
    "[| n \<in> nat; env \<in> list(A); Transset(A)|]
-    ==> sats(A, empty_fm(n), env) <-> nth(n,env) = 0"
+    ==> sats(A, empty_fm(n), env) \<longleftrightarrow> nth(n,env) = 0"
 apply (simp add: empty_fm_def empty_def Transset_def, auto)
 apply (case_tac "n < length(env)")
 apply (frule nth_type, assumption+, blast)
@@ -344,20 +344,20 @@
 
 lemma sats_upair_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, upair_fm(x,y,z), env) <->
+    ==> sats(A, upair_fm(x,y,z), env) \<longleftrightarrow>
             upair(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: upair_fm_def upair_def)
 
 lemma upair_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> upair(##A, x, y, z) <-> sats(A, upair_fm(i,j,k), env)"
+       ==> upair(##A, x, y, z) \<longleftrightarrow> sats(A, upair_fm(i,j,k), env)"
 by (simp add: sats_upair_fm)
 
 text{*Useful? At least it refers to "real" unordered pairs*}
 lemma sats_upair_fm2 [simp]:
    "[| x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A); Transset(A)|]
-    ==> sats(A, upair_fm(x,y,z), env) <->
+    ==> sats(A, upair_fm(x,y,z), env) \<longleftrightarrow>
         nth(z,env) = {nth(x,env), nth(y,env)}"
 apply (frule lt_length_in_nat, assumption)
 apply (simp add: upair_fm_def Transset_def, auto)
@@ -386,14 +386,14 @@
 
 lemma sats_pair_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, pair_fm(x,y,z), env) <->
+    ==> sats(A, pair_fm(x,y,z), env) \<longleftrightarrow>
         pair(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: pair_fm_def pair_def)
 
 lemma pair_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> pair(##A, x, y, z) <-> sats(A, pair_fm(i,j,k), env)"
+       ==> pair(##A, x, y, z) \<longleftrightarrow> sats(A, pair_fm(i,j,k), env)"
 by (simp add: sats_pair_fm)
 
 theorem pair_reflection:
@@ -418,14 +418,14 @@
 
 lemma sats_union_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, union_fm(x,y,z), env) <->
+    ==> sats(A, union_fm(x,y,z), env) \<longleftrightarrow>
         union(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: union_fm_def union_def)
 
 lemma union_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> union(##A, x, y, z) <-> sats(A, union_fm(i,j,k), env)"
+       ==> union(##A, x, y, z) \<longleftrightarrow> sats(A, union_fm(i,j,k), env)"
 by (simp add: sats_union_fm)
 
 theorem union_reflection:
@@ -451,14 +451,14 @@
 
 lemma sats_cons_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, cons_fm(x,y,z), env) <->
+    ==> sats(A, cons_fm(x,y,z), env) \<longleftrightarrow>
         is_cons(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: cons_fm_def is_cons_def)
 
 lemma cons_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_cons(##A, x, y, z) <-> sats(A, cons_fm(i,j,k), env)"
+       ==> is_cons(##A, x, y, z) \<longleftrightarrow> sats(A, cons_fm(i,j,k), env)"
 by simp
 
 theorem cons_reflection:
@@ -481,14 +481,14 @@
 
 lemma sats_succ_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, succ_fm(x,y), env) <->
+    ==> sats(A, succ_fm(x,y), env) \<longleftrightarrow>
         successor(##A, nth(x,env), nth(y,env))"
 by (simp add: succ_fm_def successor_def)
 
 lemma successor_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> successor(##A, x, y) <-> sats(A, succ_fm(i,j), env)"
+       ==> successor(##A, x, y) \<longleftrightarrow> sats(A, succ_fm(i,j), env)"
 by simp
 
 theorem successor_reflection:
@@ -512,13 +512,13 @@
 
 lemma sats_number1_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, number1_fm(x), env) <-> number1(##A, nth(x,env))"
+    ==> sats(A, number1_fm(x), env) \<longleftrightarrow> number1(##A, nth(x,env))"
 by (simp add: number1_fm_def number1_def)
 
 lemma number1_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; env \<in> list(A)|]
-       ==> number1(##A, x) <-> sats(A, number1_fm(i), env)"
+       ==> number1(##A, x) \<longleftrightarrow> sats(A, number1_fm(i), env)"
 by simp
 
 theorem number1_reflection:
@@ -531,7 +531,7 @@
 
 subsubsection{*Big Union, Internalized*}
 
-(*  "big_union(M,A,z) == \<forall>x[M]. x \<in> z <-> (\<exists>y[M]. y\<in>A & x \<in> y)" *)
+(*  "big_union(M,A,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A & x \<in> y)" *)
 definition
   big_union_fm :: "[i,i]=>i" where
     "big_union_fm(A,z) ==
@@ -544,14 +544,14 @@
 
 lemma sats_big_union_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, big_union_fm(x,y), env) <->
+    ==> sats(A, big_union_fm(x,y), env) \<longleftrightarrow>
         big_union(##A, nth(x,env), nth(y,env))"
 by (simp add: big_union_fm_def big_union_def)
 
 lemma big_union_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> big_union(##A, x, y) <-> sats(A, big_union_fm(i,j), env)"
+       ==> big_union(##A, x, y) \<longleftrightarrow> sats(A, big_union_fm(i,j), env)"
 by simp
 
 theorem big_union_reflection:
@@ -572,7 +572,7 @@
 
 lemma sats_subset_fm':
    "[|x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, subset_fm(x,y), env) <-> subset(##A, nth(x,env), nth(y,env))"
+    ==> sats(A, subset_fm(x,y), env) \<longleftrightarrow> subset(##A, nth(x,env), nth(y,env))"
 by (simp add: subset_fm_def Relative.subset_def)
 
 theorem subset_reflection:
@@ -584,7 +584,7 @@
 
 lemma sats_transset_fm':
    "[|x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, transset_fm(x), env) <-> transitive_set(##A, nth(x,env))"
+    ==> sats(A, transset_fm(x), env) \<longleftrightarrow> transitive_set(##A, nth(x,env))"
 by (simp add: sats_subset_fm' transset_fm_def transitive_set_def)
 
 theorem transitive_set_reflection:
@@ -596,12 +596,12 @@
 
 lemma sats_ordinal_fm':
    "[|x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, ordinal_fm(x), env) <-> ordinal(##A,nth(x,env))"
+    ==> sats(A, ordinal_fm(x), env) \<longleftrightarrow> ordinal(##A,nth(x,env))"
 by (simp add: sats_transset_fm' ordinal_fm_def ordinal_def)
 
 lemma ordinal_iff_sats:
       "[| nth(i,env) = x;  i \<in> nat; env \<in> list(A)|]
-       ==> ordinal(##A, x) <-> sats(A, ordinal_fm(i), env)"
+       ==> ordinal(##A, x) \<longleftrightarrow> sats(A, ordinal_fm(i), env)"
 by (simp add: sats_ordinal_fm')
 
 theorem ordinal_reflection:
@@ -628,14 +628,14 @@
 
 lemma sats_Memrel_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, Memrel_fm(x,y), env) <->
+    ==> sats(A, Memrel_fm(x,y), env) \<longleftrightarrow>
         membership(##A, nth(x,env), nth(y,env))"
 by (simp add: Memrel_fm_def membership_def)
 
 lemma Memrel_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> membership(##A, x, y) <-> sats(A, Memrel_fm(i,j), env)"
+       ==> membership(##A, x, y) \<longleftrightarrow> sats(A, Memrel_fm(i,j), env)"
 by simp
 
 theorem membership_reflection:
@@ -663,14 +663,14 @@
 
 lemma sats_pred_set_fm [simp]:
    "[| U \<in> nat; x \<in> nat; r \<in> nat; B \<in> nat; env \<in> list(A)|]
-    ==> sats(A, pred_set_fm(U,x,r,B), env) <->
+    ==> sats(A, pred_set_fm(U,x,r,B), env) \<longleftrightarrow>
         pred_set(##A, nth(U,env), nth(x,env), nth(r,env), nth(B,env))"
 by (simp add: pred_set_fm_def pred_set_def)
 
 lemma pred_set_iff_sats:
       "[| nth(i,env) = U; nth(j,env) = x; nth(k,env) = r; nth(l,env) = B;
           i \<in> nat; j \<in> nat; k \<in> nat; l \<in> nat; env \<in> list(A)|]
-       ==> pred_set(##A,U,x,r,B) <-> sats(A, pred_set_fm(i,j,k,l), env)"
+       ==> pred_set(##A,U,x,r,B) \<longleftrightarrow> sats(A, pred_set_fm(i,j,k,l), env)"
 by (simp add: sats_pred_set_fm)
 
 theorem pred_set_reflection:
@@ -685,7 +685,7 @@
 subsubsection{*Domain of a Relation, Internalized*}
 
 (* "is_domain(M,r,z) ==
-        \<forall>x[M]. (x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w))))" *)
+        \<forall>x[M]. (x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w))))" *)
 definition
   domain_fm :: "[i,i]=>i" where
     "domain_fm(r,z) ==
@@ -699,14 +699,14 @@
 
 lemma sats_domain_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, domain_fm(x,y), env) <->
+    ==> sats(A, domain_fm(x,y), env) \<longleftrightarrow>
         is_domain(##A, nth(x,env), nth(y,env))"
 by (simp add: domain_fm_def is_domain_def)
 
 lemma domain_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> is_domain(##A, x, y) <-> sats(A, domain_fm(i,j), env)"
+       ==> is_domain(##A, x, y) \<longleftrightarrow> sats(A, domain_fm(i,j), env)"
 by simp
 
 theorem domain_reflection:
@@ -720,7 +720,7 @@
 subsubsection{*Range of a Relation, Internalized*}
 
 (* "is_range(M,r,z) ==
-        \<forall>y[M]. (y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w))))" *)
+        \<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w))))" *)
 definition
   range_fm :: "[i,i]=>i" where
     "range_fm(r,z) ==
@@ -734,14 +734,14 @@
 
 lemma sats_range_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, range_fm(x,y), env) <->
+    ==> sats(A, range_fm(x,y), env) \<longleftrightarrow>
         is_range(##A, nth(x,env), nth(y,env))"
 by (simp add: range_fm_def is_range_def)
 
 lemma range_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> is_range(##A, x, y) <-> sats(A, range_fm(i,j), env)"
+       ==> is_range(##A, x, y) \<longleftrightarrow> sats(A, range_fm(i,j), env)"
 by simp
 
 theorem range_reflection:
@@ -770,14 +770,14 @@
 
 lemma sats_field_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, field_fm(x,y), env) <->
+    ==> sats(A, field_fm(x,y), env) \<longleftrightarrow>
         is_field(##A, nth(x,env), nth(y,env))"
 by (simp add: field_fm_def is_field_def)
 
 lemma field_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> is_field(##A, x, y) <-> sats(A, field_fm(i,j), env)"
+       ==> is_field(##A, x, y) \<longleftrightarrow> sats(A, field_fm(i,j), env)"
 by simp
 
 theorem field_reflection:
@@ -792,7 +792,7 @@
 subsubsection{*Image under a Relation, Internalized*}
 
 (* "image(M,r,A,z) ==
-        \<forall>y[M]. (y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w))))" *)
+        \<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w))))" *)
 definition
   image_fm :: "[i,i,i]=>i" where
     "image_fm(r,A,z) ==
@@ -807,14 +807,14 @@
 
 lemma sats_image_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, image_fm(x,y,z), env) <->
+    ==> sats(A, image_fm(x,y,z), env) \<longleftrightarrow>
         image(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: image_fm_def Relative.image_def)
 
 lemma image_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> image(##A, x, y, z) <-> sats(A, image_fm(i,j,k), env)"
+       ==> image(##A, x, y, z) \<longleftrightarrow> sats(A, image_fm(i,j,k), env)"
 by (simp add: sats_image_fm)
 
 theorem image_reflection:
@@ -828,7 +828,7 @@
 subsubsection{*Pre-Image under a Relation, Internalized*}
 
 (* "pre_image(M,r,A,z) ==
-        \<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))" *)
+        \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))" *)
 definition
   pre_image_fm :: "[i,i,i]=>i" where
     "pre_image_fm(r,A,z) ==
@@ -843,14 +843,14 @@
 
 lemma sats_pre_image_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, pre_image_fm(x,y,z), env) <->
+    ==> sats(A, pre_image_fm(x,y,z), env) \<longleftrightarrow>
         pre_image(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: pre_image_fm_def Relative.pre_image_def)
 
 lemma pre_image_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> pre_image(##A, x, y, z) <-> sats(A, pre_image_fm(i,j,k), env)"
+       ==> pre_image(##A, x, y, z) \<longleftrightarrow> sats(A, pre_image_fm(i,j,k), env)"
 by (simp add: sats_pre_image_fm)
 
 theorem pre_image_reflection:
@@ -879,14 +879,14 @@
 
 lemma sats_fun_apply_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, fun_apply_fm(x,y,z), env) <->
+    ==> sats(A, fun_apply_fm(x,y,z), env) \<longleftrightarrow>
         fun_apply(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: fun_apply_fm_def fun_apply_def)
 
 lemma fun_apply_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> fun_apply(##A, x, y, z) <-> sats(A, fun_apply_fm(i,j,k), env)"
+       ==> fun_apply(##A, x, y, z) \<longleftrightarrow> sats(A, fun_apply_fm(i,j,k), env)"
 by simp
 
 theorem fun_apply_reflection:
@@ -901,7 +901,7 @@
 subsubsection{*The Concept of Relation, Internalized*}
 
 (* "is_relation(M,r) ==
-        (\<forall>z[M]. z\<in>r --> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))" *)
+        (\<forall>z[M]. z\<in>r \<longrightarrow> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))" *)
 definition
   relation_fm :: "i=>i" where
     "relation_fm(r) ==
@@ -913,13 +913,13 @@
 
 lemma sats_relation_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, relation_fm(x), env) <-> is_relation(##A, nth(x,env))"
+    ==> sats(A, relation_fm(x), env) \<longleftrightarrow> is_relation(##A, nth(x,env))"
 by (simp add: relation_fm_def is_relation_def)
 
 lemma relation_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; env \<in> list(A)|]
-       ==> is_relation(##A, x) <-> sats(A, relation_fm(i), env)"
+       ==> is_relation(##A, x) \<longleftrightarrow> sats(A, relation_fm(i), env)"
 by simp
 
 theorem is_relation_reflection:
@@ -934,7 +934,7 @@
 
 (* "is_function(M,r) ==
         \<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
-           pair(M,x,y,p) --> pair(M,x,y',p') --> p\<in>r --> p'\<in>r --> y=y'" *)
+           pair(M,x,y,p) \<longrightarrow> pair(M,x,y',p') \<longrightarrow> p\<in>r \<longrightarrow> p'\<in>r \<longrightarrow> y=y'" *)
 definition
   function_fm :: "i=>i" where
     "function_fm(r) ==
@@ -950,13 +950,13 @@
 
 lemma sats_function_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, function_fm(x), env) <-> is_function(##A, nth(x,env))"
+    ==> sats(A, function_fm(x), env) \<longleftrightarrow> is_function(##A, nth(x,env))"
 by (simp add: function_fm_def is_function_def)
 
 lemma is_function_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; env \<in> list(A)|]
-       ==> is_function(##A, x) <-> sats(A, function_fm(i), env)"
+       ==> is_function(##A, x) \<longleftrightarrow> sats(A, function_fm(i), env)"
 by simp
 
 theorem is_function_reflection:
@@ -971,7 +971,7 @@
 
 (* "typed_function(M,A,B,r) ==
         is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
-        (\<forall>u[M]. u\<in>r --> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) --> y\<in>B))" *)
+        (\<forall>u[M]. u\<in>r \<longrightarrow> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) \<longrightarrow> y\<in>B))" *)
 
 definition
   typed_function_fm :: "[i,i,i]=>i" where
@@ -988,14 +988,14 @@
 
 lemma sats_typed_function_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, typed_function_fm(x,y,z), env) <->
+    ==> sats(A, typed_function_fm(x,y,z), env) \<longleftrightarrow>
         typed_function(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: typed_function_fm_def typed_function_def)
 
 lemma typed_function_iff_sats:
   "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
       i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-   ==> typed_function(##A, x, y, z) <-> sats(A, typed_function_fm(i,j,k), env)"
+   ==> typed_function(##A, x, y, z) \<longleftrightarrow> sats(A, typed_function_fm(i,j,k), env)"
 by simp
 
 lemmas function_reflections =
@@ -1029,7 +1029,7 @@
 subsubsection{*Composition of Relations, Internalized*}
 
 (* "composition(M,r,s,t) ==
-        \<forall>p[M]. p \<in> t <->
+        \<forall>p[M]. p \<in> t \<longleftrightarrow>
                (\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
                 pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) &
                 xy \<in> s & yz \<in> r)" *)
@@ -1049,14 +1049,14 @@
 
 lemma sats_composition_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, composition_fm(x,y,z), env) <->
+    ==> sats(A, composition_fm(x,y,z), env) \<longleftrightarrow>
         composition(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: composition_fm_def composition_def)
 
 lemma composition_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> composition(##A, x, y, z) <-> sats(A, composition_fm(i,j,k), env)"
+       ==> composition(##A, x, y, z) \<longleftrightarrow> sats(A, composition_fm(i,j,k), env)"
 by simp
 
 theorem composition_reflection:
@@ -1072,7 +1072,7 @@
 (* "injection(M,A,B,f) ==
         typed_function(M,A,B,f) &
         (\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M].
-          pair(M,x,y,p) --> pair(M,x',y,p') --> p\<in>f --> p'\<in>f --> x=x')" *)
+          pair(M,x,y,p) \<longrightarrow> pair(M,x',y,p') \<longrightarrow> p\<in>f \<longrightarrow> p'\<in>f \<longrightarrow> x=x')" *)
 definition
   injection_fm :: "[i,i,i]=>i" where
   "injection_fm(A,B,f) ==
@@ -1090,14 +1090,14 @@
 
 lemma sats_injection_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, injection_fm(x,y,z), env) <->
+    ==> sats(A, injection_fm(x,y,z), env) \<longleftrightarrow>
         injection(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: injection_fm_def injection_def)
 
 lemma injection_iff_sats:
   "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
       i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-   ==> injection(##A, x, y, z) <-> sats(A, injection_fm(i,j,k), env)"
+   ==> injection(##A, x, y, z) \<longleftrightarrow> sats(A, injection_fm(i,j,k), env)"
 by simp
 
 theorem injection_reflection:
@@ -1113,7 +1113,7 @@
 (*  surjection :: "[i=>o,i,i,i] => o"
     "surjection(M,A,B,f) ==
         typed_function(M,A,B,f) &
-        (\<forall>y[M]. y\<in>B --> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))" *)
+        (\<forall>y[M]. y\<in>B \<longrightarrow> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))" *)
 definition
   surjection_fm :: "[i,i,i]=>i" where
   "surjection_fm(A,B,f) ==
@@ -1128,14 +1128,14 @@
 
 lemma sats_surjection_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, surjection_fm(x,y,z), env) <->
+    ==> sats(A, surjection_fm(x,y,z), env) \<longleftrightarrow>
         surjection(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: surjection_fm_def surjection_def)
 
 lemma surjection_iff_sats:
   "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
       i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-   ==> surjection(##A, x, y, z) <-> sats(A, surjection_fm(i,j,k), env)"
+   ==> surjection(##A, x, y, z) \<longleftrightarrow> sats(A, surjection_fm(i,j,k), env)"
 by simp
 
 theorem surjection_reflection:
@@ -1161,14 +1161,14 @@
 
 lemma sats_bijection_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, bijection_fm(x,y,z), env) <->
+    ==> sats(A, bijection_fm(x,y,z), env) \<longleftrightarrow>
         bijection(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: bijection_fm_def bijection_def)
 
 lemma bijection_iff_sats:
   "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
       i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-   ==> bijection(##A, x, y, z) <-> sats(A, bijection_fm(i,j,k), env)"
+   ==> bijection(##A, x, y, z) \<longleftrightarrow> sats(A, bijection_fm(i,j,k), env)"
 by simp
 
 theorem bijection_reflection:
@@ -1183,7 +1183,7 @@
 
 
 (* "restriction(M,r,A,z) ==
-        \<forall>x[M]. x \<in> z <-> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))" *)
+        \<forall>x[M]. x \<in> z \<longleftrightarrow> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))" *)
 definition
   restriction_fm :: "[i,i,i]=>i" where
     "restriction_fm(r,A,z) ==
@@ -1198,14 +1198,14 @@
 
 lemma sats_restriction_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, restriction_fm(x,y,z), env) <->
+    ==> sats(A, restriction_fm(x,y,z), env) \<longleftrightarrow>
         restriction(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: restriction_fm_def restriction_def)
 
 lemma restriction_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> restriction(##A, x, y, z) <-> sats(A, restriction_fm(i,j,k), env)"
+       ==> restriction(##A, x, y, z) \<longleftrightarrow> sats(A, restriction_fm(i,j,k), env)"
 by simp
 
 theorem restriction_reflection:
@@ -1220,10 +1220,10 @@
 (*  order_isomorphism :: "[i=>o,i,i,i,i,i] => o"
    "order_isomorphism(M,A,r,B,s,f) ==
         bijection(M,A,B,f) &
-        (\<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A -->
+        (\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow>
           (\<forall>p[M]. \<forall>fx[M]. \<forall>fy[M]. \<forall>q[M].
-            pair(M,x,y,p) --> fun_apply(M,f,x,fx) --> fun_apply(M,f,y,fy) -->
-            pair(M,fx,fy,q) --> (p\<in>r <-> q\<in>s))))"
+            pair(M,x,y,p) \<longrightarrow> fun_apply(M,f,x,fx) \<longrightarrow> fun_apply(M,f,y,fy) \<longrightarrow>
+            pair(M,fx,fy,q) \<longrightarrow> (p\<in>r \<longleftrightarrow> q\<in>s))))"
   *)
 
 definition
@@ -1246,7 +1246,7 @@
 
 lemma sats_order_isomorphism_fm [simp]:
    "[| U \<in> nat; r \<in> nat; B \<in> nat; s \<in> nat; f \<in> nat; env \<in> list(A)|]
-    ==> sats(A, order_isomorphism_fm(U,r,B,s,f), env) <->
+    ==> sats(A, order_isomorphism_fm(U,r,B,s,f), env) \<longleftrightarrow>
         order_isomorphism(##A, nth(U,env), nth(r,env), nth(B,env),
                                nth(s,env), nth(f,env))"
 by (simp add: order_isomorphism_fm_def order_isomorphism_def)
@@ -1255,7 +1255,7 @@
   "[| nth(i,env) = U; nth(j,env) = r; nth(k,env) = B; nth(j',env) = s;
       nth(k',env) = f;
       i \<in> nat; j \<in> nat; k \<in> nat; j' \<in> nat; k' \<in> nat; env \<in> list(A)|]
-   ==> order_isomorphism(##A,U,r,B,s,f) <->
+   ==> order_isomorphism(##A,U,r,B,s,f) \<longleftrightarrow>
        sats(A, order_isomorphism_fm(i,j,k,j',k'), env)"
 by simp
 
@@ -1272,7 +1272,7 @@
 
 (* "limit_ordinal(M,a) ==
         ordinal(M,a) & ~ empty(M,a) &
-        (\<forall>x[M]. x\<in>a --> (\<exists>y[M]. y\<in>a & successor(M,x,y)))" *)
+        (\<forall>x[M]. x\<in>a \<longrightarrow> (\<exists>y[M]. y\<in>a & successor(M,x,y)))" *)
 
 definition
   limit_ordinal_fm :: "i=>i" where
@@ -1289,13 +1289,13 @@
 
 lemma sats_limit_ordinal_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, limit_ordinal_fm(x), env) <-> limit_ordinal(##A, nth(x,env))"
+    ==> sats(A, limit_ordinal_fm(x), env) \<longleftrightarrow> limit_ordinal(##A, nth(x,env))"
 by (simp add: limit_ordinal_fm_def limit_ordinal_def sats_ordinal_fm')
 
 lemma limit_ordinal_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; env \<in> list(A)|]
-       ==> limit_ordinal(##A, x) <-> sats(A, limit_ordinal_fm(i), env)"
+       ==> limit_ordinal(##A, x) \<longleftrightarrow> sats(A, limit_ordinal_fm(i), env)"
 by simp
 
 theorem limit_ordinal_reflection:
@@ -1310,7 +1310,7 @@
 
 (*     "finite_ordinal(M,a) == 
         ordinal(M,a) & ~ limit_ordinal(M,a) & 
-        (\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x))" *)
+        (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x))" *)
 definition
   finite_ordinal_fm :: "i=>i" where
     "finite_ordinal_fm(x) ==
@@ -1325,13 +1325,13 @@
 
 lemma sats_finite_ordinal_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, finite_ordinal_fm(x), env) <-> finite_ordinal(##A, nth(x,env))"
+    ==> sats(A, finite_ordinal_fm(x), env) \<longleftrightarrow> finite_ordinal(##A, nth(x,env))"
 by (simp add: finite_ordinal_fm_def sats_ordinal_fm' finite_ordinal_def)
 
 lemma finite_ordinal_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; env \<in> list(A)|]
-       ==> finite_ordinal(##A, x) <-> sats(A, finite_ordinal_fm(i), env)"
+       ==> finite_ordinal(##A, x) \<longleftrightarrow> sats(A, finite_ordinal_fm(i), env)"
 by simp
 
 theorem finite_ordinal_reflection:
@@ -1344,7 +1344,7 @@
 
 subsubsection{*Omega: The Set of Natural Numbers*}
 
-(* omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x)) *)
+(* omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x)) *)
 definition
   omega_fm :: "i=>i" where
     "omega_fm(x) ==
@@ -1358,13 +1358,13 @@
 
 lemma sats_omega_fm [simp]:
    "[| x \<in> nat; env \<in> list(A)|]
-    ==> sats(A, omega_fm(x), env) <-> omega(##A, nth(x,env))"
+    ==> sats(A, omega_fm(x), env) \<longleftrightarrow> omega(##A, nth(x,env))"
 by (simp add: omega_fm_def omega_def)
 
 lemma omega_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; env \<in> list(A)|]
-       ==> omega(##A, x) <-> sats(A, omega_fm(i), env)"
+       ==> omega(##A, x) \<longleftrightarrow> sats(A, omega_fm(i), env)"
 by simp
 
 theorem omega_reflection:
--- a/src/ZF/Constructible/Normal.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Normal.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -19,11 +19,11 @@
 
 definition
   Closed :: "(i=>o) => o" where
-    "Closed(P) == \<forall>I. I \<noteq> 0 --> (\<forall>i\<in>I. Ord(i) \<and> P(i)) --> P(\<Union>(I))"
+    "Closed(P) == \<forall>I. I \<noteq> 0 \<longrightarrow> (\<forall>i\<in>I. Ord(i) \<and> P(i)) \<longrightarrow> P(\<Union>(I))"
 
 definition
   Unbounded :: "(i=>o) => o" where
-    "Unbounded(P) == \<forall>i. Ord(i) --> (\<exists>j. i<j \<and> P(j))"
+    "Unbounded(P) == \<forall>i. Ord(i) \<longrightarrow> (\<exists>j. i<j \<and> P(j))"
 
 definition
   Closed_Unbounded :: "(i=>o) => o" where
@@ -188,7 +188,7 @@
 done
 
 lemma Int_iff_INT2:
-     "P(x) \<and> Q(x)  <->  (\<forall>i\<in>2. (i=0 --> P(x)) \<and> (i=1 --> Q(x)))"
+     "P(x) \<and> Q(x)  \<longleftrightarrow>  (\<forall>i\<in>2. (i=0 \<longrightarrow> P(x)) \<and> (i=1 \<longrightarrow> Q(x)))"
 by auto
 
 theorem Closed_Unbounded_Int:
@@ -203,15 +203,15 @@
 
 definition
   mono_le_subset :: "(i=>i) => o" where
-    "mono_le_subset(M) == \<forall>i j. i\<le>j --> M(i) \<subseteq> M(j)"
+    "mono_le_subset(M) == \<forall>i j. i\<le>j \<longrightarrow> M(i) \<subseteq> M(j)"
 
 definition
   mono_Ord :: "(i=>i) => o" where
-    "mono_Ord(F) == \<forall>i j. i<j --> F(i) < F(j)"
+    "mono_Ord(F) == \<forall>i j. i<j \<longrightarrow> F(i) < F(j)"
 
 definition
   cont_Ord :: "(i=>i) => o" where
-    "cont_Ord(F) == \<forall>l. Limit(l) --> F(l) = (\<Union>i<l. F(i))"
+    "cont_Ord(F) == \<forall>l. Limit(l) \<longrightarrow> F(l) = (\<Union>i<l. F(i))"
 
 definition
   Normal :: "(i=>i) => o" where
@@ -267,11 +267,11 @@
 
 text{*The following equation is taken for granted in any set theory text.*}
 lemma cont_Ord_Union:
-     "[| cont_Ord(F); mono_le_subset(F); X=0 --> F(0)=0; \<forall>x\<in>X. Ord(x) |] 
-      ==> F(Union(X)) = (\<Union>y\<in>X. F(y))"
+     "[| cont_Ord(F); mono_le_subset(F); X=0 \<longrightarrow> F(0)=0; \<forall>x\<in>X. Ord(x) |] 
+      ==> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))"
 apply (frule Ord_set_cases)
 apply (erule disjE, force) 
-apply (thin_tac "X=0 --> ?Q", auto)
+apply (thin_tac "X=0 \<longrightarrow> ?Q", auto)
  txt{*The trival case of @{term "\<Union>X \<in> X"}*}
  apply (rule equalityI, blast intro: Ord_Union_eq_succD) 
  apply (simp add: mono_le_subset_def UN_subset_iff le_subset_iff) 
@@ -293,7 +293,7 @@
 done
 
 lemma Normal_Union:
-     "[| X\<noteq>0; \<forall>x\<in>X. Ord(x); Normal(F) |] ==> F(Union(X)) = (\<Union>y\<in>X. F(y))"
+     "[| X\<noteq>0; \<forall>x\<in>X. Ord(x); Normal(F) |] ==> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))"
 apply (simp add: Normal_def) 
 apply (blast intro: mono_Ord_imp_le_subset cont_Ord_Union) 
 done
@@ -378,7 +378,7 @@
 *}
 definition
   normalize :: "[i=>i, i] => i" where
-    "normalize(F,a) == transrec2(a, F(0), \<lambda>x r. F(succ(x)) Un succ(r))"
+    "normalize(F,a) == transrec2(a, F(0), \<lambda>x r. F(succ(x)) \<union> succ(r))"
 
 
 lemma Ord_normalize [simp, intro]:
@@ -389,7 +389,7 @@
 
 lemma normalize_lemma [rule_format]:
      "[| Ord(b); !!x. Ord(x) ==> Ord(F(x)) |] 
-      ==> \<forall>a. a < b --> normalize(F, a) < normalize(F, b)"
+      ==> \<forall>a. a < b \<longrightarrow> normalize(F, a) < normalize(F, b)"
 apply (erule trans_induct3)
   apply (simp_all add: le_iff def_transrec2 [OF normalize_def])
  apply clarify
@@ -443,7 +443,7 @@
 done
 
 lemma Aleph_lemma [rule_format]:
-     "Ord(b) ==> \<forall>a. a < b --> Aleph(a) < Aleph(b)"
+     "Ord(b) ==> \<forall>a. a < b \<longrightarrow> Aleph(a) < Aleph(b)"
 apply (erule trans_induct3) 
 apply (simp_all add: le_iff def_transrec2 [OF Aleph_def])  
 apply (blast intro: lt_trans lt_csucc Card_is_Ord, clarify)
--- a/src/ZF/Constructible/Rank.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Rank.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -12,7 +12,7 @@
 locale M_ordertype = M_basic +
 assumes well_ord_iso_separation:
      "[| M(A); M(f); M(r) |]
-      ==> separation (M, \<lambda>x. x\<in>A --> (\<exists>y[M]. (\<exists>p[M].
+      ==> separation (M, \<lambda>x. x\<in>A \<longrightarrow> (\<exists>y[M]. (\<exists>p[M].
                      fun_apply(M,f,x,y) & pair(M,y,x,p) & p \<in> r)))"
   and obase_separation:
      --{*part of the order type formalization*}
@@ -22,7 +22,7 @@
              order_isomorphism(M,par,r,x,mx,g))"
   and obase_equals_separation:
      "[| M(A); M(r) |]
-      ==> separation (M, \<lambda>x. x\<in>A --> ~(\<exists>y[M]. \<exists>g[M].
+      ==> separation (M, \<lambda>x. x\<in>A \<longrightarrow> ~(\<exists>y[M]. \<exists>g[M].
                               ordinal(M,y) & (\<exists>my[M]. \<exists>pxr[M].
                               membership(M,y,my) & pred_set(M,A,x,r,pxr) &
                               order_isomorphism(M,pxr,r,y,my,g))))"
@@ -146,7 +146,7 @@
     --{*the function that maps wosets to order types*}
    "omap(M,A,r,f) == 
         \<forall>z[M].
-         z \<in> f <-> (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> & Ord(x) & 
+         z \<in> f \<longleftrightarrow> (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> & Ord(x) & 
                         g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))"
 
 definition
@@ -159,7 +159,7 @@
       is also more useful than the definition, @{text omap_def}.*}
 lemma (in M_ordertype) omap_iff:
      "[| omap(M,A,r,f); M(A); M(f) |] 
-      ==> z \<in> f <->
+      ==> z \<in> f \<longleftrightarrow>
           (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> & Ord(x) & 
                                 g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))"
 apply (simp add: omap_def Memrel_closed pred_closed) 
@@ -181,7 +181,7 @@
 
 lemma (in M_ordertype) otype_iff:
      "[| otype(M,A,r,i); M(A); M(r); M(i) |] 
-      ==> x \<in> i <-> 
+      ==> x \<in> i \<longleftrightarrow> 
           (M(x) & Ord(x) & 
            (\<exists>a\<in>A. \<exists>g[M]. g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x))))"
 apply (auto simp add: omap_iff otype_def)
@@ -324,7 +324,7 @@
  apply (frule obase_equals_separation [of A r], assumption) 
  apply (simp, clarify) 
 apply (rename_tac b) 
-apply (subgoal_tac "Order.pred(A,b,r) <= obase(M,A,r)") 
+apply (subgoal_tac "Order.pred(A,b,r) \<subseteq> obase(M,A,r)") 
  apply (blast intro!: restrict_omap_ord_iso Ord_omap_image_pred)
 apply (force simp add: pred_iff obase_def)  
 done
@@ -394,7 +394,7 @@
 
 text{*(a) The notion of Wellordering is absolute*}
 theorem (in M_ordertype) well_ord_abs [simp]: 
-     "[| M(A); M(r) |] ==> wellordered(M,A,r) <-> well_ord(A,r)" 
+     "[| M(A); M(r) |] ==> wellordered(M,A,r) \<longleftrightarrow> well_ord(A,r)" 
 by (blast intro: well_ord_imp_relativized relativized_imp_well_ord)  
 
 
@@ -420,8 +420,8 @@
 definition
   is_oadd_fun :: "[i=>o,i,i,i,i] => o" where
     "is_oadd_fun(M,i,j,x,f) == 
-       (\<forall>sj msj. M(sj) --> M(msj) --> 
-                 successor(M,j,sj) --> membership(M,sj,msj) --> 
+       (\<forall>sj msj. M(sj) \<longrightarrow> M(msj) \<longrightarrow> 
+                 successor(M,j,sj) \<longrightarrow> membership(M,sj,msj) \<longrightarrow> 
                  M_is_recfun(M, 
                      %x g y. \<exists>gx[M]. image(M,g,x,gx) & union(M,i,gx,y),
                      msj, x, f))"
@@ -442,9 +442,9 @@
   omult_eqns :: "[i,i,i,i] => o" where
     "omult_eqns(i,x,g,z) ==
             Ord(x) & 
-            (x=0 --> z=0) &
-            (\<forall>j. x = succ(j) --> z = g`j ++ i) &
-            (Limit(x) --> z = \<Union>(g``x))"
+            (x=0 \<longrightarrow> z=0) &
+            (\<forall>j. x = succ(j) \<longrightarrow> z = g`j ++ i) &
+            (Limit(x) \<longrightarrow> z = \<Union>(g``x))"
 
 definition
   is_omult_fun :: "[i=>o,i,i,i] => o" where
@@ -467,7 +467,7 @@
     strong_replacement(M, 
          \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & 
                   (\<exists>f[M]. \<exists>fx[M]. is_oadd_fun(M,i,j,x,f) & 
-                           image(M,f,x,fx) & y = i Un fx))"
+                           image(M,f,x,fx) & y = i \<union> fx))"
 
  and omult_strong_replacement':
    "[| M(i); M(j) |] ==>
@@ -481,11 +481,11 @@
 text{*@{text is_oadd_fun}: Relating the pure "language of set theory" to Isabelle/ZF*}
 lemma (in M_ord_arith) is_oadd_fun_iff:
    "[| a\<le>j; M(i); M(j); M(a); M(f) |] 
-    ==> is_oadd_fun(M,i,j,a,f) <->
-        f \<in> a \<rightarrow> range(f) & (\<forall>x. M(x) --> x < a --> f`x = i Un f``x)"
+    ==> is_oadd_fun(M,i,j,a,f) \<longleftrightarrow>
+        f \<in> a \<rightarrow> range(f) & (\<forall>x. M(x) \<longrightarrow> x < a \<longrightarrow> f`x = i \<union> f``x)"
 apply (frule lt_Ord) 
 apply (simp add: is_oadd_fun_def Memrel_closed Un_closed 
-             relation2_def is_recfun_abs [of "%x g. i Un g``x"]
+             relation2_def is_recfun_abs [of "%x g. i \<union> g``x"]
              image_closed is_recfun_iff_equation  
              Ball_def lt_trans [OF ltI, of _ a] lt_Memrel)
 apply (simp add: lt_def) 
@@ -497,17 +497,17 @@
     "[| M(i); M(j) |] ==>
      strong_replacement(M, 
             \<lambda>x z. \<exists>y[M]. z = <x,y> &
-                  (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. i Un g``x,g) & 
-                  y = i Un g``x))" 
+                  (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. i \<union> g``x,g) & 
+                  y = i \<union> g``x))" 
 apply (insert oadd_strong_replacement [of i j]) 
 apply (simp add: is_oadd_fun_def relation2_def
-                 is_recfun_abs [of "%x g. i Un g``x"])  
+                 is_recfun_abs [of "%x g. i \<union> g``x"])  
 done
 
 
 lemma (in M_ord_arith) exists_oadd:
     "[| Ord(j);  M(i);  M(j) |]
-     ==> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, %x g. i Un g``x, f)"
+     ==> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, %x g. i \<union> g``x, f)"
 apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel])
     apply (simp_all add: Memrel_type oadd_strong_replacement') 
 done 
@@ -525,7 +525,7 @@
 
 lemma (in M_ord_arith) is_oadd_fun_apply:
     "[| x < j; M(i); M(j); M(f); is_oadd_fun(M,i,j,j,f) |] 
-     ==> f`x = i Un (\<Union>k\<in>x. {f ` k})"
+     ==> f`x = i \<union> (\<Union>k\<in>x. {f ` k})"
 apply (simp add: is_oadd_fun_iff lt_Ord2, clarify) 
 apply (frule lt_closed, simp)
 apply (frule leI [THEN le_imp_subset])  
@@ -534,7 +534,7 @@
 
 lemma (in M_ord_arith) is_oadd_fun_iff_oadd [rule_format]:
     "[| is_oadd_fun(M,i,J,J,f); M(i); M(J); M(f); Ord(i); Ord(j) |] 
-     ==> j<J --> f`j = i++j"
+     ==> j<J \<longrightarrow> f`j = i++j"
 apply (erule_tac i=j in trans_induct, clarify) 
 apply (subgoal_tac "\<forall>k\<in>x. k<J")
  apply (simp (no_asm_simp) add: is_oadd_def oadd_unfold is_oadd_fun_apply)
@@ -542,13 +542,13 @@
 done
 
 lemma (in M_ord_arith) Ord_oadd_abs:
-    "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_oadd(M,i,j,k) <-> k = i++j"
+    "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_oadd(M,i,j,k) \<longleftrightarrow> k = i++j"
 apply (simp add: is_oadd_def is_oadd_fun_iff_oadd)
 apply (frule exists_oadd_fun [of j i], blast+)
 done
 
 lemma (in M_ord_arith) oadd_abs:
-    "[| M(i); M(j); M(k) |] ==> is_oadd(M,i,j,k) <-> k = i++j"
+    "[| M(i); M(j); M(k) |] ==> is_oadd(M,i,j,k) \<longleftrightarrow> k = i++j"
 apply (case_tac "Ord(i) & Ord(j)")
  apply (simp add: Ord_oadd_abs)
 apply (auto simp add: is_oadd_def oadd_eq_if_raw_oadd)
@@ -571,13 +571,13 @@
 apply (erule Ord_cases, simp_all) 
 done
 
-lemma omult_eqns_0: "omult_eqns(i,0,g,z) <-> z=0"
+lemma omult_eqns_0: "omult_eqns(i,0,g,z) \<longleftrightarrow> z=0"
 by (simp add: omult_eqns_def)
 
 lemma the_omult_eqns_0: "(THE z. omult_eqns(i,0,g,z)) = 0"
 by (simp add: omult_eqns_0)
 
-lemma omult_eqns_succ: "omult_eqns(i,succ(j),g,z) <-> Ord(j) & z = g`j ++ i"
+lemma omult_eqns_succ: "omult_eqns(i,succ(j),g,z) \<longleftrightarrow> Ord(j) & z = g`j ++ i"
 by (simp add: omult_eqns_def)
 
 lemma the_omult_eqns_succ:
@@ -585,7 +585,7 @@
 by (simp add: omult_eqns_succ) 
 
 lemma omult_eqns_Limit:
-     "Limit(x) ==> omult_eqns(i,x,g,z) <-> z = \<Union>(g``x)"
+     "Limit(x) ==> omult_eqns(i,x,g,z) \<longleftrightarrow> z = \<Union>(g``x)"
 apply (simp add: omult_eqns_def) 
 apply (blast intro: Limit_is_Ord) 
 done
@@ -649,7 +649,7 @@
 
 lemma (in M_ord_arith) is_omult_fun_eq_omult:
     "[| is_omult_fun(M,i,J,f); M(J); M(f); Ord(i); Ord(j) |] 
-     ==> j<J --> f`j = i**j"
+     ==> j<J \<longrightarrow> f`j = i**j"
 apply (erule_tac i=j in trans_induct3)
 apply (safe del: impCE)
   apply (simp add: is_omult_fun_apply_0) 
@@ -662,7 +662,7 @@
 done
 
 lemma (in M_ord_arith) omult_abs:
-    "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_omult(M,i,j,k) <-> k = i**j"
+    "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_omult(M,i,j,k) \<longleftrightarrow> k = i**j"
 apply (simp add: is_omult_def is_omult_fun_eq_omult)
 apply (frule exists_omult_fun [of j i], blast+)
 done
@@ -679,22 +679,22 @@
   assumes wfrank_separation:
      "M(r) ==>
       separation (M, \<lambda>x. 
-         \<forall>rplus[M]. tran_closure(M,r,rplus) -->
+         \<forall>rplus[M]. tran_closure(M,r,rplus) \<longrightarrow>
          ~ (\<exists>f[M]. M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f)))"
  and wfrank_strong_replacement:
      "M(r) ==>
       strong_replacement(M, \<lambda>x z. 
-         \<forall>rplus[M]. tran_closure(M,r,rplus) -->
+         \<forall>rplus[M]. tran_closure(M,r,rplus) \<longrightarrow>
          (\<exists>y[M]. \<exists>f[M]. pair(M,x,y,z)  & 
                         M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f) &
                         is_range(M,f,y)))"
  and Ord_wfrank_separation:
      "M(r) ==>
       separation (M, \<lambda>x.
-         \<forall>rplus[M]. tran_closure(M,r,rplus) --> 
+         \<forall>rplus[M]. tran_closure(M,r,rplus) \<longrightarrow> 
           ~ (\<forall>f[M]. \<forall>rangef[M]. 
-             is_range(M,f,rangef) -->
-             M_is_recfun(M, \<lambda>x f y. is_range(M,f,y), rplus, x, f) -->
+             is_range(M,f,rangef) \<longrightarrow>
+             M_is_recfun(M, \<lambda>x f y. is_range(M,f,y), rplus, x, f) \<longrightarrow>
              ordinal(M,rangef)))" 
 
 
@@ -721,7 +721,7 @@
 lemma (in M_wfrank) Ord_wfrank_separation':
      "M(r) ==>
       separation (M, \<lambda>x. 
-         ~ (\<forall>f[M]. is_recfun(r^+, x, \<lambda>x. range, f) --> Ord(range(f))))" 
+         ~ (\<forall>f[M]. is_recfun(r^+, x, \<lambda>x. range, f) \<longrightarrow> Ord(range(f))))" 
 apply (insert Ord_wfrank_separation [of r])
 apply (simp add: relation2_def is_recfun_abs [of "%x. range"])
 done
@@ -761,7 +761,7 @@
 
 lemma (in M_wfrank) Ord_wfrank_range [rule_format]:
     "[| wellfounded(M,r); a\<in>A; M(r); M(A) |]
-     ==> \<forall>f[M]. is_recfun(r^+, a, %x f. range(f), f) --> Ord(range(f))"
+     ==> \<forall>f[M]. is_recfun(r^+, a, %x f. range(f), f) \<longrightarrow> Ord(range(f))"
 apply (drule wellfounded_trancl, assumption)
 apply (rule wellfounded_induct, assumption, erule (1) transM)
   apply simp
@@ -911,7 +911,7 @@
 
 lemma (in M_wfrank) wellfounded_imp_subset_rvimage:
      "[|wellfounded(M,r); r \<subseteq> A*A; M(r); M(A)|]
-      ==> \<exists>i f. Ord(i) & r <= rvimage(A, f, Memrel(i))"
+      ==> \<exists>i f. Ord(i) & r \<subseteq> rvimage(A, f, Memrel(i))"
 apply (rule_tac x="range(wellfoundedrank(M,r,A))" in exI)
 apply (rule_tac x="wellfoundedrank(M,r,A)" in exI)
 apply (simp add: Ord_range_wellfoundedrank, clarify)
@@ -934,11 +934,11 @@
 
 
 theorem (in M_wfrank) wf_abs:
-     "[|relation(r); M(r)|] ==> wellfounded(M,r) <-> wf(r)"
+     "[|relation(r); M(r)|] ==> wellfounded(M,r) \<longleftrightarrow> wf(r)"
 by (blast intro: wellfounded_imp_wf wf_imp_relativized)
 
 theorem (in M_wfrank) wf_on_abs:
-     "[|relation(r); M(r); M(A)|] ==> wellfounded_on(M,A,r) <-> wf[A](r)"
+     "[|relation(r); M(r); M(A)|] ==> wellfounded_on(M,A,r) \<longleftrightarrow> wf[A](r)"
 by (blast intro: wellfounded_on_imp_wf_on wf_on_imp_relativized)
 
 end
\ No newline at end of file
--- a/src/ZF/Constructible/Rank_Separation.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Rank_Separation.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -17,15 +17,15 @@
 subsubsection{*Separation for Order-Isomorphisms*}
 
 lemma well_ord_iso_Reflects:
-  "REFLECTS[\<lambda>x. x\<in>A -->
+  "REFLECTS[\<lambda>x. x\<in>A \<longrightarrow>
                 (\<exists>y[L]. \<exists>p[L]. fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r),
-        \<lambda>i x. x\<in>A --> (\<exists>y \<in> Lset(i). \<exists>p \<in> Lset(i).
+        \<lambda>i x. x\<in>A \<longrightarrow> (\<exists>y \<in> Lset(i). \<exists>p \<in> Lset(i).
                 fun_apply(##Lset(i),f,x,y) & pair(##Lset(i),y,x,p) & p \<in> r)]"
 by (intro FOL_reflections function_reflections)
 
 lemma well_ord_iso_separation:
      "[| L(A); L(f); L(r) |]
-      ==> separation (L, \<lambda>x. x\<in>A --> (\<exists>y[L]. (\<exists>p[L].
+      ==> separation (L, \<lambda>x. x\<in>A \<longrightarrow> (\<exists>y[L]. (\<exists>p[L].
                      fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r)))"
 apply (rule gen_separation_multi [OF well_ord_iso_Reflects, of "{A,f,r}"], 
        auto)
@@ -60,11 +60,11 @@
 subsubsection{*Separation for a Theorem about @{term "obase"}*}
 
 lemma obase_equals_reflects:
-  "REFLECTS[\<lambda>x. x\<in>A --> ~(\<exists>y[L]. \<exists>g[L].
+  "REFLECTS[\<lambda>x. x\<in>A \<longrightarrow> ~(\<exists>y[L]. \<exists>g[L].
                 ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
                 membership(L,y,my) & pred_set(L,A,x,r,pxr) &
                 order_isomorphism(L,pxr,r,y,my,g))),
-        \<lambda>i x. x\<in>A --> ~(\<exists>y \<in> Lset(i). \<exists>g \<in> Lset(i).
+        \<lambda>i x. x\<in>A \<longrightarrow> ~(\<exists>y \<in> Lset(i). \<exists>g \<in> Lset(i).
                 ordinal(##Lset(i),y) & (\<exists>my \<in> Lset(i). \<exists>pxr \<in> Lset(i).
                 membership(##Lset(i),y,my) & pred_set(##Lset(i),A,x,r,pxr) &
                 order_isomorphism(##Lset(i),pxr,r,y,my,g)))]"
@@ -72,7 +72,7 @@
 
 lemma obase_equals_separation:
      "[| L(A); L(r) |]
-      ==> separation (L, \<lambda>x. x\<in>A --> ~(\<exists>y[L]. \<exists>g[L].
+      ==> separation (L, \<lambda>x. x\<in>A \<longrightarrow> ~(\<exists>y[L]. \<exists>g[L].
                               ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
                               membership(L,y,my) & pred_set(L,A,x,r,pxr) &
                               order_isomorphism(L,pxr,r,y,my,g))))"
@@ -132,9 +132,9 @@
 subsubsection{*Separation for @{term "wfrank"}*}
 
 lemma wfrank_Reflects:
- "REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) -->
+ "REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
               ~ (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)),
-      \<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) -->
+      \<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) \<longrightarrow>
          ~ (\<exists>f \<in> Lset(i).
             M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y),
                         rplus, x, f))]"
@@ -142,7 +142,7 @@
 
 lemma wfrank_separation:
      "L(r) ==>
-      separation (L, \<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) -->
+      separation (L, \<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
          ~ (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))"
 apply (rule gen_separation [OF wfrank_Reflects], simp)
 apply (rule_tac env="[r]" in DPow_LsetI)
@@ -154,12 +154,12 @@
 
 lemma wfrank_replacement_Reflects:
  "REFLECTS[\<lambda>z. \<exists>x[L]. x \<in> A &
-        (\<forall>rplus[L]. tran_closure(L,r,rplus) -->
+        (\<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
          (\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z)  &
                         M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) &
                         is_range(L,f,y))),
  \<lambda>i z. \<exists>x \<in> Lset(i). x \<in> A &
-      (\<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) -->
+      (\<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) \<longrightarrow>
        (\<exists>y \<in> Lset(i). \<exists>f \<in> Lset(i). pair(##Lset(i),x,y,z)  &
          M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y), rplus, x, f) &
          is_range(##Lset(i),f,y)))]"
@@ -169,7 +169,7 @@
 lemma wfrank_strong_replacement:
      "L(r) ==>
       strong_replacement(L, \<lambda>x z.
-         \<forall>rplus[L]. tran_closure(L,r,rplus) -->
+         \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
          (\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z)  &
                         M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) &
                         is_range(L,f,y)))"
@@ -185,16 +185,16 @@
 subsubsection{*Separation for Proving @{text Ord_wfrank_range}*}
 
 lemma Ord_wfrank_Reflects:
- "REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) -->
+ "REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
           ~ (\<forall>f[L]. \<forall>rangef[L].
-             is_range(L,f,rangef) -->
-             M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) -->
+             is_range(L,f,rangef) \<longrightarrow>
+             M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) \<longrightarrow>
              ordinal(L,rangef)),
-      \<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) -->
+      \<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) \<longrightarrow>
           ~ (\<forall>f \<in> Lset(i). \<forall>rangef \<in> Lset(i).
-             is_range(##Lset(i),f,rangef) -->
+             is_range(##Lset(i),f,rangef) \<longrightarrow>
              M_is_recfun(##Lset(i), \<lambda>x f y. is_range(##Lset(i),f,y),
-                         rplus, x, f) -->
+                         rplus, x, f) \<longrightarrow>
              ordinal(##Lset(i),rangef))]"
 by (intro FOL_reflections function_reflections is_recfun_reflection
           tran_closure_reflection ordinal_reflection)
@@ -202,10 +202,10 @@
 lemma  Ord_wfrank_separation:
      "L(r) ==>
       separation (L, \<lambda>x.
-         \<forall>rplus[L]. tran_closure(L,r,rplus) -->
+         \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
           ~ (\<forall>f[L]. \<forall>rangef[L].
-             is_range(L,f,rangef) -->
-             M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) -->
+             is_range(L,f,rangef) \<longrightarrow>
+             M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) \<longrightarrow>
              ordinal(L,rangef)))"
 apply (rule gen_separation [OF Ord_wfrank_Reflects], simp)
 apply (rule_tac env="[r]" in DPow_LsetI)
--- a/src/ZF/Constructible/Rec_Separation.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Rec_Separation.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -25,7 +25,7 @@
        (\<exists>f[M]. typed_function(M,n',A,f) &
         (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
           fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
-        (\<forall>j[M]. j\<in>n -->
+        (\<forall>j[M]. j\<in>n \<longrightarrow>
           (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M].
             fun_apply(M,f,j,fj) & successor(M,j,sj) &
             fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"*)
@@ -55,14 +55,14 @@
 
 lemma sats_rtran_closure_mem_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, rtran_closure_mem_fm(x,y,z), env) <->
+    ==> sats(A, rtran_closure_mem_fm(x,y,z), env) \<longleftrightarrow>
         rtran_closure_mem(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: rtran_closure_mem_fm_def rtran_closure_mem_def)
 
 lemma rtran_closure_mem_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> rtran_closure_mem(##A, x, y, z) <-> sats(A, rtran_closure_mem_fm(i,j,k), env)"
+       ==> rtran_closure_mem(##A, x, y, z) \<longleftrightarrow> sats(A, rtran_closure_mem_fm(i,j,k), env)"
 by (simp add: sats_rtran_closure_mem_fm)
 
 lemma rtran_closure_mem_reflection:
@@ -85,8 +85,8 @@
 subsubsection{*Reflexive/Transitive Closure, Internalized*}
 
 (*  "rtran_closure(M,r,s) ==
-        \<forall>A[M]. is_field(M,r,A) -->
-         (\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))" *)
+        \<forall>A[M]. is_field(M,r,A) \<longrightarrow>
+         (\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))" *)
 definition
   rtran_closure_fm :: "[i,i]=>i" where
   "rtran_closure_fm(r,s) ==
@@ -100,14 +100,14 @@
 
 lemma sats_rtran_closure_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, rtran_closure_fm(x,y), env) <->
+    ==> sats(A, rtran_closure_fm(x,y), env) \<longleftrightarrow>
         rtran_closure(##A, nth(x,env), nth(y,env))"
 by (simp add: rtran_closure_fm_def rtran_closure_def)
 
 lemma rtran_closure_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> rtran_closure(##A, x, y) <-> sats(A, rtran_closure_fm(i,j), env)"
+       ==> rtran_closure(##A, x, y) \<longleftrightarrow> sats(A, rtran_closure_fm(i,j), env)"
 by simp
 
 theorem rtran_closure_reflection:
@@ -133,14 +133,14 @@
 
 lemma sats_tran_closure_fm [simp]:
    "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
-    ==> sats(A, tran_closure_fm(x,y), env) <->
+    ==> sats(A, tran_closure_fm(x,y), env) \<longleftrightarrow>
         tran_closure(##A, nth(x,env), nth(y,env))"
 by (simp add: tran_closure_fm_def tran_closure_def)
 
 lemma tran_closure_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y;
           i \<in> nat; j \<in> nat; env \<in> list(A)|]
-       ==> tran_closure(##A, x, y) <-> sats(A, tran_closure_fm(i,j), env)"
+       ==> tran_closure(##A, x, y) \<longleftrightarrow> sats(A, tran_closure_fm(i,j), env)"
 by simp
 
 theorem tran_closure_reflection:
@@ -307,7 +307,7 @@
 
 lemma sats_nth_fm [simp]:
    "[| x < length(env); y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, nth_fm(x,y,z), env) <->
+    ==> sats(A, nth_fm(x,y,z), env) \<longleftrightarrow>
         is_nth(##A, nth(x,env), nth(y,env), nth(z,env))"
 apply (frule lt_length_in_nat, assumption)  
 apply (simp add: nth_fm_def is_nth_def sats_is_iterates_fm) 
@@ -316,7 +316,7 @@
 lemma nth_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
           i < length(env); j \<in> nat; k \<in> nat; env \<in> list(A)|]
-       ==> is_nth(##A, x, y, z) <-> sats(A, nth_fm(i,j,k), env)"
+       ==> is_nth(##A, x, y, z) \<longleftrightarrow> sats(A, nth_fm(i,j,k), env)"
 by (simp add: sats_nth_fm)
 
 theorem nth_reflection:
--- a/src/ZF/Constructible/Reflection.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Reflection.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -6,26 +6,26 @@
 
 theory Reflection imports Normal begin
 
-lemma all_iff_not_ex_not: "(\<forall>x. P(x)) <-> (~ (\<exists>x. ~ P(x)))";
+lemma all_iff_not_ex_not: "(\<forall>x. P(x)) \<longleftrightarrow> (~ (\<exists>x. ~ P(x)))";
 by blast
 
-lemma ball_iff_not_bex_not: "(\<forall>x\<in>A. P(x)) <-> (~ (\<exists>x\<in>A. ~ P(x)))";
+lemma ball_iff_not_bex_not: "(\<forall>x\<in>A. P(x)) \<longleftrightarrow> (~ (\<exists>x\<in>A. ~ P(x)))";
 by blast
 
-text{*From the notes of A. S. Kechris, page 6, and from 
+text{*From the notes of A. S. Kechris, page 6, and from
       Andrzej Mostowski, \emph{Constructible Sets with Applications},
       North-Holland, 1969, page 23.*}
 
 
 subsection{*Basic Definitions*}
 
-text{*First part: the cumulative hierarchy defining the class @{text M}.  
+text{*First part: the cumulative hierarchy defining the class @{text M}.
 To avoid handling multiple arguments, we assume that @{text "Mset(l)"} is
 closed under ordered pairing provided @{text l} is limit.  Possibly this
-could be avoided: the induction hypothesis @{term Cl_reflects} 
+could be avoided: the induction hypothesis @{term Cl_reflects}
 (in locale @{text ex_reflection}) could be weakened to
-@{term "\<forall>y\<in>Mset(a). \<forall>z\<in>Mset(a). P(<y,z>) <-> Q(a,<y,z>)"}, removing most
-uses of @{term Pair_in_Mset}.  But there isn't much point in doing so, since 
+@{term "\<forall>y\<in>Mset(a). \<forall>z\<in>Mset(a). P(<y,z>) \<longleftrightarrow> Q(a,<y,z>)"}, removing most
+uses of @{term Pair_in_Mset}.  But there isn't much point in doing so, since
 ultimately the @{text ex_reflection} proof is packaged up using the
 predicate @{text Reflects}.
 *}
@@ -33,29 +33,29 @@
   fixes Mset and M and Reflects
   assumes Mset_mono_le : "mono_le_subset(Mset)"
       and Mset_cont    : "cont_Ord(Mset)"
-      and Pair_in_Mset : "[| x \<in> Mset(a); y \<in> Mset(a); Limit(a) |] 
+      and Pair_in_Mset : "[| x \<in> Mset(a); y \<in> Mset(a); Limit(a) |]
                           ==> <x,y> \<in> Mset(a)"
   defines "M(x) == \<exists>a. Ord(a) & x \<in> Mset(a)"
       and "Reflects(Cl,P,Q) == Closed_Unbounded(Cl) &
-                              (\<forall>a. Cl(a) --> (\<forall>x\<in>Mset(a). P(x) <-> Q(a,x)))"
+                              (\<forall>a. Cl(a) \<longrightarrow> (\<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)))"
   fixes F0 --{*ordinal for a specific value @{term y}*}
   fixes FF --{*sup over the whole level, @{term "y\<in>Mset(a)"}*}
   fixes ClEx --{*Reflecting ordinals for the formula @{term "\<exists>z. P"}*}
-  defines "F0(P,y) == \<mu> b. (\<exists>z. M(z) & P(<y,z>)) --> 
+  defines "F0(P,y) == \<mu> b. (\<exists>z. M(z) & P(<y,z>)) \<longrightarrow>
                                (\<exists>z\<in>Mset(b). P(<y,z>))"
       and "FF(P)   == \<lambda>a. \<Union>y\<in>Mset(a). F0(P,y)"
       and "ClEx(P,a) == Limit(a) & normalize(FF(P),a) = a"
 
-lemma (in reflection) Mset_mono: "i\<le>j ==> Mset(i) <= Mset(j)"
-apply (insert Mset_mono_le) 
-apply (simp add: mono_le_subset_def leI) 
+lemma (in reflection) Mset_mono: "i\<le>j ==> Mset(i) \<subseteq> Mset(j)"
+apply (insert Mset_mono_le)
+apply (simp add: mono_le_subset_def leI)
 done
 
 text{*Awkward: we need a version of @{text ClEx_def} as an equality
       at the level of classes, which do not really exist*}
 lemma (in reflection) ClEx_eq:
      "ClEx(P) == \<lambda>a. Limit(a) & normalize(FF(P),a) = a"
-by (simp add: ClEx_def [symmetric]) 
+by (simp add: ClEx_def [symmetric])
 
 
 subsection{*Easy Cases of the Reflection Theorem*}
@@ -66,33 +66,33 @@
 
 theorem (in reflection) Not_reflection [intro]:
      "Reflects(Cl,P,Q) ==> Reflects(Cl, \<lambda>x. ~P(x), \<lambda>a x. ~Q(a,x))"
-by (simp add: Reflects_def) 
+by (simp add: Reflects_def)
 
 theorem (in reflection) And_reflection [intro]:
-     "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] 
-      ==> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) & P'(x), 
+     "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |]
+      ==> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) & P'(x),
                                       \<lambda>a x. Q(a,x) & Q'(a,x))"
 by (simp add: Reflects_def Closed_Unbounded_Int, blast)
 
 theorem (in reflection) Or_reflection [intro]:
-     "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] 
-      ==> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) | P'(x), 
+     "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |]
+      ==> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) | P'(x),
                                       \<lambda>a x. Q(a,x) | Q'(a,x))"
 by (simp add: Reflects_def Closed_Unbounded_Int, blast)
 
 theorem (in reflection) Imp_reflection [intro]:
-     "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] 
-      ==> Reflects(\<lambda>a. Cl(a) & C'(a), 
-                   \<lambda>x. P(x) --> P'(x), 
-                   \<lambda>a x. Q(a,x) --> Q'(a,x))"
+     "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |]
+      ==> Reflects(\<lambda>a. Cl(a) & C'(a),
+                   \<lambda>x. P(x) \<longrightarrow> P'(x),
+                   \<lambda>a x. Q(a,x) \<longrightarrow> Q'(a,x))"
 by (simp add: Reflects_def Closed_Unbounded_Int, blast)
 
 theorem (in reflection) Iff_reflection [intro]:
-     "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] 
-      ==> Reflects(\<lambda>a. Cl(a) & C'(a), 
-                   \<lambda>x. P(x) <-> P'(x), 
-                   \<lambda>a x. Q(a,x) <-> Q'(a,x))"
-by (simp add: Reflects_def Closed_Unbounded_Int, blast) 
+     "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |]
+      ==> Reflects(\<lambda>a. Cl(a) & C'(a),
+                   \<lambda>x. P(x) \<longleftrightarrow> P'(x),
+                   \<lambda>a x. Q(a,x) \<longleftrightarrow> Q'(a,x))"
+by (simp add: Reflects_def Closed_Unbounded_Int, blast)
 
 subsection{*Reflection for Existential Quantifiers*}
 
@@ -115,20 +115,20 @@
 apply (simp add: cont_Ord_def FF_def, blast)
 done
 
-text{*Recall that @{term F0} depends upon @{term "y\<in>Mset(a)"}, 
+text{*Recall that @{term F0} depends upon @{term "y\<in>Mset(a)"},
 while @{term FF} depends only upon @{term a}. *}
 lemma (in reflection) FF_works:
      "[| M(z); y\<in>Mset(a); P(<y,z>); Ord(a) |] ==> \<exists>z\<in>Mset(FF(P,a)). P(<y,z>)"
 apply (simp add: FF_def)
-apply (simp_all add: cont_Ord_Union [of concl: Mset] 
+apply (simp_all add: cont_Ord_Union [of concl: Mset]
                      Mset_cont Mset_mono_le not_emptyI Ord_F0)
-apply (blast intro: F0_works)  
+apply (blast intro: F0_works)
 done
 
 lemma (in reflection) FFN_works:
-     "[| M(z); y\<in>Mset(a); P(<y,z>); Ord(a) |] 
+     "[| M(z); y\<in>Mset(a); P(<y,z>); Ord(a) |]
       ==> \<exists>z\<in>Mset(normalize(FF(P),a)). P(<y,z>)"
-apply (drule FF_works [of concl: P], assumption+) 
+apply (drule FF_works [of concl: P], assumption+)
 apply (blast intro: cont_Ord_FF le_normalize [THEN Mset_mono, THEN subsetD])
 done
 
@@ -139,20 +139,20 @@
   fixes P  --"the original formula"
   fixes Q  --"the reflected formula"
   fixes Cl --"the class of reflecting ordinals"
-  assumes Cl_reflects: "[| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) <-> Q(a,x)"
+  assumes Cl_reflects: "[| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)"
 
 lemma (in ex_reflection) ClEx_downward:
-     "[| M(z); y\<in>Mset(a); P(<y,z>); Cl(a); ClEx(P,a) |] 
+     "[| M(z); y\<in>Mset(a); P(<y,z>); Cl(a); ClEx(P,a) |]
       ==> \<exists>z\<in>Mset(a). Q(a,<y,z>)"
-apply (simp add: ClEx_def, clarify) 
-apply (frule Limit_is_Ord) 
-apply (frule FFN_works [of concl: P], assumption+) 
-apply (drule Cl_reflects, assumption+) 
+apply (simp add: ClEx_def, clarify)
+apply (frule Limit_is_Ord)
+apply (frule FFN_works [of concl: P], assumption+)
+apply (drule Cl_reflects, assumption+)
 apply (auto simp add: Limit_is_Ord Pair_in_Mset)
 done
 
 lemma (in ex_reflection) ClEx_upward:
-     "[| z\<in>Mset(a); y\<in>Mset(a); Q(a,<y,z>); Cl(a); ClEx(P,a) |] 
+     "[| z\<in>Mset(a); y\<in>Mset(a); Q(a,<y,z>); Cl(a); ClEx(P,a) |]
       ==> \<exists>z. M(z) & P(<y,z>)"
 apply (simp add: ClEx_def M_def)
 apply (blast dest: Cl_reflects
@@ -161,9 +161,9 @@
 
 text{*Class @{text ClEx} indeed consists of reflecting ordinals...*}
 lemma (in ex_reflection) ZF_ClEx_iff:
-     "[| y\<in>Mset(a); Cl(a); ClEx(P,a) |] 
-      ==> (\<exists>z. M(z) & P(<y,z>)) <-> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
-by (blast intro: dest: ClEx_downward ClEx_upward) 
+     "[| y\<in>Mset(a); Cl(a); ClEx(P,a) |]
+      ==> (\<exists>z. M(z) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
+by (blast intro: dest: ClEx_downward ClEx_upward)
 
 text{*...and it is closed and unbounded*}
 lemma (in ex_reflection) ZF_Closed_Unbounded_ClEx:
@@ -178,8 +178,8 @@
 text{*Class @{text ClEx} indeed consists of reflecting ordinals...*}
 lemma (in reflection) ClEx_iff:
      "[| y\<in>Mset(a); Cl(a); ClEx(P,a);
-        !!a. [| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) <-> Q(a,x) |] 
-      ==> (\<exists>z. M(z) & P(<y,z>)) <-> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
+        !!a. [| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x) |]
+      ==> (\<exists>z. M(z) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
 apply (unfold ClEx_def FF_def F0_def M_def)
 apply (rule ex_reflection.ZF_ClEx_iff
   [OF ex_reflection.intro, OF reflection.intro ex_reflection_axioms.intro,
@@ -195,9 +195,9 @@
 *)
 
 lemma (in reflection) Closed_Unbounded_ClEx:
-     "(!!a. [| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) <-> Q(a,x))
+     "(!!a. [| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x))
       ==> Closed_Unbounded(ClEx(P))"
-apply (unfold ClEx_eq FF_def F0_def M_def) 
+apply (unfold ClEx_eq FF_def F0_def M_def)
 apply (rule ex_reflection.ZF_Closed_Unbounded_ClEx [of Mset _ _ Cl])
 apply (rule ex_reflection.intro, rule reflection_axioms)
 apply (blast intro: ex_reflection_axioms.intro)
@@ -206,58 +206,58 @@
 subsection{*Packaging the Quantifier Reflection Rules*}
 
 lemma (in reflection) Ex_reflection_0:
-     "Reflects(Cl,P0,Q0) 
-      ==> Reflects(\<lambda>a. Cl(a) & ClEx(P0,a), 
-                   \<lambda>x. \<exists>z. M(z) & P0(<x,z>), 
-                   \<lambda>a x. \<exists>z\<in>Mset(a). Q0(a,<x,z>))" 
-apply (simp add: Reflects_def) 
+     "Reflects(Cl,P0,Q0)
+      ==> Reflects(\<lambda>a. Cl(a) & ClEx(P0,a),
+                   \<lambda>x. \<exists>z. M(z) & P0(<x,z>),
+                   \<lambda>a x. \<exists>z\<in>Mset(a). Q0(a,<x,z>))"
+apply (simp add: Reflects_def)
 apply (intro conjI Closed_Unbounded_Int)
-  apply blast 
- apply (rule Closed_Unbounded_ClEx [of Cl P0 Q0], blast, clarify) 
-apply (rule_tac Cl=Cl in  ClEx_iff, assumption+, blast) 
+  apply blast
+ apply (rule Closed_Unbounded_ClEx [of Cl P0 Q0], blast, clarify)
+apply (rule_tac Cl=Cl in  ClEx_iff, assumption+, blast)
 done
 
 lemma (in reflection) All_reflection_0:
-     "Reflects(Cl,P0,Q0) 
-      ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x.~P0(x), a), 
-                   \<lambda>x. \<forall>z. M(z) --> P0(<x,z>), 
-                   \<lambda>a x. \<forall>z\<in>Mset(a). Q0(a,<x,z>))" 
-apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not) 
-apply (rule Not_reflection, drule Not_reflection, simp) 
+     "Reflects(Cl,P0,Q0)
+      ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x.~P0(x), a),
+                   \<lambda>x. \<forall>z. M(z) \<longrightarrow> P0(<x,z>),
+                   \<lambda>a x. \<forall>z\<in>Mset(a). Q0(a,<x,z>))"
+apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not)
+apply (rule Not_reflection, drule Not_reflection, simp)
 apply (erule Ex_reflection_0)
 done
 
 theorem (in reflection) Ex_reflection [intro]:
-     "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) 
-      ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a), 
-                   \<lambda>x. \<exists>z. M(z) & P(x,z), 
+     "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
+      ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a),
+                   \<lambda>x. \<exists>z. M(z) & P(x,z),
                    \<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))"
-by (rule Ex_reflection_0 [of _ " \<lambda>x. P(fst(x),snd(x))" 
+by (rule Ex_reflection_0 [of _ " \<lambda>x. P(fst(x),snd(x))"
                                "\<lambda>a x. Q(a,fst(x),snd(x))", simplified])
 
 theorem (in reflection) All_reflection [intro]:
      "Reflects(Cl,  \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
-      ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. ~P(fst(x),snd(x)), a), 
-                   \<lambda>x. \<forall>z. M(z) --> P(x,z), 
-                   \<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))" 
-by (rule All_reflection_0 [of _ "\<lambda>x. P(fst(x),snd(x))" 
+      ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. ~P(fst(x),snd(x)), a),
+                   \<lambda>x. \<forall>z. M(z) \<longrightarrow> P(x,z),
+                   \<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))"
+by (rule All_reflection_0 [of _ "\<lambda>x. P(fst(x),snd(x))"
                                 "\<lambda>a x. Q(a,fst(x),snd(x))", simplified])
 
 text{*And again, this time using class-bounded quantifiers*}
 
 theorem (in reflection) Rex_reflection [intro]:
-     "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) 
-      ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a), 
-                   \<lambda>x. \<exists>z[M]. P(x,z), 
+     "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
+      ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a),
+                   \<lambda>x. \<exists>z[M]. P(x,z),
                    \<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))"
-by (unfold rex_def, blast) 
+by (unfold rex_def, blast)
 
 theorem (in reflection) Rall_reflection [intro]:
      "Reflects(Cl,  \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
-      ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. ~P(fst(x),snd(x)), a), 
-                   \<lambda>x. \<forall>z[M]. P(x,z), 
-                   \<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))" 
-by (unfold rall_def, blast) 
+      ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. ~P(fst(x),snd(x)), a),
+                   \<lambda>x. \<forall>z[M]. P(x,z),
+                   \<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))"
+by (unfold rall_def, blast)
 
 
 text{*No point considering bounded quantifiers, where reflection is trivial.*}
@@ -266,62 +266,62 @@
 subsection{*Simple Examples of Reflection*}
 
 text{*Example 1: reflecting a simple formula.  The reflecting class is first
-given as the variable @{text ?Cl} and later retrieved from the final 
+given as the variable @{text ?Cl} and later retrieved from the final
 proof state.*}
-schematic_lemma (in reflection) 
+schematic_lemma (in reflection)
      "Reflects(?Cl,
-               \<lambda>x. \<exists>y. M(y) & x \<in> y, 
+               \<lambda>x. \<exists>y. M(y) & x \<in> y,
                \<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)"
 by fast
 
 text{*Problem here: there needs to be a conjunction (class intersection)
 in the class of reflecting ordinals.  The @{term "Ord(a)"} is redundant,
 though harmless.*}
-lemma (in reflection) 
-     "Reflects(\<lambda>a. Ord(a) & ClEx(\<lambda>x. fst(x) \<in> snd(x), a),   
-               \<lambda>x. \<exists>y. M(y) & x \<in> y, 
-               \<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)" 
+lemma (in reflection)
+     "Reflects(\<lambda>a. Ord(a) & ClEx(\<lambda>x. fst(x) \<in> snd(x), a),
+               \<lambda>x. \<exists>y. M(y) & x \<in> y,
+               \<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)"
 by fast
 
 
 text{*Example 2*}
-schematic_lemma (in reflection) 
+schematic_lemma (in reflection)
      "Reflects(?Cl,
-               \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> z \<subseteq> x --> z \<in> y), 
-               \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x --> z \<in> y)" 
+               \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
+               \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)"
 by fast
 
 text{*Example 2'.  We give the reflecting class explicitly. *}
-lemma (in reflection) 
+lemma (in reflection)
   "Reflects
     (\<lambda>a. (Ord(a) &
-          ClEx(\<lambda>x. ~ (snd(x) \<subseteq> fst(fst(x)) --> snd(x) \<in> snd(fst(x))), a)) &
-          ClEx(\<lambda>x. \<forall>z. M(z) --> z \<subseteq> fst(x) --> z \<in> snd(x), a),
-            \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> z \<subseteq> x --> z \<in> y), 
-            \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x --> z \<in> y)" 
+          ClEx(\<lambda>x. ~ (snd(x) \<subseteq> fst(fst(x)) \<longrightarrow> snd(x) \<in> snd(fst(x))), a)) &
+          ClEx(\<lambda>x. \<forall>z. M(z) \<longrightarrow> z \<subseteq> fst(x) \<longrightarrow> z \<in> snd(x), a),
+            \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
+            \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)"
 by fast
 
 text{*Example 2''.  We expand the subset relation.*}
-schematic_lemma (in reflection) 
+schematic_lemma (in reflection)
   "Reflects(?Cl,
-        \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> (\<forall>w. M(w) --> w\<in>z --> w\<in>x) --> z\<in>y),
-        \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). (\<forall>w\<in>Mset(a). w\<in>z --> w\<in>x) --> z\<in>y)"
+        \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> (\<forall>w. M(w) \<longrightarrow> w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y),
+        \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). (\<forall>w\<in>Mset(a). w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y)"
 by fast
 
 text{*Example 2'''.  Single-step version, to reveal the reflecting class.*}
-schematic_lemma (in reflection) 
+schematic_lemma (in reflection)
      "Reflects(?Cl,
-               \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> z \<subseteq> x --> z \<in> y), 
-               \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x --> z \<in> y)" 
-apply (rule Ex_reflection) 
+               \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
+               \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)"
+apply (rule Ex_reflection)
 txt{*
 @{goals[display,indent=0,margin=60]}
 *}
-apply (rule All_reflection) 
+apply (rule All_reflection)
 txt{*
 @{goals[display,indent=0,margin=60]}
 *}
-apply (rule Triv_reflection) 
+apply (rule Triv_reflection)
 txt{*
 @{goals[display,indent=0,margin=60]}
 *}
@@ -329,21 +329,21 @@
 
 text{*Example 3.  Warning: the following examples make sense only
 if @{term P} is quantifier-free, since it is not being relativized.*}
-schematic_lemma (in reflection) 
+schematic_lemma (in reflection)
      "Reflects(?Cl,
-               \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> z \<in> y <-> z \<in> x & P(z)), 
-               \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<in> y <-> z \<in> x & P(z))"
+               \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<in> y \<longleftrightarrow> z \<in> x & P(z)),
+               \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<in> y \<longleftrightarrow> z \<in> x & P(z))"
 by fast
 
 text{*Example 3'*}
-schematic_lemma (in reflection) 
+schematic_lemma (in reflection)
      "Reflects(?Cl,
                \<lambda>x. \<exists>y. M(y) & y = Collect(x,P),
                \<lambda>a x. \<exists>y\<in>Mset(a). y = Collect(x,P))";
 by fast
 
 text{*Example 3''*}
-schematic_lemma (in reflection) 
+schematic_lemma (in reflection)
      "Reflects(?Cl,
                \<lambda>x. \<exists>y. M(y) & y = Replace(x,P),
                \<lambda>a x. \<exists>y\<in>Mset(a). y = Replace(x,P))";
@@ -351,10 +351,10 @@
 
 text{*Example 4: Axiom of Choice.  Possibly wrong, since @{text \<Pi>} needs
 to be relativized.*}
-schematic_lemma (in reflection) 
+schematic_lemma (in reflection)
      "Reflects(?Cl,
-               \<lambda>A. 0\<notin>A --> (\<exists>f. M(f) & f \<in> (\<Pi> X \<in> A. X)),
-               \<lambda>a A. 0\<notin>A --> (\<exists>f\<in>Mset(a). f \<in> (\<Pi> X \<in> A. X)))"
+               \<lambda>A. 0\<notin>A \<longrightarrow> (\<exists>f. M(f) & f \<in> (\<Pi> X \<in> A. X)),
+               \<lambda>a A. 0\<notin>A \<longrightarrow> (\<exists>f\<in>Mset(a). f \<in> (\<Pi> X \<in> A. X)))"
 by fast
 
 end
--- a/src/ZF/Constructible/Relative.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Relative.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -14,11 +14,11 @@
 
 definition
   subset :: "[i=>o,i,i] => o" where
-    "subset(M,A,B) == \<forall>x[M]. x\<in>A --> x \<in> B"
+    "subset(M,A,B) == \<forall>x[M]. x\<in>A \<longrightarrow> x \<in> B"
 
 definition
   upair :: "[i=>o,i,i,i] => o" where
-    "upair(M,a,b,z) == a \<in> z & b \<in> z & (\<forall>x[M]. x\<in>z --> x = a | x = b)"
+    "upair(M,a,b,z) == a \<in> z & b \<in> z & (\<forall>x[M]. x\<in>z \<longrightarrow> x = a | x = b)"
 
 definition
   pair :: "[i=>o,i,i,i] => o" where
@@ -28,7 +28,7 @@
 
 definition
   union :: "[i=>o,i,i,i] => o" where
-    "union(M,a,b,z) == \<forall>x[M]. x \<in> z <-> x \<in> a | x \<in> b"
+    "union(M,a,b,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a | x \<in> b"
 
 definition
   is_cons :: "[i=>o,i,i,i] => o" where
@@ -52,38 +52,38 @@
 
 definition
   powerset :: "[i=>o,i,i] => o" where
-    "powerset(M,A,z) == \<forall>x[M]. x \<in> z <-> subset(M,x,A)"
+    "powerset(M,A,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> subset(M,x,A)"
 
 definition
   is_Collect :: "[i=>o,i,i=>o,i] => o" where
-    "is_Collect(M,A,P,z) == \<forall>x[M]. x \<in> z <-> x \<in> A & P(x)"
+    "is_Collect(M,A,P,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A & P(x)"
 
 definition
   is_Replace :: "[i=>o,i,[i,i]=>o,i] => o" where
-    "is_Replace(M,A,P,z) == \<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & P(x,u))"
+    "is_Replace(M,A,P,z) == \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,u))"
 
 definition
   inter :: "[i=>o,i,i,i] => o" where
-    "inter(M,a,b,z) == \<forall>x[M]. x \<in> z <-> x \<in> a & x \<in> b"
+    "inter(M,a,b,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a & x \<in> b"
 
 definition
   setdiff :: "[i=>o,i,i,i] => o" where
-    "setdiff(M,a,b,z) == \<forall>x[M]. x \<in> z <-> x \<in> a & x \<notin> b"
+    "setdiff(M,a,b,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a & x \<notin> b"
 
 definition
   big_union :: "[i=>o,i,i] => o" where
-    "big_union(M,A,z) == \<forall>x[M]. x \<in> z <-> (\<exists>y[M]. y\<in>A & x \<in> y)"
+    "big_union(M,A,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A & x \<in> y)"
 
 definition
   big_inter :: "[i=>o,i,i] => o" where
     "big_inter(M,A,z) ==
-             (A=0 --> z=0) &
-             (A\<noteq>0 --> (\<forall>x[M]. x \<in> z <-> (\<forall>y[M]. y\<in>A --> x \<in> y)))"
+             (A=0 \<longrightarrow> z=0) &
+             (A\<noteq>0 \<longrightarrow> (\<forall>x[M]. x \<in> z \<longleftrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> x \<in> y)))"
 
 definition
   cartprod :: "[i=>o,i,i,i] => o" where
     "cartprod(M,A,B,z) ==
-        \<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))"
+        \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))"
 
 definition
   is_sum :: "[i=>o,i,i,i] => o" where
@@ -103,23 +103,23 @@
 definition
   is_converse :: "[i=>o,i,i] => o" where
     "is_converse(M,r,z) ==
-        \<forall>x[M]. x \<in> z <->
+        \<forall>x[M]. x \<in> z \<longleftrightarrow>
              (\<exists>w[M]. w\<in>r & (\<exists>u[M]. \<exists>v[M]. pair(M,u,v,w) & pair(M,v,u,x)))"
 
 definition
   pre_image :: "[i=>o,i,i,i] => o" where
     "pre_image(M,r,A,z) ==
-        \<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))"
+        \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))"
 
 definition
   is_domain :: "[i=>o,i,i] => o" where
     "is_domain(M,r,z) ==
-        \<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w)))"
+        \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w)))"
 
 definition
   image :: "[i=>o,i,i,i] => o" where
     "image(M,r,A,z) ==
-        \<forall>y[M]. y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w)))"
+        \<forall>y[M]. y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w)))"
 
 definition
   is_range :: "[i=>o,i,i] => o" where
@@ -128,7 +128,7 @@
       unfortunately needs an instance of separation in order to prove
         @{term "M(converse(r))"}.*}
     "is_range(M,r,z) ==
-        \<forall>y[M]. y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w)))"
+        \<forall>y[M]. y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w)))"
 
 definition
   is_field :: "[i=>o,i,i] => o" where
@@ -139,13 +139,13 @@
 definition
   is_relation :: "[i=>o,i] => o" where
     "is_relation(M,r) ==
-        (\<forall>z[M]. z\<in>r --> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))"
+        (\<forall>z[M]. z\<in>r \<longrightarrow> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))"
 
 definition
   is_function :: "[i=>o,i] => o" where
     "is_function(M,r) ==
         \<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
-           pair(M,x,y,p) --> pair(M,x,y',p') --> p\<in>r --> p'\<in>r --> y=y'"
+           pair(M,x,y,p) \<longrightarrow> pair(M,x,y',p') \<longrightarrow> p\<in>r \<longrightarrow> p'\<in>r \<longrightarrow> y=y'"
 
 definition
   fun_apply :: "[i=>o,i,i,i] => o" where
@@ -157,17 +157,17 @@
   typed_function :: "[i=>o,i,i,i] => o" where
     "typed_function(M,A,B,r) ==
         is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
-        (\<forall>u[M]. u\<in>r --> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) --> y\<in>B))"
+        (\<forall>u[M]. u\<in>r \<longrightarrow> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) \<longrightarrow> y\<in>B))"
 
 definition
   is_funspace :: "[i=>o,i,i,i] => o" where
     "is_funspace(M,A,B,F) ==
-        \<forall>f[M]. f \<in> F <-> typed_function(M,A,B,f)"
+        \<forall>f[M]. f \<in> F \<longleftrightarrow> typed_function(M,A,B,f)"
 
 definition
   composition :: "[i=>o,i,i,i] => o" where
     "composition(M,r,s,t) ==
-        \<forall>p[M]. p \<in> t <->
+        \<forall>p[M]. p \<in> t \<longleftrightarrow>
                (\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
                 pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) &
                 xy \<in> s & yz \<in> r)"
@@ -177,13 +177,13 @@
     "injection(M,A,B,f) ==
         typed_function(M,A,B,f) &
         (\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M].
-          pair(M,x,y,p) --> pair(M,x',y,p') --> p\<in>f --> p'\<in>f --> x=x')"
+          pair(M,x,y,p) \<longrightarrow> pair(M,x',y,p') \<longrightarrow> p\<in>f \<longrightarrow> p'\<in>f \<longrightarrow> x=x')"
 
 definition
   surjection :: "[i=>o,i,i,i] => o" where
     "surjection(M,A,B,f) ==
         typed_function(M,A,B,f) &
-        (\<forall>y[M]. y\<in>B --> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))"
+        (\<forall>y[M]. y\<in>B \<longrightarrow> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))"
 
 definition
   bijection :: "[i=>o,i,i,i] => o" where
@@ -192,23 +192,23 @@
 definition
   restriction :: "[i=>o,i,i,i] => o" where
     "restriction(M,r,A,z) ==
-        \<forall>x[M]. x \<in> z <-> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))"
+        \<forall>x[M]. x \<in> z \<longleftrightarrow> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))"
 
 definition
   transitive_set :: "[i=>o,i] => o" where
-    "transitive_set(M,a) == \<forall>x[M]. x\<in>a --> subset(M,x,a)"
+    "transitive_set(M,a) == \<forall>x[M]. x\<in>a \<longrightarrow> subset(M,x,a)"
 
 definition
   ordinal :: "[i=>o,i] => o" where
      --{*an ordinal is a transitive set of transitive sets*}
-    "ordinal(M,a) == transitive_set(M,a) & (\<forall>x[M]. x\<in>a --> transitive_set(M,x))"
+    "ordinal(M,a) == transitive_set(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> transitive_set(M,x))"
 
 definition
   limit_ordinal :: "[i=>o,i] => o" where
     --{*a limit ordinal is a non-empty, successor-closed ordinal*}
     "limit_ordinal(M,a) ==
         ordinal(M,a) & ~ empty(M,a) &
-        (\<forall>x[M]. x\<in>a --> (\<exists>y[M]. y\<in>a & successor(M,x,y)))"
+        (\<forall>x[M]. x\<in>a \<longrightarrow> (\<exists>y[M]. y\<in>a & successor(M,x,y)))"
 
 definition
   successor_ordinal :: "[i=>o,i] => o" where
@@ -221,12 +221,12 @@
     --{*an ordinal is finite if neither it nor any of its elements are limit*}
     "finite_ordinal(M,a) ==
         ordinal(M,a) & ~ limit_ordinal(M,a) &
-        (\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x))"
+        (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x))"
 
 definition
   omega :: "[i=>o,i] => o" where
     --{*omega is a limit ordinal none of whose elements are limit*}
-    "omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x))"
+    "omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x))"
 
 definition
   is_quasinat :: "[i=>o,i] => o" where
@@ -235,44 +235,44 @@
 definition
   is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" where
     "is_nat_case(M, a, is_b, k, z) ==
-       (empty(M,k) --> z=a) &
-       (\<forall>m[M]. successor(M,m,k) --> is_b(m,z)) &
+       (empty(M,k) \<longrightarrow> z=a) &
+       (\<forall>m[M]. successor(M,m,k) \<longrightarrow> is_b(m,z)) &
        (is_quasinat(M,k) | empty(M,z))"
 
 definition
   relation1 :: "[i=>o, [i,i]=>o, i=>i] => o" where
-    "relation1(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. is_f(x,y) <-> y = f(x)"
+    "relation1(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. is_f(x,y) \<longleftrightarrow> y = f(x)"
 
 definition
   Relation1 :: "[i=>o, i, [i,i]=>o, i=>i] => o" where
     --{*as above, but typed*}
     "Relation1(M,A,is_f,f) ==
-        \<forall>x[M]. \<forall>y[M]. x\<in>A --> is_f(x,y) <-> y = f(x)"
+        \<forall>x[M]. \<forall>y[M]. x\<in>A \<longrightarrow> is_f(x,y) \<longleftrightarrow> y = f(x)"
 
 definition
   relation2 :: "[i=>o, [i,i,i]=>o, [i,i]=>i] => o" where
-    "relation2(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. is_f(x,y,z) <-> z = f(x,y)"
+    "relation2(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. is_f(x,y,z) \<longleftrightarrow> z = f(x,y)"
 
 definition
   Relation2 :: "[i=>o, i, i, [i,i,i]=>o, [i,i]=>i] => o" where
     "Relation2(M,A,B,is_f,f) ==
-        \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. x\<in>A --> y\<in>B --> is_f(x,y,z) <-> z = f(x,y)"
+        \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. x\<in>A \<longrightarrow> y\<in>B \<longrightarrow> is_f(x,y,z) \<longleftrightarrow> z = f(x,y)"
 
 definition
   relation3 :: "[i=>o, [i,i,i,i]=>o, [i,i,i]=>i] => o" where
     "relation3(M,is_f,f) ==
-       \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M]. is_f(x,y,z,u) <-> u = f(x,y,z)"
+       \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M]. is_f(x,y,z,u) \<longleftrightarrow> u = f(x,y,z)"
 
 definition
   Relation3 :: "[i=>o, i, i, i, [i,i,i,i]=>o, [i,i,i]=>i] => o" where
     "Relation3(M,A,B,C,is_f,f) ==
        \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M].
-         x\<in>A --> y\<in>B --> z\<in>C --> is_f(x,y,z,u) <-> u = f(x,y,z)"
+         x\<in>A \<longrightarrow> y\<in>B \<longrightarrow> z\<in>C \<longrightarrow> is_f(x,y,z,u) \<longleftrightarrow> u = f(x,y,z)"
 
 definition
   relation4 :: "[i=>o, [i,i,i,i,i]=>o, [i,i,i,i]=>i] => o" where
     "relation4(M,is_f,f) ==
-       \<forall>u[M]. \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>a[M]. is_f(u,x,y,z,a) <-> a = f(u,x,y,z)"
+       \<forall>u[M]. \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>a[M]. is_f(u,x,y,z,a) \<longleftrightarrow> a = f(u,x,y,z)"
 
 
 text{*Useful when absoluteness reasoning has replaced the predicates by terms*}
@@ -290,7 +290,7 @@
 definition
   extensionality :: "(i=>o) => o" where
     "extensionality(M) ==
-        \<forall>x[M]. \<forall>y[M]. (\<forall>z[M]. z \<in> x <-> z \<in> y) --> x=y"
+        \<forall>x[M]. \<forall>y[M]. (\<forall>z[M]. z \<in> x \<longleftrightarrow> z \<in> y) \<longrightarrow> x=y"
 
 definition
   separation :: "[i=>o, i=>o] => o" where
@@ -299,7 +299,7 @@
         to @{text M}.  We do not have separation as a scheme; every instance
         that we need must be assumed (and later proved) separately.*}
     "separation(M,P) ==
-        \<forall>z[M]. \<exists>y[M]. \<forall>x[M]. x \<in> y <-> x \<in> z & P(x)"
+        \<forall>z[M]. \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z & P(x)"
 
 definition
   upair_ax :: "(i=>o) => o" where
@@ -316,24 +316,24 @@
 definition
   univalent :: "[i=>o, i, [i,i]=>o] => o" where
     "univalent(M,A,P) ==
-        \<forall>x[M]. x\<in>A --> (\<forall>y[M]. \<forall>z[M]. P(x,y) & P(x,z) --> y=z)"
+        \<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. \<forall>z[M]. P(x,y) & P(x,z) \<longrightarrow> y=z)"
 
 definition
   replacement :: "[i=>o, [i,i]=>o] => o" where
     "replacement(M,P) ==
-      \<forall>A[M]. univalent(M,A,P) -->
-      (\<exists>Y[M]. \<forall>b[M]. (\<exists>x[M]. x\<in>A & P(x,b)) --> b \<in> Y)"
+      \<forall>A[M]. univalent(M,A,P) \<longrightarrow>
+      (\<exists>Y[M]. \<forall>b[M]. (\<exists>x[M]. x\<in>A & P(x,b)) \<longrightarrow> b \<in> Y)"
 
 definition
   strong_replacement :: "[i=>o, [i,i]=>o] => o" where
     "strong_replacement(M,P) ==
-      \<forall>A[M]. univalent(M,A,P) -->
-      (\<exists>Y[M]. \<forall>b[M]. b \<in> Y <-> (\<exists>x[M]. x\<in>A & P(x,b)))"
+      \<forall>A[M]. univalent(M,A,P) \<longrightarrow>
+      (\<exists>Y[M]. \<forall>b[M]. b \<in> Y \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,b)))"
 
 definition
   foundation_ax :: "(i=>o) => o" where
     "foundation_ax(M) ==
-        \<forall>x[M]. (\<exists>y[M]. y\<in>x) --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & z \<in> y))"
+        \<forall>x[M]. (\<exists>y[M]. y\<in>x) \<longrightarrow> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & z \<in> y))"
 
 
 subsection{*A trivial consistency proof for $V_\omega$ *}
@@ -348,22 +348,22 @@
 done
 
 lemma univ0_Ball_abs [simp]:
-     "A \<in> univ(0) ==> (\<forall>x\<in>A. x \<in> univ(0) --> P(x)) <-> (\<forall>x\<in>A. P(x))"
+     "A \<in> univ(0) ==> (\<forall>x\<in>A. x \<in> univ(0) \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(x))"
 by (blast intro: univ0_downwards_mem)
 
 lemma univ0_Bex_abs [simp]:
-     "A \<in> univ(0) ==> (\<exists>x\<in>A. x \<in> univ(0) & P(x)) <-> (\<exists>x\<in>A. P(x))"
+     "A \<in> univ(0) ==> (\<exists>x\<in>A. x \<in> univ(0) & P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
 by (blast intro: univ0_downwards_mem)
 
 text{*Congruence rule for separation: can assume the variable is in @{text M}*}
 lemma separation_cong [cong]:
-     "(!!x. M(x) ==> P(x) <-> P'(x))
-      ==> separation(M, %x. P(x)) <-> separation(M, %x. P'(x))"
+     "(!!x. M(x) ==> P(x) \<longleftrightarrow> P'(x))
+      ==> separation(M, %x. P(x)) \<longleftrightarrow> separation(M, %x. P'(x))"
 by (simp add: separation_def)
 
 lemma univalent_cong [cong]:
-     "[| A=A'; !!x y. [| x\<in>A; M(x); M(y) |] ==> P(x,y) <-> P'(x,y) |]
-      ==> univalent(M, A, %x y. P(x,y)) <-> univalent(M, A', %x y. P'(x,y))"
+     "[| A=A'; !!x y. [| x\<in>A; M(x); M(y) |] ==> P(x,y) \<longleftrightarrow> P'(x,y) |]
+      ==> univalent(M, A, %x y. P(x,y)) \<longleftrightarrow> univalent(M, A', %x y. P'(x,y))"
 by (simp add: univalent_def)
 
 lemma univalent_triv [intro,simp]:
@@ -376,8 +376,8 @@
 
 text{*Congruence rule for replacement*}
 lemma strong_replacement_cong [cong]:
-     "[| !!x y. [| M(x); M(y) |] ==> P(x,y) <-> P'(x,y) |]
-      ==> strong_replacement(M, %x y. P(x,y)) <->
+     "[| !!x y. [| M(x); M(y) |] ==> P(x,y) \<longleftrightarrow> P'(x,y) |]
+      ==> strong_replacement(M, %x y. P(x,y)) \<longleftrightarrow>
           strong_replacement(M, %x y. P'(x,y))"
 by (simp add: strong_replacement_def)
 
@@ -457,13 +457,13 @@
 subsection{*Lemmas Needed to Reduce Some Set Constructions to Instances
       of Separation*}
 
-lemma image_iff_Collect: "r `` A = {y \<in> Union(Union(r)). \<exists>p\<in>r. \<exists>x\<in>A. p=<x,y>}"
+lemma image_iff_Collect: "r `` A = {y \<in> \<Union>(\<Union>(r)). \<exists>p\<in>r. \<exists>x\<in>A. p=<x,y>}"
 apply (rule equalityI, auto)
 apply (simp add: Pair_def, blast)
 done
 
 lemma vimage_iff_Collect:
-     "r -`` A = {x \<in> Union(Union(r)). \<exists>p\<in>r. \<exists>y\<in>A. p=<x,y>}"
+     "r -`` A = {x \<in> \<Union>(\<Union>(r)). \<exists>p\<in>r. \<exists>y\<in>A. p=<x,y>}"
 apply (rule equalityI, auto)
 apply (simp add: Pair_def, blast)
 done
@@ -485,16 +485,16 @@
 
 lemma replacementD:
     "[| replacement(M,P); M(A);  univalent(M,A,P) |]
-     ==> \<exists>Y[M]. (\<forall>b[M]. ((\<exists>x[M]. x\<in>A & P(x,b)) --> b \<in> Y))"
+     ==> \<exists>Y[M]. (\<forall>b[M]. ((\<exists>x[M]. x\<in>A & P(x,b)) \<longrightarrow> b \<in> Y))"
 by (simp add: replacement_def)
 
 lemma strong_replacementD:
     "[| strong_replacement(M,P); M(A);  univalent(M,A,P) |]
-     ==> \<exists>Y[M]. (\<forall>b[M]. (b \<in> Y <-> (\<exists>x[M]. x\<in>A & P(x,b))))"
+     ==> \<exists>Y[M]. (\<forall>b[M]. (b \<in> Y \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,b))))"
 by (simp add: strong_replacement_def)
 
 lemma separationD:
-    "[| separation(M,P); M(z) |] ==> \<exists>y[M]. \<forall>x[M]. x \<in> y <-> x \<in> z & P(x)"
+    "[| separation(M,P); M(z) |] ==> \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z & P(x)"
 by (simp add: separation_def)
 
 
@@ -504,20 +504,20 @@
   order_isomorphism :: "[i=>o,i,i,i,i,i] => o" where
     "order_isomorphism(M,A,r,B,s,f) ==
         bijection(M,A,B,f) &
-        (\<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A -->
+        (\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow>
           (\<forall>p[M]. \<forall>fx[M]. \<forall>fy[M]. \<forall>q[M].
-            pair(M,x,y,p) --> fun_apply(M,f,x,fx) --> fun_apply(M,f,y,fy) -->
-            pair(M,fx,fy,q) --> (p\<in>r <-> q\<in>s))))"
+            pair(M,x,y,p) \<longrightarrow> fun_apply(M,f,x,fx) \<longrightarrow> fun_apply(M,f,y,fy) \<longrightarrow>
+            pair(M,fx,fy,q) \<longrightarrow> (p\<in>r \<longleftrightarrow> q\<in>s))))"
 
 definition
   pred_set :: "[i=>o,i,i,i,i] => o" where
     "pred_set(M,A,x,r,B) ==
-        \<forall>y[M]. y \<in> B <-> (\<exists>p[M]. p\<in>r & y \<in> A & pair(M,y,x,p))"
+        \<forall>y[M]. y \<in> B \<longleftrightarrow> (\<exists>p[M]. p\<in>r & y \<in> A & pair(M,y,x,p))"
 
 definition
   membership :: "[i=>o,i,i] => o" where --{*membership relation*}
     "membership(M,A,r) ==
-        \<forall>p[M]. p \<in> r <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>A & x\<in>y & pair(M,x,y,p)))"
+        \<forall>p[M]. p \<in> r \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>A & x\<in>y & pair(M,x,y,p)))"
 
 
 subsection{*Introducing a Transitive Class Model*}
@@ -540,65 +540,65 @@
 by (blast intro: transM)
 
 lemma (in M_trivial) rall_abs [simp]:
-     "M(A) ==> (\<forall>x[M]. x\<in>A --> P(x)) <-> (\<forall>x\<in>A. P(x))"
+     "M(A) ==> (\<forall>x[M]. x\<in>A \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(x))"
 by (blast intro: transM)
 
 lemma (in M_trivial) rex_abs [simp]:
-     "M(A) ==> (\<exists>x[M]. x\<in>A & P(x)) <-> (\<exists>x\<in>A. P(x))"
+     "M(A) ==> (\<exists>x[M]. x\<in>A & P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
 by (blast intro: transM)
 
 lemma (in M_trivial) ball_iff_equiv:
-     "M(A) ==> (\<forall>x[M]. (x\<in>A <-> P(x))) <->
-               (\<forall>x\<in>A. P(x)) & (\<forall>x. P(x) --> M(x) --> x\<in>A)"
+     "M(A) ==> (\<forall>x[M]. (x\<in>A \<longleftrightarrow> P(x))) \<longleftrightarrow>
+               (\<forall>x\<in>A. P(x)) & (\<forall>x. P(x) \<longrightarrow> M(x) \<longrightarrow> x\<in>A)"
 by (blast intro: transM)
 
 text{*Simplifies proofs of equalities when there's an iff-equality
-      available for rewriting, universally quantified over M.  
+      available for rewriting, universally quantified over M.
       But it's not the only way to prove such equalities: its
       premises @{term "M(A)"} and  @{term "M(B)"} can be too strong.*}
 lemma (in M_trivial) M_equalityI:
-     "[| !!x. M(x) ==> x\<in>A <-> x\<in>B; M(A); M(B) |] ==> A=B"
+     "[| !!x. M(x) ==> x\<in>A \<longleftrightarrow> x\<in>B; M(A); M(B) |] ==> A=B"
 by (blast intro!: equalityI dest: transM)
 
 
 subsubsection{*Trivial Absoluteness Proofs: Empty Set, Pairs, etc.*}
 
 lemma (in M_trivial) empty_abs [simp]:
-     "M(z) ==> empty(M,z) <-> z=0"
+     "M(z) ==> empty(M,z) \<longleftrightarrow> z=0"
 apply (simp add: empty_def)
 apply (blast intro: transM)
 done
 
 lemma (in M_trivial) subset_abs [simp]:
-     "M(A) ==> subset(M,A,B) <-> A \<subseteq> B"
+     "M(A) ==> subset(M,A,B) \<longleftrightarrow> A \<subseteq> B"
 apply (simp add: subset_def)
 apply (blast intro: transM)
 done
 
 lemma (in M_trivial) upair_abs [simp]:
-     "M(z) ==> upair(M,a,b,z) <-> z={a,b}"
+     "M(z) ==> upair(M,a,b,z) \<longleftrightarrow> z={a,b}"
 apply (simp add: upair_def)
 apply (blast intro: transM)
 done
 
 lemma (in M_trivial) upair_in_M_iff [iff]:
-     "M({a,b}) <-> M(a) & M(b)"
+     "M({a,b}) \<longleftrightarrow> M(a) & M(b)"
 apply (insert upair_ax, simp add: upair_ax_def)
 apply (blast intro: transM)
 done
 
 lemma (in M_trivial) singleton_in_M_iff [iff]:
-     "M({a}) <-> M(a)"
+     "M({a}) \<longleftrightarrow> M(a)"
 by (insert upair_in_M_iff [of a a], simp)
 
 lemma (in M_trivial) pair_abs [simp]:
-     "M(z) ==> pair(M,a,b,z) <-> z=<a,b>"
+     "M(z) ==> pair(M,a,b,z) \<longleftrightarrow> z=<a,b>"
 apply (simp add: pair_def ZF.Pair_def)
 apply (blast intro: transM)
 done
 
 lemma (in M_trivial) pair_in_M_iff [iff]:
-     "M(<a,b>) <-> M(a) & M(b)"
+     "M(<a,b>) \<longleftrightarrow> M(a) & M(b)"
 by (simp add: ZF.Pair_def)
 
 lemma (in M_trivial) pair_components_in_M:
@@ -608,7 +608,7 @@
 done
 
 lemma (in M_trivial) cartprod_abs [simp]:
-     "[| M(A); M(B); M(z) |] ==> cartprod(M,A,B,z) <-> z = A*B"
+     "[| M(A); M(B); M(z) |] ==> cartprod(M,A,B,z) \<longleftrightarrow> z = A*B"
 apply (simp add: cartprod_def)
 apply (rule iffI)
  apply (blast intro!: equalityI intro: transM dest!: rspec)
@@ -618,35 +618,35 @@
 subsubsection{*Absoluteness for Unions and Intersections*}
 
 lemma (in M_trivial) union_abs [simp]:
-     "[| M(a); M(b); M(z) |] ==> union(M,a,b,z) <-> z = a Un b"
+     "[| M(a); M(b); M(z) |] ==> union(M,a,b,z) \<longleftrightarrow> z = a \<union> b"
 apply (simp add: union_def)
 apply (blast intro: transM)
 done
 
 lemma (in M_trivial) inter_abs [simp]:
-     "[| M(a); M(b); M(z) |] ==> inter(M,a,b,z) <-> z = a Int b"
+     "[| M(a); M(b); M(z) |] ==> inter(M,a,b,z) \<longleftrightarrow> z = a \<inter> b"
 apply (simp add: inter_def)
 apply (blast intro: transM)
 done
 
 lemma (in M_trivial) setdiff_abs [simp]:
-     "[| M(a); M(b); M(z) |] ==> setdiff(M,a,b,z) <-> z = a-b"
+     "[| M(a); M(b); M(z) |] ==> setdiff(M,a,b,z) \<longleftrightarrow> z = a-b"
 apply (simp add: setdiff_def)
 apply (blast intro: transM)
 done
 
 lemma (in M_trivial) Union_abs [simp]:
-     "[| M(A); M(z) |] ==> big_union(M,A,z) <-> z = Union(A)"
+     "[| M(A); M(z) |] ==> big_union(M,A,z) \<longleftrightarrow> z = \<Union>(A)"
 apply (simp add: big_union_def)
 apply (blast intro!: equalityI dest: transM)
 done
 
 lemma (in M_trivial) Union_closed [intro,simp]:
-     "M(A) ==> M(Union(A))"
+     "M(A) ==> M(\<Union>(A))"
 by (insert Union_ax, simp add: Union_ax_def)
 
 lemma (in M_trivial) Un_closed [intro,simp]:
-     "[| M(A); M(B) |] ==> M(A Un B)"
+     "[| M(A); M(B) |] ==> M(A \<union> B)"
 by (simp only: Un_eq_Union, blast)
 
 lemma (in M_trivial) cons_closed [intro,simp]:
@@ -654,15 +654,15 @@
 by (subst cons_eq [symmetric], blast)
 
 lemma (in M_trivial) cons_abs [simp]:
-     "[| M(b); M(z) |] ==> is_cons(M,a,b,z) <-> z = cons(a,b)"
+     "[| M(b); M(z) |] ==> is_cons(M,a,b,z) \<longleftrightarrow> z = cons(a,b)"
 by (simp add: is_cons_def, blast intro: transM)
 
 lemma (in M_trivial) successor_abs [simp]:
-     "[| M(a); M(z) |] ==> successor(M,a,z) <-> z = succ(a)"
+     "[| M(a); M(z) |] ==> successor(M,a,z) \<longleftrightarrow> z = succ(a)"
 by (simp add: successor_def, blast)
 
 lemma (in M_trivial) succ_in_M_iff [iff]:
-     "M(succ(a)) <-> M(a)"
+     "M(succ(a)) \<longleftrightarrow> M(a)"
 apply (simp add: succ_def)
 apply (blast intro: transM)
 done
@@ -678,11 +678,11 @@
 done
 
 lemma separation_iff:
-     "separation(M,P) <-> (\<forall>z[M]. \<exists>y[M]. is_Collect(M,z,P,y))"
+     "separation(M,P) \<longleftrightarrow> (\<forall>z[M]. \<exists>y[M]. is_Collect(M,z,P,y))"
 by (simp add: separation_def is_Collect_def)
 
 lemma (in M_trivial) Collect_abs [simp]:
-     "[| M(A); M(z) |] ==> is_Collect(M,A,P,z) <-> z = Collect(A,P)"
+     "[| M(A); M(z) |] ==> is_Collect(M,A,P,z) \<longleftrightarrow> z = Collect(A,P)"
 apply (simp add: is_Collect_def)
 apply (blast intro!: equalityI dest: transM)
 done
@@ -703,16 +703,16 @@
 
 lemma is_Replace_cong [cong]:
      "[| A=A';
-         !!x y. [| M(x); M(y) |] ==> P(x,y) <-> P'(x,y);
+         !!x y. [| M(x); M(y) |] ==> P(x,y) \<longleftrightarrow> P'(x,y);
          z=z' |]
-      ==> is_Replace(M, A, %x y. P(x,y), z) <->
+      ==> is_Replace(M, A, %x y. P(x,y), z) \<longleftrightarrow>
           is_Replace(M, A', %x y. P'(x,y), z')"
 by (simp add: is_Replace_def)
 
 lemma (in M_trivial) univalent_Replace_iff:
      "[| M(A); univalent(M,A,P);
          !!x y. [| x\<in>A; P(x,y) |] ==> M(y) |]
-      ==> u \<in> Replace(A,P) <-> (\<exists>x. x\<in>A & P(x,u))"
+      ==> u \<in> Replace(A,P) \<longleftrightarrow> (\<exists>x. x\<in>A & P(x,u))"
 apply (simp add: Replace_iff univalent_def)
 apply (blast dest: transM)
 done
@@ -731,13 +731,13 @@
 done
 
 lemma (in M_trivial) Replace_abs:
-     "[| M(A); M(z); univalent(M,A,P); 
+     "[| M(A); M(z); univalent(M,A,P);
          !!x y. [| x\<in>A; P(x,y) |] ==> M(y)  |]
-      ==> is_Replace(M,A,P,z) <-> z = Replace(A,P)"
+      ==> is_Replace(M,A,P,z) \<longleftrightarrow> z = Replace(A,P)"
 apply (simp add: is_Replace_def)
 apply (rule iffI)
  apply (rule equality_iffI)
- apply (simp_all add: univalent_Replace_iff) 
+ apply (simp_all add: univalent_Replace_iff)
  apply (blast dest: transM)+
 done
 
@@ -747,7 +747,7 @@
   Let K be a nonconstructible subset of nat and define
   f(x) = x if x:K and f(x)=0 otherwise.  Then RepFun(nat,f) = cons(0,K), a
   nonconstructible set.  So we cannot assume that M(X) implies M(RepFun(X,f))
-  even for f : M -> M.
+  even for f \<in> M -> M.
 *)
 lemma (in M_trivial) RepFun_closed:
      "[| strong_replacement(M, \<lambda>x y. y = f(x)); M(A); \<forall>x\<in>A. M(f(x)) |]
@@ -775,7 +775,7 @@
 definition
  is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" where
     "is_lambda(M, A, is_b, z) ==
-       \<forall>p[M]. p \<in> z <->
+       \<forall>p[M]. p \<in> z \<longleftrightarrow>
         (\<exists>u[M]. \<exists>v[M]. u\<in>A & pair(M,u,v,p) & is_b(u,v))"
 
 lemma (in M_trivial) lam_closed:
@@ -786,33 +786,33 @@
 text{*Better than @{text lam_closed}: has the formula @{term "x\<in>A"}*}
 lemma (in M_trivial) lam_closed2:
   "[|strong_replacement(M, \<lambda>x y. x\<in>A & y = \<langle>x, b(x)\<rangle>);
-     M(A); \<forall>m[M]. m\<in>A --> M(b(m))|] ==> M(Lambda(A,b))"
+     M(A); \<forall>m[M]. m\<in>A \<longrightarrow> M(b(m))|] ==> M(Lambda(A,b))"
 apply (simp add: lam_def)
 apply (blast intro: RepFun_closed2 dest: transM)
 done
 
 lemma (in M_trivial) lambda_abs2:
-     "[| Relation1(M,A,is_b,b); M(A); \<forall>m[M]. m\<in>A --> M(b(m)); M(z) |]
-      ==> is_lambda(M,A,is_b,z) <-> z = Lambda(A,b)"
+     "[| Relation1(M,A,is_b,b); M(A); \<forall>m[M]. m\<in>A \<longrightarrow> M(b(m)); M(z) |]
+      ==> is_lambda(M,A,is_b,z) \<longleftrightarrow> z = Lambda(A,b)"
 apply (simp add: Relation1_def is_lambda_def)
 apply (rule iffI)
  prefer 2 apply (simp add: lam_def)
 apply (rule equality_iffI)
-apply (simp add: lam_def) 
-apply (rule iffI) 
- apply (blast dest: transM) 
-apply (auto simp add: transM [of _ A]) 
+apply (simp add: lam_def)
+apply (rule iffI)
+ apply (blast dest: transM)
+apply (auto simp add: transM [of _ A])
 done
 
 lemma is_lambda_cong [cong]:
      "[| A=A';  z=z';
-         !!x y. [| x\<in>A; M(x); M(y) |] ==> is_b(x,y) <-> is_b'(x,y) |]
-      ==> is_lambda(M, A, %x y. is_b(x,y), z) <->
+         !!x y. [| x\<in>A; M(x); M(y) |] ==> is_b(x,y) \<longleftrightarrow> is_b'(x,y) |]
+      ==> is_lambda(M, A, %x y. is_b(x,y), z) \<longleftrightarrow>
           is_lambda(M, A', %x y. is_b'(x,y), z')"
 by (simp add: is_lambda_def)
 
 lemma (in M_trivial) image_abs [simp]:
-     "[| M(r); M(A); M(z) |] ==> image(M,r,A,z) <-> z = r``A"
+     "[| M(r); M(A); M(z) |] ==> image(M,r,A,z) \<longleftrightarrow> z = r``A"
 apply (simp add: image_def)
 apply (rule iffI)
  apply (blast intro!: equalityI dest: transM, blast)
@@ -828,7 +828,7 @@
 text{*But we can't prove that the powerset in @{text M} includes the
       real powerset.*}
 lemma (in M_trivial) powerset_imp_subset_Pow:
-     "[| powerset(M,x,y); M(y) |] ==> y <= Pow(x)"
+     "[| powerset(M,x,y); M(y) |] ==> y \<subseteq> Pow(x)"
 apply (simp add: powerset_def)
 apply (blast dest: transM)
 done
@@ -847,12 +847,12 @@
 done
 
 lemma (in M_trivial) quasinat_abs [simp]:
-     "M(z) ==> is_quasinat(M,z) <-> quasinat(z)"
+     "M(z) ==> is_quasinat(M,z) \<longleftrightarrow> quasinat(z)"
 by (auto simp add: is_quasinat_def quasinat_def)
 
 lemma (in M_trivial) nat_case_abs [simp]:
      "[| relation1(M,is_b,b); M(k); M(z) |]
-      ==> is_nat_case(M,a,is_b,k,z) <-> z = nat_case(a,b,k)"
+      ==> is_nat_case(M,a,is_b,k,z) \<longleftrightarrow> z = nat_case(a,b,k)"
 apply (case_tac "quasinat(k)")
  prefer 2
  apply (simp add: is_nat_case_def non_nat_case)
@@ -867,8 +867,8 @@
   is_nat_case by nat_case before attempting congruence reasoning.*)
 lemma is_nat_case_cong:
      "[| a = a'; k = k';  z = z';  M(z');
-       !!x y. [| M(x); M(y) |] ==> is_b(x,y) <-> is_b'(x,y) |]
-      ==> is_nat_case(M, a, is_b, k, z) <-> is_nat_case(M, a', is_b', k', z')"
+       !!x y. [| M(x); M(y) |] ==> is_b(x,y) \<longleftrightarrow> is_b'(x,y) |]
+      ==> is_nat_case(M, a, is_b, k, z) \<longleftrightarrow> is_nat_case(M, a', is_b', k', z')"
 by (simp add: is_nat_case_def)
 
 
@@ -880,22 +880,22 @@
 by (blast dest: ltD intro: transM)
 
 lemma (in M_trivial) transitive_set_abs [simp]:
-     "M(a) ==> transitive_set(M,a) <-> Transset(a)"
+     "M(a) ==> transitive_set(M,a) \<longleftrightarrow> Transset(a)"
 by (simp add: transitive_set_def Transset_def)
 
 lemma (in M_trivial) ordinal_abs [simp]:
-     "M(a) ==> ordinal(M,a) <-> Ord(a)"
+     "M(a) ==> ordinal(M,a) \<longleftrightarrow> Ord(a)"
 by (simp add: ordinal_def Ord_def)
 
 lemma (in M_trivial) limit_ordinal_abs [simp]:
-     "M(a) ==> limit_ordinal(M,a) <-> Limit(a)"
+     "M(a) ==> limit_ordinal(M,a) \<longleftrightarrow> Limit(a)"
 apply (unfold Limit_def limit_ordinal_def)
 apply (simp add: Ord_0_lt_iff)
 apply (simp add: lt_def, blast)
 done
 
 lemma (in M_trivial) successor_ordinal_abs [simp]:
-     "M(a) ==> successor_ordinal(M,a) <-> Ord(a) & (\<exists>b[M]. a = succ(b))"
+     "M(a) ==> successor_ordinal(M,a) \<longleftrightarrow> Ord(a) & (\<exists>b[M]. a = succ(b))"
 apply (simp add: successor_ordinal_def, safe)
 apply (drule Ord_cases_disj, auto)
 done
@@ -905,7 +905,7 @@
 by (induct a rule: trans_induct3, simp_all)
 
 lemma (in M_trivial) finite_ordinal_abs [simp]:
-     "M(a) ==> finite_ordinal(M,a) <-> a \<in> nat"
+     "M(a) ==> finite_ordinal(M,a) \<longleftrightarrow> a \<in> nat"
 apply (simp add: finite_ordinal_def)
 apply (blast intro: finite_Ord_is_nat intro: nat_into_Ord
              dest: Ord_trans naturals_not_limit)
@@ -921,21 +921,21 @@
 done
 
 lemma (in M_trivial) omega_abs [simp]:
-     "M(a) ==> omega(M,a) <-> a = nat"
+     "M(a) ==> omega(M,a) \<longleftrightarrow> a = nat"
 apply (simp add: omega_def)
 apply (blast intro: Limit_non_Limit_implies_nat dest: naturals_not_limit)
 done
 
 lemma (in M_trivial) number1_abs [simp]:
-     "M(a) ==> number1(M,a) <-> a = 1"
+     "M(a) ==> number1(M,a) \<longleftrightarrow> a = 1"
 by (simp add: number1_def)
 
 lemma (in M_trivial) number2_abs [simp]:
-     "M(a) ==> number2(M,a) <-> a = succ(1)"
+     "M(a) ==> number2(M,a) \<longleftrightarrow> a = succ(1)"
 by (simp add: number2_def)
 
 lemma (in M_trivial) number3_abs [simp]:
-     "M(a) ==> number3(M,a) <-> a = succ(succ(1))"
+     "M(a) ==> number3(M,a) \<longleftrightarrow> a = succ(succ(1))"
 by (simp add: number3_def)
 
 text{*Kunen continued to 20...*}
@@ -965,7 +965,7 @@
 
 locale M_basic = M_trivial +
 assumes Inter_separation:
-     "M(A) ==> separation(M, \<lambda>x. \<forall>y[M]. y\<in>A --> x\<in>y)"
+     "M(A) ==> separation(M, \<lambda>x. \<forall>y[M]. y\<in>A \<longrightarrow> x\<in>y)"
   and Diff_separation:
      "M(B) ==> separation(M, \<lambda>x. x \<notin> B)"
   and cartprod_separation:
@@ -1003,7 +1003,7 @@
                                    fx \<noteq> gx))"
 
 lemma (in M_basic) cartprod_iff_lemma:
-     "[| M(C);  \<forall>u[M]. u \<in> C <-> (\<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}});
+     "[| M(C);  \<forall>u[M]. u \<in> C \<longleftrightarrow> (\<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}});
          powerset(M, A \<union> B, p1); powerset(M, p1, p2);  M(p2) |]
        ==> C = {u \<in> p2 . \<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}}}"
 apply (simp add: powerset_def)
@@ -1017,8 +1017,8 @@
 
 lemma (in M_basic) cartprod_iff:
      "[| M(A); M(B); M(C) |]
-      ==> cartprod(M,A,B,C) <->
-          (\<exists>p1[M]. \<exists>p2[M]. powerset(M,A Un B,p1) & powerset(M,p1,p2) &
+      ==> cartprod(M,A,B,C) \<longleftrightarrow>
+          (\<exists>p1[M]. \<exists>p2[M]. powerset(M,A \<union> B,p1) & powerset(M,p1,p2) &
                    C = {z \<in> p2. \<exists>x\<in>A. \<exists>y\<in>B. z = <x,y>})"
 apply (simp add: Pair_def cartprod_def, safe)
 defer 1
@@ -1026,7 +1026,7 @@
  apply blast
 txt{*Final, difficult case: the left-to-right direction of the theorem.*}
 apply (insert power_ax, simp add: power_ax_def)
-apply (frule_tac x="A Un B" and P="\<lambda>x. rex(M,?Q(x))" in rspec)
+apply (frule_tac x="A \<union> B" and P="\<lambda>x. rex(M,?Q(x))" in rspec)
 apply (blast, clarify)
 apply (drule_tac x=z and P="\<lambda>x. rex(M,?Q(x))" in rspec)
 apply assumption
@@ -1037,7 +1037,7 @@
      "[| M(A); M(B) |] ==> \<exists>C[M]. cartprod(M,A,B,C)"
 apply (simp del: cartprod_abs add: cartprod_iff)
 apply (insert power_ax, simp add: power_ax_def)
-apply (frule_tac x="A Un B" and P="\<lambda>x. rex(M,?Q(x))" in rspec)
+apply (frule_tac x="A \<union> B" and P="\<lambda>x. rex(M,?Q(x))" in rspec)
 apply (blast, clarify)
 apply (drule_tac x=z and P="\<lambda>x. rex(M,?Q(x))" in rspec, auto)
 apply (intro rexI conjI, simp+)
@@ -1055,23 +1055,23 @@
 by (simp add: sum_def)
 
 lemma (in M_basic) sum_abs [simp]:
-     "[| M(A); M(B); M(Z) |] ==> is_sum(M,A,B,Z) <-> (Z = A+B)"
+     "[| M(A); M(B); M(Z) |] ==> is_sum(M,A,B,Z) \<longleftrightarrow> (Z = A+B)"
 by (simp add: is_sum_def sum_def singleton_0 nat_into_M)
 
 lemma (in M_trivial) Inl_in_M_iff [iff]:
-     "M(Inl(a)) <-> M(a)"
+     "M(Inl(a)) \<longleftrightarrow> M(a)"
 by (simp add: Inl_def)
 
 lemma (in M_trivial) Inl_abs [simp]:
-     "M(Z) ==> is_Inl(M,a,Z) <-> (Z = Inl(a))"
+     "M(Z) ==> is_Inl(M,a,Z) \<longleftrightarrow> (Z = Inl(a))"
 by (simp add: is_Inl_def Inl_def)
 
 lemma (in M_trivial) Inr_in_M_iff [iff]:
-     "M(Inr(a)) <-> M(a)"
+     "M(Inr(a)) \<longleftrightarrow> M(a)"
 by (simp add: Inr_def)
 
 lemma (in M_trivial) Inr_abs [simp]:
-     "M(Z) ==> is_Inr(M,a,Z) <-> (Z = Inr(a))"
+     "M(Z) ==> is_Inr(M,a,Z) \<longleftrightarrow> (Z = Inr(a))"
 by (simp add: is_Inr_def Inr_def)
 
 
@@ -1080,7 +1080,7 @@
 lemma (in M_basic) M_converse_iff:
      "M(r) ==>
       converse(r) =
-      {z \<in> Union(Union(r)) * Union(Union(r)).
+      {z \<in> \<Union>(\<Union>(r)) * \<Union>(\<Union>(r)).
        \<exists>p\<in>r. \<exists>x[M]. \<exists>y[M]. p = \<langle>x,y\<rangle> & z = \<langle>y,x\<rangle>}"
 apply (rule equalityI)
  prefer 2 apply (blast dest: transM, clarify, simp)
@@ -1095,7 +1095,7 @@
 done
 
 lemma (in M_basic) converse_abs [simp]:
-     "[| M(r); M(z) |] ==> is_converse(M,r,z) <-> z = converse(r)"
+     "[| M(r); M(z) |] ==> is_converse(M,r,z) \<longleftrightarrow> z = converse(r)"
 apply (simp add: is_converse_def)
 apply (rule iffI)
  prefer 2 apply blast
@@ -1114,7 +1114,7 @@
 done
 
 lemma (in M_basic) vimage_abs [simp]:
-     "[| M(r); M(A); M(z) |] ==> pre_image(M,r,A,z) <-> z = r-``A"
+     "[| M(r); M(A); M(z) |] ==> pre_image(M,r,A,z) \<longleftrightarrow> z = r-``A"
 apply (simp add: pre_image_def)
 apply (rule iffI)
  apply (blast intro!: equalityI dest: transM, blast)
@@ -1128,7 +1128,7 @@
 subsubsection{*Domain, range and field*}
 
 lemma (in M_basic) domain_abs [simp]:
-     "[| M(r); M(z) |] ==> is_domain(M,r,z) <-> z = domain(r)"
+     "[| M(r); M(z) |] ==> is_domain(M,r,z) \<longleftrightarrow> z = domain(r)"
 apply (simp add: is_domain_def)
 apply (blast intro!: equalityI dest: transM)
 done
@@ -1139,7 +1139,7 @@
 done
 
 lemma (in M_basic) range_abs [simp]:
-     "[| M(r); M(z) |] ==> is_range(M,r,z) <-> z = range(r)"
+     "[| M(r); M(z) |] ==> is_range(M,r,z) \<longleftrightarrow> z = range(r)"
 apply (simp add: is_range_def)
 apply (blast intro!: equalityI dest: transM)
 done
@@ -1150,7 +1150,7 @@
 done
 
 lemma (in M_basic) field_abs [simp]:
-     "[| M(r); M(z) |] ==> is_field(M,r,z) <-> z = field(r)"
+     "[| M(r); M(z) |] ==> is_field(M,r,z) \<longleftrightarrow> z = field(r)"
 by (simp add: domain_closed range_closed is_field_def field_def)
 
 lemma (in M_basic) field_closed [intro,simp]:
@@ -1161,13 +1161,13 @@
 subsubsection{*Relations, functions and application*}
 
 lemma (in M_basic) relation_abs [simp]:
-     "M(r) ==> is_relation(M,r) <-> relation(r)"
+     "M(r) ==> is_relation(M,r) \<longleftrightarrow> relation(r)"
 apply (simp add: is_relation_def relation_def)
 apply (blast dest!: bspec dest: pair_components_in_M)+
 done
 
 lemma (in M_basic) function_abs [simp]:
-     "M(r) ==> is_function(M,r) <-> function(r)"
+     "M(r) ==> is_function(M,r) \<longleftrightarrow> function(r)"
 apply (simp add: is_function_def function_def, safe)
    apply (frule transM, assumption)
   apply (blast dest: pair_components_in_M)+
@@ -1178,28 +1178,28 @@
 by (simp add: apply_def)
 
 lemma (in M_basic) apply_abs [simp]:
-     "[| M(f); M(x); M(y) |] ==> fun_apply(M,f,x,y) <-> f`x = y"
+     "[| M(f); M(x); M(y) |] ==> fun_apply(M,f,x,y) \<longleftrightarrow> f`x = y"
 apply (simp add: fun_apply_def apply_def, blast)
 done
 
 lemma (in M_basic) typed_function_abs [simp]:
-     "[| M(A); M(f) |] ==> typed_function(M,A,B,f) <-> f \<in> A -> B"
+     "[| M(A); M(f) |] ==> typed_function(M,A,B,f) \<longleftrightarrow> f \<in> A -> B"
 apply (auto simp add: typed_function_def relation_def Pi_iff)
 apply (blast dest: pair_components_in_M)+
 done
 
 lemma (in M_basic) injection_abs [simp]:
-     "[| M(A); M(f) |] ==> injection(M,A,B,f) <-> f \<in> inj(A,B)"
+     "[| M(A); M(f) |] ==> injection(M,A,B,f) \<longleftrightarrow> f \<in> inj(A,B)"
 apply (simp add: injection_def apply_iff inj_def apply_closed)
 apply (blast dest: transM [of _ A])
 done
 
 lemma (in M_basic) surjection_abs [simp]:
-     "[| M(A); M(B); M(f) |] ==> surjection(M,A,B,f) <-> f \<in> surj(A,B)"
+     "[| M(A); M(B); M(f) |] ==> surjection(M,A,B,f) \<longleftrightarrow> f \<in> surj(A,B)"
 by (simp add: surjection_def surj_def)
 
 lemma (in M_basic) bijection_abs [simp]:
-     "[| M(A); M(B); M(f) |] ==> bijection(M,A,B,f) <-> f \<in> bij(A,B)"
+     "[| M(A); M(B); M(f) |] ==> bijection(M,A,B,f) \<longleftrightarrow> f \<in> bij(A,B)"
 by (simp add: bijection_def bij_def)
 
 
@@ -1224,7 +1224,7 @@
 done
 
 lemma (in M_basic) composition_abs [simp]:
-     "[| M(r); M(s); M(t) |] ==> composition(M,r,s,t) <-> t = r O s"
+     "[| M(r); M(s); M(t) |] ==> composition(M,r,s,t) \<longleftrightarrow> t = r O s"
 apply safe
  txt{*Proving @{term "composition(M, r, s, r O s)"}*}
  prefer 2
@@ -1246,7 +1246,7 @@
 
 lemma (in M_basic) restriction_abs [simp]:
      "[| M(f); M(A); M(z) |]
-      ==> restriction(M,f,A,z) <-> z = restrict(f,A)"
+      ==> restriction(M,f,A,z) \<longleftrightarrow> z = restrict(f,A)"
 apply (simp add: ball_iff_equiv restriction_def restrict_def)
 apply (blast intro!: equalityI dest: transM)
 done
@@ -1263,17 +1263,17 @@
 done
 
 lemma (in M_basic) Inter_abs [simp]:
-     "[| M(A); M(z) |] ==> big_inter(M,A,z) <-> z = Inter(A)"
+     "[| M(A); M(z) |] ==> big_inter(M,A,z) \<longleftrightarrow> z = \<Inter>(A)"
 apply (simp add: big_inter_def Inter_def)
 apply (blast intro!: equalityI dest: transM)
 done
 
 lemma (in M_basic) Inter_closed [intro,simp]:
-     "M(A) ==> M(Inter(A))"
+     "M(A) ==> M(\<Inter>(A))"
 by (insert Inter_separation, simp add: Inter_def)
 
 lemma (in M_basic) Int_closed [intro,simp]:
-     "[| M(A); M(B) |] ==> M(A Int B)"
+     "[| M(A); M(B) |] ==> M(A \<inter> B)"
 apply (subgoal_tac "M({A,B})")
 apply (frule Inter_closed, force+)
 done
@@ -1291,7 +1291,7 @@
 
 (*???equalities*)
 lemma Collect_Un_Collect_eq:
-     "Collect(A,P) Un Collect(A,Q) = Collect(A, %x. P(x) | Q(x))"
+     "Collect(A,P) \<union> Collect(A,Q) = Collect(A, %x. P(x) | Q(x))"
 by blast
 
 lemma Diff_Collect_eq:
@@ -1299,7 +1299,7 @@
 by blast
 
 lemma (in M_trivial) Collect_rall_eq:
-     "M(Y) ==> Collect(A, %x. \<forall>y[M]. y\<in>Y --> P(x,y)) =
+     "M(Y) ==> Collect(A, %x. \<forall>y[M]. y\<in>Y \<longrightarrow> P(x,y)) =
                (if Y=0 then A else (\<Inter>y \<in> Y. {x \<in> A. P(x,y)}))"
 apply simp
 apply (blast intro!: equalityI dest: transM)
@@ -1317,7 +1317,7 @@
 
 lemma (in M_basic) separation_imp:
      "[|separation(M,P); separation(M,Q)|]
-      ==> separation(M, \<lambda>z. P(z) --> Q(z))"
+      ==> separation(M, \<lambda>z. P(z) \<longrightarrow> Q(z))"
 by (simp add: separation_neg separation_disj not_disj_iff_imp [symmetric])
 
 text{*This result is a hint of how little can be done without the Reflection
@@ -1326,7 +1326,7 @@
 lemma (in M_basic) separation_rall:
      "[|M(Y); \<forall>y[M]. separation(M, \<lambda>x. P(x,y));
         \<forall>z[M]. strong_replacement(M, \<lambda>x y. y = {u \<in> z . P(u,x)})|]
-      ==> separation(M, \<lambda>x. \<forall>y[M]. y\<in>Y --> P(x,y))"
+      ==> separation(M, \<lambda>x. \<forall>y[M]. y\<in>Y \<longrightarrow> P(x,y))"
 apply (simp del: separation_closed rall_abs
          add: separation_iff Collect_rall_eq)
 apply (blast intro!: Inter_closed RepFun_closed dest: transM)
@@ -1338,7 +1338,7 @@
 text{*The assumption @{term "M(A->B)"} is unusual, but essential: in
 all but trivial cases, A->B cannot be expected to belong to @{term M}.*}
 lemma (in M_basic) is_funspace_abs [simp]:
-     "[|M(A); M(B); M(F); M(A->B)|] ==> is_funspace(M,A,B,F) <-> F = A->B";
+     "[|M(A); M(B); M(F); M(A->B)|] ==> is_funspace(M,A,B,F) \<longleftrightarrow> F = A->B";
 apply (simp add: is_funspace_def)
 apply (rule iffI)
  prefer 2 apply blast
@@ -1392,20 +1392,20 @@
                       (~number1(M,a) & z=b)"
 
 lemma (in M_trivial) bool_of_o_abs [simp]:
-     "M(z) ==> is_bool_of_o(M,P,z) <-> z = bool_of_o(P)"
+     "M(z) ==> is_bool_of_o(M,P,z) \<longleftrightarrow> z = bool_of_o(P)"
 by (simp add: is_bool_of_o_def bool_of_o_def)
 
 
 lemma (in M_trivial) not_abs [simp]:
-     "[| M(a); M(z)|] ==> is_not(M,a,z) <-> z = not(a)"
+     "[| M(a); M(z)|] ==> is_not(M,a,z) \<longleftrightarrow> z = not(a)"
 by (simp add: Bool.not_def cond_def is_not_def)
 
 lemma (in M_trivial) and_abs [simp]:
-     "[| M(a); M(b); M(z)|] ==> is_and(M,a,b,z) <-> z = a and b"
+     "[| M(a); M(b); M(z)|] ==> is_and(M,a,b,z) \<longleftrightarrow> z = a and b"
 by (simp add: Bool.and_def cond_def is_and_def)
 
 lemma (in M_trivial) or_abs [simp]:
-     "[| M(a); M(b); M(z)|] ==> is_or(M,a,b,z) <-> z = a or b"
+     "[| M(a); M(b); M(z)|] ==> is_or(M,a,b,z) \<longleftrightarrow> z = a or b"
 by (simp add: Bool.or_def cond_def is_or_def)
 
 
@@ -1442,14 +1442,14 @@
 lemma (in M_trivial) Nil_in_M [intro,simp]: "M(Nil)"
 by (simp add: Nil_def)
 
-lemma (in M_trivial) Nil_abs [simp]: "M(Z) ==> is_Nil(M,Z) <-> (Z = Nil)"
+lemma (in M_trivial) Nil_abs [simp]: "M(Z) ==> is_Nil(M,Z) \<longleftrightarrow> (Z = Nil)"
 by (simp add: is_Nil_def Nil_def)
 
-lemma (in M_trivial) Cons_in_M_iff [iff]: "M(Cons(a,l)) <-> M(a) & M(l)"
+lemma (in M_trivial) Cons_in_M_iff [iff]: "M(Cons(a,l)) \<longleftrightarrow> M(a) & M(l)"
 by (simp add: Cons_def)
 
 lemma (in M_trivial) Cons_abs [simp]:
-     "[|M(a); M(l); M(Z)|] ==> is_Cons(M,a,l,Z) <-> (Z = Cons(a,l))"
+     "[|M(a); M(l); M(Z)|] ==> is_Cons(M,a,l,Z) \<longleftrightarrow> (Z = Cons(a,l))"
 by (simp add: is_Cons_def Cons_def)
 
 
@@ -1471,8 +1471,8 @@
   is_list_case :: "[i=>o, i, [i,i,i]=>o, i, i] => o" where
     --{*Returns 0 for non-lists*}
     "is_list_case(M, a, is_b, xs, z) ==
-       (is_Nil(M,xs) --> z=a) &
-       (\<forall>x[M]. \<forall>l[M]. is_Cons(M,x,l,xs) --> is_b(x,l,z)) &
+       (is_Nil(M,xs) \<longrightarrow> z=a) &
+       (\<forall>x[M]. \<forall>l[M]. is_Cons(M,x,l,xs) \<longrightarrow> is_b(x,l,z)) &
        (is_quasilist(M,xs) | empty(M,z))"
 
 definition
@@ -1490,7 +1490,7 @@
      --{* @{term "hd([]) = 0"} no constraints if not a list.
           Avoiding implication prevents the simplifier's looping.*}
     "is_hd(M,xs,H) ==
-       (is_Nil(M,xs) --> empty(M,H)) &
+       (is_Nil(M,xs) \<longrightarrow> empty(M,H)) &
        (\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | H=x) &
        (is_quasilist(M,xs) | empty(M,H))"
 
@@ -1498,7 +1498,7 @@
   is_tl :: "[i=>o,i,i] => o" where
      --{* @{term "tl([]) = []"}; see comments about @{term is_hd}*}
     "is_tl(M,xs,T) ==
-       (is_Nil(M,xs) --> T=xs) &
+       (is_Nil(M,xs) \<longrightarrow> T=xs) &
        (\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | T=l) &
        (is_quasilist(M,xs) | empty(M,T))"
 
@@ -1536,12 +1536,12 @@
 done
 
 lemma (in M_trivial) quasilist_abs [simp]:
-     "M(z) ==> is_quasilist(M,z) <-> quasilist(z)"
+     "M(z) ==> is_quasilist(M,z) \<longleftrightarrow> quasilist(z)"
 by (auto simp add: is_quasilist_def quasilist_def)
 
 lemma (in M_trivial) list_case_abs [simp]:
      "[| relation2(M,is_b,b); M(k); M(z) |]
-      ==> is_list_case(M,a,is_b,k,z) <-> z = list_case'(a,b,k)"
+      ==> is_list_case(M,a,is_b,k,z) \<longleftrightarrow> z = list_case'(a,b,k)"
 apply (case_tac "quasilist(k)")
  prefer 2
  apply (simp add: is_list_case_def non_list_case)
@@ -1554,15 +1554,15 @@
 
 subsubsection{*The Modified Operators @{term hd'} and @{term tl'}*}
 
-lemma (in M_trivial) is_hd_Nil: "is_hd(M,[],Z) <-> empty(M,Z)"
+lemma (in M_trivial) is_hd_Nil: "is_hd(M,[],Z) \<longleftrightarrow> empty(M,Z)"
 by (simp add: is_hd_def)
 
 lemma (in M_trivial) is_hd_Cons:
-     "[|M(a); M(l)|] ==> is_hd(M,Cons(a,l),Z) <-> Z = a"
+     "[|M(a); M(l)|] ==> is_hd(M,Cons(a,l),Z) \<longleftrightarrow> Z = a"
 by (force simp add: is_hd_def)
 
 lemma (in M_trivial) hd_abs [simp]:
-     "[|M(x); M(y)|] ==> is_hd(M,x,y) <-> y = hd'(x)"
+     "[|M(x); M(y)|] ==> is_hd(M,x,y) \<longleftrightarrow> y = hd'(x)"
 apply (simp add: hd'_def)
 apply (intro impI conjI)
  prefer 2 apply (force simp add: is_hd_def)
@@ -1570,15 +1570,15 @@
 apply (elim disjE exE, auto)
 done
 
-lemma (in M_trivial) is_tl_Nil: "is_tl(M,[],Z) <-> Z = []"
+lemma (in M_trivial) is_tl_Nil: "is_tl(M,[],Z) \<longleftrightarrow> Z = []"
 by (simp add: is_tl_def)
 
 lemma (in M_trivial) is_tl_Cons:
-     "[|M(a); M(l)|] ==> is_tl(M,Cons(a,l),Z) <-> Z = l"
+     "[|M(a); M(l)|] ==> is_tl(M,Cons(a,l),Z) \<longleftrightarrow> Z = l"
 by (force simp add: is_tl_def)
 
 lemma (in M_trivial) tl_abs [simp]:
-     "[|M(x); M(y)|] ==> is_tl(M,x,y) <-> y = tl'(x)"
+     "[|M(x); M(y)|] ==> is_tl(M,x,y) \<longleftrightarrow> y = tl'(x)"
 apply (simp add: tl'_def)
 apply (intro impI conjI)
  prefer 2 apply (force simp add: is_tl_def)
--- a/src/ZF/Constructible/Satisfies_absolute.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Satisfies_absolute.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -31,7 +31,7 @@
 
 lemma sats_depth_fm [simp]:
    "[| x \<in> nat; y < length(env); env \<in> list(A)|]
-    ==> sats(A, depth_fm(x,y), env) <->
+    ==> sats(A, depth_fm(x,y), env) \<longleftrightarrow>
         is_depth(##A, nth(x,env), nth(y,env))"
 apply (frule_tac x=y in lt_length_in_nat, assumption)  
 apply (simp add: depth_fm_def is_depth_def) 
@@ -40,7 +40,7 @@
 lemma depth_iff_sats:
       "[| nth(i,env) = x; nth(j,env) = y; 
           i \<in> nat; j < length(env); env \<in> list(A)|]
-       ==> is_depth(##A, x, y) <-> sats(A, depth_fm(i,j), env)"
+       ==> is_depth(##A, x, y) \<longleftrightarrow> sats(A, depth_fm(i,j), env)"
 by (simp add: sats_depth_fm)
 
 theorem depth_reflection:
@@ -60,11 +60,11 @@
 (* is_formula_case :: 
     "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o"
   "is_formula_case(M, is_a, is_b, is_c, is_d, v, z) == 
-      (\<forall>x[M]. \<forall>y[M]. x\<in>nat --> y\<in>nat --> is_Member(M,x,y,v) --> is_a(x,y,z)) &
-      (\<forall>x[M]. \<forall>y[M]. x\<in>nat --> y\<in>nat --> is_Equal(M,x,y,v) --> is_b(x,y,z)) &
-      (\<forall>x[M]. \<forall>y[M]. x\<in>formula --> y\<in>formula --> 
-                     is_Nand(M,x,y,v) --> is_c(x,y,z)) &
-      (\<forall>x[M]. x\<in>formula --> is_Forall(M,x,v) --> is_d(x,z))" *)
+      (\<forall>x[M]. \<forall>y[M]. x\<in>nat \<longrightarrow> y\<in>nat \<longrightarrow> is_Member(M,x,y,v) \<longrightarrow> is_a(x,y,z)) &
+      (\<forall>x[M]. \<forall>y[M]. x\<in>nat \<longrightarrow> y\<in>nat \<longrightarrow> is_Equal(M,x,y,v) \<longrightarrow> is_b(x,y,z)) &
+      (\<forall>x[M]. \<forall>y[M]. x\<in>formula \<longrightarrow> y\<in>formula \<longrightarrow> 
+                     is_Nand(M,x,y,v) \<longrightarrow> is_c(x,y,z)) &
+      (\<forall>x[M]. x\<in>formula \<longrightarrow> is_Forall(M,x,v) \<longrightarrow> is_d(x,z))" *)
 
 definition
   formula_case_fm :: "[i, i, i, i, i, i]=>i" where
@@ -96,22 +96,22 @@
   assumes is_a_iff_sats: 
       "!!a0 a1 a2. 
         [|a0\<in>A; a1\<in>A; a2\<in>A|]  
-        ==> ISA(a2, a1, a0) <-> sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))"
+        ==> ISA(a2, a1, a0) \<longleftrightarrow> sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))"
   and is_b_iff_sats: 
       "!!a0 a1 a2. 
         [|a0\<in>A; a1\<in>A; a2\<in>A|]  
-        ==> ISB(a2, a1, a0) <-> sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))"
+        ==> ISB(a2, a1, a0) \<longleftrightarrow> sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))"
   and is_c_iff_sats: 
       "!!a0 a1 a2. 
         [|a0\<in>A; a1\<in>A; a2\<in>A|]  
-        ==> ISC(a2, a1, a0) <-> sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))"
+        ==> ISC(a2, a1, a0) \<longleftrightarrow> sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))"
   and is_d_iff_sats: 
       "!!a0 a1. 
         [|a0\<in>A; a1\<in>A|]  
-        ==> ISD(a1, a0) <-> sats(A, is_d, Cons(a0,Cons(a1,env)))"
+        ==> ISD(a1, a0) \<longleftrightarrow> sats(A, is_d, Cons(a0,Cons(a1,env)))"
   shows 
       "[|x \<in> nat; y < length(env); env \<in> list(A)|]
-       ==> sats(A, formula_case_fm(is_a,is_b,is_c,is_d,x,y), env) <->
+       ==> sats(A, formula_case_fm(is_a,is_b,is_c,is_d,x,y), env) \<longleftrightarrow>
            is_formula_case(##A, ISA, ISB, ISC, ISD, nth(x,env), nth(y,env))"
 apply (frule_tac x=y in lt_length_in_nat, assumption)  
 apply (simp add: formula_case_fm_def is_formula_case_def 
@@ -123,23 +123,23 @@
   assumes is_a_iff_sats: 
       "!!a0 a1 a2. 
         [|a0\<in>A; a1\<in>A; a2\<in>A|]  
-        ==> ISA(a2, a1, a0) <-> sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))"
+        ==> ISA(a2, a1, a0) \<longleftrightarrow> sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))"
   and is_b_iff_sats: 
       "!!a0 a1 a2. 
         [|a0\<in>A; a1\<in>A; a2\<in>A|]  
-        ==> ISB(a2, a1, a0) <-> sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))"
+        ==> ISB(a2, a1, a0) \<longleftrightarrow> sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))"
   and is_c_iff_sats: 
       "!!a0 a1 a2. 
         [|a0\<in>A; a1\<in>A; a2\<in>A|]  
-        ==> ISC(a2, a1, a0) <-> sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))"
+        ==> ISC(a2, a1, a0) \<longleftrightarrow> sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))"
   and is_d_iff_sats: 
       "!!a0 a1. 
         [|a0\<in>A; a1\<in>A|]  
-        ==> ISD(a1, a0) <-> sats(A, is_d, Cons(a0,Cons(a1,env)))"
+        ==> ISD(a1, a0) \<longleftrightarrow> sats(A, is_d, Cons(a0,Cons(a1,env)))"
   shows 
       "[|nth(i,env) = x; nth(j,env) = y; 
       i \<in> nat; j < length(env); env \<in> list(A)|]
-       ==> is_formula_case(##A, ISA, ISB, ISC, ISD, x, y) <->
+       ==> is_formula_case(##A, ISA, ISB, ISC, ISD, x, y) \<longleftrightarrow>
            sats(A, formula_case_fm(is_a,is_b,is_c,is_d,i,j), env)"
 by (simp add: sats_formula_case_fm [OF is_a_iff_sats is_b_iff_sats 
                                        is_c_iff_sats is_d_iff_sats])
@@ -184,7 +184,7 @@
 
 lemma (in M_datatypes) is_depth_apply_abs [simp]:
      "[|M(h); p \<in> formula; M(z)|] 
-      ==> is_depth_apply(M,h,p,z) <-> z = h ` succ(depth(p)) ` p"
+      ==> is_depth_apply(M,h,p,z) \<longleftrightarrow> z = h ` succ(depth(p)) ` p"
 by (simp add: is_depth_apply_def formula_into_M depth_type eq_commute)
 
 
@@ -202,7 +202,7 @@
 definition
   satisfies_is_a :: "[i=>o,i,i,i,i]=>o" where
    "satisfies_is_a(M,A) == 
-    \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) -->
+    \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
                 \<lambda>env z. is_bool_of_o(M, 
                       \<exists>nx[M]. \<exists>ny[M]. 
@@ -219,7 +219,7 @@
    --{*We simplify the formula to have just @{term nx} rather than 
        introducing @{term ny} with  @{term "nx=ny"} *}
   "satisfies_is_b(M,A) == 
-    \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) -->
+    \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
                 \<lambda>env z. is_bool_of_o(M, 
                       \<exists>nx[M]. is_nth(M,x,env,nx) & is_nth(M,y,env,nx), z),
@@ -232,7 +232,7 @@
 definition
   satisfies_is_c :: "[i=>o,i,i,i,i,i]=>o" where
    "satisfies_is_c(M,A,h) == 
-    \<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) -->
+    \<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M]. 
                  (\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & 
                  (\<exists>rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & 
@@ -247,13 +247,13 @@
 definition
   satisfies_is_d :: "[i=>o,i,i,i,i]=>o" where
    "satisfies_is_d(M,A,h) == 
-    \<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) -->
+    \<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
                 \<lambda>env z. \<exists>rp[M]. is_depth_apply(M,h,p,rp) & 
                     is_bool_of_o(M, 
                            \<forall>x[M]. \<forall>xenv[M]. \<forall>hp[M]. 
-                              x\<in>A --> is_Cons(M,x,env,xenv) --> 
-                              fun_apply(M,rp,xenv,hp) --> number1(M,hp),
+                              x\<in>A \<longrightarrow> is_Cons(M,x,env,xenv) \<longrightarrow> 
+                              fun_apply(M,rp,xenv,hp) \<longrightarrow> number1(M,hp),
                   z),
                zz)"
 
@@ -263,7 +263,7 @@
         the correct arity.*}
   "satisfies_MH == 
     \<lambda>M A u f z. 
-         \<forall>fml[M]. is_formula(M,fml) -->
+         \<forall>fml[M]. is_formula(M,fml) \<longrightarrow>
              is_lambda (M, fml, 
                is_formula_case (M, satisfies_is_a(M,A), 
                                 satisfies_is_b(M,A), 
@@ -321,8 +321,8 @@
               env \<in> list(A) & 
               is_bool_of_o (M, 
                             \<forall>a[M]. \<forall>co[M]. \<forall>rpco[M]. 
-                               a\<in>A --> is_Cons(M,a,env,co) -->
-                               fun_apply(M,rp,co,rpco) --> number1(M, rpco), 
+                               a\<in>A \<longrightarrow> is_Cons(M,a,env,co) \<longrightarrow>
+                               fun_apply(M,rp,co,rpco) \<longrightarrow> number1(M, rpco), 
                             bo) &
               pair(M,env,bo,z))"
  and
@@ -500,7 +500,7 @@
 
 lemma (in M_satisfies) satisfies_abs:
   "[|M(A); M(z); p \<in> formula|] 
-   ==> is_satisfies(M,A,p,z) <-> z = satisfies(A,p)"
+   ==> is_satisfies(M,A,p,z) \<longleftrightarrow> z = satisfies(A,p)"
 by (simp only: Formula_Rec.formula_rec_abs [OF Formula_Rec_M]  
                satisfies_eq is_satisfies_def satisfies_MH_def)
 
@@ -529,14 +529,14 @@
 
 lemma sats_depth_apply_fm [simp]:
    "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, depth_apply_fm(x,y,z), env) <->
+    ==> sats(A, depth_apply_fm(x,y,z), env) \<longleftrightarrow>
         is_depth_apply(##A, nth(x,env), nth(y,env), nth(z,env))"
 by (simp add: depth_apply_fm_def is_depth_apply_def)
 
 lemma depth_apply_iff_sats:
     "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
         i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
-     ==> is_depth_apply(##A, x, y, z) <-> sats(A, depth_apply_fm(i,j,k), env)"
+     ==> is_depth_apply(##A, x, y, z) \<longleftrightarrow> sats(A, depth_apply_fm(i,j,k), env)"
 by simp
 
 lemma depth_apply_reflection:
@@ -551,7 +551,7 @@
 subsubsection{*The Operator @{term satisfies_is_a}, Internalized*}
 
 (* satisfies_is_a(M,A) == 
-    \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) -->
+    \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
                 \<lambda>env z. is_bool_of_o(M, 
                       \<exists>nx[M]. \<exists>ny[M]. 
@@ -577,7 +577,7 @@
 
 lemma sats_satisfies_is_a_fm [simp]:
    "[| u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, satisfies_is_a_fm(u,x,y,z), env) <->
+    ==> sats(A, satisfies_is_a_fm(u,x,y,z), env) \<longleftrightarrow>
         satisfies_is_a(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))"
 apply (frule_tac x=x in lt_length_in_nat, assumption)  
 apply (frule_tac x=y in lt_length_in_nat, assumption)  
@@ -588,7 +588,7 @@
 lemma satisfies_is_a_iff_sats:
   "[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
       u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|]
-   ==> satisfies_is_a(##A,nu,nx,ny,nz) <->
+   ==> satisfies_is_a(##A,nu,nx,ny,nz) \<longleftrightarrow>
        sats(A, satisfies_is_a_fm(u,x,y,z), env)"
 by simp
 
@@ -604,7 +604,7 @@
 subsubsection{*The Operator @{term satisfies_is_b}, Internalized*}
 
 (* satisfies_is_b(M,A) == 
-    \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) -->
+    \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
                 \<lambda>env z. is_bool_of_o(M, 
                       \<exists>nx[M]. is_nth(M,x,env,nx) & is_nth(M,y,env,nx), z),
@@ -626,7 +626,7 @@
 
 lemma sats_satisfies_is_b_fm [simp]:
    "[| u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, satisfies_is_b_fm(u,x,y,z), env) <->
+    ==> sats(A, satisfies_is_b_fm(u,x,y,z), env) \<longleftrightarrow>
         satisfies_is_b(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))"
 apply (frule_tac x=x in lt_length_in_nat, assumption)  
 apply (frule_tac x=y in lt_length_in_nat, assumption)  
@@ -637,7 +637,7 @@
 lemma satisfies_is_b_iff_sats:
   "[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
       u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|]
-   ==> satisfies_is_b(##A,nu,nx,ny,nz) <->
+   ==> satisfies_is_b(##A,nu,nx,ny,nz) \<longleftrightarrow>
        sats(A, satisfies_is_b_fm(u,x,y,z), env)"
 by simp
 
@@ -653,7 +653,7 @@
 subsubsection{*The Operator @{term satisfies_is_c}, Internalized*}
 
 (* satisfies_is_c(M,A,h) == 
-    \<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) -->
+    \<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M]. 
                  (\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & 
                  (\<exists>rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & 
@@ -679,7 +679,7 @@
 
 lemma sats_satisfies_is_c_fm [simp]:
    "[| u \<in> nat; v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, satisfies_is_c_fm(u,v,x,y,z), env) <->
+    ==> sats(A, satisfies_is_c_fm(u,v,x,y,z), env) \<longleftrightarrow>
         satisfies_is_c(##A, nth(u,env), nth(v,env), nth(x,env), 
                             nth(y,env), nth(z,env))"  
 by (simp add: satisfies_is_c_fm_def satisfies_is_c_def sats_lambda_fm)
@@ -688,7 +688,7 @@
   "[| nth(u,env) = nu; nth(v,env) = nv; nth(x,env) = nx; nth(y,env) = ny; 
       nth(z,env) = nz;
       u \<in> nat; v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-   ==> satisfies_is_c(##A,nu,nv,nx,ny,nz) <->
+   ==> satisfies_is_c(##A,nu,nv,nx,ny,nz) \<longleftrightarrow>
        sats(A, satisfies_is_c_fm(u,v,x,y,z), env)"
 by simp
 
@@ -704,13 +704,13 @@
 subsubsection{*The Operator @{term satisfies_is_d}, Internalized*}
 
 (* satisfies_is_d(M,A,h) == 
-    \<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) -->
+    \<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
                 \<lambda>env z. \<exists>rp[M]. is_depth_apply(M,h,p,rp) & 
                     is_bool_of_o(M, 
                            \<forall>x[M]. \<forall>xenv[M]. \<forall>hp[M]. 
-                              x\<in>A --> is_Cons(M,x,env,xenv) --> 
-                              fun_apply(M,rp,xenv,hp) --> number1(M,hp),
+                              x\<in>A \<longrightarrow> is_Cons(M,x,env,xenv) \<longrightarrow> 
+                              fun_apply(M,rp,xenv,hp) \<longrightarrow> number1(M,hp),
                   z),
                zz) *)
 
@@ -736,7 +736,7 @@
 
 lemma sats_satisfies_is_d_fm [simp]:
    "[| u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, satisfies_is_d_fm(u,x,y,z), env) <->
+    ==> sats(A, satisfies_is_d_fm(u,x,y,z), env) \<longleftrightarrow>
         satisfies_is_d(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))"  
 by (simp add: satisfies_is_d_fm_def satisfies_is_d_def sats_lambda_fm
               sats_bool_of_o_fm)
@@ -744,7 +744,7 @@
 lemma satisfies_is_d_iff_sats:
   "[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
       u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-   ==> satisfies_is_d(##A,nu,nx,ny,nz) <->
+   ==> satisfies_is_d(##A,nu,nx,ny,nz) \<longleftrightarrow>
        sats(A, satisfies_is_d_fm(u,x,y,z), env)"
 by simp
 
@@ -762,7 +762,7 @@
 
 (* satisfies_MH == 
     \<lambda>M A u f zz. 
-         \<forall>fml[M]. is_formula(M,fml) -->
+         \<forall>fml[M]. is_formula(M,fml) \<longrightarrow>
              is_lambda (M, fml, 
                is_formula_case (M, satisfies_is_a(M,A), 
                                 satisfies_is_b(M,A), 
@@ -789,7 +789,7 @@
 
 lemma sats_satisfies_MH_fm [simp]:
    "[| u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-    ==> sats(A, satisfies_MH_fm(u,x,y,z), env) <->
+    ==> sats(A, satisfies_MH_fm(u,x,y,z), env) \<longleftrightarrow>
         satisfies_MH(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))"  
 by (simp add: satisfies_MH_fm_def satisfies_MH_def sats_lambda_fm
               sats_formula_case_fm)
@@ -797,7 +797,7 @@
 lemma satisfies_MH_iff_sats:
   "[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
       u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
-   ==> satisfies_MH(##A,nu,nx,ny,nz) <->
+   ==> satisfies_MH(##A,nu,nx,ny,nz) \<longleftrightarrow>
        sats(A, satisfies_MH_fm(u,x,y,z), env)"
 by simp 
 
@@ -936,8 +936,8 @@
               env \<in> list(A) & 
               is_bool_of_o (L, 
                             \<forall>a[L]. \<forall>co[L]. \<forall>rpco[L]. 
-                               a\<in>A --> is_Cons(L,a,env,co) -->
-                               fun_apply(L,rp,co,rpco) --> number1(L, rpco), 
+                               a\<in>A \<longrightarrow> is_Cons(L,a,env,co) \<longrightarrow>
+                               fun_apply(L,rp,co,rpco) \<longrightarrow> number1(L, rpco), 
                             bo) &
               pair(L,env,bo,z))"
 apply (rule strong_replacementI)
--- a/src/ZF/Constructible/Separation.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Separation.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -38,7 +38,7 @@
 
 text{*Reduces the original comprehension to the reflected one*}
 lemma reflection_imp_L_separation:
-      "[| \<forall>x\<in>Lset(j). P(x) <-> Q(x);
+      "[| \<forall>x\<in>Lset(j). P(x) \<longleftrightarrow> Q(x);
           {x \<in> Lset(j) . Q(x)} \<in> DPow(Lset(j));
           Ord(j);  z \<in> Lset(j)|] ==> L({x \<in> z . P(x)})"
 apply (rule_tac i = "succ(j)" in L_I)
@@ -85,12 +85,12 @@
 subsection{*Separation for Intersection*}
 
 lemma Inter_Reflects:
-     "REFLECTS[\<lambda>x. \<forall>y[L]. y\<in>A --> x \<in> y,
-               \<lambda>i x. \<forall>y\<in>Lset(i). y\<in>A --> x \<in> y]"
+     "REFLECTS[\<lambda>x. \<forall>y[L]. y\<in>A \<longrightarrow> x \<in> y,
+               \<lambda>i x. \<forall>y\<in>Lset(i). y\<in>A \<longrightarrow> x \<in> y]"
 by (intro FOL_reflections)
 
 lemma Inter_separation:
-     "L(A) ==> separation(L, \<lambda>x. \<forall>y[L]. y\<in>A --> x\<in>y)"
+     "L(A) ==> separation(L, \<lambda>x. \<forall>y[L]. y\<in>A \<longrightarrow> x\<in>y)"
 apply (rule gen_separation [OF Inter_Reflects], simp)
 apply (rule DPow_LsetI)
  txt{*I leave this one example of a manual proof.  The tedium of manually
--- a/src/ZF/Constructible/WF_absolute.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/WF_absolute.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -18,7 +18,7 @@
 lemma alt_rtrancl_lemma1 [rule_format]:
     "n \<in> nat
      ==> \<forall>f \<in> succ(n) -> field(r).
-         (\<forall>i\<in>n. \<langle>f`i, f ` succ(i)\<rangle> \<in> r) --> \<langle>f`0, f`n\<rangle> \<in> r^*"
+         (\<forall>i\<in>n. \<langle>f`i, f ` succ(i)\<rangle> \<in> r) \<longrightarrow> \<langle>f`0, f`n\<rangle> \<in> r^*"
 apply (induct_tac n)
 apply (simp_all add: apply_funtype rtrancl_refl, clarify)
 apply (rename_tac n f)
@@ -29,24 +29,24 @@
 apply (simp add: Ord_succ_mem_iff nat_0_le [THEN ltD] leI [THEN ltD] ltI)
 done
 
-lemma rtrancl_alt_subset_rtrancl: "rtrancl_alt(field(r),r) <= r^*"
+lemma rtrancl_alt_subset_rtrancl: "rtrancl_alt(field(r),r) \<subseteq> r^*"
 apply (simp add: rtrancl_alt_def)
 apply (blast intro: alt_rtrancl_lemma1)
 done
 
-lemma rtrancl_subset_rtrancl_alt: "r^* <= rtrancl_alt(field(r),r)"
+lemma rtrancl_subset_rtrancl_alt: "r^* \<subseteq> rtrancl_alt(field(r),r)"
 apply (simp add: rtrancl_alt_def, clarify)
 apply (frule rtrancl_type [THEN subsetD], clarify, simp)
 apply (erule rtrancl_induct)
  txt{*Base case, trivial*}
  apply (rule_tac x=0 in bexI)
-  apply (rule_tac x="lam x:1. xa" in bexI)
+  apply (rule_tac x="\<lambda>x\<in>1. xa" in bexI)
    apply simp_all
 txt{*Inductive step*}
 apply clarify
 apply (rename_tac n f)
 apply (rule_tac x="succ(n)" in bexI)
- apply (rule_tac x="lam i:succ(succ(n)). if i=succ(n) then z else f`i" in bexI)
+ apply (rule_tac x="\<lambda>i\<in>succ(succ(n)). if i=succ(n) then z else f`i" in bexI)
   apply (simp add: Ord_succ_mem_iff nat_0_le [THEN ltD] leI [THEN ltD] ltI)
   apply (blast intro: mem_asym)
  apply typecheck
@@ -67,7 +67,7 @@
                (\<exists>f[M]. typed_function(M,n',A,f) &
                 (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
                   fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
-                  (\<forall>j[M]. j\<in>n --> 
+                  (\<forall>j[M]. j\<in>n \<longrightarrow> 
                     (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. 
                       fun_apply(M,f,j,fj) & successor(M,j,sj) &
                       fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"
@@ -75,8 +75,8 @@
 definition
   rtran_closure :: "[i=>o,i,i] => o" where
     "rtran_closure(M,r,s) == 
-        \<forall>A[M]. is_field(M,r,A) -->
-         (\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))"
+        \<forall>A[M]. is_field(M,r,A) \<longrightarrow>
+         (\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))"
 
 definition
   tran_closure :: "[i=>o,i,i] => o" where
@@ -85,7 +85,7 @@
 
 lemma (in M_basic) rtran_closure_mem_iff:
      "[|M(A); M(r); M(p)|]
-      ==> rtran_closure_mem(M,A,r,p) <->
+      ==> rtran_closure_mem(M,A,r,p) \<longleftrightarrow>
           (\<exists>n[M]. n\<in>nat & 
            (\<exists>f[M]. f \<in> succ(n) -> A &
             (\<exists>x[M]. \<exists>y[M]. p = <x,y> & f`0 = x & f`n = y) &
@@ -118,7 +118,7 @@
 done
 
 lemma (in M_trancl) rtrancl_abs [simp]:
-     "[| M(r); M(z) |] ==> rtran_closure(M,r,z) <-> z = rtrancl(r)"
+     "[| M(r); M(z) |] ==> rtran_closure(M,r,z) \<longleftrightarrow> z = rtrancl(r)"
 apply (rule iffI)
  txt{*Proving the right-to-left implication*}
  prefer 2 apply (blast intro: rtran_closure_rtrancl)
@@ -133,7 +133,7 @@
 by (simp add: trancl_def comp_closed rtrancl_closed)
 
 lemma (in M_trancl) trancl_abs [simp]:
-     "[| M(r); M(z) |] ==> tran_closure(M,r,z) <-> z = trancl(r)"
+     "[| M(r); M(z) |] ==> tran_closure(M,r,z) \<longleftrightarrow> z = trancl(r)"
 by (simp add: tran_closure_def trancl_def)
 
 lemma (in M_trancl) wellfounded_trancl_separation':
@@ -142,7 +142,7 @@
 
 text{*Alternative proof of @{text wf_on_trancl}; inspiration for the
       relativized version.  Original version is on theory WF.*}
-lemma "[| wf[A](r);  r-``A <= A |] ==> wf[A](r^+)"
+lemma "[| wf[A](r);  r-``A \<subseteq> A |] ==> wf[A](r^+)"
 apply (simp add: wf_on_def wf_def)
 apply (safe intro!: equalityI)
 apply (drule_tac x = "{x\<in>A. \<exists>w. \<langle>w,x\<rangle> \<in> r^+ & w \<in> Z}" in spec)
@@ -150,7 +150,7 @@
 done
 
 lemma (in M_trancl) wellfounded_on_trancl:
-     "[| wellfounded_on(M,A,r);  r-``A <= A; M(r); M(A) |]
+     "[| wellfounded_on(M,A,r);  r-``A \<subseteq> A; M(r); M(A) |]
       ==> wellfounded_on(M,A,r^+)"
 apply (simp add: wellfounded_on_def)
 apply (safe intro!: equalityI)
@@ -186,8 +186,8 @@
           pair(M,x,y,z) & 
           is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
           y = H(x, restrict(g, r -`` {x}))); 
-     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
-   ==> wfrec(r,a,H) = z <-> 
+     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
+   ==> wfrec(r,a,H) = z \<longleftrightarrow> 
        (\<exists>f[M]. is_recfun(r^+, a, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & 
             z = H(a,restrict(f,r-``{a})))"
 apply (frule wf_trancl) 
@@ -206,8 +206,8 @@
 theorem (in M_trancl) trans_wfrec_relativize:
   "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);
      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
-     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
-   ==> wfrec(r,a,H) = z <-> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))" 
+     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
+   ==> wfrec(r,a,H) = z \<longleftrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))" 
 apply (frule wfrec_replacement', assumption+) 
 apply (simp cong: is_recfun_cong
            add: wfrec_relativize trancl_eq_r
@@ -217,16 +217,16 @@
 theorem (in M_trancl) trans_wfrec_abs:
   "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);  M(z);
      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
-     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
-   ==> is_wfrec(M,MH,r,a,z) <-> z=wfrec(r,a,H)" 
+     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
+   ==> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> z=wfrec(r,a,H)" 
 by (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast) 
 
 
 lemma (in M_trancl) trans_eq_pair_wfrec_iff:
   "[|wf(r);  trans(r); relation(r); M(r);  M(y); 
      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
-     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
-   ==> y = <x, wfrec(r, x, H)> <-> 
+     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
+   ==> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
        (\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
 apply safe 
  apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) 
@@ -242,8 +242,8 @@
 lemma (in M_trancl) wfrec_closed_lemma [rule_format]:
      "[|wf(r); M(r); 
         strong_replacement(M, \<lambda>x y. y = \<langle>x, wfrec(r, x, H)\<rangle>);
-        \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] 
-      ==> M(a) --> M(wfrec(r,a,H))"
+        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
+      ==> M(a) \<longrightarrow> M(wfrec(r,a,H))"
 apply (rule_tac a=a in wf_induct, assumption+)
 apply (subst wfrec, assumption, clarify)
 apply (drule_tac x1=x and x="\<lambda>x\<in>r -`` {x}. wfrec(r, x, H)" 
@@ -255,7 +255,7 @@
 text{*Eliminates one instance of replacement.*}
 lemma (in M_trancl) wfrec_replacement_iff:
      "strong_replacement(M, \<lambda>x z. 
-          \<exists>y[M]. pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))) <->
+          \<exists>y[M]. pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))) \<longleftrightarrow>
       strong_replacement(M, 
            \<lambda>x y. \<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
 apply simp 
@@ -266,7 +266,7 @@
 theorem (in M_trancl) trans_wfrec_closed:
      "[|wf(r); trans(r); relation(r); M(r); M(a);
        wfrec_replacement(M,MH,r);  relation2(M,MH,H);
-        \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] 
+        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
       ==> M(wfrec(r,a,H))"
 apply (frule wfrec_replacement', assumption+) 
 apply (frule wfrec_replacement_iff [THEN iffD1]) 
@@ -281,8 +281,8 @@
           pair(M,x,y,z) & 
           is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
           y = H(x, restrict(g, r -`` {x}))); 
-     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
-   ==> y = <x, wfrec(r, x, H)> <-> 
+     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
+   ==> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
        (\<exists>f[M]. is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & 
             y = <x, H(x,restrict(f,r-``{x}))>)"
 apply safe  
@@ -297,7 +297,7 @@
      "[|wf(r); M(r); M(a);
         wfrec_replacement(M,MH,r^+);  
         relation2(M,MH, \<lambda>x f. H(x, restrict(f, r -`` {x})));
-        \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] 
+        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
       ==> M(wfrec(r,a,H))"
 apply (frule wfrec_replacement' 
                [of MH "r^+" "\<lambda>x f. H(x, restrict(f, r -`` {x}))"])
--- a/src/ZF/Constructible/WFrec.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/WFrec.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -20,7 +20,7 @@
 
 text{*Expresses @{text is_recfun} as a recursion equation*}
 lemma is_recfun_iff_equation:
-     "is_recfun(r,a,H,f) <->
+     "is_recfun(r,a,H,f) \<longleftrightarrow>
            f \<in> r -`` {a} \<rightarrow> range(f) &
            (\<forall>x \<in> r-``{a}. f`x = H(x, restrict(f, r-``{x})))"  
 apply (rule iffI) 
@@ -47,7 +47,7 @@
 
 lemma is_recfun_cong_lemma:
   "[| is_recfun(r,a,H,f); r = r'; a = a'; f = f'; 
-      !!x g. [| <x,a'> \<in> r'; relation(g); domain(g) <= r' -``{x} |] 
+      !!x g. [| <x,a'> \<in> r'; relation(g); domain(g) \<subseteq> r' -``{x} |] 
              ==> H(x,g) = H'(x,g) |]
    ==> is_recfun(r',a',H',f')"
 apply (simp add: is_recfun_def) 
@@ -60,9 +60,9 @@
       whose domains are initial segments of @{term r}.*}
 lemma is_recfun_cong:
   "[| r = r'; a = a'; f = f'; 
-      !!x g. [| <x,a'> \<in> r'; relation(g); domain(g) <= r' -``{x} |] 
+      !!x g. [| <x,a'> \<in> r'; relation(g); domain(g) \<subseteq> r' -``{x} |] 
              ==> H(x,g) = H'(x,g) |]
-   ==> is_recfun(r,a,H,f) <-> is_recfun(r',a',H',f')"
+   ==> is_recfun(r,a,H,f) \<longleftrightarrow> is_recfun(r',a',H',f')"
 apply (rule iffI)
 txt{*Messy: fast and blast don't work for some reason*}
 apply (erule is_recfun_cong_lemma, auto) 
@@ -91,7 +91,7 @@
     "[|is_recfun(r,a,H,f);  is_recfun(r,b,H,g);  
        wellfounded(M,r);  trans(r);
        M(f); M(g); M(r); M(x); M(a); M(b) |] 
-     ==> <x,a> \<in> r --> <x,b> \<in> r --> f`x=g`x"
+     ==> <x,a> \<in> r \<longrightarrow> <x,b> \<in> r \<longrightarrow> f`x=g`x"
 apply (frule_tac f=f in is_recfun_type) 
 apply (frule_tac f=g in is_recfun_type) 
 apply (simp add: is_recfun_def)
@@ -104,7 +104,7 @@
 apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def)
 apply (rename_tac x1)
 apply (rule_tac t="%z. H(x1,z)" in subst_context) 
-apply (subgoal_tac "\<forall>y \<in> r-``{x1}. ALL z. <y,z>\<in>f <-> <y,z>\<in>g")
+apply (subgoal_tac "\<forall>y \<in> r-``{x1}. \<forall>z. <y,z>\<in>f \<longleftrightarrow> <y,z>\<in>g")
  apply (blast intro: transD) 
 apply (simp add: apply_iff) 
 apply (blast intro: transD sym) 
@@ -132,9 +132,9 @@
 
 text{*Tells us that @{text is_recfun} can (in principle) be relativized.*}
 lemma (in M_basic) is_recfun_relativize:
-  "[| M(r); M(f); \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] 
-   ==> is_recfun(r,a,H,f) <->
-       (\<forall>z[M]. z \<in> f <-> 
+  "[| M(r); M(f); \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
+   ==> is_recfun(r,a,H,f) \<longleftrightarrow>
+       (\<forall>z[M]. z \<in> f \<longleftrightarrow> 
         (\<exists>x[M]. <x,a> \<in> r & z = <x, H(x, restrict(f, r-``{x}))>))";
 apply (simp add: is_recfun_def lam_def)
 apply (safe intro!: equalityI) 
@@ -155,7 +155,7 @@
 lemma (in M_basic) is_recfun_restrict:
      "[| wellfounded(M,r); trans(r); is_recfun(r,x,H,f); \<langle>y,x\<rangle> \<in> r; 
        M(r); M(f); 
-       \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |]
+       \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |]
        ==> is_recfun(r, y, H, restrict(f, r -`` {y}))"
 apply (frule pair_components_in_M, assumption, clarify) 
 apply (simp (no_asm_simp) add: is_recfun_relativize restrict_iff
@@ -171,14 +171,14 @@
  
 lemma (in M_basic) restrict_Y_lemma:
    "[| wellfounded(M,r); trans(r); M(r);
-       \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g));  M(Y);
+       \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g));  M(Y);
        \<forall>b[M]. 
-           b \<in> Y <->
+           b \<in> Y \<longleftrightarrow>
            (\<exists>x[M]. <x,a1> \<in> r &
             (\<exists>y[M]. b = \<langle>x,y\<rangle> & (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g))));
           \<langle>x,a1\<rangle> \<in> r; is_recfun(r,x,H,f); M(f) |]
        ==> restrict(Y, r -`` {x}) = f"
-apply (subgoal_tac "\<forall>y \<in> r-``{x}. \<forall>z. <y,z>:Y <-> <y,z>:f") 
+apply (subgoal_tac "\<forall>y \<in> r-``{x}. \<forall>z. <y,z>:Y \<longleftrightarrow> <y,z>:f") 
  apply (simp (no_asm_simp) add: restrict_def) 
  apply (thin_tac "rall(M,?P)")+  --{*essential for efficiency*}
  apply (frule is_recfun_type [THEN fun_is_rel], blast)
@@ -210,11 +210,11 @@
 text{*Proof of the inductive step for @{text exists_is_recfun}, since
       we must prove two versions.*}
 lemma (in M_basic) exists_is_recfun_indstep:
-    "[|\<forall>y. \<langle>y, a1\<rangle> \<in> r --> (\<exists>f[M]. is_recfun(r, y, H, f)); 
+    "[|\<forall>y. \<langle>y, a1\<rangle> \<in> r \<longrightarrow> (\<exists>f[M]. is_recfun(r, y, H, f)); 
        wellfounded(M,r); trans(r); M(r); M(a1);
        strong_replacement(M, \<lambda>x z. 
               \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
-       \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|]   
+       \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]   
       ==> \<exists>f[M]. is_recfun(r,a1,H,f)"
 apply (drule_tac A="r-``{a1}" in strong_replacementD)
   apply blast 
@@ -248,7 +248,7 @@
        strong_replacement(M, \<lambda>x z. 
           \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
        M(r);  M(a);  
-       \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |]   
+       \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |]   
       ==> \<exists>f[M]. is_recfun(r,a,H,f)"
 apply (rule wellfounded_induct, assumption+, clarify)
 apply (rule exists_is_recfun_indstep, assumption+)
@@ -258,8 +258,8 @@
     "[|wf(r);  trans(r);  M(r);  
        strong_replacement(M, \<lambda>x z. 
          \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
-       \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |]   
-      ==> M(a) --> (\<exists>f[M]. is_recfun(r,a,H,f))"
+       \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |]   
+      ==> M(a) \<longrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f))"
 apply (rule wf_induct, assumption+)
 apply (frule wf_imp_relativized)
 apply (intro impI)
@@ -273,7 +273,7 @@
 definition
   M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where
    "M_is_recfun(M,MH,r,a,f) == 
-     \<forall>z[M]. z \<in> f <-> 
+     \<forall>z[M]. z \<in> f \<longleftrightarrow> 
             (\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M]. 
                pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) &
                pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) &
@@ -291,9 +291,9 @@
              \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))"
 
 lemma (in M_basic) is_recfun_abs:
-     "[| \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g));  M(r); M(a); M(f); 
+     "[| \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g));  M(r); M(a); M(f); 
          relation2(M,MH,H) |] 
-      ==> M_is_recfun(M,MH,r,a,f) <-> is_recfun(r,a,H,f)"
+      ==> M_is_recfun(M,MH,r,a,f) \<longleftrightarrow> is_recfun(r,a,H,f)"
 apply (simp add: M_is_recfun_def relation2_def is_recfun_relativize)
 apply (rule rall_cong)
 apply (blast dest: transM)
@@ -301,30 +301,30 @@
 
 lemma M_is_recfun_cong [cong]:
      "[| r = r'; a = a'; f = f'; 
-       !!x g y. [| M(x); M(g); M(y) |] ==> MH(x,g,y) <-> MH'(x,g,y) |]
-      ==> M_is_recfun(M,MH,r,a,f) <-> M_is_recfun(M,MH',r',a',f')"
+       !!x g y. [| M(x); M(g); M(y) |] ==> MH(x,g,y) \<longleftrightarrow> MH'(x,g,y) |]
+      ==> M_is_recfun(M,MH,r,a,f) \<longleftrightarrow> M_is_recfun(M,MH',r',a',f')"
 by (simp add: M_is_recfun_def) 
 
 lemma (in M_basic) is_wfrec_abs:
-     "[| \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)); 
+     "[| \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)); 
          relation2(M,MH,H);  M(r); M(a); M(z) |]
-      ==> is_wfrec(M,MH,r,a,z) <-> 
+      ==> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> 
           (\<exists>g[M]. is_recfun(r,a,H,g) & z = H(a,g))"
 by (simp add: is_wfrec_def relation2_def is_recfun_abs)
 
 text{*Relating @{term wfrec_replacement} to native constructs*}
 lemma (in M_basic) wfrec_replacement':
   "[|wfrec_replacement(M,MH,r);
-     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)); 
+     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)); 
      relation2(M,MH,H);  M(r)|] 
    ==> strong_replacement(M, \<lambda>x z. \<exists>y[M]. 
                 pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g)))"
 by (simp add: wfrec_replacement_def is_wfrec_abs) 
 
 lemma wfrec_replacement_cong [cong]:
-     "[| !!x y z. [| M(x); M(y); M(z) |] ==> MH(x,y,z) <-> MH'(x,y,z);
+     "[| !!x y z. [| M(x); M(y); M(z) |] ==> MH(x,y,z) \<longleftrightarrow> MH'(x,y,z);
          r=r' |] 
-      ==> wfrec_replacement(M, %x y. MH(x,y), r) <-> 
+      ==> wfrec_replacement(M, %x y. MH(x,y), r) \<longleftrightarrow> 
           wfrec_replacement(M, %x y. MH'(x,y), r')" 
 by (simp add: is_wfrec_def wfrec_replacement_def) 
 
--- a/src/ZF/Constructible/Wellorderings.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Constructible/Wellorderings.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -17,29 +17,29 @@
 
 definition
   irreflexive :: "[i=>o,i,i]=>o" where
-    "irreflexive(M,A,r) == \<forall>x[M]. x\<in>A --> <x,x> \<notin> r"
+    "irreflexive(M,A,r) == \<forall>x[M]. x\<in>A \<longrightarrow> <x,x> \<notin> r"
   
 definition
   transitive_rel :: "[i=>o,i,i]=>o" where
     "transitive_rel(M,A,r) == 
-        \<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A --> (\<forall>z[M]. z\<in>A --> 
-                          <x,y>\<in>r --> <y,z>\<in>r --> <x,z>\<in>r))"
+        \<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> (\<forall>z[M]. z\<in>A \<longrightarrow> 
+                          <x,y>\<in>r \<longrightarrow> <y,z>\<in>r \<longrightarrow> <x,z>\<in>r))"
 
 definition
   linear_rel :: "[i=>o,i,i]=>o" where
     "linear_rel(M,A,r) == 
-        \<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A --> <x,y>\<in>r | x=y | <y,x>\<in>r)"
+        \<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> <x,y>\<in>r | x=y | <y,x>\<in>r)"
 
 definition
   wellfounded :: "[i=>o,i]=>o" where
     --{*EVERY non-empty set has an @{text r}-minimal element*}
     "wellfounded(M,r) == 
-        \<forall>x[M]. x\<noteq>0 --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
+        \<forall>x[M]. x\<noteq>0 \<longrightarrow> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
 definition
   wellfounded_on :: "[i=>o,i,i]=>o" where
     --{*every non-empty SUBSET OF @{text A} has an @{text r}-minimal element*}
     "wellfounded_on(M,A,r) == 
-        \<forall>x[M]. x\<noteq>0 --> x\<subseteq>A --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
+        \<forall>x[M]. x\<noteq>0 \<longrightarrow> x\<subseteq>A \<longrightarrow> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
 
 definition
   wellordered :: "[i=>o,i,i]=>o" where
@@ -51,15 +51,15 @@
 subsubsection {*Trivial absoluteness proofs*}
 
 lemma (in M_basic) irreflexive_abs [simp]: 
-     "M(A) ==> irreflexive(M,A,r) <-> irrefl(A,r)"
+     "M(A) ==> irreflexive(M,A,r) \<longleftrightarrow> irrefl(A,r)"
 by (simp add: irreflexive_def irrefl_def)
 
 lemma (in M_basic) transitive_rel_abs [simp]: 
-     "M(A) ==> transitive_rel(M,A,r) <-> trans[A](r)"
+     "M(A) ==> transitive_rel(M,A,r) \<longleftrightarrow> trans[A](r)"
 by (simp add: transitive_rel_def trans_on_def)
 
 lemma (in M_basic) linear_rel_abs [simp]: 
-     "M(A) ==> linear_rel(M,A,r) <-> linear(A,r)"
+     "M(A) ==> linear_rel(M,A,r) \<longleftrightarrow> linear(A,r)"
 by (simp add: linear_rel_def linear_def)
 
 lemma (in M_basic) wellordered_is_trans_on: 
@@ -86,7 +86,7 @@
 subsubsection {*Well-founded relations*}
 
 lemma  (in M_basic) wellfounded_on_iff_wellfounded:
-     "wellfounded_on(M,A,r) <-> wellfounded(M, r \<inter> A*A)"
+     "wellfounded_on(M,A,r) \<longleftrightarrow> wellfounded(M, r \<inter> A*A)"
 apply (simp add: wellfounded_on_def wellfounded_def, safe)
  apply force
 apply (drule_tac x=x in rspec, assumption, blast) 
@@ -101,14 +101,14 @@
 by (simp add: wellfounded_def wellfounded_on_iff_wellfounded, fast)
 
 lemma (in M_basic) wellfounded_iff_wellfounded_on_field:
-     "M(r) ==> wellfounded(M,r) <-> wellfounded_on(M, field(r), r)"
+     "M(r) ==> wellfounded(M,r) \<longleftrightarrow> wellfounded_on(M, field(r), r)"
 by (blast intro: wellfounded_imp_wellfounded_on
                  wellfounded_on_field_imp_wellfounded)
 
 (*Consider the least z in domain(r) such that P(z) does not hold...*)
 lemma (in M_basic) wellfounded_induct: 
      "[| wellfounded(M,r); M(a); M(r); separation(M, \<lambda>x. ~P(x));  
-         \<forall>x. M(x) & (\<forall>y. <y,x> \<in> r --> P(y)) --> P(x) |]
+         \<forall>x. M(x) & (\<forall>y. <y,x> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x) |]
       ==> P(a)";
 apply (simp (no_asm_use) add: wellfounded_def)
 apply (drule_tac x="{z \<in> domain(r). ~P(z)}" in rspec)
@@ -117,11 +117,11 @@
 
 lemma (in M_basic) wellfounded_on_induct: 
      "[| a\<in>A;  wellfounded_on(M,A,r);  M(A);  
-       separation(M, \<lambda>x. x\<in>A --> ~P(x));  
-       \<forall>x\<in>A. M(x) & (\<forall>y\<in>A. <y,x> \<in> r --> P(y)) --> P(x) |]
+       separation(M, \<lambda>x. x\<in>A \<longrightarrow> ~P(x));  
+       \<forall>x\<in>A. M(x) & (\<forall>y\<in>A. <y,x> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x) |]
       ==> P(a)";
 apply (simp (no_asm_use) add: wellfounded_on_def)
-apply (drule_tac x="{z\<in>A. z\<in>A --> ~P(z)}" in rspec)
+apply (drule_tac x="{z\<in>A. z\<in>A \<longrightarrow> ~P(z)}" in rspec)
 apply (blast intro: transM)+
 done
 
@@ -158,11 +158,11 @@
 
 lemma (in M_basic) order_isomorphism_abs [simp]: 
      "[| M(A); M(B); M(f) |] 
-      ==> order_isomorphism(M,A,r,B,s,f) <-> f \<in> ord_iso(A,r,B,s)"
+      ==> order_isomorphism(M,A,r,B,s,f) \<longleftrightarrow> f \<in> ord_iso(A,r,B,s)"
 by (simp add: apply_closed order_isomorphism_def ord_iso_def)
 
 lemma (in M_basic) pred_set_abs [simp]: 
-     "[| M(r); M(B) |] ==> pred_set(M,A,x,r,B) <-> B = Order.pred(A,x,r)"
+     "[| M(r); M(B) |] ==> pred_set(M,A,x,r,B) \<longleftrightarrow> B = Order.pred(A,x,r)"
 apply (simp add: pred_set_def Order.pred_def)
 apply (blast dest: transM) 
 done
@@ -174,7 +174,7 @@
 done
 
 lemma (in M_basic) membership_abs [simp]: 
-     "[| M(r); M(A) |] ==> membership(M,A,r) <-> r = Memrel(A)"
+     "[| M(r); M(A) |] ==> membership(M,A,r) \<longleftrightarrow> r = Memrel(A)"
 apply (simp add: membership_def Memrel_def, safe)
   apply (rule equalityI) 
    apply clarify 
--- a/src/ZF/Resid/Confluence.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Resid/Confluence.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -8,12 +8,12 @@
 definition
   confluence    :: "i=>o"  where
     "confluence(R) ==   
-       \<forall>x y. <x,y> \<in> R --> (\<forall>z.<x,z> \<in> R --> (\<exists>u.<y,u> \<in> R & <z,u> \<in> R))"
+       \<forall>x y. <x,y> \<in> R \<longrightarrow> (\<forall>z.<x,z> \<in> R \<longrightarrow> (\<exists>u.<y,u> \<in> R & <z,u> \<in> R))"
 
 definition
   strip         :: "o"  where
-    "strip == \<forall>x y. (x ===> y) --> 
-                    (\<forall>z.(x =1=> z) --> (\<exists>u.(y =1=> u) & (z===>u)))" 
+    "strip == \<forall>x y. (x ===> y) \<longrightarrow> 
+                    (\<forall>z.(x =1=> z) \<longrightarrow> (\<exists>u.(y =1=> u) & (z===>u)))" 
 
 
 (* ------------------------------------------------------------------------- *)
@@ -72,11 +72,11 @@
   "a<-1->b == <a,b> \<in> Sconv1"
 
 abbreviation
-  Sconv_rel (infixl "<--->" 50) where
-  "a<--->b == <a,b> \<in> Sconv"
+  Sconv_rel (infixl "<-\<longrightarrow>" 50) where
+  "a<-\<longrightarrow>b == <a,b> \<in> Sconv"
   
 inductive
-  domains       "Sconv1" <= "lambda*lambda"
+  domains       "Sconv1" \<subseteq> "lambda*lambda"
   intros
     red1:        "m -1-> n ==> m<-1->n"
     expl:        "n -1-> m ==> m<-1->n"
@@ -85,16 +85,16 @@
 declare Sconv1.intros [intro]
 
 inductive
-  domains       "Sconv" <= "lambda*lambda"
+  domains       "Sconv" \<subseteq> "lambda*lambda"
   intros
-    one_step:    "m<-1->n  ==> m<--->n"
-    refl:        "m \<in> lambda ==> m<--->m"
-    trans:       "[|m<--->n; n<--->p|] ==> m<--->p"
+    one_step:    "m<-1->n  ==> m<-\<longrightarrow>n"
+    refl:        "m \<in> lambda ==> m<-\<longrightarrow>m"
+    trans:       "[|m<-\<longrightarrow>n; n<-\<longrightarrow>p|] ==> m<-\<longrightarrow>p"
   type_intros    Sconv1.dom_subset [THEN subsetD] lambda.intros bool_typechecks
 
 declare Sconv.intros [intro]
 
-lemma conv_sym: "m<--->n ==> n<--->m"
+lemma conv_sym: "m<-\<longrightarrow>n ==> n<-\<longrightarrow>m"
 apply (erule Sconv.induct)
 apply (erule Sconv1.induct, blast+)
 done
@@ -103,7 +103,7 @@
 (*      Church_Rosser Theorem                                                *)
 (* ------------------------------------------------------------------------- *)
 
-lemma Church_Rosser: "m<--->n ==> \<exists>p.(m --->p) & (n ---> p)"
+lemma Church_Rosser: "m<-\<longrightarrow>n ==> \<exists>p.(m -\<longrightarrow>p) & (n -\<longrightarrow> p)"
 apply (erule Sconv.induct)
 apply (erule Sconv1.induct)
 apply (blast intro: red1D1 redD2)
--- a/src/ZF/Resid/Redex.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Resid/Redex.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -53,7 +53,7 @@
 
 
 inductive
-  domains       "Ssub" <= "redexes*redexes"
+  domains       "Ssub" \<subseteq> "redexes*redexes"
   intros
     Sub_Var:     "n \<in> nat ==> Var(n)<== Var(n)"
     Sub_Fun:     "[|u<== v|]==> Fun(u)<== Fun(v)"
@@ -63,7 +63,7 @@
   type_intros    redexes.intros bool_typechecks
 
 inductive
-  domains       "Scomp" <= "redexes*redexes"
+  domains       "Scomp" \<subseteq> "redexes*redexes"
   intros
     Comp_Var:    "n \<in> nat ==> Var(n) ~ Var(n)"
     Comp_Fun:    "[|u ~ v|]==> Fun(u) ~ Fun(v)"
@@ -72,7 +72,7 @@
   type_intros    redexes.intros bool_typechecks
 
 inductive
-  domains       "Sreg" <= redexes
+  domains       "Sreg" \<subseteq> redexes
   intros
     Reg_Var:     "n \<in> nat ==> regular(Var(n))"
     Reg_Fun:     "[|regular(u)|]==> regular(Fun(u))"
@@ -146,10 +146,10 @@
 lemma comp_sym: "u ~ v ==> v ~ u"
 by (erule Scomp.induct, blast+)
 
-lemma comp_sym_iff: "u ~ v <-> v ~ u"
+lemma comp_sym_iff: "u ~ v \<longleftrightarrow> v ~ u"
 by (blast intro: comp_sym)
 
-lemma comp_trans [rule_format]: "u ~ v ==> \<forall>w. v ~ w-->u ~ w"
+lemma comp_trans [rule_format]: "u ~ v ==> \<forall>w. v ~ w\<longrightarrow>u ~ w"
 by (erule Scomp.induct, blast+)
 
 (* ------------------------------------------------------------------------- *)
@@ -174,7 +174,7 @@
 (* ------------------------------------------------------------------------- *)
 
 lemma union_preserve_regular [rule_format]:
-     "u ~ v ==> regular(u)-->regular(v)-->regular(u un v)"
+     "u ~ v ==> regular(u)\<longrightarrow>regular(v)\<longrightarrow>regular(u un v)"
   by (erule Scomp.induct, auto)
 
 end
--- a/src/ZF/Resid/Reduction.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Resid/Reduction.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -16,7 +16,7 @@
   "Apl(n,m) == App(0,n,m)"
   
 inductive
-  domains       "lambda" <= redexes
+  domains       "lambda" \<subseteq> redexes
   intros
     Lambda_Var:  "               n \<in> nat ==>     Var(n) \<in> lambda"
     Lambda_Fun:  "            u \<in> lambda ==>     Fun(u) \<in> lambda"
@@ -78,8 +78,8 @@
   "a -1-> b == <a,b> \<in> Sred1"
 
 abbreviation
-  Sred_rel (infixl "--->" 50) where
-  "a ---> b == <a,b> \<in> Sred"
+  Sred_rel (infixl "-\<longrightarrow>" 50) where
+  "a -\<longrightarrow> b == <a,b> \<in> Sred"
 
 abbreviation
   Spar_red1_rel (infixl "=1=>" 50) where
@@ -91,7 +91,7 @@
   
   
 inductive
-  domains       "Sred1" <= "lambda*lambda"
+  domains       "Sred1" \<subseteq> "lambda*lambda"
   intros
     beta:       "[|m \<in> lambda; n \<in> lambda|] ==> Apl(Fun(m),n) -1-> n/m"
     rfun:       "[|m -1-> n|] ==> Fun(m) -1-> Fun(n)"
@@ -102,18 +102,18 @@
 declare Sred1.intros [intro, simp]
 
 inductive
-  domains       "Sred" <= "lambda*lambda"
+  domains       "Sred" \<subseteq> "lambda*lambda"
   intros
-    one_step:   "m-1->n ==> m--->n"
-    refl:       "m \<in> lambda==>m --->m"
-    trans:      "[|m--->n; n--->p|] ==>m--->p"
+    one_step:   "m-1->n ==> m-\<longrightarrow>n"
+    refl:       "m \<in> lambda==>m -\<longrightarrow>m"
+    trans:      "[|m-\<longrightarrow>n; n-\<longrightarrow>p|] ==>m-\<longrightarrow>p"
   type_intros    Sred1.dom_subset [THEN subsetD] red_typechecks
 
 declare Sred.one_step [intro, simp]
 declare Sred.refl     [intro, simp]
 
 inductive
-  domains       "Spar_red1" <= "lambda*lambda"
+  domains       "Spar_red1" \<subseteq> "lambda*lambda"
   intros
     beta:       "[|m =1=> m'; n =1=> n'|] ==> Apl(Fun(m),n) =1=> n'/m'"
     rvar:       "n \<in> nat ==> Var(n) =1=> Var(n)"
@@ -124,7 +124,7 @@
 declare Spar_red1.intros [intro, simp]
 
 inductive
-  domains "Spar_red" <= "lambda*lambda"
+  domains "Spar_red" \<subseteq> "lambda*lambda"
   intros
     one_step:   "m =1=> n ==> m ===> n"
     trans:      "[|m===>n; n===>p|] ==> m===>p"
@@ -158,28 +158,28 @@
 (*     Lemmas for reduction                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-lemma red_Fun: "m--->n ==> Fun(m) ---> Fun(n)"
+lemma red_Fun: "m-\<longrightarrow>n ==> Fun(m) -\<longrightarrow> Fun(n)"
 apply (erule Sred.induct)
 apply (rule_tac [3] Sred.trans, simp_all)
 done
 
-lemma red_Apll: "[|n \<in> lambda; m ---> m'|] ==> Apl(m,n)--->Apl(m',n)"
+lemma red_Apll: "[|n \<in> lambda; m -\<longrightarrow> m'|] ==> Apl(m,n)-\<longrightarrow>Apl(m',n)"
 apply (erule Sred.induct)
 apply (rule_tac [3] Sred.trans, simp_all)
 done
 
-lemma red_Aplr: "[|n \<in> lambda; m ---> m'|] ==> Apl(n,m)--->Apl(n,m')"
+lemma red_Aplr: "[|n \<in> lambda; m -\<longrightarrow> m'|] ==> Apl(n,m)-\<longrightarrow>Apl(n,m')"
 apply (erule Sred.induct)
 apply (rule_tac [3] Sred.trans, simp_all)
 done
 
-lemma red_Apl: "[|m ---> m'; n--->n'|] ==> Apl(m,n)--->Apl(m',n')"
+lemma red_Apl: "[|m -\<longrightarrow> m'; n-\<longrightarrow>n'|] ==> Apl(m,n)-\<longrightarrow>Apl(m',n')"
 apply (rule_tac n = "Apl (m',n) " in Sred.trans)
 apply (simp_all add: red_Apll red_Aplr)
 done
 
-lemma red_beta: "[|m \<in> lambda; m':lambda; n \<in> lambda; n':lambda; m ---> m'; n--->n'|] ==>  
-               Apl(Fun(m),n)---> n'/m'"
+lemma red_beta: "[|m \<in> lambda; m':lambda; n \<in> lambda; n':lambda; m -\<longrightarrow> m'; n-\<longrightarrow>n'|] ==>  
+               Apl(Fun(m),n)-\<longrightarrow> n'/m'"
 apply (rule_tac n = "Apl (Fun (m'),n') " in Sred.trans)
 apply (simp_all add: red_Apl red_Fun)
 done
@@ -196,13 +196,13 @@
 lemma red1_par_red1: "m-1->n ==> m=1=>n"
 by (erule Sred1.induct, simp_all add: refl_par_red1)
 
-lemma red_par_red: "m--->n ==> m===>n"
+lemma red_par_red: "m-\<longrightarrow>n ==> m===>n"
 apply (erule Sred.induct)
 apply (rule_tac [3] Spar_red.trans)
 apply (simp_all add: refl_par_red1 red1_par_red1)
 done
 
-lemma par_red_red: "m===>n ==> m--->n"
+lemma par_red_red: "m===>n ==> m-\<longrightarrow>n"
 apply (erule Spar_red.induct)
 apply (erule Spar_red1.induct)
 apply (rule_tac [5] Sred.trans)
@@ -237,7 +237,7 @@
 (* ------------------------------------------------------------------------- *)
 
 lemma completeness_l [rule_format]:
-     "u~v ==> regular(v) --> unmark(u) =1=> unmark(u|>v)"
+     "u~v ==> regular(v) \<longrightarrow> unmark(u) =1=> unmark(u|>v)"
 apply (erule Scomp.induct)
 apply (auto simp add: unmmark_subst_rec)
 done
--- a/src/ZF/Resid/Residuals.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Resid/Residuals.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -12,7 +12,7 @@
   "residuals(u,v,w) == <u,v,w> \<in> Sres"
 
 inductive
-  domains       "Sres" <= "redexes*redexes*redexes"
+  domains       "Sres" \<subseteq> "redexes*redexes*redexes"
   intros
     Res_Var:    "n \<in> nat ==> residuals(Var(n),Var(n),Var(n))"
     Res_Fun:    "[|residuals(u,v,w)|]==>   
@@ -64,11 +64,11 @@
 subsection{*residuals is a  partial function*}
 
 lemma residuals_function [rule_format]:
-     "residuals(u,v,w) ==> \<forall>w1. residuals(u,v,w1) --> w1 = w"
+     "residuals(u,v,w) ==> \<forall>w1. residuals(u,v,w1) \<longrightarrow> w1 = w"
 by (erule Sres.induct, force+)
 
 lemma residuals_intro [rule_format]:
-     "u~v ==> regular(v) --> (\<exists>w. residuals(u,v,w))"
+     "u~v ==> regular(v) \<longrightarrow> (\<exists>w. residuals(u,v,w))"
 by (erule Scomp.induct, force+)
 
 lemma comp_resfuncD:
@@ -105,7 +105,7 @@
 done
 
 lemma resfunc_type [simp]:
-     "[|s~t; regular(t)|]==> regular(t) --> s |> t \<in> redexes"
+     "[|s~t; regular(t)|]==> regular(t) \<longrightarrow> s |> t \<in> redexes"
   by (erule Scomp.induct, auto)
 
 subsection{*Commutation theorem*}
@@ -114,17 +114,17 @@
 by (erule Ssub.induct, simp_all)
 
 lemma sub_preserve_reg [rule_format, simp]:
-     "u<==v  ==> regular(v) --> regular(u)"
+     "u<==v  ==> regular(v) \<longrightarrow> regular(u)"
 by (erule Ssub.induct, auto)
 
-lemma residuals_lift_rec: "[|u~v; k \<in> nat|]==> regular(v)--> (\<forall>n \<in> nat.   
+lemma residuals_lift_rec: "[|u~v; k \<in> nat|]==> regular(v)\<longrightarrow> (\<forall>n \<in> nat.   
          lift_rec(u,n) |> lift_rec(v,n) = lift_rec(u |> v,n))"
 apply (erule Scomp.induct, safe)
 apply (simp_all add: lift_rec_Var subst_Var lift_subst)
 done
 
 lemma residuals_subst_rec:
-     "u1~u2 ==>  \<forall>v1 v2. v1~v2 --> regular(v2) --> regular(u2) --> 
+     "u1~u2 ==>  \<forall>v1 v2. v1~v2 \<longrightarrow> regular(v2) \<longrightarrow> regular(u2) \<longrightarrow> 
                   (\<forall>n \<in> nat. subst_rec(v1,u1,n) |> subst_rec(v2,u2,n) =  
                     subst_rec(v1 |> v2, u1 |> u2,n))"
 apply (erule Scomp.induct, safe)
@@ -143,11 +143,11 @@
 subsection{*Residuals are comp and regular*}
 
 lemma residuals_preserve_comp [rule_format, simp]:
-     "u~v ==> \<forall>w. u~w --> v~w --> regular(w) --> (u|>w) ~ (v|>w)"
+     "u~v ==> \<forall>w. u~w \<longrightarrow> v~w \<longrightarrow> regular(w) \<longrightarrow> (u|>w) ~ (v|>w)"
 by (erule Scomp.induct, force+)
 
 lemma residuals_preserve_reg [rule_format, simp]:
-     "u~v ==> regular(u) --> regular(v) --> regular(u|>v)"
+     "u~v ==> regular(u) \<longrightarrow> regular(v) \<longrightarrow> regular(u|>v)"
 apply (erule Scomp.induct, auto)
 done
 
@@ -157,7 +157,7 @@
 by (erule Scomp.induct, simp_all)
 
 lemma preservation [rule_format]:
-     "u ~ v ==> regular(v) --> u|>v = (u un v)|>v"
+     "u ~ v ==> regular(v) \<longrightarrow> u|>v = (u un v)|>v"
 apply (erule Scomp.induct, safe)
 apply (drule_tac [3] psi = "Fun (?u) |> ?v = ?w" in asm_rl)
 apply (auto simp add: union_preserve_comp comp_sym_iff)
@@ -171,7 +171,7 @@
 (* Having more assumptions than needed -- removed below  *)
 lemma prism_l [rule_format]:
      "v<==u ==>  
-       regular(u) --> (\<forall>w. w~v --> w~u -->   
+       regular(u) \<longrightarrow> (\<forall>w. w~v \<longrightarrow> w~u \<longrightarrow>   
                             w |> u = (w|>v) |> (u|>v))"
 by (erule Ssub.induct, force+)
 
--- a/src/ZF/Resid/Substitution.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/Resid/Substitution.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -54,7 +54,7 @@
 lemma gt_not_eq: "p < n ==> n\<noteq>p"
 by blast
 
-lemma succ_pred [rule_format, simp]: "j \<in> nat ==> i < j --> succ(j #- 1) = j"
+lemma succ_pred [rule_format, simp]: "j \<in> nat ==> i < j \<longrightarrow> succ(j #- 1) = j"
 by (induct_tac "j", auto)
 
 lemma lt_pred: "[|succ(x)<n; n \<in> nat|] ==> x < n #- 1 "
@@ -148,7 +148,7 @@
 (*The i\<in>nat is redundant*)
 lemma lift_lift_rec [rule_format]:
      "u \<in> redexes
-      ==> \<forall>n \<in> nat. \<forall>i \<in> nat. i\<le>n -->
+      ==> \<forall>n \<in> nat. \<forall>i \<in> nat. i\<le>n \<longrightarrow>
            (lift_rec(lift_rec(u,i),succ(n)) = lift_rec(lift_rec(u,n),i))"
 apply (erule redexes.induct, auto)
 apply (case_tac "n < i")
@@ -166,7 +166,7 @@
 
 lemma lift_rec_subst_rec [rule_format]:
      "v \<in> redexes ==>
-       \<forall>n \<in> nat. \<forall>m \<in> nat. \<forall>u \<in> redexes. n\<le>m-->
+       \<forall>n \<in> nat. \<forall>m \<in> nat. \<forall>u \<in> redexes. n\<le>m\<longrightarrow>
           lift_rec(subst_rec(u,v,n),m) =
                subst_rec(lift_rec(u,m),lift_rec(v,succ(m)),n)"
 apply (erule redexes.induct, simp_all (no_asm_simp) add: lift_lift)
@@ -188,7 +188,7 @@
 
 lemma lift_rec_subst_rec_lt [rule_format]:
      "v \<in> redexes ==>
-       \<forall>n \<in> nat. \<forall>m \<in> nat. \<forall>u \<in> redexes. m\<le>n-->
+       \<forall>n \<in> nat. \<forall>m \<in> nat. \<forall>u \<in> redexes. m\<le>n\<longrightarrow>
           lift_rec(subst_rec(u,v,n),m) =
                subst_rec(lift_rec(u,m),lift_rec(v,m),succ(n))"
 apply (erule redexes.induct, simp_all (no_asm_simp) add: lift_lift)
@@ -212,7 +212,7 @@
 
 lemma subst_rec_subst_rec [rule_format]:
      "v \<in> redexes ==>
-        \<forall>m \<in> nat. \<forall>n \<in> nat. \<forall>u \<in> redexes. \<forall>w \<in> redexes. m\<le>n -->
+        \<forall>m \<in> nat. \<forall>n \<in> nat. \<forall>u \<in> redexes. \<forall>w \<in> redexes. m\<le>n \<longrightarrow>
           subst_rec(subst_rec(w,u,n),subst_rec(lift_rec(w,m),v,succ(n)),m) =
           subst_rec(w,subst_rec(u,v,m),n)"
 apply (erule redexes.induct)
@@ -254,7 +254,7 @@
 
 lemma subst_rec_preserve_comp [rule_format, simp]:
      "u2 ~ v2 ==> \<forall>m \<in> nat. \<forall>u1 \<in> redexes. \<forall>v1 \<in> redexes.
-                  u1 ~ v1--> subst_rec(u1,u2,m) ~ subst_rec(v1,v2,m)"
+                  u1 ~ v1\<longrightarrow> subst_rec(u1,u2,m) ~ subst_rec(v1,v2,m)"
 by (erule Scomp.induct,
     simp_all add: subst_Var lift_rec_preserve_comp comp_refl)
 
@@ -264,7 +264,7 @@
 
 lemma subst_rec_preserve_reg [simp]:
      "regular(v) ==>
-        \<forall>m \<in> nat. \<forall>u \<in> redexes. regular(u)-->regular(subst_rec(u,v,m))"
+        \<forall>m \<in> nat. \<forall>u \<in> redexes. regular(u)\<longrightarrow>regular(subst_rec(u,v,m))"
 by (erule Sreg.induct, simp_all add: subst_Var lift_rec_preserve_reg)
 
 end
--- a/src/ZF/UNITY/AllocBase.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/AllocBase.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -86,18 +86,18 @@
 by (cut_tac Nclients_pos NbT_pos, auto)
 
 lemma INT_Nclient_iff [iff]:
-     "b\<in>Inter(RepFun(Nclients, B)) <-> (\<forall>x\<in>Nclients. b\<in>B(x))"
+     "b\<in>\<Inter>(RepFun(Nclients, B)) \<longleftrightarrow> (\<forall>x\<in>Nclients. b\<in>B(x))"
 by (force simp add: INT_iff)
 
 lemma setsum_fun_mono [rule_format]:
      "n\<in>nat ==>  
-      (\<forall>i\<in>nat. i<n --> f(i) $<= g(i)) -->  
+      (\<forall>i\<in>nat. i<n \<longrightarrow> f(i) $<= g(i)) \<longrightarrow>  
       setsum(f, n) $<= setsum(g,n)"
 apply (induct_tac "n", simp_all)
 apply (subgoal_tac "Finite(x) & x\<notin>x")
  prefer 2 apply (simp add: nat_into_Finite mem_not_refl, clarify)
 apply (simp (no_asm_simp) add: succ_def)
-apply (subgoal_tac "\<forall>i\<in>nat. i<x--> f(i) $<= g(i) ")
+apply (subgoal_tac "\<forall>i\<in>nat. i<x\<longrightarrow> f(i) $<= g(i) ")
  prefer 2 apply (force dest: leI) 
 apply (rule zadd_zle_mono, simp_all)
 done
@@ -107,7 +107,7 @@
 
 lemma tokens_mono_aux [rule_format]:
      "xs\<in>list(A) ==> \<forall>ys\<in>list(A). <xs, ys>\<in>prefix(A)  
-   --> tokens(xs) \<le> tokens(ys)"
+   \<longrightarrow> tokens(xs) \<le> tokens(ys)"
 apply (induct_tac "xs")
 apply (auto dest: gen_prefix.dom_subset [THEN subsetD] simp add: prefix_def)
 done
@@ -148,7 +148,7 @@
 
 lemma bag_of_mono_aux [rule_format]:
      "xs\<in>list(A) ==> \<forall>ys\<in>list(A). <xs, ys>\<in>prefix(A)  
-      --> <bag_of(xs), bag_of(ys)>\<in>MultLe(A, r)"
+      \<longrightarrow> <bag_of(xs), bag_of(ys)>\<in>MultLe(A, r)"
 apply (induct_tac "xs", simp_all, clarify) 
 apply (frule_tac l = ys in bag_of_multiset)
 apply (auto intro: empty_le_MultLe simp add: prefix_def)
@@ -172,7 +172,7 @@
 
 lemmas nat_into_Fin = eqpoll_refl [THEN [2] Fin_lemma]
 
-lemma list_Int_length_Fin: "l \<in> list(A) ==> C Int length(l) \<in> Fin(length(l))"
+lemma list_Int_length_Fin: "l \<in> list(A) ==> C \<inter> length(l) \<in> Fin(length(l))"
 apply (drule length_type)
 apply (rule Fin_subset)
 apply (rule Int_lower2)
@@ -186,7 +186,7 @@
 by (simp add: ltI)
 
 lemma Int_succ_right:
-     "A Int succ(k) = (if k : A then cons(k, A Int k) else A Int k)"
+     "A \<inter> succ(k) = (if k \<in> A then cons(k, A \<inter> k) else A \<inter> k)"
 by auto
 
 
@@ -210,9 +210,9 @@
 
 lemma bag_of_sublist_lemma2:
      "l\<in>list(A) ==>  
-  C <= nat ==>  
+  C \<subseteq> nat ==>  
   bag_of(sublist(l, C)) =  
-      msetsum(%i. {#nth(i, l)#}, C Int length(l), A)"
+      msetsum(%i. {#nth(i, l)#}, C \<inter> length(l), A)"
 apply (erule list_append_induct)
 apply (simp (no_asm))
 apply (simp (no_asm_simp) add: sublist_append nth_append bag_of_sublist_lemma munion_commute bag_of_sublist_lemma msetsum_multiset munion_0)
@@ -227,17 +227,17 @@
 (*eliminating the assumption C<=nat*)
 lemma bag_of_sublist:
      "l\<in>list(A) ==>  
-  bag_of(sublist(l, C)) = msetsum(%i. {#nth(i, l)#}, C Int length(l), A)"
-apply (subgoal_tac " bag_of (sublist (l, C Int nat)) = msetsum (%i. {#nth (i, l) #}, C Int length (l), A) ")
+  bag_of(sublist(l, C)) = msetsum(%i. {#nth(i, l)#}, C \<inter> length(l), A)"
+apply (subgoal_tac " bag_of (sublist (l, C \<inter> nat)) = msetsum (%i. {#nth (i, l) #}, C \<inter> length (l), A) ")
 apply (simp add: sublist_Int_eq)
 apply (simp add: bag_of_sublist_lemma2 Int_lower2 Int_assoc nat_Int_length_eq)
 done
 
 lemma bag_of_sublist_Un_Int: 
 "l\<in>list(A) ==>  
-  bag_of(sublist(l, B Un C)) +# bag_of(sublist(l, B Int C)) =  
+  bag_of(sublist(l, B \<union> C)) +# bag_of(sublist(l, B \<inter> C)) =  
       bag_of(sublist(l, B)) +# bag_of(sublist(l, C))"
-apply (subgoal_tac "B Int C Int length (l) = (B Int length (l)) Int (C Int length (l))")
+apply (subgoal_tac "B \<inter> C \<inter> length (l) = (B \<inter> length (l)) \<inter> (C \<inter> length (l))")
 prefer 2 apply blast
 apply (simp (no_asm_simp) add: bag_of_sublist Int_Un_distrib2 msetsum_Un_Int)
 apply (rule msetsum_Un_Int)
@@ -247,14 +247,14 @@
 
 
 lemma bag_of_sublist_Un_disjoint:
-     "[| l\<in>list(A); B Int C = 0  |] 
-      ==> bag_of(sublist(l, B Un C)) =  
+     "[| l\<in>list(A); B \<inter> C = 0  |] 
+      ==> bag_of(sublist(l, B \<union> C)) =  
           bag_of(sublist(l, B)) +# bag_of(sublist(l, C))"
 by (simp add: bag_of_sublist_Un_Int [symmetric] bag_of_multiset)
 
 
 lemma bag_of_sublist_UN_disjoint [rule_format]:
-     "[|Finite(I); \<forall>i\<in>I. \<forall>j\<in>I. i\<noteq>j --> A(i) \<inter> A(j) = 0;  
+     "[|Finite(I); \<forall>i\<in>I. \<forall>j\<in>I. i\<noteq>j \<longrightarrow> A(i) \<inter> A(j) = 0;  
         l\<in>list(B) |]  
       ==> bag_of(sublist(l, \<Union>i\<in>I. A(i))) =   
           (msetsum(%i. bag_of(sublist(l, A(i))), I, B)) "
@@ -278,8 +278,8 @@
 by (unfold all_distinct_def, auto)
 
 lemma all_distinct_Cons [simp]: 
-    "all_distinct(Cons(a,l)) <->  
-      (a\<in>set_of_list(l) --> False) & (a \<notin> set_of_list(l) --> all_distinct(l))"
+    "all_distinct(Cons(a,l)) \<longleftrightarrow>  
+      (a\<in>set_of_list(l) \<longrightarrow> False) & (a \<notin> set_of_list(l) \<longrightarrow> all_distinct(l))"
 apply (unfold all_distinct_def)
 apply (auto elim: list.cases)
 done
@@ -359,7 +359,7 @@
 done
 
 lemma Inter_Diff_var_iff:
-     "Finite(A) ==> b\<in>(Inter(RepFun(var-A, B))) <-> (\<forall>x\<in>var-A. b\<in>B(x))"
+     "Finite(A) ==> b\<in>(\<Inter>(RepFun(var-A, B))) \<longleftrightarrow> (\<forall>x\<in>var-A. b\<in>B(x))"
 apply (subgoal_tac "\<exists>x. x\<in>var-A", auto)
 apply (subgoal_tac "~Finite (var-A) ")
 apply (drule not_Finite_imp_exist, auto)
@@ -369,16 +369,16 @@
 done
 
 lemma Inter_var_DiffD:
-     "[| b\<in>Inter(RepFun(var-A, B)); Finite(A); x\<in>var-A |] ==> b\<in>B(x)"
+     "[| b\<in>\<Inter>(RepFun(var-A, B)); Finite(A); x\<in>var-A |] ==> b\<in>B(x)"
 by (simp add: Inter_Diff_var_iff)
 
-(* [| Finite(A); (\<forall>x\<in>var-A. b\<in>B(x)) |] ==> b\<in>Inter(RepFun(var-A, B)) *)
+(* [| Finite(A); (\<forall>x\<in>var-A. b\<in>B(x)) |] ==> b\<in>\<Inter>(RepFun(var-A, B)) *)
 lemmas Inter_var_DiffI = Inter_Diff_var_iff [THEN iffD2]
 
 declare Inter_var_DiffI [intro!]
 
 lemma Acts_subset_Int_Pow_simp [simp]:
-     "Acts(F)<= A \<inter> Pow(state*state)  <-> Acts(F)<=A"
+     "Acts(F)<= A \<inter> Pow(state*state)  \<longleftrightarrow> Acts(F)<=A"
 by (insert Acts_type [of F], auto)
 
 lemma setsum_nsetsum_eq: 
--- a/src/ZF/UNITY/AllocImpl.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/AllocImpl.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -29,7 +29,7 @@
         \<exists>k. k = length(s`giv) &
             t = s(giv := s`giv @ [nth(k, s`ask)],
                   available_tok := s`available_tok #- nth(k, s`ask)) &
-            k < length(s`ask) & nth(k, s`ask) le s`available_tok}"
+            k < length(s`ask) & nth(k, s`ask) \<le> s`available_tok}"
 
 definition
   "alloc_rel_act ==
@@ -73,7 +73,7 @@
 
 
 lemma alloc_prog_ok_iff:
-"\<forall>G \<in> program. (alloc_prog ok G) <->
+"\<forall>G \<in> program. (alloc_prog ok G) \<longleftrightarrow>
      (G \<in> preserves(lift(giv)) & G \<in> preserves(lift(available_tok)) &
        G \<in> preserves(lift(NbR)) &  alloc_prog \<in> Allowed(G))"
 by (auto simp add: ok_iff_Allowed alloc_prog_def [THEN def_prg_Allowed])
@@ -167,7 +167,7 @@
 apply (drule_tac f = "lift (rel) " in preserves_imp_eq)
 apply assumption+
 apply (force dest: ActsD)
-apply (erule_tac V = "\<forall>x \<in> Acts (alloc_prog) Un Acts (G). ?P(x)" in thin_rl)
+apply (erule_tac V = "\<forall>x \<in> Acts (alloc_prog) \<union> Acts (G). ?P(x)" in thin_rl)
 apply (erule_tac V = "alloc_prog \<in> stable (?u)" in thin_rl)
 apply (drule_tac a = "xc`rel" and f = "lift (rel)" in Increasing_imp_Stable)
 apply (auto simp add: Stable_def Constrains_def constrains_def)
@@ -327,7 +327,7 @@
   alloc_prog \<squnion> G \<in>
   {s\<in>state. nth(length(s`giv), s`ask) \<le> s`available_tok} \<inter>
   {s\<in>state.  k < length(s`ask)} \<inter> {s\<in>state. length(s`giv)=k}
-  Ensures {s\<in>state. ~ k <length(s`ask)} Un {s\<in>state. length(s`giv) \<noteq> k}"
+  Ensures {s\<in>state. ~ k <length(s`ask)} \<union> {s\<in>state. length(s`giv) \<noteq> k}"
 apply (rule EnsuresI, auto)
 apply (erule_tac [2] V = "G\<notin>?u" in thin_rl)
 apply (rule_tac [2] act = alloc_giv_act in transientI)
@@ -339,7 +339,7 @@
 apply (rule_tac [2] ReplaceI)
 apply (rule_tac [2] x = "x (giv := x ` giv @ [nth (length(x`giv), x ` ask) ], available_tok := x ` available_tok #- nth (length(x`giv), x ` ask))" in exI)
 apply (auto intro!: state_update_type simp add: app_type)
-apply (rule_tac A = "{s\<in>state . nth (length(s ` giv), s ` ask) \<le> s ` available_tok} \<inter> {s\<in>state . k < length(s ` ask) } \<inter> {s\<in>state. length(s`giv) =k}" and A' = "{s\<in>state . nth (length(s ` giv), s ` ask) \<le> s ` available_tok} Un {s\<in>state. ~ k < length(s`ask) } Un {s\<in>state . length(s ` giv) \<noteq> k}" in Constrains_weaken)
+apply (rule_tac A = "{s\<in>state . nth (length(s ` giv), s ` ask) \<le> s ` available_tok} \<inter> {s\<in>state . k < length(s ` ask) } \<inter> {s\<in>state. length(s`giv) =k}" and A' = "{s\<in>state . nth (length(s ` giv), s ` ask) \<le> s ` available_tok} \<union> {s\<in>state. ~ k < length(s`ask) } \<union> {s\<in>state . length(s ` giv) \<noteq> k}" in Constrains_weaken)
 apply (auto dest: ActsD simp add: Constrains_def constrains_def alloc_prog_def [THEN def_prg_Acts] alloc_prog_ok_iff)
 apply (subgoal_tac "length(xa ` giv @ [nth (length(xa ` giv), xa ` ask) ]) = length(xa ` giv) #+ 1")
 apply (rule_tac [2] trans)
@@ -377,7 +377,7 @@
         {s\<in>state.  k < length(s`ask)} \<inter>
         {s\<in>state. length(s`giv) = k}
         LeadsTo {s\<in>state. k < length(s`giv)}"
-apply (subgoal_tac "alloc_prog \<squnion> G \<in> {s\<in>state. nth (length(s`giv), s`ask) \<le> s`available_tok} \<inter> {s\<in>state. k < length(s`ask) } \<inter> {s\<in>state. length(s`giv) = k} LeadsTo {s\<in>state. ~ k <length(s`ask) } Un {s\<in>state. length(s`giv) \<noteq> k}")
+apply (subgoal_tac "alloc_prog \<squnion> G \<in> {s\<in>state. nth (length(s`giv), s`ask) \<le> s`available_tok} \<inter> {s\<in>state. k < length(s`ask) } \<inter> {s\<in>state. length(s`giv) = k} LeadsTo {s\<in>state. ~ k <length(s`ask) } \<union> {s\<in>state. length(s`giv) \<noteq> k}")
 prefer 2 apply (blast intro: alloc_prog_giv_Ensures_lemma [THEN LeadsTo_Basis])
 apply (subgoal_tac "alloc_prog \<squnion> G \<in> Stable ({s\<in>state. k < length(s`ask) }) ")
 apply (drule PSP_Stable, assumption)
@@ -401,7 +401,7 @@
     alloc_prog \<squnion> G \<in> Incr(lift(ask));
     alloc_prog \<squnion> G \<in> Incr(lift(rel)) |]
   ==> alloc_prog \<squnion> G \<in>
-        Always({s\<in>state. tokens(s`giv) \<le> tokens(take(s`NbR, s`rel)) -->
+        Always({s\<in>state. tokens(s`giv) \<le> tokens(take(s`NbR, s`rel)) \<longrightarrow>
                 NbT \<le> s`available_tok})"
 apply (subgoal_tac
        "alloc_prog \<squnion> G
@@ -429,8 +429,8 @@
 
 lemma PSP_StableI:
 "[| F \<in> Stable(C); F \<in> A - C LeadsTo B;
-   F \<in> A \<inter> C LeadsTo B Un (state - C) |] ==> F \<in> A LeadsTo  B"
-apply (rule_tac A = " (A-C) Un (A \<inter> C)" in LeadsTo_weaken_L)
+   F \<in> A \<inter> C LeadsTo B \<union> (state - C) |] ==> F \<in> A LeadsTo  B"
+apply (rule_tac A = " (A-C) \<union> (A \<inter> C)" in LeadsTo_weaken_L)
  prefer 2 apply blast
 apply (rule LeadsTo_Un, assumption)
 apply (blast intro: LeadsTo_weaken dest: PSP_Stable)
--- a/src/ZF/UNITY/ClientImpl.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/ClientImpl.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -91,7 +91,7 @@
 declare  client_ask_act_def [THEN def_act_simp, simp]
 
 lemma client_prog_ok_iff:
-  "\<forall>G \<in> program. (client_prog ok G) <->  
+  "\<forall>G \<in> program. (client_prog ok G) \<longleftrightarrow>  
    (G \<in> preserves(lift(rel)) & G \<in> preserves(lift(ask)) &  
     G \<in> preserves(lift(tok)) &  client_prog \<in> Allowed(G))"
 by (auto simp add: ok_iff_Allowed client_prog_def [THEN def_prg_Allowed])
@@ -117,7 +117,7 @@
 
 (*Safety property 1: ask, rel are increasing: (24) *)
 lemma client_prog_Increasing_ask_rel: 
-"client_prog: program guarantees Incr(lift(ask)) Int Incr(lift(rel))"
+"client_prog: program guarantees Incr(lift(ask)) \<inter> Incr(lift(rel))"
 apply (unfold guar_def)
 apply (auto intro!: increasing_imp_Increasing 
             simp add: client_prog_ok_iff Increasing.increasing_def preserves_imp_prefix)
@@ -137,7 +137,7 @@
 lemma ask_Bounded_lemma: 
 "[| client_prog ok G; G \<in> program |] 
       ==> client_prog \<squnion> G \<in>    
-              Always({s \<in> state. s`tok \<le> NbT}  Int   
+              Always({s \<in> state. s`tok \<le> NbT}  \<inter>   
                       {s \<in> state. \<forall>elt \<in> set_of_list(s`ask). elt \<le> NbT})"
 apply (rotate_tac -1)
 apply (auto simp add: client_prog_ok_iff)
@@ -209,7 +209,7 @@
 done
 
 lemma strict_prefix_is_prefix: 
-    "<xs, ys> \<in> strict_prefix(A) <->  <xs, ys> \<in> prefix(A) & xs\<noteq>ys"
+    "<xs, ys> \<in> strict_prefix(A) \<longleftrightarrow>  <xs, ys> \<in> prefix(A) & xs\<noteq>ys"
 apply (unfold strict_prefix_def id_def lam_def)
 apply (auto dest: prefix_type [THEN subsetD])
 done
@@ -298,7 +298,7 @@
 
 lemma client_prog_Allowed:
      "Allowed(client_prog) =  
-      preserves(lift(rel)) Int preserves(lift(ask)) Int preserves(lift(tok))"
+      preserves(lift(rel)) \<inter> preserves(lift(ask)) \<inter> preserves(lift(tok))"
 apply (cut_tac v = "lift (ask)" in preserves_type)
 apply (auto simp add: Allowed_def client_prog_def [THEN def_prg_Allowed] 
                       cons_Int_distrib safety_prop_Acts_iff)
--- a/src/ZF/UNITY/Comp.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/Comp.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -19,25 +19,25 @@
 
 definition
   component :: "[i,i]=>o"  (infixl "component" 65)  where
-  "F component H == (EX G. F Join G = H)"
+  "F component H == (\<exists>G. F Join G = H)"
 
 definition
   strict_component :: "[i,i]=>o" (infixl "strict'_component" 65)  where
-  "F strict_component H == F component H & F~=H"
+  "F strict_component H == F component H & F\<noteq>H"
 
 definition
   (* A stronger form of the component relation *)
   component_of :: "[i,i]=>o"   (infixl "component'_of" 65)  where
-  "F component_of H  == (EX G. F ok G & F Join G = H)"
+  "F component_of H  == (\<exists>G. F ok G & F Join G = H)"
   
 definition
   strict_component_of :: "[i,i]=>o" (infixl "strict'_component'_of" 65)  where
-  "F strict_component_of H  == F component_of H  & F~=H"
+  "F strict_component_of H  == F component_of H  & F\<noteq>H"
 
 definition
   (* Preserves a state function f, in particular a variable *)
   preserves :: "(i=>i)=>i"  where
-  "preserves(f) == {F:program. ALL z. F: stable({s:state. f(s) = z})}"
+  "preserves(f) == {F:program. \<forall>z. F: stable({s:state. f(s) = z})}"
 
 definition
   fun_pair :: "[i=>i, i =>i] =>(i=>i)"  where
@@ -46,7 +46,7 @@
 definition
   localize  :: "[i=>i, i] => i"  where
   "localize(f,F) == mk_program(Init(F), Acts(F),
-                       AllowedActs(F) Int (UN G:preserves(f). Acts(G)))"
+                       AllowedActs(F) \<inter> (\<Union>G\<in>preserves(f). Acts(G)))"
 
   
 (*** component and strict_component relations ***)
@@ -60,8 +60,8 @@
 done
 
 lemma component_eq_subset: 
-     "G \<in> program ==> (F component G) <->  
-   (Init(G) <= Init(F) & Acts(F) <= Acts(G) & AllowedActs(G) <= AllowedActs(F))"
+     "G \<in> program ==> (F component G) \<longleftrightarrow>  
+   (Init(G) \<subseteq> Init(F) & Acts(F) \<subseteq> Acts(G) & AllowedActs(G) \<subseteq> AllowedActs(F))"
 apply (unfold component_def, auto)
 apply (rule exI)
 apply (rule program_equalityI, auto)
@@ -100,7 +100,7 @@
 by (auto simp add: Join_ac component_def)
 
 lemma JN_component_iff:
-     "H \<in> program==>(JOIN(I,F) component H) <-> (\<forall>i \<in> I. F(i) component H)"
+     "H \<in> program==>(JOIN(I,F) component H) \<longleftrightarrow> (\<forall>i \<in> I. F(i) component H)"
 apply (case_tac "I=0", force)
 apply (simp (no_asm_simp) add: component_eq_subset)
 apply auto
@@ -121,7 +121,7 @@
 done
 
 lemma component_antisym:
-     "[| F \<in> program; G \<in> program |] ==>(F component G & G  component F) --> F = G"
+     "[| F \<in> program; G \<in> program |] ==>(F component G & G  component F) \<longrightarrow> F = G"
 apply (simp (no_asm_simp) add: component_eq_subset)
 apply clarify
 apply (rule program_equalityI, auto)
@@ -129,7 +129,7 @@
 
 lemma Join_component_iff:
      "H \<in> program ==> 
-      ((F Join G) component H) <-> (F component H & G component H)"
+      ((F Join G) component H) \<longleftrightarrow> (F component H & G component H)"
 apply (simp (no_asm_simp) add: component_eq_subset)
 apply blast
 done
@@ -163,12 +163,12 @@
 done
 
 lemma Join_preserves [iff]: 
-"(F Join G \<in> preserves(v)) <->   
+"(F Join G \<in> preserves(v)) \<longleftrightarrow>   
       (programify(F) \<in> preserves(v) & programify(G) \<in> preserves(v))"
 by (auto simp add: preserves_def INT_iff)
  
 lemma JN_preserves [iff]:
-     "(JOIN(I,F): preserves(v)) <-> (\<forall>i \<in> I. programify(F(i)):preserves(v))"
+     "(JOIN(I,F): preserves(v)) \<longleftrightarrow> (\<forall>i \<in> I. programify(F(i)):preserves(v))"
 by (auto simp add: JN_stable preserves_def INT_iff)
 
 lemma SKIP_preserves [iff]: "SKIP \<in> preserves(v)"
@@ -180,13 +180,13 @@
 done
 
 lemma preserves_fun_pair:
-     "preserves(fun_pair(v,w)) = preserves(v) Int preserves(w)"
+     "preserves(fun_pair(v,w)) = preserves(v) \<inter> preserves(w)"
 apply (rule equalityI)
 apply (auto simp add: preserves_def stable_def constrains_def, blast+)
 done
 
 lemma preserves_fun_pair_iff [iff]:
-     "F \<in> preserves(fun_pair(v, w))  <-> F \<in> preserves(v) Int preserves(w)"
+     "F \<in> preserves(fun_pair(v, w))  \<longleftrightarrow> F \<in> preserves(v) \<inter> preserves(w)"
 by (simp add: preserves_fun_pair)
 
 lemma fun_pair_comp_distrib:
@@ -202,7 +202,7 @@
 lemma preserves_into_program [TC]: "F \<in> preserves(f) ==> F \<in> program"
 by (blast intro: preserves_type [THEN subsetD])
 
-lemma subset_preserves_comp: "preserves(f) <= preserves(g comp f)"
+lemma subset_preserves_comp: "preserves(f) \<subseteq> preserves(g comp f)"
 apply (auto simp add: preserves_def stable_def constrains_def)
 apply (drule_tac x = "f (xb)" in spec)
 apply (drule_tac x = act in bspec, auto)
@@ -211,7 +211,7 @@
 lemma imp_preserves_comp: "F \<in> preserves(f) ==> F \<in> preserves(g comp f)"
 by (blast intro: subset_preserves_comp [THEN subsetD])
 
-lemma preserves_subset_stable: "preserves(f) <= stable({s \<in> state. P(f(s))})"
+lemma preserves_subset_stable: "preserves(f) \<subseteq> stable({s \<in> state. P(f(s))})"
 apply (auto simp add: preserves_def stable_def constrains_def)
 apply (rename_tac s' s)
 apply (subgoal_tac "f (s) = f (s') ")
@@ -230,7 +230,7 @@
 done
 
 lemma preserves_id_subset_stable: 
- "st_set(A) ==> preserves(%x. x) <= stable(A)"
+ "st_set(A) ==> preserves(%x. x) \<subseteq> stable(A)"
 apply (unfold preserves_def stable_def constrains_def, auto)
 apply (drule_tac x = xb in spec)
 apply (drule_tac x = act in bspec)
@@ -275,7 +275,7 @@
 by (unfold localize_def, simp)
 
 lemma localize_AllowedActs_eq [simp]: 
- "AllowedActs(localize(v,F)) = AllowedActs(F) Int (\<Union>G \<in> preserves(v). Acts(G))"
+ "AllowedActs(localize(v,F)) = AllowedActs(F) \<inter> (\<Union>G \<in> preserves(v). Acts(G))"
 apply (unfold localize_def)
 apply (rule equalityI)
 apply (auto dest: Acts_type [THEN subsetD])
@@ -325,10 +325,10 @@
      G \<in> preserves(f); \<forall>s \<in> state. f(s):A|]
   ==> F Join G \<in> Stable({s \<in> state. P(f(s), g(s))})"
 apply (unfold stable_def Stable_def preserves_def)
-apply (rule_tac A = "(\<Union>k \<in> A. {s \<in> state. f(s)=k} Int {s \<in> state. P (f (s), g (s))})" in Constrains_weaken_L)
+apply (rule_tac A = "(\<Union>k \<in> A. {s \<in> state. f(s)=k} \<inter> {s \<in> state. P (f (s), g (s))})" in Constrains_weaken_L)
 prefer 2 apply blast
 apply (rule Constrains_UN_left, auto)
-apply (rule_tac A = "{s \<in> state. f(s)=k} Int {s \<in> state. P (f (s), g (s))} Int {s \<in> state. P (k, g (s))}" and A' = "({s \<in> state. f(s)=k} Un {s \<in> state. P (f (s), g (s))}) Int {s \<in> state. P (k, g (s))}" in Constrains_weaken)
+apply (rule_tac A = "{s \<in> state. f(s)=k} \<inter> {s \<in> state. P (f (s), g (s))} \<inter> {s \<in> state. P (k, g (s))}" and A' = "({s \<in> state. f(s)=k} \<union> {s \<in> state. P (f (s), g (s))}) \<inter> {s \<in> state. P (k, g (s))}" in Constrains_weaken)
  prefer 2 apply blast 
  prefer 2 apply blast 
 apply (rule Constrains_Int)
--- a/src/ZF/UNITY/Constrains.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/Constrains.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -15,13 +15,13 @@
 inductive 
   domains 
      "traces(init, acts)" <=
-         "(init Un (UN act:acts. field(act)))*list(UN act:acts. field(act))"
+         "(init \<union> (\<Union>act\<in>acts. field(act)))*list(\<Union>act\<in>acts. field(act))"
   intros 
          (*Initial trace is empty*)
-    Init: "s: init ==> <s,[]> : traces(init,acts)"
+    Init: "s: init ==> <s,[]> \<in> traces(init,acts)"
 
-    Acts: "[| act:acts;  <s,evs> : traces(init,acts);  <s,s'>: act |]
-           ==> <s', Cons(s,evs)> : traces(init, acts)"
+    Acts: "[| act:acts;  <s,evs> \<in> traces(init,acts);  <s,s'>: act |]
+           ==> <s', Cons(s,evs)> \<in> traces(init, acts)"
   
   type_intros list.intros UnI1 UnI2 UN_I fieldI2 fieldI1
 
@@ -29,7 +29,7 @@
 consts reachable :: "i=>i"
 inductive
   domains
-  "reachable(F)" <= "Init(F) Un (UN act:Acts(F). field(act))"
+  "reachable(F)" \<subseteq> "Init(F) \<union> (\<Union>act\<in>Acts(F). field(act))"
   intros 
     Init: "s:Init(F) ==> s:reachable(F)"
 
@@ -41,11 +41,11 @@
   
 definition
   Constrains :: "[i,i] => i"  (infixl "Co" 60)  where
-  "A Co B == {F:program. F:(reachable(F) Int A) co B}"
+  "A Co B == {F:program. F:(reachable(F) \<inter> A) co B}"
 
 definition
   op_Unless  :: "[i, i] => i"  (infixl "Unless" 60)  where
-  "A Unless B == (A-B) Co (A Un B)"
+  "A Unless B == (A-B) Co (A \<union> B)"
 
 definition
   Stable     :: "i => i"  where
@@ -54,12 +54,12 @@
 definition
   (*Always is the weak form of "invariant"*)
   Always :: "i => i"  where
-  "Always(A) == initially(A) Int Stable(A)"
+  "Always(A) == initially(A) \<inter> Stable(A)"
 
 
 (*** traces and reachable ***)
 
-lemma reachable_type: "reachable(F) <= state"
+lemma reachable_type: "reachable(F) \<subseteq> state"
 apply (cut_tac F = F in Init_type)
 apply (cut_tac F = F in Acts_type)
 apply (cut_tac F = F in reachable.dom_subset, blast)
@@ -71,11 +71,11 @@
 done
 declare st_set_reachable [iff]
 
-lemma reachable_Int_state: "reachable(F) Int state = reachable(F)"
+lemma reachable_Int_state: "reachable(F) \<inter> state = reachable(F)"
 by (cut_tac reachable_type, auto)
 declare reachable_Int_state [iff]
 
-lemma state_Int_reachable: "state Int reachable(F) = reachable(F)"
+lemma state_Int_reachable: "state \<inter> reachable(F) = reachable(F)"
 by (cut_tac reachable_type, auto)
 declare state_Int_reachable [iff]
 
@@ -88,11 +88,11 @@
 apply (blast intro: reachable.intros traces.intros)+
 done
 
-lemma Init_into_reachable: "Init(F) <= reachable(F)"
+lemma Init_into_reachable: "Init(F) \<subseteq> reachable(F)"
 by (blast intro: reachable.intros)
 
 lemma stable_reachable: "[| F \<in> program; G \<in> program;  
-    Acts(G) <= Acts(F)  |] ==> G \<in> stable(reachable(F))"
+    Acts(G) \<subseteq> Acts(F)  |] ==> G \<in> stable(reachable(F))"
 apply (blast intro: stableI constrainsI st_setI
              reachable_type [THEN subsetD] reachable.intros)
 done
@@ -108,7 +108,7 @@
 done
 
 (*...in fact the strongest invariant!*)
-lemma invariant_includes_reachable: "F \<in> invariant(A) ==> reachable(F) <= A"
+lemma invariant_includes_reachable: "F \<in> invariant(A) ==> reachable(F) \<subseteq> A"
 apply (cut_tac F = F in Acts_type)
 apply (cut_tac F = F in Init_type)
 apply (cut_tac F = F in reachable_type)
@@ -120,7 +120,7 @@
 
 (*** Co ***)
 
-lemma constrains_reachable_Int: "F \<in> B co B'==>F:(reachable(F) Int B) co (reachable(F) Int B')"
+lemma constrains_reachable_Int: "F \<in> B co B'==>F:(reachable(F) \<inter> B) co (reachable(F) \<inter> B')"
 apply (frule constrains_type [THEN subsetD])
 apply (frule stable_reachable [OF _ _ subset_refl])
 apply (simp_all add: stable_def constrains_Int)
@@ -128,7 +128,7 @@
 
 (*Resembles the previous definition of Constrains*)
 lemma Constrains_eq_constrains: 
-"A Co B = {F \<in> program. F:(reachable(F) Int A) co (reachable(F)  Int  B)}"
+"A Co B = {F \<in> program. F:(reachable(F) \<inter> A) co (reachable(F)  \<inter>  B)}"
 apply (unfold Constrains_def)
 apply (blast dest: constrains_reachable_Int constrains_type [THEN subsetD]
              intro: constrains_weaken)
@@ -150,16 +150,16 @@
 done
 
 lemma Constrains_type: 
- "A Co B <= program"
+ "A Co B \<subseteq> program"
 apply (unfold Constrains_def, blast)
 done
 
-lemma Constrains_empty: "F \<in> 0 Co B <-> F \<in> program"
+lemma Constrains_empty: "F \<in> 0 Co B \<longleftrightarrow> F \<in> program"
 by (auto dest: Constrains_type [THEN subsetD]
             intro: constrains_imp_Constrains)
 declare Constrains_empty [iff]
 
-lemma Constrains_state: "F \<in> A Co state <-> F \<in> program"
+lemma Constrains_state: "F \<in> A Co state \<longleftrightarrow> F \<in> program"
 apply (unfold Constrains_def)
 apply (auto dest: Constrains_type [THEN subsetD] intro: constrains_imp_Constrains)
 done
@@ -185,7 +185,7 @@
 
 (** Union **)
 lemma Constrains_Un: 
-    "[| F \<in> A Co A'; F \<in> B Co B' |] ==> F \<in> (A Un B) Co (A' Un B')"
+    "[| F \<in> A Co A'; F \<in> B Co B' |] ==> F \<in> (A \<union> B) Co (A' \<union> B')"
 apply (unfold Constrains_def2, auto)
 apply (simp add: Int_Un_distrib)
 apply (blast intro: constrains_Un)
@@ -201,9 +201,9 @@
 (** Intersection **)
 
 lemma Constrains_Int: 
-    "[| F \<in> A Co A'; F \<in> B Co B'|]==> F:(A Int B) Co (A' Int B')"
+    "[| F \<in> A Co A'; F \<in> B Co B'|]==> F:(A \<inter> B) Co (A' \<inter> B')"
 apply (unfold Constrains_def)
-apply (subgoal_tac "reachable (F) Int (A Int B) = (reachable (F) Int A) Int (reachable (F) Int B) ")
+apply (subgoal_tac "reachable (F) \<inter> (A \<inter> B) = (reachable (F) \<inter> A) \<inter> (reachable (F) \<inter> B) ")
 apply (auto intro: constrains_Int)
 done
 
@@ -215,7 +215,7 @@
 apply (auto simp add: Constrains_def)
 done
 
-lemma Constrains_imp_subset: "F \<in> A Co A' ==> reachable(F) Int A <= A'"
+lemma Constrains_imp_subset: "F \<in> A Co A' ==> reachable(F) \<inter> A \<subseteq> A'"
 apply (unfold Constrains_def)
 apply (blast dest: constrains_imp_subset)
 done
@@ -227,7 +227,7 @@
 done
 
 lemma Constrains_cancel: 
-"[| F \<in> A Co (A' Un B); F \<in> B Co B' |] ==> F \<in> A Co (A' Un B')"
+"[| F \<in> A Co (A' \<union> B); F \<in> B Co B' |] ==> F \<in> A Co (A' \<union> B')"
 apply (unfold Constrains_def2)
 apply (simp (no_asm_use) add: Int_Un_distrib)
 apply (blast intro: constrains_cancel)
@@ -247,7 +247,7 @@
 by blast
 
 lemma Stable_eq_stable: 
-"F \<in> Stable(A) <->  (F \<in> stable(reachable(F) Int A))"
+"F \<in> Stable(A) \<longleftrightarrow>  (F \<in> stable(reachable(F) \<inter> A))"
 apply (auto dest: constrainsD2 simp add: Stable_def stable_def Constrains_def2)
 done
 
@@ -258,27 +258,27 @@
 by (unfold Stable_def, assumption)
 
 lemma Stable_Un: 
-    "[| F \<in> Stable(A); F \<in> Stable(A') |] ==> F \<in> Stable(A Un A')"
+    "[| F \<in> Stable(A); F \<in> Stable(A') |] ==> F \<in> Stable(A \<union> A')"
 apply (unfold Stable_def)
 apply (blast intro: Constrains_Un)
 done
 
 lemma Stable_Int: 
-    "[| F \<in> Stable(A); F \<in> Stable(A') |] ==> F \<in> Stable (A Int A')"
+    "[| F \<in> Stable(A); F \<in> Stable(A') |] ==> F \<in> Stable (A \<inter> A')"
 apply (unfold Stable_def)
 apply (blast intro: Constrains_Int)
 done
 
 lemma Stable_Constrains_Un: 
-    "[| F \<in> Stable(C); F \<in> A Co (C Un A') |]    
-     ==> F \<in> (C Un A) Co (C Un A')"
+    "[| F \<in> Stable(C); F \<in> A Co (C \<union> A') |]    
+     ==> F \<in> (C \<union> A) Co (C \<union> A')"
 apply (unfold Stable_def)
 apply (blast intro: Constrains_Un [THEN Constrains_weaken_R])
 done
 
 lemma Stable_Constrains_Int: 
-    "[| F \<in> Stable(C); F \<in> (C Int A) Co A' |]    
-     ==> F \<in> (C Int A) Co (C Int A')"
+    "[| F \<in> Stable(C); F \<in> (C \<inter> A) Co A' |]    
+     ==> F \<in> (C \<inter> A) Co (C \<inter> A')"
 apply (unfold Stable_def)
 apply (blast intro: Constrains_Int [THEN Constrains_weaken])
 done
@@ -301,7 +301,7 @@
 apply (simp (no_asm_simp) add: Stable_eq_stable Int_absorb)
 done
 
-lemma Stable_type: "Stable(A) <= program"
+lemma Stable_type: "Stable(A) \<subseteq> program"
 apply (unfold Stable_def)
 apply (rule Constrains_type)
 done
@@ -314,7 +314,7 @@
     "[| \<forall>m \<in> M. F \<in> ({s \<in> A. x(s) = m}) Co (B(m)); F \<in> program |]  
      ==> F \<in> ({s \<in> A. x(s):M}) Co (\<Union>m \<in> M. B(m))"
 apply (unfold Constrains_def, auto)
-apply (rule_tac A1 = "reachable (F) Int A" 
+apply (rule_tac A1 = "reachable (F) \<inter> A" 
         in UNITY.elimination [THEN constrains_weaken_L])
 apply (auto intro: constrains_weaken_L)
 done
@@ -351,7 +351,7 @@
 lemmas Always_imp_Stable = AlwaysD [THEN conjunct2]
 
 (*The set of all reachable states is Always*)
-lemma Always_includes_reachable: "F \<in> Always(A) ==> reachable(F) <= A"
+lemma Always_includes_reachable: "F \<in> Always(A) ==> reachable(F) \<subseteq> A"
 apply (simp (no_asm_use) add: Stable_def Constrains_def constrains_def Always_def initially_def)
 apply (rule subsetI)
 apply (erule reachable.induct)
@@ -366,20 +366,20 @@
 
 lemmas Always_reachable = invariant_reachable [THEN invariant_imp_Always]
 
-lemma Always_eq_invariant_reachable: "Always(A) = {F \<in> program. F \<in> invariant(reachable(F) Int A)}"
+lemma Always_eq_invariant_reachable: "Always(A) = {F \<in> program. F \<in> invariant(reachable(F) \<inter> A)}"
 apply (simp (no_asm) add: Always_def invariant_def Stable_def Constrains_def2 stable_def initially_def)
 apply (rule equalityI, auto) 
 apply (blast intro: reachable.intros reachable_type)
 done
 
 (*the RHS is the traditional definition of the "always" operator*)
-lemma Always_eq_includes_reachable: "Always(A) = {F \<in> program. reachable(F) <= A}"
+lemma Always_eq_includes_reachable: "Always(A) = {F \<in> program. reachable(F) \<subseteq> A}"
 apply (rule equalityI, safe)
 apply (auto dest: invariant_includes_reachable 
    simp add: subset_Int_iff invariant_reachable Always_eq_invariant_reachable)
 done
 
-lemma Always_type: "Always(A) <= program"
+lemma Always_type: "Always(A) \<subseteq> program"
 by (unfold Always_def initially_def, auto)
 
 lemma Always_state_eq: "Always(state) = program"
@@ -401,37 +401,37 @@
              dest: invariant_type [THEN subsetD])+
 done
 
-lemma Always_weaken: "[| F \<in> Always(A); A <= B |] ==> F \<in> Always(B)"
+lemma Always_weaken: "[| F \<in> Always(A); A \<subseteq> B |] ==> F \<in> Always(B)"
 by (auto simp add: Always_eq_includes_reachable)
 
 
 (*** "Co" rules involving Always ***)
 lemmas Int_absorb2 = subset_Int_iff [unfolded iff_def, THEN conjunct1, THEN mp]
 
-lemma Always_Constrains_pre: "F \<in> Always(I) ==> (F:(I Int A) Co A') <-> (F \<in> A Co A')"
+lemma Always_Constrains_pre: "F \<in> Always(I) ==> (F:(I \<inter> A) Co A') \<longleftrightarrow> (F \<in> A Co A')"
 apply (simp (no_asm_simp) add: Always_includes_reachable [THEN Int_absorb2] Constrains_def Int_assoc [symmetric])
 done
 
-lemma Always_Constrains_post: "F \<in> Always(I) ==> (F \<in> A Co (I Int A')) <->(F \<in> A Co A')"
+lemma Always_Constrains_post: "F \<in> Always(I) ==> (F \<in> A Co (I \<inter> A')) \<longleftrightarrow>(F \<in> A Co A')"
 apply (simp (no_asm_simp) add: Always_includes_reachable [THEN Int_absorb2] Constrains_eq_constrains Int_assoc [symmetric])
 done
 
-lemma Always_ConstrainsI: "[| F \<in> Always(I);  F \<in> (I Int A) Co A' |] ==> F \<in> A Co A'"
+lemma Always_ConstrainsI: "[| F \<in> Always(I);  F \<in> (I \<inter> A) Co A' |] ==> F \<in> A Co A'"
 by (blast intro: Always_Constrains_pre [THEN iffD1])
 
-(* [| F \<in> Always(I);  F \<in> A Co A' |] ==> F \<in> A Co (I Int A') *)
+(* [| F \<in> Always(I);  F \<in> A Co A' |] ==> F \<in> A Co (I \<inter> A') *)
 lemmas Always_ConstrainsD = Always_Constrains_post [THEN iffD2]
 
 (*The analogous proof of Always_LeadsTo_weaken doesn't terminate*)
 lemma Always_Constrains_weaken: 
-"[|F \<in> Always(C); F \<in> A Co A'; C Int B<=A; C Int A'<=B'|]==>F \<in> B Co B'"
+"[|F \<in> Always(C); F \<in> A Co A'; C \<inter> B<=A; C \<inter> A'<=B'|]==>F \<in> B Co B'"
 apply (rule Always_ConstrainsI)
 apply (drule_tac [2] Always_ConstrainsD, simp_all) 
 apply (blast intro: Constrains_weaken)
 done
 
 (** Conjoining Always properties **)
-lemma Always_Int_distrib: "Always(A Int B) = Always(A) Int Always(B)"
+lemma Always_Int_distrib: "Always(A \<inter> B) = Always(A) \<inter> Always(B)"
 by (auto simp add: Always_eq_includes_reachable)
 
 (* the premise i \<in> I is need since \<Inter>is formally not defined for I=0 *)
@@ -441,12 +441,12 @@
 done
 
 
-lemma Always_Int_I: "[| F \<in> Always(A);  F \<in> Always(B) |] ==> F \<in> Always(A Int B)"
+lemma Always_Int_I: "[| F \<in> Always(A);  F \<in> Always(B) |] ==> F \<in> Always(A \<inter> B)"
 apply (simp (no_asm_simp) add: Always_Int_distrib)
 done
 
 (*Allows a kind of "implication introduction"*)
-lemma Always_Diff_Un_eq: "[| F \<in> Always(A) |] ==> (F \<in> Always(C-A Un B)) <-> (F \<in> Always(B))"
+lemma Always_Diff_Un_eq: "[| F \<in> Always(A) |] ==> (F \<in> Always(C-A \<union> B)) \<longleftrightarrow> (F \<in> Always(B))"
 by (auto simp add: Always_eq_includes_reachable)
 
 (*Delete the nearest invariance assumption (which will be the second one
--- a/src/ZF/UNITY/Distributor.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/Distributor.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -79,7 +79,7 @@
 
 lemma (in distr) D_ok_iff:
      "G \<in> program ==>
-        D ok G <-> ((\<forall>n \<in> nat. G \<in> preserves(lift(Out(n)))) & D \<in> Allowed(G))"
+        D ok G \<longleftrightarrow> ((\<forall>n \<in> nat. G \<in> preserves(lift(Out(n)))) & D \<in> Allowed(G))"
 apply (cut_tac distr_spec)
 apply (auto simp add: INT_iff distr_spec_def distr_allowed_acts_def
                       Allowed_def ok_iff_Allowed)
@@ -131,7 +131,7 @@
 apply (subgoal_tac "length (s ` iIn) \<in> nat")
 apply typecheck
 apply (subgoal_tac "m \<in> nat")
-apply (drule_tac x = "nth(m, s`iIn) " and P = "%elt. ?X (elt) --> elt<Nclients" in bspec)
+apply (drule_tac x = "nth(m, s`iIn) " and P = "%elt. ?X (elt) \<longrightarrow> elt<Nclients" in bspec)
 apply (simp add: ltI)
 apply (simp_all add: Ord_mem_iff_lt)
 apply (blast dest: ltD)
--- a/src/ZF/UNITY/FP.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/FP.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -13,7 +13,7 @@
 
 definition   
   FP_Orig :: "i=>i"  where
-    "FP_Orig(F) == Union({A \<in> Pow(state). \<forall>B. F \<in> stable(A Int B)})"
+    "FP_Orig(F) == \<Union>({A \<in> Pow(state). \<forall>B. F \<in> stable(A \<inter> B)})"
 
 definition
   FP :: "i=>i"  where
@@ -36,19 +36,19 @@
 apply (rule FP_type)
 done
 
-lemma stable_FP_Orig_Int: "F \<in> program ==> F \<in> stable(FP_Orig(F) Int B)"
+lemma stable_FP_Orig_Int: "F \<in> program ==> F \<in> stable(FP_Orig(F) \<inter> B)"
 apply (simp only: FP_Orig_def stable_def Int_Union2)
 apply (blast intro: constrains_UN)
 done
 
 lemma FP_Orig_weakest2: 
-    "[| \<forall>B. F \<in> stable (A Int B); st_set(A) |]  ==> A \<subseteq> FP_Orig(F)"
+    "[| \<forall>B. F \<in> stable (A \<inter> B); st_set(A) |]  ==> A \<subseteq> FP_Orig(F)"
 by (unfold FP_Orig_def stable_def st_set_def, blast)
 
 lemmas FP_Orig_weakest = allI [THEN FP_Orig_weakest2]
 
-lemma stable_FP_Int: "F \<in> program ==> F \<in> stable (FP(F) Int B)"
-apply (subgoal_tac "FP (F) Int B = (\<Union>x\<in>B. FP (F) Int {x}) ")
+lemma stable_FP_Int: "F \<in> program ==> F \<in> stable (FP(F) \<inter> B)"
+apply (subgoal_tac "FP (F) \<inter> B = (\<Union>x\<in>B. FP (F) \<inter> {x}) ")
  prefer 2 apply blast
 apply (simp (no_asm_simp) add: Int_cons_right)
 apply (unfold FP_def stable_def)
@@ -71,7 +71,7 @@
 by (blast intro!: FP_Orig_subset_FP FP_subset_FP_Orig)
 
 lemma FP_weakest [rule_format]:
-     "[| \<forall>B. F \<in> stable(A Int B); F \<in> program; st_set(A) |] ==> A \<subseteq> FP(F)"
+     "[| \<forall>B. F \<in> stable(A \<inter> B); F \<in> program; st_set(A) |] ==> A \<subseteq> FP(F)"
 by (simp add: FP_equivalence FP_Orig_weakest)
 
 
--- a/src/ZF/UNITY/Follows.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/Follows.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -121,7 +121,7 @@
 apply (force simp add: part_order_def refl_def)
 apply (force simp add: part_order_def refl_def)
 apply (drule_tac x = "f1 (x) " and x1 = "f (sa) " and P2 = "%x y. \<forall>u\<in>B. ?P (x,y,u) " in bspec [THEN bspec])
-apply (drule_tac [3] x = "g (x) " and x1 = "g (sa) " and P2 = "%x y. ?P (x,y) --> ?d (x,y) \<in> t" in bspec [THEN bspec])
+apply (drule_tac [3] x = "g (x) " and x1 = "g (sa) " and P2 = "%x y. ?P (x,y) \<longrightarrow> ?d (x,y) \<in> t" in bspec [THEN bspec])
 apply auto
 apply (rule_tac b = "h (f (sa), g (sa))" and A = C in trans_onD)
 apply (auto simp add: part_order_def)
@@ -144,7 +144,7 @@
 apply (force simp add: part_order_def refl_def)
 apply (force simp add: part_order_def refl_def)
 apply (drule_tac x = "f (x) " and x1 = "f (sa) " in bspec [THEN bspec])
-apply (drule_tac [3] x = "g1 (x) " and x1 = "g (sa) " and P2 = "%x y. ?P (x,y) --> ?d (x,y) \<in> t" in bspec [THEN bspec])
+apply (drule_tac [3] x = "g1 (x) " and x1 = "g (sa) " and P2 = "%x y. ?P (x,y) \<longrightarrow> ?d (x,y) \<in> t" in bspec [THEN bspec])
 apply auto
 apply (rule_tac b = "h (f (sa), g (sa))" and A = C in trans_onD)
 apply (auto simp add: part_order_def)
@@ -152,7 +152,7 @@
 
 (**  This general result is used to prove Follows Un, munion, etc. **)
 lemma imp_LeadsTo_comp2:
-"[| F \<in> Increasing(A, r, f1) Int  Increasing(B, s, g);
+"[| F \<in> Increasing(A, r, f1) \<inter>  Increasing(B, s, g);
   \<forall>j \<in> A. F: {s \<in> state. <j, f(s)> \<in> r} LeadsTo {s \<in> state. <j,f1(s)> \<in> r};
   \<forall>j \<in> B. F: {x \<in> state. <j, g(x)> \<in> s} LeadsTo {x \<in> state. <j,g1(x)> \<in> s};
   mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t);
@@ -187,7 +187,7 @@
 
 lemma subset_Follows_comp:
 "[| mono1(A, r, B, s, h); refl(A, r); trans[B](s) |]
-   ==> Follows(A, r, f, g) <= Follows(B, s,  h comp f, h comp g)"
+   ==> Follows(A, r, f, g) \<subseteq> Follows(B, s,  h comp f, h comp g)"
 apply (unfold Follows_def, clarify)
 apply (frule_tac f = g in IncreasingD)
 apply (frule_tac f = f in IncreasingD)
@@ -324,25 +324,25 @@
 lemma increasing_Un:
      "[| F \<in> Increasing.increasing(Pow(A), SetLe(A), f);
          F \<in> Increasing.increasing(Pow(A), SetLe(A), g) |]
-     ==> F \<in> Increasing.increasing(Pow(A), SetLe(A), %x. f(x) Un g(x))"
+     ==> F \<in> Increasing.increasing(Pow(A), SetLe(A), %x. f(x) \<union> g(x))"
 by (rule_tac h = "op Un" in imp_increasing_comp2, auto)
 
 lemma Increasing_Un:
      "[| F \<in> Increasing(Pow(A), SetLe(A), f);
          F \<in> Increasing(Pow(A), SetLe(A), g) |]
-     ==> F \<in> Increasing(Pow(A), SetLe(A), %x. f(x) Un g(x))"
+     ==> F \<in> Increasing(Pow(A), SetLe(A), %x. f(x) \<union> g(x))"
 by (rule_tac h = "op Un" in imp_Increasing_comp2, auto)
 
 lemma Always_Un:
-     "[| F \<in> Always({s \<in> state. f1(s) <= f(s)});
-     F \<in> Always({s \<in> state. g1(s) <= g(s)}) |]
-      ==> F \<in> Always({s \<in> state. f1(s) Un g1(s) <= f(s) Un g(s)})"
+     "[| F \<in> Always({s \<in> state. f1(s) \<subseteq> f(s)});
+     F \<in> Always({s \<in> state. g1(s) \<subseteq> g(s)}) |]
+      ==> F \<in> Always({s \<in> state. f1(s) \<union> g1(s) \<subseteq> f(s) \<union> g(s)})"
 by (simp add: Always_eq_includes_reachable, blast)
 
 lemma Follows_Un:
 "[| F \<in> Follows(Pow(A), SetLe(A), f1, f);
      F \<in> Follows(Pow(A), SetLe(A), g1, g) |]
-     ==> F \<in> Follows(Pow(A), SetLe(A), %s. f1(s) Un g1(s), %s. f(s) Un g(s))"
+     ==> F \<in> Follows(Pow(A), SetLe(A), %s. f1(s) \<union> g1(s), %s. f(s) \<union> g(s))"
 by (rule_tac h = "op Un" in imp_Follows_comp2, auto)
 
 (** Multiset union properties (with the MultLe ordering) **)
--- a/src/ZF/UNITY/GenPrefix.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/GenPrefix.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -25,7 +25,7 @@
 inductive
   (* Parameter A is the domain of zs's elements *)
 
-  domains "gen_prefix(A, r)" <= "list(A)*list(A)"
+  domains "gen_prefix(A, r)" \<subseteq> "list(A)*list(A)"
 
   intros
     Nil:     "<[],[]>:gen_prefix(A, r)"
@@ -79,7 +79,7 @@
 
 lemma Cons_gen_prefix_aux:
   "[| <xs', ys'> \<in> gen_prefix(A, r) |]
-   ==> (\<forall>x xs. x \<in> A --> xs'= Cons(x,xs) -->
+   ==> (\<forall>x xs. x \<in> A \<longrightarrow> xs'= Cons(x,xs) \<longrightarrow>
        (\<exists>y ys. y \<in> A & ys' = Cons(y,ys) &
        <x,y>:r & <xs, ys> \<in> gen_prefix(A, r)))"
 apply (erule gen_prefix.induct)
@@ -97,14 +97,14 @@
 
 lemma Cons_gen_prefix_Cons:
 "(<Cons(x,xs),Cons(y,ys)> \<in> gen_prefix(A, r))
-  <-> (x \<in> A & y \<in> A & <x,y>:r & <xs,ys> \<in> gen_prefix(A, r))"
+  \<longleftrightarrow> (x \<in> A & y \<in> A & <x,y>:r & <xs,ys> \<in> gen_prefix(A, r))"
 apply (auto intro: gen_prefix.prepend)
 done
 declare Cons_gen_prefix_Cons [iff]
 
 (** Monotonicity of gen_prefix **)
 
-lemma gen_prefix_mono2: "r<=s ==> gen_prefix(A, r) <= gen_prefix(A, s)"
+lemma gen_prefix_mono2: "r<=s ==> gen_prefix(A, r) \<subseteq> gen_prefix(A, s)"
 apply clarify
 apply (frule gen_prefix.dom_subset [THEN subsetD], clarify)
 apply (erule rev_mp)
@@ -112,7 +112,7 @@
 apply (auto intro: gen_prefix.append)
 done
 
-lemma gen_prefix_mono1: "A<=B ==>gen_prefix(A, r) <= gen_prefix(B, r)"
+lemma gen_prefix_mono1: "A<=B ==>gen_prefix(A, r) \<subseteq> gen_prefix(B, r)"
 apply clarify
 apply (frule gen_prefix.dom_subset [THEN subsetD], clarify)
 apply (erule rev_mp)
@@ -126,7 +126,7 @@
             intro: gen_prefix.append list_mono [THEN subsetD])
 done
 
-lemma gen_prefix_mono: "[| A <= B; r <= s |] ==> gen_prefix(A, r) <= gen_prefix(B, s)"
+lemma gen_prefix_mono: "[| A \<subseteq> B; r \<subseteq> s |] ==> gen_prefix(A, r) \<subseteq> gen_prefix(B, s)"
 apply (rule subset_trans)
 apply (rule gen_prefix_mono1)
 apply (rule_tac [2] gen_prefix_mono2, auto)
@@ -145,7 +145,7 @@
 (* A lemma for proving gen_prefix_trans_comp *)
 
 lemma append_gen_prefix [rule_format (no_asm)]: "xs \<in> list(A) ==>
-   \<forall>zs. <xs @ ys, zs> \<in> gen_prefix(A, r) --> <xs, zs>: gen_prefix(A, r)"
+   \<forall>zs. <xs @ ys, zs> \<in> gen_prefix(A, r) \<longrightarrow> <xs, zs>: gen_prefix(A, r)"
 apply (erule list.induct)
 apply (auto dest: gen_prefix.dom_subset [THEN subsetD])
 done
@@ -154,7 +154,7 @@
 
 lemma gen_prefix_trans_comp [rule_format (no_asm)]:
      "<x, y>: gen_prefix(A, r) ==>
-   (\<forall>z \<in> list(A). <y,z> \<in> gen_prefix(A, s)--><x, z> \<in> gen_prefix(A, s O r))"
+   (\<forall>z \<in> list(A). <y,z> \<in> gen_prefix(A, s)\<longrightarrow><x, z> \<in> gen_prefix(A, s O r))"
 apply (erule gen_prefix.induct)
 apply (auto elim: ConsE simp add: Nil_gen_prefix)
 apply (subgoal_tac "ys \<in> list (A) ")
@@ -162,7 +162,7 @@
 apply (drule_tac xs = ys and r = s in append_gen_prefix, auto)
 done
 
-lemma trans_comp_subset: "trans(r) ==> r O r <= r"
+lemma trans_comp_subset: "trans(r) ==> r O r \<subseteq> r"
 by (auto dest: transD)
 
 lemma trans_gen_prefix: "trans(r) ==> trans(gen_prefix(A,r))"
@@ -198,7 +198,7 @@
 
 (** Antisymmetry **)
 
-lemma nat_le_lemma [rule_format]: "n \<in> nat ==> \<forall>b \<in> nat. n #+ b \<le> n --> b = 0"
+lemma nat_le_lemma [rule_format]: "n \<in> nat ==> \<forall>b \<in> nat. n #+ b \<le> n \<longrightarrow> b = 0"
 by (induct_tac "n", auto)
 
 lemma antisym_gen_prefix: "antisym(r) ==> antisym(gen_prefix(A, r))"
@@ -225,13 +225,13 @@
 
 (*** recursion equations ***)
 
-lemma gen_prefix_Nil: "xs \<in> list(A) ==> <xs, []> \<in> gen_prefix(A,r) <-> (xs = [])"
+lemma gen_prefix_Nil: "xs \<in> list(A) ==> <xs, []> \<in> gen_prefix(A,r) \<longleftrightarrow> (xs = [])"
 by (induct_tac "xs", auto)
 declare gen_prefix_Nil [simp]
 
 lemma same_gen_prefix_gen_prefix:
  "[| refl(A, r);  xs \<in> list(A) |] ==>
-    <xs@ys, xs@zs>: gen_prefix(A, r) <-> <ys,zs> \<in> gen_prefix(A, r)"
+    <xs@ys, xs@zs>: gen_prefix(A, r) \<longleftrightarrow> <ys,zs> \<in> gen_prefix(A, r)"
 apply (unfold refl_def)
 apply (induct_tac "xs")
 apply (simp_all (no_asm_simp))
@@ -239,7 +239,7 @@
 declare same_gen_prefix_gen_prefix [simp]
 
 lemma gen_prefix_Cons: "[| xs \<in> list(A); ys \<in> list(A); y \<in> A |] ==>
-    <xs, Cons(y,ys)> \<in> gen_prefix(A,r)  <->
+    <xs, Cons(y,ys)> \<in> gen_prefix(A,r)  \<longleftrightarrow>
       (xs=[] | (\<exists>z zs. xs=Cons(z,zs) & z \<in> A & <z,y>:r & <zs,ys> \<in> gen_prefix(A,r)))"
 apply (induct_tac "xs", auto)
 done
@@ -268,7 +268,7 @@
 
 lemma append_one_gen_prefix_lemma [rule_format]:
      "[| <xs,ys> \<in> gen_prefix(A, r);  refl(A, r) |]
-      ==> length(xs) < length(ys) -->
+      ==> length(xs) < length(ys) \<longrightarrow>
           <xs @ [nth(length(xs), ys)], ys> \<in> gen_prefix(A, r)"
 apply (erule gen_prefix.induct, blast)
 apply (frule gen_prefix.dom_subset [THEN subsetD], clarify)
@@ -281,7 +281,7 @@
 apply (simp_all add: nth_append length_type length_app)
 apply (rule conjI)
 apply (blast intro: gen_prefix.append)
-apply (erule_tac V = "length (xs) < length (ys) -->?u" in thin_rl)
+apply (erule_tac V = "length (xs) < length (ys) \<longrightarrow>?u" in thin_rl)
 apply (erule_tac a = zs in list.cases, auto)
 apply (rule_tac P1 = "%x. <?u (x), ?v>:?w" in nat_diff_split [THEN iffD2])
 apply auto
@@ -300,7 +300,7 @@
 
 lemma gen_prefix_imp_nth_lemma [rule_format]: "xs \<in> list(A) ==>
   \<forall>ys \<in> list(A). \<forall>i \<in> nat. i < length(xs)
-          --> <xs, ys>: gen_prefix(A, r) --> <nth(i, xs), nth(i, ys)>:r"
+          \<longrightarrow> <xs, ys>: gen_prefix(A, r) \<longrightarrow> <nth(i, xs), nth(i, ys)>:r"
 apply (induct_tac "xs", simp, clarify)
 apply simp
 apply (erule natE, auto)
@@ -315,8 +315,8 @@
 
 lemma nth_imp_gen_prefix [rule_format]: "xs \<in> list(A) ==>
   \<forall>ys \<in> list(A). length(xs) \<le> length(ys)
-      --> (\<forall>i. i < length(xs) --> <nth(i, xs), nth(i,ys)>:r)
-      --> <xs, ys> \<in> gen_prefix(A, r)"
+      \<longrightarrow> (\<forall>i. i < length(xs) \<longrightarrow> <nth(i, xs), nth(i,ys)>:r)
+      \<longrightarrow> <xs, ys> \<in> gen_prefix(A, r)"
 apply (induct_tac "xs")
 apply (simp_all (no_asm_simp))
 apply clarify
@@ -324,9 +324,9 @@
 apply (force intro!: nat_0_le simp add: lt_nat_in_nat)
 done
 
-lemma gen_prefix_iff_nth: "(<xs,ys> \<in> gen_prefix(A,r)) <->
+lemma gen_prefix_iff_nth: "(<xs,ys> \<in> gen_prefix(A,r)) \<longleftrightarrow>
       (xs \<in> list(A) & ys \<in> list(A) & length(xs) \<le> length(ys) &
-      (\<forall>i. i < length(xs) --> <nth(i,xs), nth(i, ys)>: r))"
+      (\<forall>i. i < length(xs) \<longrightarrow> <nth(i,xs), nth(i, ys)>: r))"
 apply (rule iffI)
 apply (frule gen_prefix.dom_subset [THEN subsetD])
 apply (frule gen_prefix_length_le, auto)
@@ -363,7 +363,7 @@
 (* Monotonicity of "set" operator WRT prefix *)
 
 lemma set_of_list_prefix_mono:
-"<xs,ys> \<in> prefix(A) ==> set_of_list(xs) <= set_of_list(ys)"
+"<xs,ys> \<in> prefix(A) ==> set_of_list(xs) \<subseteq> set_of_list(ys)"
 
 apply (unfold prefix_def)
 apply (erule gen_prefix.induct)
@@ -382,7 +382,7 @@
 declare Nil_prefix [simp]
 
 
-lemma prefix_Nil: "<xs, []> \<in> prefix(A) <-> (xs = [])"
+lemma prefix_Nil: "<xs, []> \<in> prefix(A) \<longleftrightarrow> (xs = [])"
 
 apply (unfold prefix_def, auto)
 apply (frule gen_prefix.dom_subset [THEN subsetD])
@@ -392,13 +392,13 @@
 declare prefix_Nil [iff]
 
 lemma Cons_prefix_Cons:
-"<Cons(x,xs), Cons(y,ys)> \<in> prefix(A) <-> (x=y & <xs,ys> \<in> prefix(A) & y \<in> A)"
+"<Cons(x,xs), Cons(y,ys)> \<in> prefix(A) \<longleftrightarrow> (x=y & <xs,ys> \<in> prefix(A) & y \<in> A)"
 apply (unfold prefix_def, auto)
 done
 declare Cons_prefix_Cons [iff]
 
 lemma same_prefix_prefix:
-"xs \<in> list(A)==> <xs@ys,xs@zs> \<in> prefix(A) <-> (<ys,zs> \<in> prefix(A))"
+"xs \<in> list(A)==> <xs@ys,xs@zs> \<in> prefix(A) \<longleftrightarrow> (<ys,zs> \<in> prefix(A))"
 apply (unfold prefix_def)
 apply (subgoal_tac "refl (A,id (A))")
 apply (simp (no_asm_simp))
@@ -406,8 +406,8 @@
 done
 declare same_prefix_prefix [simp]
 
-lemma same_prefix_prefix_Nil: "xs \<in> list(A) ==> <xs@ys,xs> \<in> prefix(A) <-> (<ys,[]> \<in> prefix(A))"
-apply (rule_tac P = "%x. <?u, x>:?v <-> ?w (x) " in app_right_Nil [THEN subst])
+lemma same_prefix_prefix_Nil: "xs \<in> list(A) ==> <xs@ys,xs> \<in> prefix(A) \<longleftrightarrow> (<ys,[]> \<in> prefix(A))"
+apply (rule_tac P = "%x. <?u, x>:?v \<longleftrightarrow> ?w (x) " in app_right_Nil [THEN subst])
 apply (rule_tac [2] same_prefix_prefix, auto)
 done
 declare same_prefix_prefix_Nil [simp]
@@ -421,7 +421,7 @@
 
 lemma prefix_Cons:
 "[| xs \<in> list(A); ys \<in> list(A); y \<in> A |] ==>
-  <xs,Cons(y,ys)> \<in> prefix(A) <->
+  <xs,Cons(y,ys)> \<in> prefix(A) \<longleftrightarrow>
   (xs=[] | (\<exists>zs. xs=Cons(y,zs) & <zs,ys> \<in> prefix(A)))"
 apply (unfold prefix_def)
 apply (auto simp add: gen_prefix_Cons)
@@ -448,13 +448,13 @@
 done
 
 lemma strict_prefix_type:
-"strict_prefix(A) <= list(A)*list(A)"
+"strict_prefix(A) \<subseteq> list(A)*list(A)"
 apply (unfold strict_prefix_def)
 apply (blast intro!: prefix_type [THEN subsetD])
 done
 
 lemma strict_prefix_length_lt_aux:
-     "<xs,ys> \<in> prefix(A) ==> xs\<noteq>ys --> length(xs) < length(ys)"
+     "<xs,ys> \<in> prefix(A) ==> xs\<noteq>ys \<longrightarrow> length(xs) < length(ys)"
 apply (unfold prefix_def)
 apply (erule gen_prefix.induct, clarify)
 apply (subgoal_tac [!] "ys \<in> list(A) & xs \<in> list(A)")
@@ -475,7 +475,7 @@
 
 (*Equivalence to the definition used in Lex/Prefix.thy*)
 lemma prefix_iff:
-    "<xs,zs> \<in> prefix(A) <-> (\<exists>ys \<in> list(A). zs = xs@ys) & xs \<in> list(A)"
+    "<xs,zs> \<in> prefix(A) \<longleftrightarrow> (\<exists>ys \<in> list(A). zs = xs@ys) & xs \<in> list(A)"
 apply (unfold prefix_def)
 apply (auto simp add: gen_prefix_iff_nth lt_nat_in_nat nth_append nth_type app_type length_app)
 apply (subgoal_tac "drop (length (xs), zs) \<in> list (A) ")
@@ -494,7 +494,7 @@
 
 lemma prefix_snoc:
 "[|xs \<in> list(A); ys \<in> list(A); y \<in> A |] ==>
-   <xs, ys@[y]> \<in> prefix(A) <-> (xs = ys@[y] | <xs,ys> \<in> prefix(A))"
+   <xs, ys@[y]> \<in> prefix(A) \<longleftrightarrow> (xs = ys@[y] | <xs,ys> \<in> prefix(A))"
 apply (simp (no_asm) add: prefix_iff)
 apply (rule iffI, clarify)
 apply (erule_tac xs = ysa in rev_list_elim, simp)
@@ -504,7 +504,7 @@
 declare prefix_snoc [simp]
 
 lemma prefix_append_iff [rule_format]: "zs \<in> list(A) ==> \<forall>xs \<in> list(A). \<forall>ys \<in> list(A).
-   (<xs, ys@zs> \<in> prefix(A)) <->
+   (<xs, ys@zs> \<in> prefix(A)) \<longleftrightarrow>
   (<xs,ys> \<in> prefix(A) | (\<exists>us. xs = ys@us & <us,zs> \<in> prefix(A)))"
 apply (erule list_append_induct, force, clarify)
 apply (rule iffI)
@@ -519,8 +519,8 @@
 (*Although the prefix ordering is not linear, the prefixes of a list
   are linearly ordered.*)
 lemma common_prefix_linear_lemma [rule_format]: "[| zs \<in> list(A); xs \<in> list(A); ys \<in> list(A) |]
-   ==> <xs, zs> \<in> prefix(A) --> <ys,zs> \<in> prefix(A)
-  --><xs,ys> \<in> prefix(A) | <ys,xs> \<in> prefix(A)"
+   ==> <xs, zs> \<in> prefix(A) \<longrightarrow> <ys,zs> \<in> prefix(A)
+  \<longrightarrow><xs,ys> \<in> prefix(A) | <ys,xs> \<in> prefix(A)"
 apply (erule list_append_induct, auto)
 done
 
@@ -646,7 +646,7 @@
 apply (erule natE, auto)
 done
 
-lemma prefix_take_iff: "<xs,ys> \<in> prefix(A) <-> (xs=take(length(xs), ys) & xs \<in> list(A) & ys \<in> list(A))"
+lemma prefix_take_iff: "<xs,ys> \<in> prefix(A) \<longleftrightarrow> (xs=take(length(xs), ys) & xs \<in> list(A) & ys \<in> list(A))"
 apply (rule iffI)
 apply (frule prefix_type [THEN subsetD])
 apply (blast intro: prefix_imp_take, clarify)
--- a/src/ZF/UNITY/Guar.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/Guar.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -25,7 +25,7 @@
 
 
 (* To be moved to theory WFair???? *)
-lemma leadsTo_Basis': "[| F \<in> A co A Un B; F \<in> transient(A); st_set(B) |] ==> F \<in> A leadsTo B"
+lemma leadsTo_Basis': "[| F \<in> A co A \<union> B; F \<in> transient(A); st_set(B) |] ==> F \<in> A leadsTo B"
 apply (frule constrainsD2)
 apply (drule_tac B = "A-B" in constrains_weaken_L, blast) 
 apply (drule_tac B = "A-B" in transient_strengthen, blast) 
@@ -39,37 +39,37 @@
 definition
    ex_prop :: "i => o"  where
    "ex_prop(X) == X<=program &
-    (\<forall>F \<in> program. \<forall>G \<in> program. F ok G --> F \<in> X | G \<in> X --> (F Join G) \<in> X)"
+    (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<in> X | G \<in> X \<longrightarrow> (F Join G) \<in> X)"
 
 definition
   strict_ex_prop  :: "i => o"  where
   "strict_ex_prop(X) == X<=program &
-   (\<forall>F \<in> program. \<forall>G \<in> program. F ok G --> (F \<in> X | G \<in> X) <-> (F Join G \<in> X))"
+   (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> (F \<in> X | G \<in> X) \<longleftrightarrow> (F Join G \<in> X))"
 
 definition
   uv_prop  :: "i => o"  where
    "uv_prop(X) == X<=program &
-    (SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G --> F \<in> X & G \<in> X --> (F Join G) \<in> X))"
+    (SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<in> X & G \<in> X \<longrightarrow> (F Join G) \<in> X))"
   
 definition
  strict_uv_prop  :: "i => o"  where
    "strict_uv_prop(X) == X<=program &
-    (SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G -->(F \<in> X & G \<in> X) <-> (F Join G \<in> X)))"
+    (SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow>(F \<in> X & G \<in> X) \<longleftrightarrow> (F Join G \<in> X)))"
 
 definition
   guar :: "[i, i] => i" (infixl "guarantees" 55)  where
               (*higher than membership, lower than Co*)
-  "X guarantees Y == {F \<in> program. \<forall>G \<in> program. F ok G --> F Join G \<in> X --> F Join G \<in> Y}"
+  "X guarantees Y == {F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F Join G \<in> X \<longrightarrow> F Join G \<in> Y}"
   
 definition
   (* Weakest guarantees *)
    wg :: "[i,i] => i"  where
-  "wg(F,Y) == Union({X \<in> Pow(program). F:(X guarantees Y)})"
+  "wg(F,Y) == \<Union>({X \<in> Pow(program). F:(X guarantees Y)})"
 
 definition
   (* Weakest existential property stronger than X *)
    wx :: "i =>i"  where
-   "wx(X) == Union({Y \<in> Pow(program). Y<=X & ex_prop(Y)})"
+   "wx(X) == \<Union>({Y \<in> Pow(program). Y<=X & ex_prop(Y)})"
 
 definition
   (*Ill-defined programs can arise through "Join"*)
@@ -79,12 +79,12 @@
 definition
   refines :: "[i, i, i] => o" ("(3_ refines _ wrt _)" [10,10,10] 10)  where
   "G refines F wrt X ==
-   \<forall>H \<in> program. (F ok H  & G ok H & F Join H \<in> welldef Int X)
-    --> (G Join H \<in> welldef Int X)"
+   \<forall>H \<in> program. (F ok H  & G ok H & F Join H \<in> welldef \<inter> X)
+    \<longrightarrow> (G Join H \<in> welldef \<inter> X)"
 
 definition
   iso_refines :: "[i,i, i] => o"  ("(3_ iso'_refines _ wrt _)" [10,10,10] 10)  where
-  "G iso_refines F wrt X ==  F \<in> welldef Int X --> G \<in> welldef Int X"
+  "G iso_refines F wrt X ==  F \<in> welldef \<inter> X \<longrightarrow> G \<in> welldef \<inter> X"
 
 
 (*** existential properties ***)
@@ -94,7 +94,7 @@
 
 lemma ex1 [rule_format]: 
  "GG \<in> Fin(program) 
-  ==> ex_prop(X) --> GG Int X\<noteq>0 --> OK(GG, (%G. G)) -->(\<Squnion>G \<in> GG. G) \<in> X"
+  ==> ex_prop(X) \<longrightarrow> GG \<inter> X\<noteq>0 \<longrightarrow> OK(GG, (%G. G)) \<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X"
 apply (unfold ex_prop_def)
 apply (erule Fin_induct)
 apply (simp_all add: OK_cons_iff)
@@ -103,8 +103,8 @@
 
 lemma ex2 [rule_format]: 
 "X \<subseteq> program ==>  
- (\<forall>GG \<in> Fin(program). GG Int X \<noteq> 0 --> OK(GG,(%G. G))-->(\<Squnion>G \<in> GG. G) \<in> X) 
-   --> ex_prop(X)"
+ (\<forall>GG \<in> Fin(program). GG \<inter> X \<noteq> 0 \<longrightarrow> OK(GG,(%G. G))\<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X) 
+   \<longrightarrow> ex_prop(X)"
 apply (unfold ex_prop_def, clarify)
 apply (drule_tac x = "{F,G}" in bspec)
 apply (simp_all add: OK_iff_ok)
@@ -114,16 +114,16 @@
 (*Chandy & Sanders take this as a definition*)
 
 lemma ex_prop_finite:
-     "ex_prop(X) <-> (X\<subseteq>program &  
-  (\<forall>GG \<in> Fin(program). GG Int X \<noteq> 0 & OK(GG,(%G. G))-->(\<Squnion>G \<in> GG. G) \<in> X))"
+     "ex_prop(X) \<longleftrightarrow> (X\<subseteq>program &  
+  (\<forall>GG \<in> Fin(program). GG \<inter> X \<noteq> 0 & OK(GG,(%G. G))\<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X))"
 apply auto
 apply (blast intro: ex1 ex2 dest: ex_imp_subset_program)+
 done
 
 (* Equivalent definition of ex_prop given at the end of section 3*)
 lemma ex_prop_equiv: 
-"ex_prop(X) <->  
-  X\<subseteq>program & (\<forall>G \<in> program. (G \<in> X <-> (\<forall>H \<in> program. (G component_of H) --> H \<in> X)))"
+"ex_prop(X) \<longleftrightarrow>  
+  X\<subseteq>program & (\<forall>G \<in> program. (G \<in> X \<longleftrightarrow> (\<forall>H \<in> program. (G component_of H) \<longrightarrow> H \<in> X)))"
 apply (unfold ex_prop_def component_of_def, safe, force, force, blast) 
 apply (subst Join_commute)
 apply (blast intro: ok_sym) 
@@ -138,7 +138,7 @@
 
 lemma uv1 [rule_format]: 
      "GG \<in> Fin(program) ==>  
-      (uv_prop(X)--> GG \<subseteq> X & OK(GG, (%G. G)) --> (\<Squnion>G \<in> GG. G) \<in> X)"
+      (uv_prop(X)\<longrightarrow> GG \<subseteq> X & OK(GG, (%G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)"
 apply (unfold uv_prop_def)
 apply (erule Fin_induct)
 apply (auto simp add: OK_cons_iff)
@@ -146,8 +146,8 @@
 
 lemma uv2 [rule_format]: 
      "X\<subseteq>program  ==> 
-      (\<forall>GG \<in> Fin(program). GG \<subseteq> X & OK(GG,(%G. G)) --> (\<Squnion>G \<in> GG. G) \<in> X)
-      --> uv_prop(X)"
+      (\<forall>GG \<in> Fin(program). GG \<subseteq> X & OK(GG,(%G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)
+      \<longrightarrow> uv_prop(X)"
 apply (unfold uv_prop_def, auto)
 apply (drule_tac x = 0 in bspec, simp+) 
 apply (drule_tac x = "{F,G}" in bspec, simp) 
@@ -156,8 +156,8 @@
 
 (*Chandy & Sanders take this as a definition*)
 lemma uv_prop_finite: 
-"uv_prop(X) <-> X\<subseteq>program &  
-  (\<forall>GG \<in> Fin(program). GG \<subseteq> X & OK(GG, %G. G) --> (\<Squnion>G \<in> GG. G) \<in>  X)"
+"uv_prop(X) \<longleftrightarrow> X\<subseteq>program &  
+  (\<forall>GG \<in> Fin(program). GG \<subseteq> X & OK(GG, %G. G) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in>  X)"
 apply auto
 apply (blast dest: uv_imp_subset_program)
 apply (blast intro: uv1)
@@ -224,7 +224,7 @@
 apply (force elim: equalityE)+
 done
 
-lemma ex_prop_equiv2: "(ex_prop(Y)) <-> (Y = program guarantees Y)"
+lemma ex_prop_equiv2: "(ex_prop(Y)) \<longleftrightarrow> (Y = program guarantees Y)"
 by (blast intro: ex_prop_imp guarantees_imp)
 
 (** Distributive laws.  Re-orient to perform miniscoping **)
@@ -238,7 +238,7 @@
 done
 
 lemma guarantees_Un_left: 
-     "(X Un Y) guarantees Z = (X guarantees Z) Int (Y guarantees Z)"
+     "(X \<union> Y) guarantees Z = (X guarantees Z) \<inter> (Y guarantees Z)"
 apply (unfold guar_def)
 apply (rule equalityI, safe, blast+)
 done
@@ -250,21 +250,21 @@
 done
 
 lemma guarantees_Int_right: 
-     "Z guarantees (X Int Y) = (Z guarantees X) Int (Z guarantees Y)"
+     "Z guarantees (X \<inter> Y) = (Z guarantees X) \<inter> (Z guarantees Y)"
 by (unfold guar_def, blast)
 
 lemma guarantees_Int_right_I:
      "[| F \<in> Z guarantees X;  F \<in> Z guarantees Y |]  
-     ==> F \<in> Z guarantees (X Int Y)"
+     ==> F \<in> Z guarantees (X \<inter> Y)"
 by (simp (no_asm_simp) add: guarantees_Int_right)
 
 lemma guarantees_INT_right_iff:
-     "i \<in> I==> (F \<in> X guarantees (\<Inter>i \<in> I. Y(i))) <->  
+     "i \<in> I==> (F \<in> X guarantees (\<Inter>i \<in> I. Y(i))) \<longleftrightarrow>  
       (\<forall>i \<in> I. F \<in> X guarantees Y(i))"
 by (simp add: guarantees_INT_right INT_iff, blast)
 
 
-lemma shunting: "(X guarantees Y) = (program guarantees ((program-X) Un Y))"
+lemma shunting: "(X guarantees Y) = (program guarantees ((program-X) \<union> Y))"
 by (unfold guar_def, auto)
 
 lemma contrapositive:
@@ -276,13 +276,13 @@
 **)
 
 lemma combining1: 
-    "[| F \<in> V guarantees X;  F \<in> (X Int Y) guarantees Z |] 
-     ==> F \<in> (V Int Y) guarantees Z"
+    "[| F \<in> V guarantees X;  F \<in> (X \<inter> Y) guarantees Z |] 
+     ==> F \<in> (V \<inter> Y) guarantees Z"
 by (unfold guar_def, blast)
 
 lemma combining2: 
-    "[| F \<in> V guarantees (X Un Y);  F \<in> Y guarantees Z |] 
-     ==> F \<in> V guarantees (X Un Z)"
+    "[| F \<in> V guarantees (X \<union> Y);  F \<in> Y guarantees Z |] 
+     ==> F \<in> V guarantees (X \<union> Z)"
 by (unfold guar_def, blast)
 
 
@@ -305,7 +305,7 @@
 
 lemma guarantees_Join_Int: 
     "[| F \<in> U guarantees V;  G \<in> X guarantees Y; F ok G |]  
-     ==> F Join G: (U Int X) guarantees (V Int Y)"
+     ==> F Join G: (U \<inter> X) guarantees (V \<inter> Y)"
 
 apply (unfold guar_def)
 apply (simp (no_asm))
@@ -318,7 +318,7 @@
 
 lemma guarantees_Join_Un: 
     "[| F \<in> U guarantees V;  G \<in> X guarantees Y; F ok G |]   
-     ==> F Join G: (U Un X) guarantees (V Un Y)"
+     ==> F Join G: (U \<union> X) guarantees (V \<union> Y)"
 apply (unfold guar_def)
 apply (simp (no_asm))
 apply safe
@@ -402,8 +402,8 @@
 done
 
 lemma guarantees_equiv: 
-"(F \<in> X guarantees Y) <->  
-   F \<in> program & (\<forall>H \<in> program. H \<in> X --> (F component_of H --> H \<in> Y))"
+"(F \<in> X guarantees Y) \<longleftrightarrow>  
+   F \<in> program & (\<forall>H \<in> program. H \<in> X \<longrightarrow> (F component_of H \<longrightarrow> H \<in> Y))"
 by (unfold guar_def component_of_def, force) 
 
 lemma wg_weakest:
@@ -414,8 +414,8 @@
 by (unfold wg_def guar_def, blast)
 
 lemma wg_equiv:
-     "H \<in> wg(F,X) <-> 
-      ((F component_of H --> H \<in> X) & F \<in> program & H \<in> program)"
+     "H \<in> wg(F,X) \<longleftrightarrow> 
+      ((F component_of H \<longrightarrow> H \<in> X) & F \<in> program & H \<in> program)"
 apply (simp add: wg_def guarantees_equiv)
 apply (rule iffI, safe)
 apply (rule_tac [4] x = "{H}" in bexI)
@@ -423,12 +423,12 @@
 done
 
 lemma component_of_wg: 
-    "F component_of H ==> H \<in> wg(F,X) <-> (H \<in> X & F \<in> program & H \<in> program)"
+    "F component_of H ==> H \<in> wg(F,X) \<longleftrightarrow> (H \<in> X & F \<in> program & H \<in> program)"
 by (simp (no_asm_simp) add: wg_equiv)
 
 lemma wg_finite [rule_format]: 
-"\<forall>FF \<in> Fin(program). FF Int X \<noteq> 0 --> OK(FF, %F. F)  
-   --> (\<forall>F \<in> FF. ((\<Squnion>F \<in> FF. F) \<in>  wg(F,X)) <-> ((\<Squnion>F \<in> FF. F) \<in> X))"
+"\<forall>FF \<in> Fin(program). FF \<inter> X \<noteq> 0 \<longrightarrow> OK(FF, %F. F)  
+   \<longrightarrow> (\<forall>F \<in> FF. ((\<Squnion>F \<in> FF. F) \<in>  wg(F,X)) \<longleftrightarrow> ((\<Squnion>F \<in> FF. F) \<in> X))"
 apply clarify
 apply (subgoal_tac "F component_of (\<Squnion>F \<in> FF. F) ")
 apply (drule_tac X = X in component_of_wg)
@@ -439,7 +439,7 @@
 done
 
 lemma wg_ex_prop:
-     "ex_prop(X) ==> (F \<in> X) <-> (\<forall>H \<in> program. H \<in> wg(F,X) & F \<in> program)"
+     "ex_prop(X) ==> (F \<in> X) \<longleftrightarrow> (\<forall>H \<in> program. H \<in> wg(F,X) & F \<in> program)"
 apply (simp (no_asm_use) add: ex_prop_equiv wg_equiv)
 apply blast
 done
@@ -456,12 +456,12 @@
 apply (rule_tac x=x in bexI, force, simp)+
 done
 
-lemma wx_weakest: "\<forall>Z. Z\<subseteq>program --> Z\<subseteq> X --> ex_prop(Z) --> Z \<subseteq> wx(X)"
+lemma wx_weakest: "\<forall>Z. Z\<subseteq>program \<longrightarrow> Z\<subseteq> X \<longrightarrow> ex_prop(Z) \<longrightarrow> Z \<subseteq> wx(X)"
 by (unfold wx_def, auto)
 
 (* Proposition 6 *)
 lemma wx'_ex_prop: 
- "ex_prop({F \<in> program. \<forall>G \<in> program. F ok G --> F Join G \<in> X})"
+ "ex_prop({F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F Join G \<in> X})"
 apply (unfold ex_prop_def, safe)
  apply (drule_tac x = "G Join Ga" in bspec)
   apply (simp (no_asm))
@@ -479,12 +479,12 @@
 (* Equivalence with the other definition of wx *)
 
 lemma wx_equiv: 
-     "wx(X) = {F \<in> program. \<forall>G \<in> program. F ok G --> (F Join G) \<in> X}"
+     "wx(X) = {F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> (F Join G) \<in> X}"
 apply (unfold wx_def)
 apply (rule equalityI, safe, blast)
  apply (simp (no_asm_use) add: ex_prop_def)
 apply blast 
-apply (rule_tac B = "{F \<in> program. \<forall>G \<in> program. F ok G --> F Join G \<in> X}" 
+apply (rule_tac B = "{F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F Join G \<in> X}" 
          in UnionI, 
        safe)
 apply (rule_tac [2] wx'_ex_prop)
@@ -498,7 +498,7 @@
 
 (* Proposition 12 *)
 (* Main result of the paper *)
-lemma guarantees_wx_eq: "(X guarantees Y) = wx((program-X) Un Y)"
+lemma guarantees_wx_eq: "(X guarantees Y) = wx((program-X) \<union> Y)"
 by (auto simp add: guar_def wx_equiv)
 
 (* 
@@ -521,7 +521,7 @@
 
 lemma constrains_guarantees_leadsTo:
      "[| F \<in> transient(A); st_set(B) |] 
-      ==> F: (A co A Un B) guarantees (A leadsTo (B-A))"
+      ==> F: (A co A \<union> B) guarantees (A leadsTo (B-A))"
 apply (rule guaranteesI)
  prefer 2 apply (blast dest: transient_type [THEN subsetD])
 apply (rule leadsTo_Basis')
--- a/src/ZF/UNITY/Increasing.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/Increasing.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -29,7 +29,7 @@
 
 (** increasing **)
 
-lemma increasing_type: "increasing[A](r, f) <= program"
+lemma increasing_type: "increasing[A](r, f) \<subseteq> program"
 by (unfold increasing_def, blast)
 
 lemma increasing_into_program: "F \<in> increasing[A](r, f) ==> F \<in> program"
@@ -47,7 +47,7 @@
 done
 
 lemma increasing_constant [simp]: 
- "F \<in> increasing[A](r, %s. c) <-> F \<in> program & c \<in> A"
+ "F \<in> increasing[A](r, %s. c) \<longleftrightarrow> F \<in> program & c \<in> A"
 apply (unfold increasing_def stable_def)
 apply (subgoal_tac "\<exists>x. x \<in> state")
 apply (auto dest: stable_type [THEN subsetD] intro: st0_in_state)
@@ -55,7 +55,7 @@
 
 lemma subset_increasing_comp: 
 "[| mono1(A, r, B, s, g); refl(A, r); trans[B](s)  |] ==>  
-   increasing[A](r, f) <= increasing[B](s, g comp f)"
+   increasing[A](r, f) \<subseteq> increasing[B](s, g comp f)"
 apply (unfold increasing_def stable_def part_order_def 
        constrains_def mono1_def metacomp_def, clarify, simp)
 apply clarify
@@ -79,11 +79,11 @@
 by (rule subset_increasing_comp [THEN subsetD], auto)
 
 lemma strict_increasing: 
-   "increasing[nat](Le, f) <= increasing[nat](Lt, f)"
+   "increasing[nat](Le, f) \<subseteq> increasing[nat](Lt, f)"
 by (unfold increasing_def Lt_def, auto)
 
 lemma strict_gt_increasing: 
-   "increasing[nat](Ge, f) <= increasing[nat](Gt, f)"
+   "increasing[nat](Ge, f) \<subseteq> increasing[nat](Gt, f)"
 apply (unfold increasing_def Gt_def Ge_def, auto)
 apply (erule natE)
 apply (auto simp add: stable_def)
@@ -98,7 +98,7 @@
 apply (auto intro: stable_imp_Stable)
 done
 
-lemma Increasing_type: "Increasing[A](r, f) <= program"
+lemma Increasing_type: "Increasing[A](r, f) \<subseteq> program"
 by (unfold Increasing_def, auto)
 
 lemma Increasing_into_program: "F \<in> Increasing[A](r, f) ==> F \<in> program"
@@ -116,14 +116,14 @@
 done
 
 lemma Increasing_constant [simp]: 
-     "F \<in> Increasing[A](r, %s. c) <-> F \<in> program & (c \<in> A)"
+     "F \<in> Increasing[A](r, %s. c) \<longleftrightarrow> F \<in> program & (c \<in> A)"
 apply (subgoal_tac "\<exists>x. x \<in> state")
 apply (auto dest!: IncreasingD intro: st0_in_state increasing_imp_Increasing)
 done
 
 lemma subset_Increasing_comp: 
 "[| mono1(A, r, B, s, g); refl(A, r); trans[B](s) |] ==>  
-   Increasing[A](r, f) <= Increasing[B](s, g comp f)"
+   Increasing[A](r, f) \<subseteq> Increasing[B](s, g comp f)"
 apply (unfold Increasing_def Stable_def Constrains_def part_order_def 
        constrains_def mono1_def metacomp_def, safe)
 apply (simp_all add: ActsD)
@@ -149,7 +149,7 @@
 apply (rule subset_Increasing_comp [THEN subsetD], auto)
 done
   
-lemma strict_Increasing: "Increasing[nat](Le, f) <= Increasing[nat](Lt, f)"
+lemma strict_Increasing: "Increasing[nat](Le, f) \<subseteq> Increasing[nat](Lt, f)"
 by (unfold Increasing_def Lt_def, auto)
 
 lemma strict_gt_Increasing: "Increasing[nat](Ge, f)<= Increasing[nat](Gt, f)"
--- a/src/ZF/UNITY/Merge.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/Merge.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -107,14 +107,14 @@
 done
 
 lemma (in merge) merge_Allowed: 
-     "Allowed(M) = (preserves(lift(Out)) Int preserves(lift(iOut)))"
+     "Allowed(M) = (preserves(lift(Out)) \<inter> preserves(lift(iOut)))"
 apply (insert merge_spec preserves_type [of "lift (Out)"])
 apply (auto simp add: merge_spec_def merge_allowed_acts_def Allowed_def safety_prop_Acts_iff)
 done
 
 lemma (in merge) M_ok_iff: 
      "G \<in> program ==>  
-       M ok G <-> (G \<in> preserves(lift(Out)) &   
+       M ok G \<longleftrightarrow> (G \<in> preserves(lift(Out)) &   
                    G \<in> preserves(lift(iOut)) & M \<in> Allowed(G))"
 apply (cut_tac merge_spec)
 apply (auto simp add: merge_Allowed ok_iff_Allowed)
@@ -166,7 +166,7 @@
 apply (subgoal_tac "xa \<in> nat")
 apply (simp_all add: Ord_mem_iff_lt)
 prefer 2 apply (blast intro: lt_trans)
-apply (drule_tac x = "nth (xa, x`iOut)" and P = "%elt. ?X (elt) --> elt<Nclients" in bspec)
+apply (drule_tac x = "nth (xa, x`iOut)" and P = "%elt. ?X (elt) \<longrightarrow> elt<Nclients" in bspec)
 apply (simp add: ltI nat_into_Ord)
 apply (blast dest: ltD)
 done
--- a/src/ZF/UNITY/Monotonicity.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/Monotonicity.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -14,7 +14,7 @@
 definition
   mono1 :: "[i, i, i, i, i=>i] => o"  where
   "mono1(A, r, B, s, f) ==
-    (\<forall>x \<in> A. \<forall>y \<in> A. <x,y> \<in> r --> <f(x), f(y)> \<in> s) & (\<forall>x \<in> A. f(x) \<in> B)"
+    (\<forall>x \<in> A. \<forall>y \<in> A. <x,y> \<in> r \<longrightarrow> <f(x), f(y)> \<in> s) & (\<forall>x \<in> A. f(x) \<in> B)"
 
   (* monotonicity of a 2-place meta-function f *)
 
@@ -22,18 +22,18 @@
   mono2 :: "[i, i, i, i, i, i, [i,i]=>i] => o"  where
   "mono2(A, r, B, s, C, t, f) == 
     (\<forall>x \<in> A. \<forall>y \<in> A. \<forall>u \<in> B. \<forall>v \<in> B.
-              <x,y> \<in> r & <u,v> \<in> s --> <f(x,u), f(y,v)> \<in> t) &
+              <x,y> \<in> r & <u,v> \<in> s \<longrightarrow> <f(x,u), f(y,v)> \<in> t) &
     (\<forall>x \<in> A. \<forall>y \<in> B. f(x,y) \<in> C)"
 
  (* Internalized relations on sets and multisets *)
 
 definition
   SetLe :: "i =>i"  where
-  "SetLe(A) == {<x,y> \<in> Pow(A)*Pow(A). x <= y}"
+  "SetLe(A) == {<x,y> \<in> Pow(A)*Pow(A). x \<subseteq> y}"
 
 definition
   MultLe :: "[i,i] =>i"  where
-  "MultLe(A, r) == multirel(A, r - id(A)) Un id(Mult(A))"
+  "MultLe(A, r) == multirel(A, r - id(A)) \<union> id(Mult(A))"
 
 
 lemma mono1D: 
@@ -50,24 +50,24 @@
 (** Monotonicity of take **)
 
 lemma take_mono_left_lemma:
-     "[| i le j; xs \<in> list(A); i \<in> nat; j \<in> nat |] 
+     "[| i \<le> j; xs \<in> list(A); i \<in> nat; j \<in> nat |] 
       ==> <take(i, xs), take(j, xs)> \<in> prefix(A)"
-apply (case_tac "length (xs) le i")
- apply (subgoal_tac "length (xs) le j")
+apply (case_tac "length (xs) \<le> i")
+ apply (subgoal_tac "length (xs) \<le> j")
   apply (simp)
  apply (blast intro: le_trans)
 apply (drule not_lt_imp_le, auto)
-apply (case_tac "length (xs) le j")
+apply (case_tac "length (xs) \<le> j")
  apply (auto simp add: take_prefix)
 apply (drule not_lt_imp_le, auto)
 apply (drule_tac m = i in less_imp_succ_add, auto)
-apply (subgoal_tac "i #+ k le length (xs) ")
+apply (subgoal_tac "i #+ k \<le> length (xs) ")
  apply (simp add: take_add prefix_iff take_type drop_type)
 apply (blast intro: leI)
 done
 
 lemma take_mono_left:
-     "[| i le j; xs \<in> list(A); j \<in> nat |]
+     "[| i \<le> j; xs \<in> list(A); j \<in> nat |]
       ==> <take(i, xs), take(j, xs)> \<in> prefix(A)"
 by (blast intro: le_in_nat take_mono_left_lemma) 
 
@@ -77,7 +77,7 @@
 by (auto simp add: prefix_iff)
 
 lemma take_mono:
-     "[| i le j; <xs, ys> \<in> prefix(A); j \<in> nat |]
+     "[| i \<le> j; <xs, ys> \<in> prefix(A); j \<in> nat |]
       ==> <take(i, xs), take(j, ys)> \<in> prefix(A)"
 apply (rule_tac b = "take (j, xs) " in prefix_trans)
 apply (auto dest: prefix_type [THEN subsetD] intro: take_mono_left take_mono_right)
@@ -99,7 +99,7 @@
 apply (auto dest: prefix_length_le simp add: Le_def)
 done
 
-(** Monotonicity of Un **)
+(** Monotonicity of \<union> **)
 
 lemma mono_Un [iff]: 
      "mono2(Pow(A), SetLe(A), Pow(A), SetLe(A), Pow(A), SetLe(A), op Un)"
--- a/src/ZF/UNITY/MultisetSum.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/MultisetSum.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -89,16 +89,16 @@
 (*The reversed orientation looks more natural, but LOOPS as a simprule!*)
 lemma msetsum_Un_Int: 
      "[| C \<in> Fin(A); D \<in> Fin(A); \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B |]
-      ==> msetsum(g, C Un D, B) +# msetsum(g, C Int D, B)  
+      ==> msetsum(g, C \<union> D, B) +# msetsum(g, C \<inter> D, B)  
         = msetsum(g, C, B) +# msetsum(g, D, B)"
 apply (erule Fin_induct)
-apply (subgoal_tac [2] "cons (x, y) Un D = cons (x, y Un D) ")
+apply (subgoal_tac [2] "cons (x, y) \<union> D = cons (x, y \<union> D) ")
 apply (auto simp add: msetsum_multiset)
-apply (subgoal_tac "y Un D \<in> Fin (A) & y Int D \<in> Fin (A) ")
+apply (subgoal_tac "y \<union> D \<in> Fin (A) & y \<inter> D \<in> Fin (A) ")
 apply clarify
 apply (case_tac "x \<in> D")
-apply (subgoal_tac [2] "cons (x, y) Int D = y Int D")
-apply (subgoal_tac "cons (x, y) Int D = cons (x, y Int D) ")
+apply (subgoal_tac [2] "cons (x, y) \<inter> D = y \<inter> D")
+apply (subgoal_tac "cons (x, y) \<inter> D = cons (x, y \<inter> D) ")
 apply (simp_all (no_asm_simp) add: cons_absorb munion_assoc msetsum_multiset)
 apply (simp (no_asm_simp) add: munion_lcommute msetsum_multiset)
 apply auto
@@ -106,26 +106,26 @@
 
 
 lemma msetsum_Un_disjoint:
-     "[| C \<in> Fin(A); D \<in> Fin(A); C Int D = 0;  
+     "[| C \<in> Fin(A); D \<in> Fin(A); C \<inter> D = 0;  
          \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B |]  
-      ==> msetsum(g, C Un D, B) = msetsum(g, C, B) +# msetsum(g,D, B)"
+      ==> msetsum(g, C \<union> D, B) = msetsum(g, C, B) +# msetsum(g,D, B)"
 apply (subst msetsum_Un_Int [symmetric])
 apply (auto simp add: msetsum_multiset)
 done
 
 lemma UN_Fin_lemma [rule_format (no_asm)]:
-     "I \<in> Fin(A) ==> (\<forall>i \<in> I. C(i) \<in> Fin(B)) --> (\<Union>i \<in> I. C(i)):Fin(B)"
+     "I \<in> Fin(A) ==> (\<forall>i \<in> I. C(i) \<in> Fin(B)) \<longrightarrow> (\<Union>i \<in> I. C(i)):Fin(B)"
 by (erule Fin_induct, auto)
  
 lemma msetsum_UN_disjoint [rule_format (no_asm)]:
      "[| I \<in> Fin(K); \<forall>i \<in> K. C(i) \<in> Fin(A) |] ==>  
-      (\<forall>x \<in> A. multiset(f(x)) & mset_of(f(x))\<subseteq>B) -->   
-      (\<forall>i \<in> I. \<forall>j \<in> I. i\<noteq>j --> C(i) Int C(j) = 0) -->  
+      (\<forall>x \<in> A. multiset(f(x)) & mset_of(f(x))\<subseteq>B) \<longrightarrow>   
+      (\<forall>i \<in> I. \<forall>j \<in> I. i\<noteq>j \<longrightarrow> C(i) \<inter> C(j) = 0) \<longrightarrow>  
         msetsum(f, \<Union>i \<in> I. C(i), B) = msetsum (%i. msetsum(f, C(i),B), I, B)"
 apply (erule Fin_induct, auto)
 apply (subgoal_tac "\<forall>i \<in> y. x \<noteq> i")
  prefer 2 apply blast
-apply (subgoal_tac "C(x) Int (\<Union>i \<in> y. C (i)) = 0")
+apply (subgoal_tac "C(x) \<inter> (\<Union>i \<in> y. C (i)) = 0")
  prefer 2 apply blast
 apply (subgoal_tac " (\<Union>i \<in> y. C (i)):Fin (A) & C(x) :Fin (A) ")
 prefer 2 apply (blast intro: UN_Fin_lemma dest: FinD, clarify)
@@ -157,9 +157,9 @@
 
 lemma msetsum_Un: "[| C \<in> Fin(A); D \<in> Fin(A);  
   \<forall>x \<in> A. multiset(f(x)) & mset_of(f(x)) \<subseteq> B  |]  
-   ==> msetsum(f, C Un D, B) =  
-          msetsum(f, C, B) +# msetsum(f, D, B) -# msetsum(f, C Int D, B)"
-apply (subgoal_tac "C Un D \<in> Fin (A) & C Int D \<in> Fin (A) ")
+   ==> msetsum(f, C \<union> D, B) =  
+          msetsum(f, C, B) +# msetsum(f, D, B) -# msetsum(f, C \<inter> D, B)"
+apply (subgoal_tac "C \<union> D \<in> Fin (A) & C \<inter> D \<in> Fin (A) ")
 apply clarify
 apply (subst msetsum_Un_Int [symmetric])
 apply (simp_all (no_asm_simp) add: msetsum_multiset multiset_union_diff)
--- a/src/ZF/UNITY/Mutex.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/Mutex.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -75,18 +75,18 @@
   (** The correct invariants **)
 
 definition
-  "IU == {s:state. (s`u = 1<->(#1 $<= s`m & s`m $<= #3))
-                     & (s`m = #3 --> s`p=0)}"
+  "IU == {s:state. (s`u = 1\<longleftrightarrow>(#1 $<= s`m & s`m $<= #3))
+                     & (s`m = #3 \<longrightarrow> s`p=0)}"
 
 definition
-  "IV == {s:state. (s`v = 1 <-> (#1 $<= s`n & s`n $<= #3))
-                      & (s`n = #3 --> s`p=1)}"
+  "IV == {s:state. (s`v = 1 \<longleftrightarrow> (#1 $<= s`n & s`n $<= #3))
+                      & (s`n = #3 \<longrightarrow> s`p=1)}"
 
   (** The faulty invariant (for U alone) **)
 
 definition
-  "bad_IU == {s:state. (s`u = 1 <-> (#1 $<= s`m & s`m  $<= #3))&
-                   (#3 $<= s`m & s`m $<= #4 --> s`p=0)}"
+  "bad_IU == {s:state. (s`u = 1 \<longleftrightarrow> (#1 $<= s`m & s`m  $<= #3))&
+                   (#3 $<= s`m & s`m $<= #4 \<longrightarrow> s`p=0)}"
 
 
 (** Variables' types **)
@@ -232,7 +232,7 @@
 lemma U_lemma1: "Mutex \<in> {s \<in> state. s`m = #1} LeadsTo {s \<in> state. s`p =1}"
 by (rule LeadsTo_Trans [OF U_F1 [THEN LeadsTo_weaken_R] U_lemma2], blast)
 
-lemma eq_123: "i \<in> int ==> (#1 $<= i & i $<= #3) <-> (i=#1 | i=#2 | i=#3)"
+lemma eq_123: "i \<in> int ==> (#1 $<= i & i $<= #3) \<longleftrightarrow> (i=#1 | i=#2 | i=#3)"
 apply auto
 apply (auto simp add: neq_iff_zless)
 apply (drule_tac [4] j = "#3" and i = i in zle_zless_trans)
--- a/src/ZF/UNITY/State.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/State.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -51,19 +51,19 @@
 
 (* Union *)
 
-lemma st_set_Un_iff [iff]: "st_set(A Un B) <-> st_set(A) & st_set(B)"
+lemma st_set_Un_iff [iff]: "st_set(A \<union> B) \<longleftrightarrow> st_set(A) & st_set(B)"
 by (simp add: st_set_def, auto)
 
-lemma st_set_Union_iff [iff]: "st_set(Union(S)) <-> (\<forall>A \<in> S. st_set(A))"
+lemma st_set_Union_iff [iff]: "st_set(\<Union>(S)) \<longleftrightarrow> (\<forall>A \<in> S. st_set(A))"
 by (simp add: st_set_def, auto)
 
 (* Intersection *)
 
-lemma st_set_Int [intro!]: "st_set(A) | st_set(B) ==> st_set(A Int B)"
+lemma st_set_Int [intro!]: "st_set(A) | st_set(B) ==> st_set(A \<inter> B)"
 by (simp add: st_set_def, auto)
 
 lemma st_set_Inter [intro!]: 
-   "(S=0) | (\<exists>A \<in> S. st_set(A)) ==> st_set(Inter(S))"
+   "(S=0) | (\<exists>A \<in> S. st_set(A)) ==> st_set(\<Inter>(S))"
 apply (simp add: st_set_def Inter_def, auto)
 done
 
@@ -71,16 +71,16 @@
 lemma st_set_DiffI [intro!]: "st_set(A) ==> st_set(A - B)"
 by (simp add: st_set_def, auto)
 
-lemma Collect_Int_state [simp]: "Collect(state,P) Int state = Collect(state,P)"
+lemma Collect_Int_state [simp]: "Collect(state,P) \<inter> state = Collect(state,P)"
 by auto
 
-lemma state_Int_Collect [simp]: "state Int Collect(state,P) = Collect(state,P)"
+lemma state_Int_Collect [simp]: "state \<inter> Collect(state,P) = Collect(state,P)"
 by auto
 
 
 (* Introduction and destruction rules for st_set *)
 
-lemma st_setI: "A <= state ==> st_set(A)"
+lemma st_setI: "A \<subseteq> state ==> st_set(A)"
 by (simp add: st_set_def)
 
 lemma st_setD: "st_set(A) ==> A<=state"
@@ -99,7 +99,7 @@
 lemma st_set_compl [simp]: "st_set(st_compl(A))"
 by (simp add: st_compl_def, auto)
 
-lemma st_compl_iff [simp]: "x \<in> st_compl(A) <-> x \<in> state & x \<notin> A"
+lemma st_compl_iff [simp]: "x \<in> st_compl(A) \<longleftrightarrow> x \<in> state & x \<notin> A"
 by (simp add: st_compl_def)
 
 lemma st_compl_Collect [simp]:
--- a/src/ZF/UNITY/SubstAx.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/SubstAx.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -14,11 +14,11 @@
 definition
   (* The definitions below are not `conventional', but yield simpler rules *)
   Ensures :: "[i,i] => i"            (infixl "Ensures" 60)  where
-  "A Ensures B == {F:program. F : (reachable(F) Int A) ensures (reachable(F) Int B) }"
+  "A Ensures B == {F:program. F \<in> (reachable(F) \<inter> A) ensures (reachable(F) \<inter> B) }"
 
 definition
   LeadsTo :: "[i, i] => i"            (infixl "LeadsTo" 60)  where
-  "A LeadsTo B == {F:program. F:(reachable(F) Int A) leadsTo (reachable(F) Int B)}"
+  "A LeadsTo B == {F:program. F:(reachable(F) \<inter> A) leadsTo (reachable(F) \<inter> B)}"
 
 notation (xsymbols)
   LeadsTo  (infixl " \<longmapsto>w " 60)
@@ -29,7 +29,7 @@
 
 (* Equivalence with the HOL-like definition *)
 lemma LeadsTo_eq: 
-"st_set(B)==> A LeadsTo B = {F \<in> program. F:(reachable(F) Int A) leadsTo B}"
+"st_set(B)==> A LeadsTo B = {F \<in> program. F:(reachable(F) \<inter> A) leadsTo B}"
 apply (unfold LeadsTo_def)
 apply (blast dest: psp_stable2 leadsToD2 constrainsD2 intro: leadsTo_weaken)
 done
@@ -40,20 +40,20 @@
 (*** Specialized laws for handling invariants ***)
 
 (** Conjoining an Always property **)
-lemma Always_LeadsTo_pre: "F \<in> Always(I) ==> (F:(I Int A) LeadsTo A') <-> (F \<in> A LeadsTo A')"
+lemma Always_LeadsTo_pre: "F \<in> Always(I) ==> (F:(I \<inter> A) LeadsTo A') \<longleftrightarrow> (F \<in> A LeadsTo A')"
 by (simp add: LeadsTo_def Always_eq_includes_reachable Int_absorb2 Int_assoc [symmetric] leadsToD2)
 
-lemma Always_LeadsTo_post: "F \<in> Always(I) ==> (F \<in> A LeadsTo (I Int A')) <-> (F \<in> A LeadsTo A')"
+lemma Always_LeadsTo_post: "F \<in> Always(I) ==> (F \<in> A LeadsTo (I \<inter> A')) \<longleftrightarrow> (F \<in> A LeadsTo A')"
 apply (unfold LeadsTo_def)
 apply (simp add: Always_eq_includes_reachable Int_absorb2 Int_assoc [symmetric] leadsToD2)
 done
 
 (* Like 'Always_LeadsTo_pre RS iffD1', but with premises in the good order *)
-lemma Always_LeadsToI: "[| F \<in> Always(C); F \<in> (C Int A) LeadsTo A' |] ==> F \<in> A LeadsTo A'"
+lemma Always_LeadsToI: "[| F \<in> Always(C); F \<in> (C \<inter> A) LeadsTo A' |] ==> F \<in> A LeadsTo A'"
 by (blast intro: Always_LeadsTo_pre [THEN iffD1])
 
 (* Like 'Always_LeadsTo_post RS iffD2', but with premises in the good order *)
-lemma Always_LeadsToD: "[| F \<in> Always(C);  F \<in> A LeadsTo A' |] ==> F \<in> A LeadsTo (C Int A')"
+lemma Always_LeadsToD: "[| F \<in> Always(C);  F \<in> A LeadsTo A' |] ==> F \<in> A LeadsTo (C \<inter> A')"
 by (blast intro: Always_LeadsTo_post [THEN iffD2])
 
 (*** Introduction rules \<in> Basis, Trans, Union ***)
@@ -68,7 +68,7 @@
 done
 
 lemma LeadsTo_Union:
-"[|(!!A. A \<in> S ==> F \<in> A LeadsTo B); F \<in> program|]==>F \<in> Union(S) LeadsTo B"
+"[|(!!A. A \<in> S ==> F \<in> A LeadsTo B); F \<in> program|]==>F \<in> \<Union>(S) LeadsTo B"
 apply (simp add: LeadsTo_def)
 apply (subst Int_Union_Union2)
 apply (rule leadsTo_UN, auto)
@@ -83,11 +83,11 @@
 done
 
 (*Useful with cancellation, disjunction*)
-lemma LeadsTo_Un_duplicate: "F \<in> A LeadsTo (A' Un A') ==> F \<in> A LeadsTo A'"
+lemma LeadsTo_Un_duplicate: "F \<in> A LeadsTo (A' \<union> A') ==> F \<in> A LeadsTo A'"
 by (simp add: Un_ac)
 
 lemma LeadsTo_Un_duplicate2:
-     "F \<in> A LeadsTo (A' Un C Un C) ==> F \<in> A LeadsTo (A' Un C)"
+     "F \<in> A LeadsTo (A' \<union> C \<union> C) ==> F \<in> A LeadsTo (A' \<union> C)"
 by (simp add: Un_ac)
 
 lemma LeadsTo_UN:
@@ -100,7 +100,7 @@
 
 (*Binary union introduction rule*)
 lemma LeadsTo_Un:
-     "[| F \<in> A LeadsTo C; F \<in> B LeadsTo C |] ==> F \<in> (A Un B) LeadsTo C"
+     "[| F \<in> A LeadsTo C; F \<in> B LeadsTo C |] ==> F \<in> (A \<union> B) LeadsTo C"
 apply (subst Un_eq_Union)
 apply (rule LeadsTo_Union)
 apply (auto dest: LeadsTo_type [THEN subsetD])
@@ -112,17 +112,17 @@
 apply (subst UN_singleton [symmetric], rule LeadsTo_UN, auto)
 done
 
-lemma subset_imp_LeadsTo: "[| A <= B; F \<in> program |] ==> F \<in> A LeadsTo B"
+lemma subset_imp_LeadsTo: "[| A \<subseteq> B; F \<in> program |] ==> F \<in> A LeadsTo B"
 apply (simp (no_asm_simp) add: LeadsTo_def)
 apply (blast intro: subset_imp_leadsTo)
 done
 
-lemma empty_LeadsTo: "F:0 LeadsTo A <-> F \<in> program"
+lemma empty_LeadsTo: "F:0 LeadsTo A \<longleftrightarrow> F \<in> program"
 by (auto dest: LeadsTo_type [THEN subsetD]
             intro: empty_subsetI [THEN subset_imp_LeadsTo])
 declare empty_LeadsTo [iff]
 
-lemma LeadsTo_state: "F \<in> A LeadsTo state <-> F \<in> program"
+lemma LeadsTo_state: "F \<in> A LeadsTo state \<longleftrightarrow> F \<in> program"
 by (auto dest: LeadsTo_type [THEN subsetD] simp add: LeadsTo_eq)
 declare LeadsTo_state [iff]
 
@@ -131,7 +131,7 @@
 apply (auto intro: leadsTo_weaken_R)
 done
 
-lemma LeadsTo_weaken_L: "[| F \<in> A LeadsTo A'; B <= A |] ==> F \<in> B LeadsTo A'"
+lemma LeadsTo_weaken_L: "[| F \<in> A LeadsTo A'; B \<subseteq> A |] ==> F \<in> B LeadsTo A'"
 apply (unfold LeadsTo_def)
 apply (auto intro: leadsTo_weaken_L)
 done
@@ -140,43 +140,43 @@
 by (blast intro: LeadsTo_weaken_R LeadsTo_weaken_L LeadsTo_Trans)
 
 lemma Always_LeadsTo_weaken: 
-"[| F \<in> Always(C);  F \<in> A LeadsTo A'; C Int B <= A;   C Int A' <= B' |]  
+"[| F \<in> Always(C);  F \<in> A LeadsTo A'; C \<inter> B \<subseteq> A;   C \<inter> A' \<subseteq> B' |]  
       ==> F \<in> B LeadsTo B'"
 apply (blast dest: Always_LeadsToI intro: LeadsTo_weaken Always_LeadsToD)
 done
 
 (** Two theorems for "proof lattices" **)
 
-lemma LeadsTo_Un_post: "F \<in> A LeadsTo B ==> F:(A Un B) LeadsTo B"
+lemma LeadsTo_Un_post: "F \<in> A LeadsTo B ==> F:(A \<union> B) LeadsTo B"
 by (blast dest: LeadsTo_type [THEN subsetD]
              intro: LeadsTo_Un subset_imp_LeadsTo)
 
 lemma LeadsTo_Trans_Un: "[| F \<in> A LeadsTo B;  F \<in> B LeadsTo C |]  
-      ==> F \<in> (A Un B) LeadsTo C"
+      ==> F \<in> (A \<union> B) LeadsTo C"
 apply (blast intro: LeadsTo_Un subset_imp_LeadsTo LeadsTo_weaken_L LeadsTo_Trans dest: LeadsTo_type [THEN subsetD])
 done
 
 (** Distributive laws **)
-lemma LeadsTo_Un_distrib: "(F \<in> (A Un B) LeadsTo C)  <-> (F \<in> A LeadsTo C & F \<in> B LeadsTo C)"
+lemma LeadsTo_Un_distrib: "(F \<in> (A \<union> B) LeadsTo C)  \<longleftrightarrow> (F \<in> A LeadsTo C & F \<in> B LeadsTo C)"
 by (blast intro: LeadsTo_Un LeadsTo_weaken_L)
 
-lemma LeadsTo_UN_distrib: "(F \<in> (\<Union>i \<in> I. A(i)) LeadsTo B) <->  (\<forall>i \<in> I. F \<in> A(i) LeadsTo B) & F \<in> program"
+lemma LeadsTo_UN_distrib: "(F \<in> (\<Union>i \<in> I. A(i)) LeadsTo B) \<longleftrightarrow>  (\<forall>i \<in> I. F \<in> A(i) LeadsTo B) & F \<in> program"
 by (blast dest: LeadsTo_type [THEN subsetD]
              intro: LeadsTo_UN LeadsTo_weaken_L)
 
-lemma LeadsTo_Union_distrib: "(F \<in> Union(S) LeadsTo B)  <->  (\<forall>A \<in> S. F \<in> A LeadsTo B) & F \<in> program"
+lemma LeadsTo_Union_distrib: "(F \<in> \<Union>(S) LeadsTo B)  \<longleftrightarrow>  (\<forall>A \<in> S. F \<in> A LeadsTo B) & F \<in> program"
 by (blast dest: LeadsTo_type [THEN subsetD]
              intro: LeadsTo_Union LeadsTo_weaken_L)
 
 (** More rules using the premise "Always(I)" **)
 
-lemma EnsuresI: "[| F:(A-B) Co (A Un B);  F \<in> transient (A-B) |] ==> F \<in> A Ensures B"
+lemma EnsuresI: "[| F:(A-B) Co (A \<union> B);  F \<in> transient (A-B) |] ==> F \<in> A Ensures B"
 apply (simp add: Ensures_def Constrains_eq_constrains)
 apply (blast intro: ensuresI constrains_weaken transient_strengthen dest: constrainsD2)
 done
 
-lemma Always_LeadsTo_Basis: "[| F \<in> Always(I); F \<in> (I Int (A-A')) Co (A Un A');  
-         F \<in> transient (I Int (A-A')) |]    
+lemma Always_LeadsTo_Basis: "[| F \<in> Always(I); F \<in> (I \<inter> (A-A')) Co (A \<union> A');  
+         F \<in> transient (I \<inter> (A-A')) |]    
   ==> F \<in> A LeadsTo A'"
 apply (rule Always_LeadsToI, assumption)
 apply (blast intro: EnsuresI LeadsTo_Basis Always_ConstrainsD [THEN Constrains_weaken] transient_strengthen)
@@ -185,7 +185,7 @@
 (*Set difference: maybe combine with leadsTo_weaken_L??
   This is the most useful form of the "disjunction" rule*)
 lemma LeadsTo_Diff:
-     "[| F \<in> (A-B) LeadsTo C;  F \<in> (A Int B) LeadsTo C |] ==> F \<in> A LeadsTo C"
+     "[| F \<in> (A-B) LeadsTo C;  F \<in> (A \<inter> B) LeadsTo C |] ==> F \<in> A LeadsTo C"
 by (blast intro: LeadsTo_Un LeadsTo_weaken)
 
 lemma LeadsTo_UN_UN:  
@@ -197,32 +197,32 @@
 
 (*Binary union version*)
 lemma LeadsTo_Un_Un:
-  "[| F \<in> A LeadsTo A'; F \<in> B LeadsTo B' |] ==> F:(A Un B) LeadsTo (A' Un B')"
+  "[| F \<in> A LeadsTo A'; F \<in> B LeadsTo B' |] ==> F:(A \<union> B) LeadsTo (A' \<union> B')"
 by (blast intro: LeadsTo_Un LeadsTo_weaken_R)
 
 (** The cancellation law **)
 
-lemma LeadsTo_cancel2: "[| F \<in> A LeadsTo(A' Un B); F \<in> B LeadsTo B' |] ==> F \<in> A LeadsTo (A' Un B')"
+lemma LeadsTo_cancel2: "[| F \<in> A LeadsTo(A' \<union> B); F \<in> B LeadsTo B' |] ==> F \<in> A LeadsTo (A' \<union> B')"
 by (blast intro: LeadsTo_Un_Un subset_imp_LeadsTo LeadsTo_Trans dest: LeadsTo_type [THEN subsetD])
 
-lemma Un_Diff: "A Un (B - A) = A Un B"
+lemma Un_Diff: "A \<union> (B - A) = A \<union> B"
 by auto
 
-lemma LeadsTo_cancel_Diff2: "[| F \<in> A LeadsTo (A' Un B); F \<in> (B-A') LeadsTo B' |] ==> F \<in> A LeadsTo (A' Un B')"
+lemma LeadsTo_cancel_Diff2: "[| F \<in> A LeadsTo (A' \<union> B); F \<in> (B-A') LeadsTo B' |] ==> F \<in> A LeadsTo (A' \<union> B')"
 apply (rule LeadsTo_cancel2)
 prefer 2 apply assumption
 apply (simp (no_asm_simp) add: Un_Diff)
 done
 
-lemma LeadsTo_cancel1: "[| F \<in> A LeadsTo (B Un A'); F \<in> B LeadsTo B' |] ==> F \<in> A LeadsTo (B' Un A')"
+lemma LeadsTo_cancel1: "[| F \<in> A LeadsTo (B \<union> A'); F \<in> B LeadsTo B' |] ==> F \<in> A LeadsTo (B' \<union> A')"
 apply (simp add: Un_commute)
 apply (blast intro!: LeadsTo_cancel2)
 done
 
-lemma Diff_Un2: "(B - A) Un A = B Un A"
+lemma Diff_Un2: "(B - A) \<union> A = B \<union> A"
 by auto
 
-lemma LeadsTo_cancel_Diff1: "[| F \<in> A LeadsTo (B Un A'); F \<in> (B-A') LeadsTo B' |] ==> F \<in> A LeadsTo (B' Un A')"
+lemma LeadsTo_cancel_Diff1: "[| F \<in> A LeadsTo (B \<union> A'); F \<in> (B-A') LeadsTo B' |] ==> F \<in> A LeadsTo (B' \<union> A')"
 apply (rule LeadsTo_cancel1)
 prefer 2 apply assumption
 apply (simp (no_asm_simp) add: Diff_Un2)
@@ -240,26 +240,26 @@
 (** PSP \<in> Progress-Safety-Progress **)
 
 (*Special case of PSP \<in> Misra's "stable conjunction"*)
-lemma PSP_Stable: "[| F \<in> A LeadsTo A';  F \<in> Stable(B) |]==> F:(A Int B) LeadsTo (A' Int B)"
+lemma PSP_Stable: "[| F \<in> A LeadsTo A';  F \<in> Stable(B) |]==> F:(A \<inter> B) LeadsTo (A' \<inter> B)"
 apply (simp add: LeadsTo_def Stable_eq_stable, clarify)
 apply (drule psp_stable, assumption)
 apply (simp add: Int_ac)
 done
 
-lemma PSP_Stable2: "[| F \<in> A LeadsTo A'; F \<in> Stable(B) |] ==> F \<in> (B Int A) LeadsTo (B Int A')"
+lemma PSP_Stable2: "[| F \<in> A LeadsTo A'; F \<in> Stable(B) |] ==> F \<in> (B \<inter> A) LeadsTo (B \<inter> A')"
 apply (simp (no_asm_simp) add: PSP_Stable Int_ac)
 done
 
-lemma PSP: "[| F \<in> A LeadsTo A'; F \<in> B Co B'|]==> F \<in> (A Int B') LeadsTo ((A' Int B) Un (B' - B))"
+lemma PSP: "[| F \<in> A LeadsTo A'; F \<in> B Co B'|]==> F \<in> (A \<inter> B') LeadsTo ((A' \<inter> B) \<union> (B' - B))"
 apply (simp (no_asm_use) add: LeadsTo_def Constrains_eq_constrains)
 apply (blast dest: psp intro: leadsTo_weaken)
 done
 
-lemma PSP2: "[| F \<in> A LeadsTo A'; F \<in> B Co B' |]==> F:(B' Int A) LeadsTo ((B Int A') Un (B' - B))"
+lemma PSP2: "[| F \<in> A LeadsTo A'; F \<in> B Co B' |]==> F:(B' \<inter> A) LeadsTo ((B \<inter> A') \<union> (B' - B))"
 by (simp (no_asm_simp) add: PSP Int_ac)
 
 lemma PSP_Unless: 
-"[| F \<in> A LeadsTo A'; F \<in> B Unless B'|]==> F:(A Int B) LeadsTo ((A' Int B) Un B')"
+"[| F \<in> A LeadsTo A'; F \<in> B Unless B'|]==> F:(A \<inter> B) LeadsTo ((A' \<inter> B) \<union> B')"
 apply (unfold op_Unless_def)
 apply (drule PSP, assumption)
 apply (blast intro: LeadsTo_Diff LeadsTo_weaken subset_imp_LeadsTo)
@@ -269,20 +269,20 @@
 
 (** Meta or object quantifier ????? **)
 lemma LeadsTo_wf_induct: "[| wf(r);      
-         \<forall>m \<in> I. F \<in> (A Int f-``{m}) LeadsTo                      
-                            ((A Int f-``(converse(r) `` {m})) Un B);  
+         \<forall>m \<in> I. F \<in> (A \<inter> f-``{m}) LeadsTo                      
+                            ((A \<inter> f-``(converse(r) `` {m})) \<union> B);  
          field(r)<=I; A<=f-``I; F \<in> program |]  
       ==> F \<in> A LeadsTo B"
 apply (simp (no_asm_use) add: LeadsTo_def)
 apply auto
 apply (erule_tac I = I and f = f in leadsTo_wf_induct, safe)
 apply (drule_tac [2] x = m in bspec, safe)
-apply (rule_tac [2] A' = "reachable (F) Int (A Int f -`` (converse (r) ``{m}) Un B) " in leadsTo_weaken_R)
+apply (rule_tac [2] A' = "reachable (F) \<inter> (A \<inter> f -`` (converse (r) ``{m}) \<union> B) " in leadsTo_weaken_R)
 apply (auto simp add: Int_assoc) 
 done
 
 
-lemma LessThan_induct: "[| \<forall>m \<in> nat. F:(A Int f-``{m}) LeadsTo ((A Int f-``m) Un B);  
+lemma LessThan_induct: "[| \<forall>m \<in> nat. F:(A \<inter> f-``{m}) LeadsTo ((A \<inter> f-``m) \<union> B);  
       A<=f-``nat; F \<in> program |] ==> F \<in> A LeadsTo B"
 apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN LeadsTo_wf_induct])
 apply (simp_all add: nat_measure_field)
@@ -301,18 +301,18 @@
 
 (*** Completion \<in> Binary and General Finite versions ***)
 
-lemma Completion: "[| F \<in> A LeadsTo (A' Un C);  F \<in> A' Co (A' Un C);  
-         F \<in> B LeadsTo (B' Un C);  F \<in> B' Co (B' Un C) |]  
-      ==> F \<in> (A Int B) LeadsTo ((A' Int B') Un C)"
+lemma Completion: "[| F \<in> A LeadsTo (A' \<union> C);  F \<in> A' Co (A' \<union> C);  
+         F \<in> B LeadsTo (B' \<union> C);  F \<in> B' Co (B' \<union> C) |]  
+      ==> F \<in> (A \<inter> B) LeadsTo ((A' \<inter> B') \<union> C)"
 apply (simp (no_asm_use) add: LeadsTo_def Constrains_eq_constrains Int_Un_distrib)
 apply (blast intro: completion leadsTo_weaken)
 done
 
 lemma Finite_completion_aux:
      "[| I \<in> Fin(X);F \<in> program |]  
-      ==> (\<forall>i \<in> I. F \<in> (A(i)) LeadsTo (A'(i) Un C)) -->   
-          (\<forall>i \<in> I. F \<in> (A'(i)) Co (A'(i) Un C)) -->  
-          F \<in> (\<Inter>i \<in> I. A(i)) LeadsTo ((\<Inter>i \<in> I. A'(i)) Un C)"
+      ==> (\<forall>i \<in> I. F \<in> (A(i)) LeadsTo (A'(i) \<union> C)) \<longrightarrow>   
+          (\<forall>i \<in> I. F \<in> (A'(i)) Co (A'(i) \<union> C)) \<longrightarrow>  
+          F \<in> (\<Inter>i \<in> I. A(i)) LeadsTo ((\<Inter>i \<in> I. A'(i)) \<union> C)"
 apply (erule Fin_induct)
 apply (auto simp del: INT_simps simp add: Inter_0)
 apply (rule Completion, auto) 
@@ -321,16 +321,16 @@
 done
 
 lemma Finite_completion: 
-     "[| I \<in> Fin(X); !!i. i \<in> I ==> F \<in> A(i) LeadsTo (A'(i) Un C);  
-         !!i. i \<in> I ==> F \<in> A'(i) Co (A'(i) Un C);  
+     "[| I \<in> Fin(X); !!i. i \<in> I ==> F \<in> A(i) LeadsTo (A'(i) \<union> C);  
+         !!i. i \<in> I ==> F \<in> A'(i) Co (A'(i) \<union> C);  
          F \<in> program |]    
-      ==> F \<in> (\<Inter>i \<in> I. A(i)) LeadsTo ((\<Inter>i \<in> I. A'(i)) Un C)"
+      ==> F \<in> (\<Inter>i \<in> I. A(i)) LeadsTo ((\<Inter>i \<in> I. A'(i)) \<union> C)"
 by (blast intro: Finite_completion_aux [THEN mp, THEN mp])
 
 lemma Stable_completion: 
      "[| F \<in> A LeadsTo A';  F \<in> Stable(A');    
          F \<in> B LeadsTo B';  F \<in> Stable(B') |]  
-    ==> F \<in> (A Int B) LeadsTo (A' Int B')"
+    ==> F \<in> (A \<inter> B) LeadsTo (A' \<inter> B')"
 apply (unfold Stable_def)
 apply (rule_tac C1 = 0 in Completion [THEN LeadsTo_weaken_R])
     prefer 5
--- a/src/ZF/UNITY/UNITY.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/UNITY.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -78,7 +78,7 @@
 
 definition
   strongest_rhs :: "[i, i] => i"  where
-  "strongest_rhs(F, A) == Inter({B \<in> Pow(state). F \<in> A co B})"
+  "strongest_rhs(F, A) == \<Inter>({B \<in> Pow(state). F \<in> A co B})"
 
 definition
   invariant :: "i => i"  where
@@ -99,7 +99,7 @@
     --{* the condition @{term "st_set(A)"} makes the definition slightly
          stronger than the HOL one *}
 
-  unless_def:    "A unless B == (A - B) co (A Un B)"
+  unless_def:    "A unless B == (A - B) co (A \<union> B)"
 
 
 text{*SKIP*}
@@ -193,15 +193,15 @@
 
 text{*But are they really needed?*}
 
-lemma state_subset_is_Init_iff [iff]: "state \<subseteq> Init(F) <-> Init(F)=state"
+lemma state_subset_is_Init_iff [iff]: "state \<subseteq> Init(F) \<longleftrightarrow> Init(F)=state"
 by (cut_tac F = F in Init_type, auto)
 
 lemma Pow_state_times_state_is_subset_Acts_iff [iff]:
-     "Pow(state*state) \<subseteq> Acts(F) <-> Acts(F)=Pow(state*state)"
+     "Pow(state*state) \<subseteq> Acts(F) \<longleftrightarrow> Acts(F)=Pow(state*state)"
 by (cut_tac F = F in Acts_type, auto)
 
 lemma Pow_state_times_state_is_subset_AllowedActs_iff [iff]:
-     "Pow(state*state) \<subseteq> AllowedActs(F) <-> AllowedActs(F)=Pow(state*state)"
+     "Pow(state*state) \<subseteq> AllowedActs(F) \<longleftrightarrow> AllowedActs(F)=Pow(state*state)"
 by (cut_tac F = F in AllowedActs_type, auto)
 
 subsubsection{*Eliminating @{text "\<inter> state"} from expressions*}
@@ -311,7 +311,7 @@
 
 
 lemma program_equality_iff:
-    "[| F \<in> program; G \<in> program |] ==>(F=G)  <->
+    "[| F \<in> program; G \<in> program |] ==>(F=G)  \<longleftrightarrow>
      (Init(F) = Init(G) & Acts(F) = Acts(G) & AllowedActs(F) = AllowedActs(G))"
 by (blast intro: program_equalityI program_equalityE)
 
@@ -342,11 +342,11 @@
 text{*An action is expanded only if a pair of states is being tested against it*}
 lemma def_act_simp:
      "[| act == {<s,s'> \<in> A*B. P(s, s')} |]
-      ==> (<s,s'> \<in> act) <-> (<s,s'> \<in> A*B & P(s, s'))"
+      ==> (<s,s'> \<in> act) \<longleftrightarrow> (<s,s'> \<in> A*B & P(s, s'))"
 by auto
 
 text{*A set is expanded only if an element is being tested against it*}
-lemma def_set_simp: "A == B ==> (x \<in> A) <-> (x \<in> B)"
+lemma def_set_simp: "A == B ==> (x \<in> A) \<longleftrightarrow> (x \<in> B)"
 by auto
 
 
@@ -367,18 +367,18 @@
 lemma constrainsD2: "F \<in> A co B ==> F \<in> program & st_set(A)"
 by (force simp add: constrains_def)
 
-lemma constrains_empty [iff]: "F \<in> 0 co B <-> F \<in> program"
+lemma constrains_empty [iff]: "F \<in> 0 co B \<longleftrightarrow> F \<in> program"
 by (force simp add: constrains_def st_set_def)
 
-lemma constrains_empty2 [iff]: "(F \<in> A co 0) <-> (A=0 & F \<in> program)"
+lemma constrains_empty2 [iff]: "(F \<in> A co 0) \<longleftrightarrow> (A=0 & F \<in> program)"
 by (force simp add: constrains_def st_set_def)
 
-lemma constrains_state [iff]: "(F \<in> state co B) <-> (state\<subseteq>B & F \<in> program)"
+lemma constrains_state [iff]: "(F \<in> state co B) \<longleftrightarrow> (state\<subseteq>B & F \<in> program)"
 apply (cut_tac F = F in Acts_type)
 apply (force simp add: constrains_def st_set_def)
 done
 
-lemma constrains_state2 [iff]: "F \<in> A co state <-> (F \<in> program & st_set(A))"
+lemma constrains_state2 [iff]: "F \<in> A co state \<longleftrightarrow> (F \<in> program & st_set(A))"
 apply (cut_tac F = F in Acts_type)
 apply (force simp add: constrains_def st_set_def)
 done
@@ -405,7 +405,7 @@
 subsection{*Constrains and Union*}
 
 lemma constrains_Un:
-    "[| F \<in> A co A'; F \<in> B co B' |] ==> F \<in> (A Un B) co (A' Un B')"
+    "[| F \<in> A co A'; F \<in> B co B' |] ==> F \<in> (A \<union> B) co (A' \<union> B')"
 by (auto simp add: constrains_def st_set_def, force)
 
 lemma constrains_UN:
@@ -414,7 +414,7 @@
 by (force simp add: constrains_def st_set_def) 
 
 lemma constrains_Un_distrib:
-     "(A Un B) co C = (A co C) \<inter> (B co C)"
+     "(A \<union> B) co C = (A co C) \<inter> (B co C)"
 by (force simp add: constrains_def st_set_def)
 
 lemma constrains_UN_distrib:
@@ -462,7 +462,7 @@
 by (unfold constrains_def st_set_def, auto, blast)
 
 lemma constrains_cancel:
-"[| F \<in> A co (A' Un B); F \<in> B co B' |] ==> F \<in> A co (A' Un B')"
+"[| F \<in> A co (A' \<union> B); F \<in> B co B' |] ==> F \<in> A co (A' \<union> B')"
 apply (drule_tac A = B in constrains_imp_subset)
 apply (blast intro: constrains_weaken_R)
 done
@@ -473,12 +473,12 @@
 lemma unless_type: "A unless B \<subseteq> program"
 by (force simp add: unless_def constrains_def) 
 
-lemma unlessI: "[| F \<in> (A-B) co (A Un B) |] ==> F \<in> A unless B"
+lemma unlessI: "[| F \<in> (A-B) co (A \<union> B) |] ==> F \<in> A unless B"
 apply (unfold unless_def)
 apply (blast dest: constrainsD2)
 done
 
-lemma unlessD: "F :A unless B ==> F \<in> (A-B) co (A Un B)"
+lemma unlessD: "F :A unless B ==> F \<in> (A-B) co (A \<union> B)"
 by (unfold unless_def, auto)
 
 
@@ -519,7 +519,7 @@
 subsection{*Union and Intersection with @{term stable}*}
 
 lemma stable_Un:
-    "[| F \<in> stable(A); F \<in> stable(A') |] ==> F \<in> stable(A Un A')"
+    "[| F \<in> stable(A); F \<in> stable(A') |] ==> F \<in> stable(A \<union> A')"
 apply (unfold stable_def)
 apply (blast intro: constrains_Un)
 done
@@ -552,7 +552,7 @@
 done
 
 lemma stable_constrains_Un:
-     "[| F \<in> stable(C); F \<in> A co (C Un A') |] ==> F \<in> (C Un A) co (C Un A')"
+     "[| F \<in> stable(C); F \<in> A co (C \<union> A') |] ==> F \<in> (C \<union> A) co (C \<union> A')"
 apply (unfold stable_def constrains_def st_set_def, auto)
 apply (blast dest!: bspec)
 done
--- a/src/ZF/UNITY/Union.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/Union.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -14,7 +14,7 @@
 begin
 
 definition
-  (*FIXME: conjoin Init(F) Int Init(G) \<noteq> 0 *) 
+  (*FIXME: conjoin Init(F) \<inter> Init(G) \<noteq> 0 *) 
   ok :: "[i, i] => o"     (infixl "ok" 65)  where
     "F ok G == Acts(F) \<subseteq> AllowedActs(G) &
                Acts(G) \<subseteq> AllowedActs(F)"
@@ -32,13 +32,13 @@
 
 definition
   Join :: "[i, i] => i"    (infixl "Join" 65)  where
-  "F Join G == mk_program (Init(F) Int Init(G), Acts(F) Un Acts(G),
-                                AllowedActs(F) Int AllowedActs(G))"
+  "F Join G == mk_program (Init(F) \<inter> Init(G), Acts(F) \<union> Acts(G),
+                                AllowedActs(F) \<inter> AllowedActs(G))"
 definition
   (*Characterizes safety properties.  Used with specifying AllowedActs*)
   safety_prop :: "i => o"  where
   "safety_prop(X) == X\<subseteq>program &
-      SKIP \<in> X & (\<forall>G \<in> program. Acts(G) \<subseteq> (\<Union>F \<in> X. Acts(F)) --> G \<in> X)"
+      SKIP \<in> X & (\<forall>G \<in> program. Acts(G) \<subseteq> (\<Union>F \<in> X. Acts(F)) \<longrightarrow> G \<in> X)"
   
 notation (xsymbols)
   SKIP  ("\<bottom>") and
@@ -64,10 +64,10 @@
 
 text{*Elimination programify from ok and Join*}
 
-lemma ok_programify_left [iff]: "programify(F) ok G <-> F ok G"
+lemma ok_programify_left [iff]: "programify(F) ok G \<longleftrightarrow> F ok G"
 by (simp add: ok_def)
 
-lemma ok_programify_right [iff]: "F ok programify(G) <-> F ok G"
+lemma ok_programify_right [iff]: "F ok programify(G) \<longleftrightarrow> F ok G"
 by (simp add: ok_def)
 
 lemma Join_programify_left [simp]: "programify(F) Join G = F Join G"
@@ -78,13 +78,13 @@
 
 subsection{*SKIP and safety properties*}
 
-lemma SKIP_in_constrains_iff [iff]: "(SKIP \<in> A co B) <-> (A\<subseteq>B & st_set(A))"
+lemma SKIP_in_constrains_iff [iff]: "(SKIP \<in> A co B) \<longleftrightarrow> (A\<subseteq>B & st_set(A))"
 by (unfold constrains_def st_set_def, auto)
 
-lemma SKIP_in_Constrains_iff [iff]: "(SKIP \<in> A Co B)<-> (state Int A\<subseteq>B)"
+lemma SKIP_in_Constrains_iff [iff]: "(SKIP \<in> A Co B)\<longleftrightarrow> (state \<inter> A\<subseteq>B)"
 by (unfold Constrains_def, auto)
 
-lemma SKIP_in_stable [iff]: "SKIP \<in> stable(A) <-> st_set(A)"
+lemma SKIP_in_stable [iff]: "SKIP \<in> stable(A) \<longleftrightarrow> st_set(A)"
 by (auto simp add: stable_def)
 
 lemma SKIP_in_Stable [iff]: "SKIP \<in> Stable(A)"
@@ -99,14 +99,14 @@
 by (unfold JOIN_def, auto)
 
 subsection{*Init, Acts, and AllowedActs of Join and JOIN*}
-lemma Init_Join [simp]: "Init(F Join G) = Init(F) Int Init(G)"
+lemma Init_Join [simp]: "Init(F Join G) = Init(F) \<inter> Init(G)"
 by (simp add: Int_assoc Join_def)
 
-lemma Acts_Join [simp]: "Acts(F Join G) = Acts(F) Un Acts(G)"
+lemma Acts_Join [simp]: "Acts(F Join G) = Acts(F) \<union> Acts(G)"
 by (simp add: Int_Un_distrib2 cons_absorb Join_def)
 
 lemma AllowedActs_Join [simp]: "AllowedActs(F Join G) =  
-  AllowedActs(F) Int AllowedActs(G)"
+  AllowedActs(F) \<inter> AllowedActs(G)"
 apply (simp add: Int_assoc cons_absorb Join_def)
 done
 
@@ -148,7 +148,7 @@
 
 subsection{*Eliminating programify form JN and OK expressions*}
 
-lemma OK_programify [iff]: "OK(I, %x. programify(F(x))) <-> OK(I, F)"
+lemma OK_programify [iff]: "OK(I, %x. programify(F(x))) \<longleftrightarrow> OK(I, F)"
 by (simp add: OK_def)
 
 lemma JN_programify [iff]: "JOIN(I, %x. programify(F(x))) = JOIN(I, F)"
@@ -196,7 +196,7 @@
 apply (auto simp add: cons_absorb)
 done
 
-lemma JN_Un: "(\<Squnion>i \<in> I Un J. F(i)) = ((\<Squnion>i \<in> I. F(i)) Join (\<Squnion>i \<in> J. F(i)))"
+lemma JN_Un: "(\<Squnion>i \<in> I \<union> J. F(i)) = ((\<Squnion>i \<in> I. F(i)) Join (\<Squnion>i \<in> J. F(i)))"
 apply (rule program_equalityI)
 apply (simp_all add: UN_Un INT_Un)
 apply (simp_all del: INT_simps add: INT_extend_simps, blast)
@@ -228,7 +228,7 @@
   I to be nonempty for other reasons anyway.*)
 
 lemma JN_constrains: 
- "i \<in> I==>(\<Squnion>i \<in> I. F(i)) \<in> A co B <-> (\<forall>i \<in> I. programify(F(i)) \<in> A co B)"
+ "i \<in> I==>(\<Squnion>i \<in> I. F(i)) \<in> A co B \<longleftrightarrow> (\<forall>i \<in> I. programify(F(i)) \<in> A co B)"
 
 apply (unfold constrains_def JOIN_def st_set_def, auto)
 prefer 2 apply blast
@@ -238,11 +238,11 @@
 done
 
 lemma Join_constrains [iff]:
-     "(F Join G \<in> A co B) <-> (programify(F) \<in> A co B & programify(G) \<in> A co B)"
+     "(F Join G \<in> A co B) \<longleftrightarrow> (programify(F) \<in> A co B & programify(G) \<in> A co B)"
 by (auto simp add: constrains_def)
 
 lemma Join_unless [iff]:
-     "(F Join G \<in> A unless B) <->  
+     "(F Join G \<in> A unless B) \<longleftrightarrow>  
     (programify(F) \<in> A unless B & programify(G) \<in> A unless B)"
 by (simp add: Join_constrains unless_def)
 
@@ -253,7 +253,7 @@
 
 lemma Join_constrains_weaken:
      "[| F \<in> A co A';  G \<in> B co B' |]  
-      ==> F Join G \<in> (A Int B) co (A' Un B')"
+      ==> F Join G \<in> (A \<inter> B) co (A' \<union> B')"
 apply (subgoal_tac "st_set (A) & st_set (B) & F \<in> program & G \<in> program")
 prefer 2 apply (blast dest: constrainsD2, simp)
 apply (blast intro: constrains_weaken)
@@ -274,7 +274,7 @@
 done
 
 lemma JN_stable:
-     "(\<Squnion>i \<in> I. F(i)) \<in>  stable(A) <-> ((\<forall>i \<in> I. programify(F(i)) \<in> stable(A)) & st_set(A))"
+     "(\<Squnion>i \<in> I. F(i)) \<in>  stable(A) \<longleftrightarrow> ((\<forall>i \<in> I. programify(F(i)) \<in> stable(A)) & st_set(A))"
 apply (auto simp add: stable_def constrains_def JOIN_def)
 apply (cut_tac F = "F (i) " in Acts_type)
 apply (drule_tac x = act in bspec, auto)
@@ -304,7 +304,7 @@
 done
 
 lemma Join_stable [iff]:
-     " (F Join G \<in> stable(A)) <->   
+     " (F Join G \<in> stable(A)) \<longleftrightarrow>   
       (programify(F) \<in> stable(A) & programify(G) \<in>  stable(A))"
 by (simp add: stable_def)
 
@@ -330,7 +330,7 @@
 
 lemma JN_transient:
      "i \<in> I ==> 
-      (\<Squnion>i \<in> I. F(i)) \<in> transient(A) <-> (\<exists>i \<in> I. programify(F(i)) \<in> transient(A))"
+      (\<Squnion>i \<in> I. F(i)) \<in> transient(A) \<longleftrightarrow> (\<exists>i \<in> I. programify(F(i)) \<in> transient(A))"
 apply (auto simp add: transient_def JOIN_def)
 apply (unfold st_set_def)
 apply (drule_tac [2] x = act in bspec)
@@ -338,7 +338,7 @@
 done
 
 lemma Join_transient [iff]:
-     "F Join G \<in> transient(A) <->  
+     "F Join G \<in> transient(A) \<longleftrightarrow>  
       (programify(F) \<in> transient(A) | programify(G) \<in> transient(A))"
 apply (auto simp add: transient_def Join_def Int_Un_distrib2)
 done
@@ -353,15 +353,15 @@
 (*If I=0 it degenerates to (SKIP \<in> A ensures B) = False, i.e. to ~(A\<subseteq>B) *)
 lemma JN_ensures:
      "i \<in> I ==>  
-      (\<Squnion>i \<in> I. F(i)) \<in> A ensures B <->  
-      ((\<forall>i \<in> I. programify(F(i)) \<in> (A-B) co (A Un B)) &   
+      (\<Squnion>i \<in> I. F(i)) \<in> A ensures B \<longleftrightarrow>  
+      ((\<forall>i \<in> I. programify(F(i)) \<in> (A-B) co (A \<union> B)) &   
       (\<exists>i \<in> I. programify(F(i)) \<in> A ensures B))"
 by (auto simp add: ensures_def JN_constrains JN_transient)
 
 
 lemma Join_ensures: 
-     "F Join G \<in> A ensures B  <->      
-      (programify(F) \<in> (A-B) co (A Un B) & programify(G) \<in> (A-B) co (A Un B) &  
+     "F Join G \<in> A ensures B  \<longleftrightarrow>      
+      (programify(F) \<in> (A-B) co (A \<union> B) & programify(G) \<in> (A-B) co (A \<union> B) &  
        (programify(F) \<in>  transient (A-B) | programify(G) \<in> transient (A-B)))"
 
 apply (unfold ensures_def)
@@ -421,37 +421,37 @@
 by (auto dest: Acts_type [THEN subsetD] simp add: ok_def)
 
 lemma ok_Join_commute:
-     "(F ok G & (F Join G) ok H) <-> (G ok H & F ok (G Join H))"
+     "(F ok G & (F Join G) ok H) \<longleftrightarrow> (G ok H & F ok (G Join H))"
 by (auto simp add: ok_def)
 
-lemma ok_commute: "(F ok G) <->(G ok F)"
+lemma ok_commute: "(F ok G) \<longleftrightarrow>(G ok F)"
 by (auto simp add: ok_def)
 
 lemmas ok_sym = ok_commute [THEN iffD1]
 
-lemma ok_iff_OK: "OK({<0,F>,<1,G>,<2,H>}, snd) <-> (F ok G & (F Join G) ok H)"
+lemma ok_iff_OK: "OK({<0,F>,<1,G>,<2,H>}, snd) \<longleftrightarrow> (F ok G & (F Join G) ok H)"
 by (simp add: ok_def Join_def OK_def Int_assoc cons_absorb
                  Int_Un_distrib2 Ball_def,  safe, force+)
 
-lemma ok_Join_iff1 [iff]: "F ok (G Join H) <-> (F ok G & F ok H)"
+lemma ok_Join_iff1 [iff]: "F ok (G Join H) \<longleftrightarrow> (F ok G & F ok H)"
 by (auto simp add: ok_def)
 
-lemma ok_Join_iff2 [iff]: "(G Join H) ok F <-> (G ok F & H ok F)"
+lemma ok_Join_iff2 [iff]: "(G Join H) ok F \<longleftrightarrow> (G ok F & H ok F)"
 by (auto simp add: ok_def)
 
 (*useful?  Not with the previous two around*)
 lemma ok_Join_commute_I: "[| F ok G; (F Join G) ok H |] ==> F ok (G Join H)"
 by (auto simp add: ok_def)
 
-lemma ok_JN_iff1 [iff]: "F ok JOIN(I,G) <-> (\<forall>i \<in> I. F ok G(i))"
+lemma ok_JN_iff1 [iff]: "F ok JOIN(I,G) \<longleftrightarrow> (\<forall>i \<in> I. F ok G(i))"
 by (force dest: Acts_type [THEN subsetD] elim!: not_emptyE simp add: ok_def)
 
-lemma ok_JN_iff2 [iff]: "JOIN(I,G) ok F   <->  (\<forall>i \<in> I. G(i) ok F)"
+lemma ok_JN_iff2 [iff]: "JOIN(I,G) ok F   \<longleftrightarrow>  (\<forall>i \<in> I. G(i) ok F)"
 apply (auto elim!: not_emptyE simp add: ok_def)
 apply (blast dest: Acts_type [THEN subsetD])
 done
 
-lemma OK_iff_ok: "OK(I,F) <-> (\<forall>i \<in> I. \<forall>j \<in> I-{i}. F(i) ok (F(j)))"
+lemma OK_iff_ok: "OK(I,F) \<longleftrightarrow> (\<forall>i \<in> I. \<forall>j \<in> I-{i}. F(i) ok (F(j)))"
 by (auto simp add: ok_def OK_def)
 
 lemma OK_imp_ok: "[| OK(I,F); i \<in> I; j \<in> I; i\<noteq>j|] ==> F(i) ok F(j)"
@@ -462,7 +462,7 @@
 by (simp add: OK_def)
 
 lemma OK_cons_iff:
-     "OK(cons(i, I), F) <->  
+     "OK(cons(i, I), F) \<longleftrightarrow>  
       (i \<in> I & OK(I, F)) | (i\<notin>I & OK(I, F) & F(i) ok JOIN(I,F))"
 apply (simp add: OK_iff_ok)
 apply (blast intro: ok_sym) 
@@ -476,7 +476,7 @@
 
 lemma Allowed_Join [simp]:
      "Allowed(F Join G) =  
-   Allowed(programify(F)) Int Allowed(programify(G))"
+   Allowed(programify(F)) \<inter> Allowed(programify(G))"
 apply (auto simp add: Allowed_def)
 done
 
@@ -487,13 +487,13 @@
 done
 
 lemma ok_iff_Allowed:
-     "F ok G <-> (programify(F) \<in> Allowed(programify(G)) &  
+     "F ok G \<longleftrightarrow> (programify(F) \<in> Allowed(programify(G)) &  
    programify(G) \<in> Allowed(programify(F)))"
 by (simp add: ok_def Allowed_def)
 
 
 lemma OK_iff_Allowed:
-     "OK(I,F) <->  
+     "OK(I,F) \<longleftrightarrow>  
   (\<forall>i \<in> I. \<forall>j \<in> I-{i}. programify(F(i)) \<in> Allowed(programify(F(j))))"
 apply (auto simp add: OK_iff_ok ok_iff_Allowed)
 done
@@ -501,7 +501,7 @@
 subsection{*safety_prop, for reasoning about given instances of "ok"*}
 
 lemma safety_prop_Acts_iff:
-     "safety_prop(X) ==> (Acts(G) \<subseteq> cons(id(state), (\<Union>F \<in> X. Acts(F)))) <-> (programify(G) \<in> X)"
+     "safety_prop(X) ==> (Acts(G) \<subseteq> cons(id(state), (\<Union>F \<in> X. Acts(F)))) \<longleftrightarrow> (programify(G) \<in> X)"
 apply (simp (no_asm_use) add: safety_prop_def)
 apply clarify
 apply (case_tac "G \<in> program", simp_all, blast, safe)
@@ -511,7 +511,7 @@
 
 lemma safety_prop_AllowedActs_iff_Allowed:
      "safety_prop(X) ==>  
-  (\<Union>G \<in> X. Acts(G)) \<subseteq> AllowedActs(F) <-> (X \<subseteq> Allowed(programify(F)))"
+  (\<Union>G \<in> X. Acts(G)) \<subseteq> AllowedActs(F) \<longleftrightarrow> (X \<subseteq> Allowed(programify(F)))"
 apply (simp add: Allowed_def safety_prop_Acts_iff [THEN iff_sym] 
                  safety_prop_def, blast) 
 done
@@ -519,7 +519,7 @@
 
 lemma Allowed_eq:
      "safety_prop(X) ==> Allowed(mk_program(init, acts, \<Union>F \<in> X. Acts(F))) = X"
-apply (subgoal_tac "cons (id (state), Union (RepFun (X, Acts)) Int Pow (state * state)) = Union (RepFun (X, Acts))")
+apply (subgoal_tac "cons (id (state), \<Union>(RepFun (X, Acts)) \<inter> Pow (state * state)) = \<Union>(RepFun (X, Acts))")
 apply (rule_tac [2] equalityI)
   apply (simp del: UN_simps add: Allowed_def safety_prop_Acts_iff safety_prop_def, auto)
 apply (force dest: Acts_type [THEN subsetD] simp add: safety_prop_def)+
@@ -532,7 +532,7 @@
 
 (*For safety_prop to hold, the property must be satisfiable!*)
 lemma safety_prop_constrains [iff]:
-     "safety_prop(A co B) <-> (A \<subseteq> B & st_set(A))"
+     "safety_prop(A co B) \<longleftrightarrow> (A \<subseteq> B & st_set(A))"
 by (simp add: safety_prop_def constrains_def st_set_def, blast)
 
 (* To be used with resolution *)
@@ -540,17 +540,17 @@
      "[| A\<subseteq>B; st_set(A) |] ==>safety_prop(A co B)"
 by auto
 
-lemma safety_prop_stable [iff]: "safety_prop(stable(A)) <-> st_set(A)"
+lemma safety_prop_stable [iff]: "safety_prop(stable(A)) \<longleftrightarrow> st_set(A)"
 by (simp add: stable_def)
 
 lemma safety_prop_stableI: "st_set(A) ==> safety_prop(stable(A))"
 by auto
 
 lemma safety_prop_Int [simp]:
-     "[| safety_prop(X) ; safety_prop(Y) |] ==> safety_prop(X Int Y)"
+     "[| safety_prop(X) ; safety_prop(Y) |] ==> safety_prop(X \<inter> Y)"
 apply (simp add: safety_prop_def, safe, blast)
-apply (drule_tac [2] B = "Union (RepFun (X Int Y, Acts))" and C = "Union (RepFun (Y, Acts))" in subset_trans)
-apply (drule_tac B = "Union (RepFun (X Int Y, Acts))" and C = "Union (RepFun (X, Acts))" in subset_trans)
+apply (drule_tac [2] B = "\<Union>(RepFun (X \<inter> Y, Acts))" and C = "\<Union>(RepFun (Y, Acts))" in subset_trans)
+apply (drule_tac B = "\<Union>(RepFun (X \<inter> Y, Acts))" and C = "\<Union>(RepFun (X, Acts))" in subset_trans)
 apply blast+
 done
 
@@ -567,13 +567,13 @@
 apply (drule_tac [2] i = xa in major)
 apply (frule_tac [4] i = xa in major)
 apply (auto simp add: safety_prop_def)
-apply (drule_tac B = "Union (RepFun (Inter (RepFun (I, X)), Acts))" and C = "Union (RepFun (X (xa), Acts))" in subset_trans)
+apply (drule_tac B = "\<Union>(RepFun (\<Inter>(RepFun (I, X)), Acts))" and C = "\<Union>(RepFun (X (xa), Acts))" in subset_trans)
 apply blast+
 done
 
 lemma def_UNION_ok_iff: 
 "[| F == mk_program(init,acts, \<Union>G \<in> X. Acts(G)); safety_prop(X) |]  
-      ==> F ok G <-> (programify(G) \<in> X & acts Int Pow(state*state) \<subseteq> AllowedActs(G))"
+      ==> F ok G \<longleftrightarrow> (programify(G) \<in> X & acts \<inter> Pow(state*state) \<subseteq> AllowedActs(G))"
 apply (unfold ok_def)
 apply (drule_tac G = G in safety_prop_Acts_iff)
 apply (cut_tac F = G in AllowedActs_type)
--- a/src/ZF/UNITY/WFair.thy	Tue Mar 06 16:46:27 2012 +0000
+++ b/src/ZF/UNITY/WFair.thy	Tue Mar 06 17:01:37 2012 +0000
@@ -17,12 +17,12 @@
   (* This definition specifies weak fairness.  The rest of the theory
     is generic to all forms of fairness.*) 
   transient :: "i=>i"  where
-  "transient(A) =={F:program. (EX act: Acts(F). A<=domain(act) &
-                               act``A <= state-A) & st_set(A)}"
+  "transient(A) =={F:program. (\<exists>act\<in>Acts(F). A<=domain(act) &
+                               act``A \<subseteq> state-A) & st_set(A)}"
 
 definition
   ensures :: "[i,i] => i"       (infixl "ensures" 60)  where
-  "A ensures B == ((A-B) co (A Un B)) Int transient(A-B)"
+  "A ensures B == ((A-B) co (A \<union> B)) \<inter> transient(A-B)"
   
 consts
 
@@ -31,12 +31,12 @@
 
 inductive 
   domains
-     "leads(D, F)" <= "Pow(D)*Pow(D)"
+     "leads(D, F)" \<subseteq> "Pow(D)*Pow(D)"
   intros 
     Basis:  "[| F:A ensures B;  A:Pow(D); B:Pow(D) |] ==> <A,B>:leads(D, F)"
-    Trans:  "[| <A,B> : leads(D, F); <B,C> : leads(D, F) |] ==>  <A,C>:leads(D, F)"
+    Trans:  "[| <A,B> \<in> leads(D, F); <B,C> \<in> leads(D, F) |] ==>  <A,C>:leads(D, F)"
     Union:   "[| S:Pow({A:S. <A, B>:leads(D, F)}); B:Pow(D); S:Pow(Pow(D)) |] ==> 
-              <Union(S),B>:leads(D, F)"
+              <\<Union>(S),B>:leads(D, F)"
 
   monos        Pow_mono
   type_intros  Union_Pow_iff [THEN iffD2] UnionI PowI
@@ -49,17 +49,17 @@
 definition
   (* wlt(F, B) is the largest set that leads to B*)
   wlt :: "[i, i] => i"  where
-    "wlt(F, B) == Union({A:Pow(state). F: A leadsTo B})"
+    "wlt(F, B) == \<Union>({A:Pow(state). F: A leadsTo B})"
 
 notation (xsymbols)
   leadsTo  (infixl "\<longmapsto>" 60)
 
 (** Ad-hoc set-theory rules **)
 
-lemma Int_Union_Union: "Union(B) Int A = (\<Union>b \<in> B. b Int A)"
+lemma Int_Union_Union: "\<Union>(B) \<inter> A = (\<Union>b \<in> B. b \<inter> A)"
 by auto
 
-lemma Int_Union_Union2: "A Int Union(B) = (\<Union>b \<in> B. A Int b)"
+lemma Int_Union_Union2: "A \<inter> \<Union>(B) = (\<Union>b \<in> B. A \<inter> b)"
 by auto
 
 (*** transient ***)
@@ -81,13 +81,13 @@
 done
 
 lemma transientI: 
-"[|act \<in> Acts(F); A <= domain(act); act``A <= state-A;  
+"[|act \<in> Acts(F); A \<subseteq> domain(act); act``A \<subseteq> state-A;  
     F \<in> program; st_set(A)|] ==> F \<in> transient(A)"
 by (simp add: transient_def, blast)
 
 lemma transientE: 
      "[| F \<in> transient(A);  
-         !!act. [| act \<in> Acts(F);  A <= domain(act); act``A <= state-A|]==>P|]
+         !!act. [| act \<in> Acts(F);  A \<subseteq> domain(act); act``A \<subseteq> state-A|]==>P|]
       ==>P"
 by (simp add: transient_def, blast)
 
@@ -118,19 +118,19 @@
 by (simp add: ensures_def constrains_def, auto)
 
 lemma ensuresI: 
-"[|F:(A-B) co (A Un B); F \<in> transient(A-B)|]==>F \<in> A ensures B"
+"[|F:(A-B) co (A \<union> B); F \<in> transient(A-B)|]==>F \<in> A ensures B"
 apply (unfold ensures_def)
 apply (auto simp add: transient_type [THEN subsetD])
 done
 
 (* Added by Sidi, from Misra's notes, Progress chapter, exercise 4 *)
-lemma ensuresI2: "[| F \<in> A co A Un B; F \<in> transient(A) |] ==> F \<in> A ensures B"
+lemma ensuresI2: "[| F \<in> A co A \<union> B; F \<in> transient(A) |] ==> F \<in> A ensures B"
 apply (drule_tac B = "A-B" in constrains_weaken_L)
 apply (drule_tac [2] B = "A-B" in transient_strengthen)
 apply (auto simp add: ensures_def transient_type [THEN subsetD])
 done
 
-lemma ensuresD: "F \<in> A ensures B ==> F:(A-B) co (A Un B) & F \<in> transient (A-B)"
+lemma ensuresD: "F \<in> A ensures B ==> F:(A-B) co (A \<union> B) & F \<in> transient (A-B)"
 by (unfold ensures_def, auto)
 
 lemma ensures_weaken_R: "[|F \<in> A ensures A'; A'<=B' |] ==> F \<in> A ensures B'"
@@ -140,20 +140,20 @@
 
 (*The L-version (precondition strengthening) fails, but we have this*) 
 lemma stable_ensures_Int: 
-     "[| F \<in> stable(C);  F \<in> A ensures B |] ==> F:(C Int A) ensures (C Int B)"
+     "[| F \<in> stable(C);  F \<in> A ensures B |] ==> F:(C \<inter> A) ensures (C \<inter> B)"
  
 apply (unfold ensures_def)
 apply (simp (no_asm) add: Int_Un_distrib [symmetric] Diff_Int_distrib [symmetric])
 apply (blast intro: transient_strengthen stable_constrains_Int constrains_weaken)
 done
 
-lemma stable_transient_ensures: "[|F \<in> stable(A);  F \<in> transient(C); A<=B Un C|] ==> F \<in> A ensures B"
+lemma stable_transient_ensures: "[|F \<in> stable(A);  F \<in> transient(C); A<=B \<union> C|] ==> F \<in> A ensures B"
 apply (frule stable_type [THEN subsetD])
 apply (simp add: ensures_def stable_def)
 apply (blast intro: transient_strengthen constrains_weaken)
 done
 
-lemma ensures_eq: "(A ensures B) = (A unless B) Int transient (A-B)"
+lemma ensures_eq: "(A ensures B) = (A unless B) \<inter> transient (A-B)"
 by (auto simp add: ensures_def unless_def)
 
 lemma subset_imp_ensures: "[| F \<in> program; A<=B  |] ==> F \<in> A ensures B"
@@ -163,7 +163,7 @@
 lemmas leads_left = leads.dom_subset [THEN subsetD, THEN SigmaD1]
 lemmas leads_right = leads.dom_subset [THEN subsetD, THEN SigmaD2]
 
-lemma leadsTo_type: "A leadsTo B <= program"
+lemma leadsTo_type: "A leadsTo B \<subseteq> program"
 by (unfold leadsTo_def, auto)
 
 lemma leadsToD2: 
@@ -196,24 +196,24 @@
 done
 
 (*Useful with cancellation, disjunction*)
-lemma leadsTo_Un_duplicate: "F \<in> A leadsTo (A' Un A') ==> F \<in> A leadsTo A'"
+lemma leadsTo_Un_duplicate: "F \<in> A leadsTo (A' \<union> A') ==> F \<in> A leadsTo A'"
 by simp
 
 lemma leadsTo_Un_duplicate2:
-     "F \<in> A leadsTo (A' Un C Un C) ==> F \<in> A leadsTo (A' Un C)"
+     "F \<in> A leadsTo (A' \<union> C \<union> C) ==> F \<in> A leadsTo (A' \<union> C)"
 by (simp add: Un_ac)
 
 (*The Union introduction rule as we should have liked to state it*)
 lemma leadsTo_Union: 
     "[|!!A. A \<in> S ==> F \<in> A leadsTo B; F \<in> program; st_set(B)|]
-     ==> F \<in> Union(S) leadsTo B"
+     ==> F \<in> \<Union>(S) leadsTo B"
 apply (unfold leadsTo_def st_set_def)
 apply (blast intro: leads.Union dest: leads_left)
 done
 
 lemma leadsTo_Union_Int: 
-    "[|!!A. A \<in> S ==>F : (A Int C) leadsTo B; F \<in> program; st_set(B)|]  
-     ==> F : (Union(S)Int C)leadsTo B"
+    "[|!!A. A \<in> S ==>F \<in> (A \<inter> C) leadsTo B; F \<in> program; st_set(B)|]  
+     ==> F \<in> (\<Union>(S)Int C)leadsTo B"
 apply (unfold leadsTo_def st_set_def)
 apply (simp only: Int_Union_Union)
 apply (blast dest: leads_left intro: leads.Union)
@@ -228,7 +228,7 @@
 
 (* Binary union introduction rule *)
 lemma leadsTo_Un:
-     "[| F \<in> A leadsTo C; F \<in> B leadsTo C |] ==> F \<in> (A Un B) leadsTo C"
+     "[| F \<in> A leadsTo C; F \<in> B leadsTo C |] ==> F \<in> (A \<union> B) leadsTo C"
 apply (subst Un_eq_Union)
 apply (blast intro: leadsTo_Union dest: leadsToD2)
 done
@@ -243,14 +243,14 @@
 lemma leadsTo_refl: "[| F \<in> program; st_set(A) |] ==> F \<in> A leadsTo A"
 by (blast intro: subset_imp_leadsTo)
 
-lemma leadsTo_refl_iff: "F \<in> A leadsTo A <-> F \<in> program & st_set(A)"
+lemma leadsTo_refl_iff: "F \<in> A leadsTo A \<longleftrightarrow> F \<in> program & st_set(A)"
 by (auto intro: leadsTo_refl dest: leadsToD2)
 
-lemma empty_leadsTo: "F \<in> 0 leadsTo B <-> (F \<in> program & st_set(B))"
+lemma empty_leadsTo: "F \<in> 0 leadsTo B \<longleftrightarrow> (F \<in> program & st_set(B))"
 by (auto intro: subset_imp_leadsTo dest: leadsToD2)
 declare empty_leadsTo [iff]
 
-lemma leadsTo_state: "F \<in> A leadsTo state <-> (F \<in> program & st_set(A))"
+lemma leadsTo_state: "F \<in> A leadsTo state \<longleftrightarrow> (F \<in> program & st_set(A))"
 by (auto intro: subset_imp_leadsTo dest: leadsToD2 st_setD)
 declare leadsTo_state [iff]
 
@@ -275,15 +275,15 @@
 done
 
 (*Distributes over binary unions*)
-lemma leadsTo_Un_distrib: "F:(A Un B) leadsTo C  <->  (F \<in> A leadsTo C & F \<in> B leadsTo C)"
+lemma leadsTo_Un_distrib: "F:(A \<union> B) leadsTo C  \<longleftrightarrow>  (F \<in> A leadsTo C & F \<in> B leadsTo C)"
 by (blast intro: leadsTo_Un leadsTo_weaken_L)
 
 lemma leadsTo_UN_distrib: 
-"(F:(\<Union>i \<in> I. A(i)) leadsTo B)<-> ((\<forall>i \<in> I. F \<in> A(i) leadsTo B) & F \<in> program & st_set(B))"
+"(F:(\<Union>i \<in> I. A(i)) leadsTo B)\<longleftrightarrow> ((\<forall>i \<in> I. F \<in> A(i) leadsTo B) & F \<in> program & st_set(B))"
 apply (blast dest: leadsToD2 intro: leadsTo_UN leadsTo_weaken_L)
 done
 
-lemma leadsTo_Union_distrib: "(F \<in> Union(S) leadsTo B) <->  (\<forall>A \<in> S. F \<in> A leadsTo B) & F \<in> program & st_set(B)"
+lemma leadsTo_Union_distrib: "(F \<in> \<Union>(S) leadsTo B) \<longleftrightarrow>  (\<forall>A \<in> S. F \<in> A leadsTo B) & F \<in> program & st_set(B)"
 by (blast dest: leadsToD2 intro: leadsTo_Union leadsTo_weaken_L)
 
 text{*Set difference: maybe combine with @{text leadsTo_weaken_L}??*}
@@ -300,33 +300,33 @@
 done
 
 (*Binary union version*)
-lemma leadsTo_Un_Un: "[| F \<in> A leadsTo A'; F \<in> B leadsTo B' |] ==> F \<in> (A Un B) leadsTo (A' Un B')"
+lemma leadsTo_Un_Un: "[| F \<in> A leadsTo A'; F \<in> B leadsTo B' |] ==> F \<in> (A \<union> B) leadsTo (A' \<union> B')"
 apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') ")
 prefer 2 apply (blast dest: leadsToD2)
 apply (blast intro: leadsTo_Un leadsTo_weaken_R)
 done
 
 (** The cancellation law **)
-lemma leadsTo_cancel2: "[|F \<in> A leadsTo (A' Un B); F \<in> B leadsTo B'|] ==> F \<in> A leadsTo (A' Un B')"
+lemma leadsTo_cancel2: "[|F \<in> A leadsTo (A' \<union> B); F \<in> B leadsTo B'|] ==> F \<in> A leadsTo (A' \<union> B')"
 apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') &F \<in> program")
 prefer 2 apply (blast dest: leadsToD2)
 apply (blast intro: leadsTo_Trans leadsTo_Un_Un leadsTo_refl)
 done
 
-lemma leadsTo_cancel_Diff2: "[|F \<in> A leadsTo (A' Un B); F \<in> (B-A') leadsTo B'|]==> F \<in> A leadsTo (A' Un B')"
+lemma leadsTo_cancel_Diff2: "[|F \<in> A leadsTo (A' \<union> B); F \<in> (B-A') leadsTo B'|]==> F \<in> A leadsTo (A' \<union> B')"
 apply (rule leadsTo_cancel2)
 prefer 2 apply assumption
 apply (blast dest: leadsToD2 intro: leadsTo_weaken_R)
 done
 
 
-lemma leadsTo_cancel1: "[| F \<in> A leadsTo (B Un A'); F \<in> B leadsTo B' |] ==> F \<in> A leadsTo (B' Un A')"
+lemma leadsTo_cancel1: "[| F \<in> A leadsTo (B \<union> A'); F \<in> B leadsTo B' |] ==> F \<in> A leadsTo (B' \<union> A')"
 apply (simp add: Un_commute)
 apply (blast intro!: leadsTo_cancel2)
 done
 
 lemma leadsTo_cancel_Diff1:
-     "[|F \<in> A leadsTo (B Un A'); F: (B-A') leadsTo B'|]==> F \<in> A leadsTo (B' Un A')"
+     "[|F \<in> A leadsTo (B \<union> A'); F: (B-A') leadsTo B'|]==> F \<in> A leadsTo (B' \<union> A')"
 apply (rule leadsTo_cancel1)
 prefer 2 apply assumption
 apply (blast intro: leadsTo_weaken_R dest: leadsToD2)
@@ -339,7 +339,7 @@
       and trans: "!!A B C. [| F \<in> A leadsTo B; P(A, B);  
                               F \<in> B leadsTo C; P(B, C) |] ==> P(A,C)"
       and union: "!!B S. [| \<forall>A \<in> S. F \<in> A leadsTo B; \<forall>A \<in> S. P(A,B); 
-                           st_set(B); \<forall>A \<in> S. st_set(A)|] ==> P(Union(S), B)"
+                           st_set(B); \<forall>A \<in> S. st_set(A)|] ==> P(\<Union>(S), B)"
   shows "P(za, zb)"
 apply (cut_tac major)
 apply (unfold leadsTo_def, clarify) 
@@ -354,24 +354,24 @@
 lemma leadsTo_induct2:
   assumes major: "F \<in> za leadsTo zb"
       and basis1: "!!A B. [| A<=B; st_set(B) |] ==> P(A, B)"
-      and basis2: "!!A B. [| F \<in> A co A Un B; F \<in> transient(A); st_set(B) |] 
+      and basis2: "!!A B. [| F \<in> A co A \<union> B; F \<in> transient(A); st_set(B) |] 
                           ==> P(A, B)"
       and trans: "!!A B C. [| F \<in> A leadsTo B; P(A, B);  
                               F \<in> B leadsTo C; P(B, C) |] ==> P(A,C)"
       and union: "!!B S. [| \<forall>A \<in> S. F \<in> A leadsTo B; \<forall>A \<in> S. P(A,B); 
-                           st_set(B); \<forall>A \<in> S. st_set(A)|] ==> P(Union(S), B)"
+                           st_set(B); \<forall>A \<in> S. st_set(A)|] ==> P(\<Union>(S), B)"
   shows "P(za, zb)"
 apply (cut_tac major)
 apply (erule leadsTo_induct)
 apply (auto intro: trans union)
 apply (simp add: ensures_def, clarify)
 apply (frule constrainsD2)
-apply (drule_tac B' = " (A-B) Un B" in constrains_weaken_R)
+apply (drule_tac B' = " (A-B) \<union> B" in constrains_weaken_R)
 apply blast
 apply (frule ensuresI2 [THEN leadsTo_Basis])
 apply (drule_tac [4] basis2, simp_all)
 apply (frule_tac A1 = A and B = B in Int_lower2 [THEN basis1])
-apply (subgoal_tac "A=Union ({A - B, A Int B}) ")
+apply (subgoal_tac "A=\<Union>({A - B, A \<inter> B}) ")
 prefer 2 apply blast
 apply (erule ssubst)
 apply (rule union)
@@ -385,10 +385,10 @@
   "[| F \<in> za leadsTo zb;   
       P(zb);  
       !!A B. [| F \<in> A ensures B;  P(B); st_set(A); st_set(B) |] ==> P(A);  
-      !!S. [| \<forall>A \<in> S. P(A); \<forall>A \<in> S. st_set(A) |] ==> P(Union(S))  
+      !!S. [| \<forall>A \<in> S. P(A); \<forall>A \<in> S. st_set(A) |] ==> P(\<Union>(S))  
    |] ==> P(za)"
 txt{*by induction on this formula*}
-apply (subgoal_tac "P (zb) --> P (za) ")
+apply (subgoal_tac "P (zb) \<longrightarrow> P (za) ")
 txt{*now solve first subgoal: this formula is sufficient*}
 apply (blast intro: leadsTo_refl)
 apply (erule leadsTo_induct)
@@ -400,7 +400,7 @@
   "[| F \<in> za leadsTo zb;   
       P(zb);  
       !!A B. [| F \<in> A ensures B;  F \<in> B leadsTo zb;  P(B); st_set(A) |] ==> P(A);  
-      !!S. \<forall>A \<in> S. F \<in> A leadsTo zb & P(A) & st_set(A) ==> P(Union(S))  
+      !!S. \<forall>A \<in> S. F \<in> A leadsTo zb & P(A) & st_set(A) ==> P(\<Union>(S))  
    |] ==> P(za)"
 apply (subgoal_tac " (F \<in> za leadsTo zb) & P (za) ")
 apply (erule conjunct2)
@@ -426,7 +426,7 @@
 text{*Special case of PSP: Misra's "stable conjunction"*}
 
 lemma psp_stable: 
-   "[| F \<in> A leadsTo A'; F \<in> stable(B) |] ==> F:(A Int B) leadsTo (A' Int B)"
+   "[| F \<in> A leadsTo A'; F \<in> stable(B) |] ==> F:(A \<inter> B) leadsTo (A' \<inter> B)"
 apply (unfold stable_def)
 apply (frule leadsToD2) 
 apply (erule leadsTo_induct)
@@ -438,12 +438,12 @@
 done
 
 
-lemma psp_stable2: "[|F \<in> A leadsTo A'; F \<in> stable(B) |]==>F: (B Int A) leadsTo (B Int A')"
+lemma psp_stable2: "[|F \<in> A leadsTo A'; F \<in> stable(B) |]==>F: (B \<inter> A) leadsTo (B \<inter> A')"
 apply (simp (no_asm_simp) add: psp_stable Int_ac)
 done
 
 lemma psp_ensures: 
-"[| F \<in> A ensures A'; F \<in> B co B' |]==> F: (A Int B') ensures ((A' Int B) Un (B' - B))"
+"[| F \<in> A ensures A'; F \<in> B co B' |]==> F: (A \<inter> B') ensures ((A' \<inter> B) \<union> (B' - B))"
 apply (unfold ensures_def constrains_def st_set_def)
 (*speeds up the proof*)
 apply clarify
@@ -451,7 +451,7 @@
 done
 
 lemma psp: 
-"[|F \<in> A leadsTo A'; F \<in> B co B'; st_set(B')|]==> F:(A Int B') leadsTo ((A' Int B) Un (B' - B))"
+"[|F \<in> A leadsTo A'; F \<in> B co B'; st_set(B')|]==> F:(A \<inter> B') leadsTo ((A' \<inter> B) \<union> (B' - B))"
 apply (subgoal_tac "F \<in> program & st_set (A) & st_set (A') & st_set (B) ")
 prefer 2 apply (blast dest!: constrainsD2 leadsToD2)
 apply (erule leadsTo_induct)
@@ -467,12 +467,12 @@
 
 
 lemma psp2: "[| F \<in> A leadsTo A'; F \<in> B co B'; st_set(B') |]  
-    ==> F \<in> (B' Int A) leadsTo ((B Int A') Un (B' - B))"
+    ==> F \<in> (B' \<inter> A) leadsTo ((B \<inter> A') \<union> (B' - B))"
 by (simp (no_asm_simp) add: psp Int_ac)
 
 lemma psp_unless: 
    "[| F \<in> A leadsTo A';  F \<in> B unless B'; st_set(B); st_set(B') |]  
-    ==> F \<in> (A Int B) leadsTo ((A' Int B) Un B')"
+    ==> F \<in> (A \<inter> B) leadsTo ((A' \<inter> B) \<union> B')"
 apply (unfold unless_def)
 apply (subgoal_tac "st_set (A) &st_set (A') ")
 prefer 2 apply (blast dest: leadsToD2)
@@ -488,11 +488,11 @@
          m \<in> I;  
          field(r)<=I;  
          F \<in> program; st_set(B); 
-         \<forall>m \<in> I. F \<in> (A Int f-``{m}) leadsTo                      
-                    ((A Int f-``(converse(r)``{m})) Un B) |]  
-      ==> F \<in> (A Int f-``{m}) leadsTo B"
+         \<forall>m \<in> I. F \<in> (A \<inter> f-``{m}) leadsTo                      
+                    ((A \<inter> f-``(converse(r)``{m})) \<union> B) |]  
+      ==> F \<in> (A \<inter> f-``{m}) leadsTo B"
 apply (erule_tac a = m in wf_induct2, simp_all)
-apply (subgoal_tac "F \<in> (A Int (f-`` (converse (r) ``{x}))) leadsTo B")
+apply (subgoal_tac "F \<in> (A \<inter> (f-`` (converse (r) ``{x}))) leadsTo B")
  apply (blast intro: leadsTo_cancel1 leadsTo_Un_duplicate)
 apply (subst vimage_eq_UN)
 apply (simp del: UN_simps add: Int_UN_distrib)
@@ -504,8 +504,8 @@
          field(r)<=I;  
          A<=f-``I;  
          F \<in> program; st_set(A); st_set(B);  
-         \<forall>m \<in> I. F \<in> (A Int f-``{m}) leadsTo                      
-                    ((A Int f-``(converse(r)``{m})) Un B) |]  
+         \<forall>m \<in> I. F \<in> (A \<inter> f-``{m}) leadsTo                      
+                    ((A \<inter> f-``(converse(r)``{m})) \<union> B) |]  
       ==> F \<in> A leadsTo B"
 apply (rule_tac b = A in subst)
  defer 1
@@ -535,11 +535,11 @@
 apply (blast intro: lt_trans)
 done
 
-(*Alternative proof is via the lemma F \<in> (A Int f-`(lessThan m)) leadsTo B*)
+(*Alternative proof is via the lemma F \<in> (A \<inter> f-`(lessThan m)) leadsTo B*)
 lemma lessThan_induct: 
  "[| A<=f-``nat;  
      F \<in> program; st_set(A); st_set(B);  
-     \<forall>m \<in> nat. F:(A Int f-``{m}) leadsTo ((A Int f -`` m) Un B) |]  
+     \<forall>m \<in> nat. F:(A \<inter> f-``{m}) leadsTo ((A \<inter> f -`` m) \<union> B) |]  
       ==> F \<in> A leadsTo B"
 apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN leadsTo_wf_induct]) 
 apply (simp_all add: nat_measure_field)
@@ -559,7 +559,7 @@
 done
 declare wlt_st_set [iff]
 
-lemma wlt_leadsTo_iff: "F \<in> wlt(F, B) leadsTo B <-> (F \<in> program & st_set(B))"
+lemma wlt_leadsTo_iff: "F \<in> wlt(F, B) leadsTo B \<longleftrightarrow> (F \<in> program & st_set(B))"
 apply (unfold wlt_def)
 apply (blast dest: leadsToD2 intro!: leadsTo_Union)
 done
@@ -567,48 +567,48 @@
 (* [| F \<in> program;  st_set(B) |] ==> F \<in> wlt(F, B) leadsTo B  *)
 lemmas wlt_leadsTo = conjI [THEN wlt_leadsTo_iff [THEN iffD2]]
 
-lemma leadsTo_subset: "F \<in> A leadsTo B ==> A <= wlt(F, B)"
+lemma leadsTo_subset: "F \<in> A leadsTo B ==> A \<subseteq> wlt(F, B)"
 apply (unfold wlt_def)
 apply (frule leadsToD2)
 apply (auto simp add: st_set_def)
 done
 
 (*Misra's property W2*)
-lemma leadsTo_eq_subset_wlt: "F \<in> A leadsTo B <-> (A <= wlt(F,B) & F \<in> program & st_set(B))"
+lemma leadsTo_eq_subset_wlt: "F \<in> A leadsTo B \<longleftrightarrow> (A \<subseteq> wlt(F,B) & F \<in> program & st_set(B))"
 apply auto
 apply (blast dest: leadsToD2 leadsTo_subset intro: leadsTo_weaken_L wlt_leadsTo)+
 done
 
 (*Misra's property W4*)
-lemma wlt_increasing: "[| F \<in> program; st_set(B) |] ==> B <= wlt(F,B)"
+lemma wlt_increasing: "[| F \<in> program; st_set(B) |] ==> B \<subseteq> wlt(F,B)"
 apply (rule leadsTo_subset)
 apply (simp (no_asm_simp) add: leadsTo_eq_subset_wlt [THEN iff_sym] subset_imp_leadsTo)
 done
 
 (*Used in the Trans case below*)
 lemma leadsTo_123_aux: 
-   "[| B <= A2;  
-       F \<in> (A1 - B) co (A1 Un B);  
-       F \<in> (A2 - C) co (A2 Un C) |]  
-    ==> F \<in> (A1 Un A2 - C) co (A1 Un A2 Un C)"
+   "[| B \<subseteq> A2;  
+       F \<in> (A1 - B) co (A1 \<union> B);  
+       F \<in> (A2 - C) co (A2 \<union> C) |]  
+    ==> F \<in> (A1 \<union> A2 - C) co (A1 \<union> A2 \<union> C)"
 apply (unfold constrains_def st_set_def, blast)
 done
 
 (*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*)
 (* slightly different from the HOL one \<in> B here is bounded *)
 lemma leadsTo_123: "F \<in> A leadsTo A'  
-      ==> \<exists>B \<in> Pow(state). A<=B & F \<in> B leadsTo A' & F \<in> (B-A') co (B Un A')"
+      ==> \<exists>B \<in> Pow(state). A<=B & F \<in> B leadsTo A' & F \<in> (B-A') co (B \<union> A')"
 apply (frule leadsToD2)
 apply (erule leadsTo_induct)
   txt{*Basis*}
   apply (blast dest: ensuresD constrainsD2 st_setD)
  txt{*Trans*}
  apply clarify
- apply (rule_tac x = "Ba Un Bb" in bexI)
+ apply (rule_tac x = "Ba \<union> Bb" in bexI)
  apply (blast intro: leadsTo_123_aux leadsTo_Un_Un leadsTo_cancel1 leadsTo_Un_duplicate, blast)
 txt{*Union*}
 apply (clarify dest!: ball_conj_distrib [THEN iffD1])
-apply (subgoal_tac "\<exists>y. y \<in> Pi (S, %A. {Ba \<in> Pow (state) . A<=Ba & F \<in> Ba leadsTo B & F \<in> Ba - B co Ba Un B}) ")
+apply (subgoal_tac "\<exists>y. y \<in> Pi (S, %A. {Ba \<in> Pow (state) . A<=Ba & F \<in> Ba leadsTo B & F \<in> Ba - B co Ba \<union> B}) ")
 defer 1
 apply (rule AC_ball_Pi, safe)
 apply (rotate_tac 1)
@@ -633,35 +633,35 @@
 
 subsection{*Completion: Binary and General Finite versions*}
 
-lemma completion_aux: "[| W = wlt(F, (B' Un C));      
-       F \<in> A leadsTo (A' Un C);  F \<in> A' co (A' Un C);    
-       F \<in> B leadsTo (B' Un C);  F \<in> B' co (B' Un C) |]  
-    ==> F \<in> (A Int B) leadsTo ((A' Int B') Un C)"
+lemma completion_aux: "[| W = wlt(F, (B' \<union> C));      
+       F \<in> A leadsTo (A' \<union> C);  F \<in> A' co (A' \<union> C);    
+       F \<in> B leadsTo (B' \<union> C);  F \<in> B' co (B' \<union> C) |]  
+    ==> F \<in> (A \<inter> B) leadsTo ((A' \<inter> B') \<union> C)"
 apply (subgoal_tac "st_set (C) &st_set (W) &st_set (W-C) &st_set (A') &st_set (A) & st_set (B) & st_set (B') & F \<in> program")
  prefer 2 
  apply simp 
  apply (blast dest!: leadsToD2)
-apply (subgoal_tac "F \<in> (W-C) co (W Un B' Un C) ")
+apply (subgoal_tac "F \<in> (W-C) co (W \<union> B' \<union> C) ")
  prefer 2
  apply (blast intro!: constrains_weaken [OF constrains_Un [OF _ wlt_constrains_wlt]])
 apply (subgoal_tac "F \<in> (W-C) co W")
  prefer 2
  apply (simp add: wlt_increasing [THEN subset_Un_iff2 [THEN iffD1]] Un_assoc)
-apply (subgoal_tac "F \<in> (A Int W - C) leadsTo (A' Int W Un C) ")
+apply (subgoal_tac "F \<in> (A \<inter> W - C) leadsTo (A' \<inter> W \<union> C) ")
  prefer 2 apply (blast intro: wlt_leadsTo psp [THEN leadsTo_weaken])
 (** step 13 **)
-apply (subgoal_tac "F \<in> (A' Int W Un C) leadsTo (A' Int B' Un C) ")
+apply (subgoal_tac "F \<in> (A' \<inter> W \<union> C) leadsTo (A' \<inter> B' \<union> C) ")
 apply (drule leadsTo_Diff)
 apply (blast intro: subset_imp_leadsTo dest: leadsToD2 constrainsD2)
 apply (force simp add: st_set_def)
-apply (subgoal_tac "A Int B <= A Int W")
+apply (subgoal_tac "A \<inter> B \<subseteq> A \<inter> W")
 prefer 2 apply (blast dest!: leadsTo_subset intro!: subset_refl [THEN Int_mono])
 apply (blast intro: leadsTo_Trans subset_imp_leadsTo)
 txt{*last subgoal*}
 apply (rule_tac leadsTo_Un_duplicate2)
 apply (rule_tac leadsTo_Un_Un)
  prefer 2 apply (blast intro: leadsTo_refl)
-apply (rule_tac A'1 = "B' Un C" in wlt_leadsTo[THEN psp2, THEN leadsTo_weaken])
+apply (rule_tac A'1 = "B' \<union> C" in wlt_leadsTo[THEN psp2, THEN leadsTo_weaken])
 apply blast+
 done
 
@@ -669,9 +669,9 @@
 
 lemma finite_completion_aux:
      "[| I \<in> Fin(X); F \<in> program; st_set(C) |] ==>  
-       (\<forall>i \<in> I. F \<in> (A(i)) leadsTo (A'(i) Un C)) -->   
-                     (\<forall>i \<in> I. F \<in> (A'(i)) co (A'(i) Un C)) -->  
-                   F \<in> (\<Inter>i \<in> I. A(i)) leadsTo ((\<Inter>i \<in> I. A'(i)) Un C)"
+       (\<forall>i \<in> I. F \<in> (A(i)) leadsTo (A'(i) \<union> C)) \<longrightarrow>   
+                     (\<forall>i \<in> I. F \<in> (A'(i)) co (A'(i) \<union> C)) \<longrightarrow>  
+                   F \<in> (\<Inter>i \<in> I. A(i)) leadsTo ((\<Inter>i \<in> I. A'(i)) \<union> C)"
 apply (erule Fin_induct)
 apply (auto simp add: Inter_0)
 apply (rule completion)
@@ -681,15 +681,15 @@
 
 lemma finite_completion: 
      "[| I \<in> Fin(X);   
-         !!i. i \<in> I ==> F \<in> A(i) leadsTo (A'(i) Un C);  
-         !!i. i \<in> I ==> F \<in> A'(i) co (A'(i) Un C); F \<in> program; st_set(C)|]    
-      ==> F \<in> (\<Inter>i \<in> I. A(i)) leadsTo ((\<Inter>i \<in> I. A'(i)) Un C)"
+         !!i. i \<in> I ==> F \<in> A(i) leadsTo (A'(i) \<union> C);  
+         !!i. i \<in> I ==> F \<in> A'(i) co (A'(i) \<union> C); F \<in> program; st_set(C)|]    
+      ==> F \<in> (\<Inter>i \<in> I. A(i)) leadsTo ((\<Inter>i \<in> I. A'(i)) \<union> C)"
 by (blast intro: finite_completion_aux [THEN mp, THEN mp])
 
 lemma stable_completion: 
      "[| F \<in> A leadsTo A';  F \<in> stable(A');    
          F \<in> B leadsTo B';  F \<in> stable(B') |]  
-    ==> F \<in> (A Int B) leadsTo (A' Int B')"
+    ==> F \<in> (A \<inter> B) leadsTo (A' \<inter> B')"
 apply (unfold stable_def)
 apply (rule_tac C1 = 0 in completion [THEN leadsTo_weaken_R], simp+)
 apply (blast dest: leadsToD2)