merged
authorhuffman
Sat, 03 Sep 2011 14:33:45 -0700
changeset 44691 42a2e1a4f04f
parent 44690 b6d8b11ed399 (current diff)
parent 44682 e5ba1c0b8cac (diff)
child 44692 ccfc7c193d2b
merged
--- a/src/HOL/Algebra/Ideal.thy	Sat Sep 03 11:10:38 2011 -0700
+++ b/src/HOL/Algebra/Ideal.thy	Sat Sep 03 14:33:45 2011 -0700
@@ -14,25 +14,24 @@
 
 locale ideal = additive_subgroup I R + ring R for I and R (structure) +
   assumes I_l_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I"
-      and I_r_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
+    and I_r_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
 
 sublocale ideal \<subseteq> abelian_subgroup I R
-apply (intro abelian_subgroupI3 abelian_group.intro)
-  apply (rule ideal.axioms, rule ideal_axioms)
- apply (rule abelian_group.axioms, rule ring.axioms, rule ideal.axioms, rule ideal_axioms)
-apply (rule abelian_group.axioms, rule ring.axioms, rule ideal.axioms, rule ideal_axioms)
-done
+  apply (intro abelian_subgroupI3 abelian_group.intro)
+    apply (rule ideal.axioms, rule ideal_axioms)
+   apply (rule abelian_group.axioms, rule ring.axioms, rule ideal.axioms, rule ideal_axioms)
+  apply (rule abelian_group.axioms, rule ring.axioms, rule ideal.axioms, rule ideal_axioms)
+  done
 
-lemma (in ideal) is_ideal:
-  "ideal I R"
-by (rule ideal_axioms)
+lemma (in ideal) is_ideal: "ideal I R"
+  by (rule ideal_axioms)
 
 lemma idealI:
   fixes R (structure)
   assumes "ring R"
   assumes a_subgroup: "subgroup I \<lparr>carrier = carrier R, mult = add R, one = zero R\<rparr>"
-      and I_l_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I"
-      and I_r_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
+    and I_l_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I"
+    and I_r_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
   shows "ideal I R"
 proof -
   interpret ring R by fact
@@ -47,19 +46,16 @@
 
 subsubsection (in ring) {* Ideals Generated by a Subset of @{term "carrier R"} *}
 
-definition
-  genideal :: "('a, 'b) ring_scheme \<Rightarrow> 'a set \<Rightarrow> 'a set"  ("Idl\<index> _" [80] 79)
+definition genideal :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set"  ("Idl\<index> _" [80] 79)
   where "genideal R S = Inter {I. ideal I R \<and> S \<subseteq> I}"
 
-
 subsubsection {* Principal Ideals *}
 
 locale principalideal = ideal +
   assumes generate: "\<exists>i \<in> carrier R. I = Idl {i}"
 
-lemma (in principalideal) is_principalideal:
-  shows "principalideal I R"
-by (rule principalideal_axioms)
+lemma (in principalideal) is_principalideal: "principalideal I R"
+  by (rule principalideal_axioms)
 
 lemma principalidealI:
   fixes R (structure)
@@ -68,7 +64,9 @@
   shows "principalideal I R"
 proof -
   interpret ideal I R by fact
-  show ?thesis  by (intro principalideal.intro principalideal_axioms.intro) (rule is_ideal, rule generate)
+  show ?thesis
+    by (intro principalideal.intro principalideal_axioms.intro)
+      (rule is_ideal, rule generate)
 qed
 
 
@@ -76,22 +74,22 @@
 
 locale maximalideal = ideal +
   assumes I_notcarr: "carrier R \<noteq> I"
-      and I_maximal: "\<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> J = I \<or> J = carrier R"
+    and I_maximal: "\<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> J = I \<or> J = carrier R"
 
-lemma (in maximalideal) is_maximalideal:
- shows "maximalideal I R"
-by (rule maximalideal_axioms)
+lemma (in maximalideal) is_maximalideal: "maximalideal I R"
+  by (rule maximalideal_axioms)
 
 lemma maximalidealI:
   fixes R
   assumes "ideal I R"
   assumes I_notcarr: "carrier R \<noteq> I"
-     and I_maximal: "\<And>J. \<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> J = I \<or> J = carrier R"
+    and I_maximal: "\<And>J. \<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> J = I \<or> J = carrier R"
   shows "maximalideal I R"
 proof -
   interpret ideal I R by fact
-  show ?thesis by (intro maximalideal.intro maximalideal_axioms.intro)
-    (rule is_ideal, rule I_notcarr, rule I_maximal)
+  show ?thesis
+    by (intro maximalideal.intro maximalideal_axioms.intro)
+      (rule is_ideal, rule I_notcarr, rule I_maximal)
 qed
 
 
@@ -99,24 +97,24 @@
 
 locale primeideal = ideal + cring +
   assumes I_notcarr: "carrier R \<noteq> I"
-      and I_prime: "\<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
+    and I_prime: "\<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
 
-lemma (in primeideal) is_primeideal:
- shows "primeideal I R"
-by (rule primeideal_axioms)
+lemma (in primeideal) is_primeideal: "primeideal I R"
+  by (rule primeideal_axioms)
 
 lemma primeidealI:
   fixes R (structure)
   assumes "ideal I R"
   assumes "cring R"
   assumes I_notcarr: "carrier R \<noteq> I"
-      and I_prime: "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
+    and I_prime: "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
   shows "primeideal I R"
 proof -
   interpret ideal I R by fact
   interpret cring R by fact
-  show ?thesis by (intro primeideal.intro primeideal_axioms.intro)
-    (rule is_ideal, rule is_cring, rule I_notcarr, rule I_prime)
+  show ?thesis
+    by (intro primeideal.intro primeideal_axioms.intro)
+      (rule is_ideal, rule is_cring, rule I_notcarr, rule I_prime)
 qed
 
 lemma primeidealI2:
@@ -124,9 +122,9 @@
   assumes "additive_subgroup I R"
   assumes "cring R"
   assumes I_l_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I"
-      and I_r_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
-      and I_notcarr: "carrier R \<noteq> I"
-      and I_prime: "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
+    and I_r_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
+    and I_notcarr: "carrier R \<noteq> I"
+    and I_prime: "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
   shows "primeideal I R"
 proof -
   interpret additive_subgroup I R by fact
@@ -144,31 +142,27 @@
 
 subsection {* Special Ideals *}
 
-lemma (in ring) zeroideal:
-  shows "ideal {\<zero>} R"
-apply (intro idealI subgroup.intro)
-      apply (rule is_ring)
-     apply simp+
-  apply (fold a_inv_def, simp)
- apply simp+
-done
+lemma (in ring) zeroideal: "ideal {\<zero>} R"
+  apply (intro idealI subgroup.intro)
+        apply (rule is_ring)
+       apply simp+
+    apply (fold a_inv_def, simp)
+   apply simp+
+  done
 
-lemma (in ring) oneideal:
-  shows "ideal (carrier R) R"
+lemma (in ring) oneideal: "ideal (carrier R) R"
   by (rule idealI) (auto intro: is_ring add.subgroupI)
 
-lemma (in "domain") zeroprimeideal:
- shows "primeideal {\<zero>} R"
-apply (intro primeidealI)
-   apply (rule zeroideal)
-  apply (rule domain.axioms, rule domain_axioms)
- defer 1
- apply (simp add: integral)
+lemma (in "domain") zeroprimeideal: "primeideal {\<zero>} R"
+  apply (intro primeidealI)
+     apply (rule zeroideal)
+    apply (rule domain.axioms, rule domain_axioms)
+   defer 1
+   apply (simp add: integral)
 proof (rule ccontr, simp)
   assume "carrier R = {\<zero>}"
-  from this have "\<one> = \<zero>" by (rule one_zeroI)
-  from this and one_not_zero
-      show "False" by simp
+  then have "\<one> = \<zero>" by (rule one_zeroI)
+  with one_not_zero show False by simp
 qed
 
 
@@ -177,22 +171,20 @@
 lemma (in ideal) one_imp_carrier:
   assumes I_one_closed: "\<one> \<in> I"
   shows "I = carrier R"
-apply (rule)
-apply (rule)
-apply (rule a_Hcarr, simp)
+  apply (rule)
+  apply (rule)
+  apply (rule a_Hcarr, simp)
 proof
   fix x
   assume xcarr: "x \<in> carrier R"
-  from I_one_closed and this
-      have "x \<otimes> \<one> \<in> I" by (intro I_l_closed)
-  from this and xcarr
-      show "x \<in> I" by simp
+  with I_one_closed have "x \<otimes> \<one> \<in> I" by (intro I_l_closed)
+  with xcarr show "x \<in> I" by simp
 qed
 
 lemma (in ideal) Icarr:
   assumes iI: "i \<in> I"
   shows "i \<in> carrier R"
-using iI by (rule a_Hcarr)
+  using iI by (rule a_Hcarr)
 
 
 subsection {* Intersection of Ideals *}
@@ -207,17 +199,17 @@
   interpret ideal I R by fact
   interpret ideal J R by fact
   show ?thesis
-apply (intro idealI subgroup.intro)
-      apply (rule is_ring)
-     apply (force simp add: a_subset)
-    apply (simp add: a_inv_def[symmetric])
-   apply simp
-  apply (simp add: a_inv_def[symmetric])
- apply (clarsimp, rule)
-  apply (fast intro: ideal.I_l_closed ideal.intro assms)+
-apply (clarsimp, rule)
- apply (fast intro: ideal.I_r_closed ideal.intro assms)+
-done
+    apply (intro idealI subgroup.intro)
+          apply (rule is_ring)
+         apply (force simp add: a_subset)
+        apply (simp add: a_inv_def[symmetric])
+       apply simp
+      apply (simp add: a_inv_def[symmetric])
+     apply (clarsimp, rule)
+      apply (fast intro: ideal.I_l_closed ideal.intro assms)+
+    apply (clarsimp, rule)
+     apply (fast intro: ideal.I_r_closed ideal.intro assms)+
+    done
 qed
 
 text {* The intersection of any Number of Ideals is again
@@ -226,26 +218,25 @@
   assumes Sideals: "\<And>I. I \<in> S \<Longrightarrow> ideal I R"
     and notempty: "S \<noteq> {}"
   shows "ideal (Inter S) R"
-apply (unfold_locales)
-apply (simp_all add: Inter_eq)
-      apply rule unfolding mem_Collect_eq defer 1
-      apply rule defer 1
-      apply rule defer 1
-      apply (fold a_inv_def, rule) defer 1
-      apply rule defer 1
-      apply rule defer 1
+  apply (unfold_locales)
+  apply (simp_all add: Inter_eq)
+        apply rule unfolding mem_Collect_eq defer 1
+        apply rule defer 1
+        apply rule defer 1
+        apply (fold a_inv_def, rule) defer 1
+        apply rule defer 1
+        apply rule defer 1
 proof -
   fix x y
   assume "\<forall>I\<in>S. x \<in> I"
-  hence xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp
+  then have xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp
   assume "\<forall>I\<in>S. y \<in> I"
-  hence yI: "\<And>I. I \<in> S \<Longrightarrow> y \<in> I" by simp
+  then have yI: "\<And>I. I \<in> S \<Longrightarrow> y \<in> I" by simp
 
   fix J
   assume JS: "J \<in> S"
   interpret ideal J R by (rule Sideals[OF JS])
-  from xI[OF JS] and yI[OF JS]
-      show "x \<oplus> y \<in> J" by (rule a_closed)
+  from xI[OF JS] and yI[OF JS] show "x \<oplus> y \<in> J" by (rule a_closed)
 next
   fix J
   assume JS: "J \<in> S"
@@ -254,50 +245,47 @@
 next
   fix x
   assume "\<forall>I\<in>S. x \<in> I"
-  hence xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp
+  then have xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp
 
   fix J
   assume JS: "J \<in> S"
   interpret ideal J R by (rule Sideals[OF JS])
 
-  from xI[OF JS]
-      show "\<ominus> x \<in> J" by (rule a_inv_closed)
+  from xI[OF JS] show "\<ominus> x \<in> J" by (rule a_inv_closed)
 next
   fix x y
   assume "\<forall>I\<in>S. x \<in> I"
-  hence xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp
+  then have xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp
   assume ycarr: "y \<in> carrier R"
 
   fix J
   assume JS: "J \<in> S"
   interpret ideal J R by (rule Sideals[OF JS])
 
-  from xI[OF JS] and ycarr
-      show "y \<otimes> x \<in> J" by (rule I_l_closed)
+  from xI[OF JS] and ycarr show "y \<otimes> x \<in> J" by (rule I_l_closed)
 next
   fix x y
   assume "\<forall>I\<in>S. x \<in> I"
-  hence xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp
+  then have xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp
   assume ycarr: "y \<in> carrier R"
 
   fix J
   assume JS: "J \<in> S"
   interpret ideal J R by (rule Sideals[OF JS])
 
-  from xI[OF JS] and ycarr
-      show "x \<otimes> y \<in> J" by (rule I_r_closed)
+  from xI[OF JS] and ycarr show "x \<otimes> y \<in> J" by (rule I_r_closed)
 next
   fix x
   assume "\<forall>I\<in>S. x \<in> I"
-  hence xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp
+  then have xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp
 
   from notempty have "\<exists>I0. I0 \<in> S" by blast
-  from this obtain I0 where I0S: "I0 \<in> S" by auto
+  then obtain I0 where I0S: "I0 \<in> S" by auto
 
   interpret ideal I0 R by (rule Sideals[OF I0S])
 
   from xI[OF I0S] have "x \<in> I0" .
-  from this and a_subset show "x \<in> carrier R" by fast
+  with a_subset show "x \<in> carrier R" by fast
 next
 
 qed
@@ -309,40 +297,41 @@
   assumes idealI: "ideal I R"
       and idealJ: "ideal J R"
   shows "ideal (I <+> J) R"
-apply (rule ideal.intro)
-  apply (rule add_additive_subgroups)
-   apply (intro ideal.axioms[OF idealI])
-  apply (intro ideal.axioms[OF idealJ])
- apply (rule is_ring)
-apply (rule ideal_axioms.intro)
- apply (simp add: set_add_defs, clarsimp) defer 1
- apply (simp add: set_add_defs, clarsimp) defer 1
+  apply (rule ideal.intro)
+    apply (rule add_additive_subgroups)
+     apply (intro ideal.axioms[OF idealI])
+    apply (intro ideal.axioms[OF idealJ])
+   apply (rule is_ring)
+  apply (rule ideal_axioms.intro)
+   apply (simp add: set_add_defs, clarsimp) defer 1
+   apply (simp add: set_add_defs, clarsimp) defer 1
 proof -
   fix x i j
   assume xcarr: "x \<in> carrier R"
-     and iI: "i \<in> I"
-     and jJ: "j \<in> J"
+    and iI: "i \<in> I"
+    and jJ: "j \<in> J"
   from xcarr ideal.Icarr[OF idealI iI] ideal.Icarr[OF idealJ jJ]
-      have c: "(i \<oplus> j) \<otimes> x = (i \<otimes> x) \<oplus> (j \<otimes> x)" by algebra
-  from xcarr and iI
-      have a: "i \<otimes> x \<in> I" by (simp add: ideal.I_r_closed[OF idealI])
-  from xcarr and jJ
-      have b: "j \<otimes> x \<in> J" by (simp add: ideal.I_r_closed[OF idealJ])
-  from a b c
-      show "\<exists>ha\<in>I. \<exists>ka\<in>J. (i \<oplus> j) \<otimes> x = ha \<oplus> ka" by fast
+  have c: "(i \<oplus> j) \<otimes> x = (i \<otimes> x) \<oplus> (j \<otimes> x)"
+    by algebra
+  from xcarr and iI have a: "i \<otimes> x \<in> I"
+    by (simp add: ideal.I_r_closed[OF idealI])
+  from xcarr and jJ have b: "j \<otimes> x \<in> J"
+    by (simp add: ideal.I_r_closed[OF idealJ])
+  from a b c show "\<exists>ha\<in>I. \<exists>ka\<in>J. (i \<oplus> j) \<otimes> x = ha \<oplus> ka"
+    by fast
 next
   fix x i j
   assume xcarr: "x \<in> carrier R"
-     and iI: "i \<in> I"
-     and jJ: "j \<in> J"
+    and iI: "i \<in> I"
+    and jJ: "j \<in> J"
   from xcarr ideal.Icarr[OF idealI iI] ideal.Icarr[OF idealJ jJ]
-      have c: "x \<otimes> (i \<oplus> j) = (x \<otimes> i) \<oplus> (x \<otimes> j)" by algebra
-  from xcarr and iI
-      have a: "x \<otimes> i \<in> I" by (simp add: ideal.I_l_closed[OF idealI])
-  from xcarr and jJ
-      have b: "x \<otimes> j \<in> J" by (simp add: ideal.I_l_closed[OF idealJ])
-  from a b c
-      show "\<exists>ha\<in>I. \<exists>ka\<in>J. x \<otimes> (i \<oplus> j) = ha \<oplus> ka" by fast
+  have c: "x \<otimes> (i \<oplus> j) = (x \<otimes> i) \<oplus> (x \<otimes> j)" by algebra
+  from xcarr and iI have a: "x \<otimes> i \<in> I"
+    by (simp add: ideal.I_l_closed[OF idealI])
+  from xcarr and jJ have b: "x \<otimes> j \<in> J"
+    by (simp add: ideal.I_l_closed[OF idealJ])
+  from a b c show "\<exists>ha\<in>I. \<exists>ka\<in>J. x \<otimes> (i \<oplus> j) = ha \<oplus> ka"
+    by fast
 qed
 
 
@@ -361,87 +350,74 @@
 lemma (in ring) genideal_self:
   assumes "S \<subseteq> carrier R"
   shows "S \<subseteq> Idl S"
-unfolding genideal_def
-by fast
+  unfolding genideal_def by fast
 
 lemma (in ring) genideal_self':
   assumes carr: "i \<in> carrier R"
   shows "i \<in> Idl {i}"
 proof -
-  from carr
-      have "{i} \<subseteq> Idl {i}" by (fast intro!: genideal_self)
-  thus "i \<in> Idl {i}" by fast
+  from carr have "{i} \<subseteq> Idl {i}" by (fast intro!: genideal_self)
+  then show "i \<in> Idl {i}" by fast
 qed
 
 text {* @{term genideal} generates the minimal ideal *}
 lemma (in ring) genideal_minimal:
   assumes a: "ideal I R"
-      and b: "S \<subseteq> I"
+    and b: "S \<subseteq> I"
   shows "Idl S \<subseteq> I"
-unfolding genideal_def
-by (rule, elim InterD, simp add: a b)
+  unfolding genideal_def by rule (elim InterD, simp add: a b)
 
 text {* Generated ideals and subsets *}
 lemma (in ring) Idl_subset_ideal:
   assumes Iideal: "ideal I R"
-      and Hcarr: "H \<subseteq> carrier R"
+    and Hcarr: "H \<subseteq> carrier R"
   shows "(Idl H \<subseteq> I) = (H \<subseteq> I)"
 proof
   assume a: "Idl H \<subseteq> I"
   from Hcarr have "H \<subseteq> Idl H" by (rule genideal_self)
-  from this and a
-      show "H \<subseteq> I" by simp
+  with a show "H \<subseteq> I" by simp
 next
   fix x
   assume HI: "H \<subseteq> I"
-
-  from Iideal and HI
-      have "I \<in> {I. ideal I R \<and> H \<subseteq> I}" by fast
-  from this
-      show "Idl H \<subseteq> I"
-      unfolding genideal_def
-      by fast
+  from Iideal and HI have "I \<in> {I. ideal I R \<and> H \<subseteq> I}" by fast
+  then show "Idl H \<subseteq> I" unfolding genideal_def by fast
 qed
 
 lemma (in ring) subset_Idl_subset:
   assumes Icarr: "I \<subseteq> carrier R"
-      and HI: "H \<subseteq> I"
+    and HI: "H \<subseteq> I"
   shows "Idl H \<subseteq> Idl I"
 proof -
-  from HI and genideal_self[OF Icarr] 
-      have HIdlI: "H \<subseteq> Idl I" by fast
+  from HI and genideal_self[OF Icarr] have HIdlI: "H \<subseteq> Idl I"
+    by fast
 
-  from Icarr
-      have Iideal: "ideal (Idl I) R" by (rule genideal_ideal)
-  from HI and Icarr
-      have "H \<subseteq> carrier R" by fast
-  from Iideal and this
-      have "(H \<subseteq> Idl I) = (Idl H \<subseteq> Idl I)"
-      by (rule Idl_subset_ideal[symmetric])
+  from Icarr have Iideal: "ideal (Idl I) R"
+    by (rule genideal_ideal)
+  from HI and Icarr have "H \<subseteq> carrier R"
+    by fast
+  with Iideal have "(H \<subseteq> Idl I) = (Idl H \<subseteq> Idl I)"
+    by (rule Idl_subset_ideal[symmetric])
 
-  from HIdlI and this
-      show "Idl H \<subseteq> Idl I" by simp
+  with HIdlI show "Idl H \<subseteq> Idl I" by simp
 qed
 
 lemma (in ring) Idl_subset_ideal':
   assumes acarr: "a \<in> carrier R" and bcarr: "b \<in> carrier R"
   shows "(Idl {a} \<subseteq> Idl {b}) = (a \<in> Idl {b})"
-apply (subst Idl_subset_ideal[OF genideal_ideal[of "{b}"], of "{a}"])
-  apply (fast intro: bcarr, fast intro: acarr)
-apply fast
-done
+  apply (subst Idl_subset_ideal[OF genideal_ideal[of "{b}"], of "{a}"])
+    apply (fast intro: bcarr, fast intro: acarr)
+  apply fast
+  done
 
-lemma (in ring) genideal_zero:
-  "Idl {\<zero>} = {\<zero>}"
-apply rule
- apply (rule genideal_minimal[OF zeroideal], simp)
-apply (simp add: genideal_self')
-done
+lemma (in ring) genideal_zero: "Idl {\<zero>} = {\<zero>}"
+  apply rule
+   apply (rule genideal_minimal[OF zeroideal], simp)
+  apply (simp add: genideal_self')
+  done
 
-lemma (in ring) genideal_one:
-  "Idl {\<one>} = carrier R"
+lemma (in ring) genideal_one: "Idl {\<one>} = carrier R"
 proof -
-  interpret ideal "Idl {\<one>}" "R" by (rule genideal_ideal, fast intro: one_closed)
+  interpret ideal "Idl {\<one>}" "R" by (rule genideal_ideal) fast
   show "Idl {\<one>} = carrier R"
   apply (rule, rule a_subset)
   apply (simp add: one_imp_carrier genideal_self')
@@ -451,77 +427,76 @@
 
 text {* Generation of Principal Ideals in Commutative Rings *}
 
-definition
-  cgenideal :: "('a, 'b) monoid_scheme \<Rightarrow> 'a \<Rightarrow> 'a set"  ("PIdl\<index> _" [80] 79)
+definition cgenideal :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set"  ("PIdl\<index> _" [80] 79)
   where "cgenideal R a = {x \<otimes>\<^bsub>R\<^esub> a | x. x \<in> carrier R}"
 
 text {* genhideal (?) really generates an ideal *}
 lemma (in cring) cgenideal_ideal:
   assumes acarr: "a \<in> carrier R"
   shows "ideal (PIdl a) R"
-apply (unfold cgenideal_def)
-apply (rule idealI[OF is_ring])
-   apply (rule subgroup.intro)
-      apply (simp_all add: monoid_record_simps)
-      apply (blast intro: acarr m_closed)
-      apply clarsimp defer 1
-      defer 1
-      apply (fold a_inv_def, clarsimp) defer 1
-      apply clarsimp defer 1
-      apply clarsimp defer 1
+  apply (unfold cgenideal_def)
+  apply (rule idealI[OF is_ring])
+     apply (rule subgroup.intro)
+        apply simp_all
+        apply (blast intro: acarr)
+        apply clarsimp defer 1
+        defer 1
+        apply (fold a_inv_def, clarsimp) defer 1
+        apply clarsimp defer 1
+        apply clarsimp defer 1
 proof -
   fix x y
   assume xcarr: "x \<in> carrier R"
-     and ycarr: "y \<in> carrier R"
+    and ycarr: "y \<in> carrier R"
   note carr = acarr xcarr ycarr
 
-  from carr
-      have "x \<otimes> a \<oplus> y \<otimes> a = (x \<oplus> y) \<otimes> a" by (simp add: l_distr)
-  from this and carr
-      show "\<exists>z. x \<otimes> a \<oplus> y \<otimes> a = z \<otimes> a \<and> z \<in> carrier R" by fast
+  from carr have "x \<otimes> a \<oplus> y \<otimes> a = (x \<oplus> y) \<otimes> a"
+    by (simp add: l_distr)
+  with carr show "\<exists>z. x \<otimes> a \<oplus> y \<otimes> a = z \<otimes> a \<and> z \<in> carrier R"
+    by fast
 next
   from l_null[OF acarr, symmetric] and zero_closed
-      show "\<exists>x. \<zero> = x \<otimes> a \<and> x \<in> carrier R" by fast
+  show "\<exists>x. \<zero> = x \<otimes> a \<and> x \<in> carrier R" by fast
 next
   fix x
   assume xcarr: "x \<in> carrier R"
   note carr = acarr xcarr
 
-  from carr
-      have "\<ominus> (x \<otimes> a) = (\<ominus> x) \<otimes> a" by (simp add: l_minus)
-  from this and carr
-      show "\<exists>z. \<ominus> (x \<otimes> a) = z \<otimes> a \<and> z \<in> carrier R" by fast
+  from carr have "\<ominus> (x \<otimes> a) = (\<ominus> x) \<otimes> a"
+    by (simp add: l_minus)
+  with carr show "\<exists>z. \<ominus> (x \<otimes> a) = z \<otimes> a \<and> z \<in> carrier R"
+    by fast
 next
   fix x y
   assume xcarr: "x \<in> carrier R"
      and ycarr: "y \<in> carrier R"
   note carr = acarr xcarr ycarr
   
-  from carr
-      have "y \<otimes> a \<otimes> x = (y \<otimes> x) \<otimes> a" by (simp add: m_assoc, simp add: m_comm)
-  from this and carr
-      show "\<exists>z. y \<otimes> a \<otimes> x = z \<otimes> a \<and> z \<in> carrier R" by fast
+  from carr have "y \<otimes> a \<otimes> x = (y \<otimes> x) \<otimes> a"
+    by (simp add: m_assoc) (simp add: m_comm)
+  with carr show "\<exists>z. y \<otimes> a \<otimes> x = z \<otimes> a \<and> z \<in> carrier R"
+    by fast
 next
   fix x y
   assume xcarr: "x \<in> carrier R"
      and ycarr: "y \<in> carrier R"
   note carr = acarr xcarr ycarr
 
-  from carr
-      have "x \<otimes> (y \<otimes> a) = (x \<otimes> y) \<otimes> a" by (simp add: m_assoc)
-  from this and carr
-      show "\<exists>z. x \<otimes> (y \<otimes> a) = z \<otimes> a \<and> z \<in> carrier R" by fast
+  from carr have "x \<otimes> (y \<otimes> a) = (x \<otimes> y) \<otimes> a"
+    by (simp add: m_assoc)
+  with carr show "\<exists>z. x \<otimes> (y \<otimes> a) = z \<otimes> a \<and> z \<in> carrier R"
+    by fast
 qed
 
 lemma (in ring) cgenideal_self:
   assumes icarr: "i \<in> carrier R"
   shows "i \<in> PIdl i"
-unfolding cgenideal_def
+  unfolding cgenideal_def
 proof simp
-  from icarr
-      have "i = \<one> \<otimes> i" by simp
-  from this and icarr
-      show "\<exists>x. i = x \<otimes> i \<and> x \<in> carrier R" by fast
+  from icarr have "i = \<one> \<otimes> i"
+    by simp
+  with icarr show "\<exists>x. i = x \<otimes> i \<and> x \<in> carrier R"
+    by fast
 qed
 
 text {* @{const "cgenideal"} is minimal *}
@@ -532,7 +507,8 @@
   shows "PIdl a \<subseteq> J"
 proof -
   interpret ideal J R by fact
-  show ?thesis unfolding cgenideal_def
+  show ?thesis
+    unfolding cgenideal_def
     apply rule
     apply clarify
     using aJ
@@ -543,30 +519,28 @@
 lemma (in cring) cgenideal_eq_genideal:
   assumes icarr: "i \<in> carrier R"
   shows "PIdl i = Idl {i}"
-apply rule
- apply (intro cgenideal_minimal)
-  apply (rule genideal_ideal, fast intro: icarr)
- apply (rule genideal_self', fast intro: icarr)
-apply (intro genideal_minimal)
- apply (rule cgenideal_ideal [OF icarr])
-apply (simp, rule cgenideal_self [OF icarr])
-done
+  apply rule
+   apply (intro cgenideal_minimal)
+    apply (rule genideal_ideal, fast intro: icarr)
+   apply (rule genideal_self', fast intro: icarr)
+  apply (intro genideal_minimal)
+   apply (rule cgenideal_ideal [OF icarr])
+  apply (simp, rule cgenideal_self [OF icarr])
+  done
 
-lemma (in cring) cgenideal_eq_rcos:
- "PIdl i = carrier R #> i"
-unfolding cgenideal_def r_coset_def
-by fast
+lemma (in cring) cgenideal_eq_rcos: "PIdl i = carrier R #> i"
+  unfolding cgenideal_def r_coset_def by fast
 
 lemma (in cring) cgenideal_is_principalideal:
   assumes icarr: "i \<in> carrier R"
   shows "principalideal (PIdl i) R"
-apply (rule principalidealI)
-apply (rule cgenideal_ideal [OF icarr])
+  apply (rule principalidealI)
+  apply (rule cgenideal_ideal [OF icarr])
 proof -
-  from icarr
-      have "PIdl i = Idl {i}" by (rule cgenideal_eq_genideal)
-  from icarr and this
-      show "\<exists>i'\<in>carrier R. PIdl i = Idl {i'}" by fast
+  from icarr have "PIdl i = Idl {i}"
+    by (rule cgenideal_eq_genideal)
+  with icarr show "\<exists>i'\<in>carrier R. PIdl i = Idl {i'}"
+    by fast
 qed
 
 
@@ -574,83 +548,79 @@
 
 lemma (in ring) union_genideal:
   assumes idealI: "ideal I R"
-      and idealJ: "ideal J R"
+    and idealJ: "ideal J R"
   shows "Idl (I \<union> J) = I <+> J"
-apply rule
- apply (rule ring.genideal_minimal)
-   apply (rule is_ring)
-  apply (rule add_ideals[OF idealI idealJ])
- apply (rule)
- apply (simp add: set_add_defs) apply (elim disjE) defer 1 defer 1
- apply (rule) apply (simp add: set_add_defs genideal_def) apply clarsimp defer 1
+  apply rule
+   apply (rule ring.genideal_minimal)
+     apply (rule is_ring)
+    apply (rule add_ideals[OF idealI idealJ])
+   apply (rule)
+   apply (simp add: set_add_defs) apply (elim disjE) defer 1 defer 1
+   apply (rule) apply (simp add: set_add_defs genideal_def) apply clarsimp defer 1
 proof -
   fix x
   assume xI: "x \<in> I"
   have ZJ: "\<zero> \<in> J"
-      by (intro additive_subgroup.zero_closed, rule ideal.axioms[OF idealJ])
-  from ideal.Icarr[OF idealI xI]
-      have "x = x \<oplus> \<zero>" by algebra
-  from xI and ZJ and this
-      show "\<exists>h\<in>I. \<exists>k\<in>J. x = h \<oplus> k" by fast
+    by (intro additive_subgroup.zero_closed) (rule ideal.axioms[OF idealJ])
+  from ideal.Icarr[OF idealI xI] have "x = x \<oplus> \<zero>"
+    by algebra
+  with xI and ZJ show "\<exists>h\<in>I. \<exists>k\<in>J. x = h \<oplus> k"
+    by fast
 next
   fix x
   assume xJ: "x \<in> J"
   have ZI: "\<zero> \<in> I"
-      by (intro additive_subgroup.zero_closed, rule ideal.axioms[OF idealI])
-  from ideal.Icarr[OF idealJ xJ]
-      have "x = \<zero> \<oplus> x" by algebra
-  from ZI and xJ and this
-      show "\<exists>h\<in>I. \<exists>k\<in>J. x = h \<oplus> k" by fast
+    by (intro additive_subgroup.zero_closed, rule ideal.axioms[OF idealI])
+  from ideal.Icarr[OF idealJ xJ] have "x = \<zero> \<oplus> x"
+    by algebra
+  with ZI and xJ show "\<exists>h\<in>I. \<exists>k\<in>J. x = h \<oplus> k"
+    by fast
 next
   fix i j K
   assume iI: "i \<in> I"
-     and jJ: "j \<in> J"
-     and idealK: "ideal K R"
-     and IK: "I \<subseteq> K"
-     and JK: "J \<subseteq> K"
-  from iI and IK
-     have iK: "i \<in> K" by fast
-  from jJ and JK
-     have jK: "j \<in> K" by fast
-  from iK and jK
-     show "i \<oplus> j \<in> K" by (intro additive_subgroup.a_closed) (rule ideal.axioms[OF idealK])
+    and jJ: "j \<in> J"
+    and idealK: "ideal K R"
+    and IK: "I \<subseteq> K"
+    and JK: "J \<subseteq> K"
+  from iI and IK have iK: "i \<in> K" by fast
+  from jJ and JK have jK: "j \<in> K" by fast
+  from iK and jK show "i \<oplus> j \<in> K"
+    by (intro additive_subgroup.a_closed) (rule ideal.axioms[OF idealK])
 qed
 
 
 subsection {* Properties of Principal Ideals *}
 
 text {* @{text "\<zero>"} generates the zero ideal *}
-lemma (in ring) zero_genideal:
-  shows "Idl {\<zero>} = {\<zero>}"
-apply rule
-apply (simp add: genideal_minimal zeroideal)
-apply (fast intro!: genideal_self)
-done
+lemma (in ring) zero_genideal: "Idl {\<zero>} = {\<zero>}"
+  apply rule
+  apply (simp add: genideal_minimal zeroideal)
+  apply (fast intro!: genideal_self)
+  done
 
 text {* @{text "\<one>"} generates the unit ideal *}
-lemma (in ring) one_genideal:
-  shows "Idl {\<one>} = carrier R"
+lemma (in ring) one_genideal: "Idl {\<one>} = carrier R"
 proof -
-  have "\<one> \<in> Idl {\<one>}" by (simp add: genideal_self')
-  thus "Idl {\<one>} = carrier R" by (intro ideal.one_imp_carrier, fast intro: genideal_ideal)
+  have "\<one> \<in> Idl {\<one>}"
+    by (simp add: genideal_self')
+  then show "Idl {\<one>} = carrier R"
+    by (intro ideal.one_imp_carrier) (fast intro: genideal_ideal)
 qed
 
 
 text {* The zero ideal is a principal ideal *}
-corollary (in ring) zeropideal:
-  shows "principalideal {\<zero>} R"
-apply (rule principalidealI)
- apply (rule zeroideal)
-apply (blast intro!: zero_closed zero_genideal[symmetric])
-done
+corollary (in ring) zeropideal: "principalideal {\<zero>} R"
+  apply (rule principalidealI)
+   apply (rule zeroideal)
+  apply (blast intro!: zero_genideal[symmetric])
+  done
 
 text {* The unit ideal is a principal ideal *}
-corollary (in ring) onepideal:
-  shows "principalideal (carrier R) R"
-apply (rule principalidealI)
- apply (rule oneideal)
-apply (blast intro!: one_closed one_genideal[symmetric])
-done
+corollary (in ring) onepideal: "principalideal (carrier R) R"
+  apply (rule principalidealI)
+   apply (rule oneideal)
+  apply (blast intro!: one_genideal[symmetric])
+  done
 
 
 text {* Every principal ideal is a right coset of the carrier *}
@@ -659,28 +629,24 @@
   shows "\<exists>x\<in>I. I = carrier R #> x"
 proof -
   interpret cring R by fact
-  from generate
-      obtain i
-        where icarr: "i \<in> carrier R"
-        and I1: "I = Idl {i}"
-      by fast+
+  from generate obtain i where icarr: "i \<in> carrier R" and I1: "I = Idl {i}"
+    by fast+
 
-  from icarr and genideal_self[of "{i}"]
-      have "i \<in> Idl {i}" by fast
-  hence iI: "i \<in> I" by (simp add: I1)
+  from icarr and genideal_self[of "{i}"] have "i \<in> Idl {i}"
+    by fast
+  then have iI: "i \<in> I" by (simp add: I1)
 
-  from I1 icarr
-      have I2: "I = PIdl i" by (simp add: cgenideal_eq_genideal)
+  from I1 icarr have I2: "I = PIdl i"
+    by (simp add: cgenideal_eq_genideal)
 
   have "PIdl i = carrier R #> i"
-      unfolding cgenideal_def r_coset_def
-      by fast
+    unfolding cgenideal_def r_coset_def by fast
 
-  from I2 and this
-      have "I = carrier R #> i" by simp
+  with I2 have "I = carrier R #> i"
+    by simp
 
-  from iI and this
-      show "\<exists>x\<in>I. I = carrier R #> x" by fast
+  with iI show "\<exists>x\<in>I. I = carrier R #> x"
+    by fast
 qed
 
 
@@ -693,16 +659,16 @@
 proof (rule ccontr, clarsimp)
   interpret cring R by fact
   assume InR: "carrier R \<noteq> I"
-     and "\<forall>a. a \<in> carrier R \<longrightarrow> (\<forall>b. a \<otimes> b \<in> I \<longrightarrow> b \<in> carrier R \<longrightarrow> a \<in> I \<or> b \<in> I)"
-  hence I_prime: "\<And> a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I" by simp
+    and "\<forall>a. a \<in> carrier R \<longrightarrow> (\<forall>b. a \<otimes> b \<in> I \<longrightarrow> b \<in> carrier R \<longrightarrow> a \<in> I \<or> b \<in> I)"
+  then have I_prime: "\<And> a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
+    by simp
   have "primeideal I R"
-      apply (rule primeideal.intro [OF is_ideal is_cring])
-      apply (rule primeideal_axioms.intro)
-       apply (rule InR)
-      apply (erule (2) I_prime)
-      done
-  from this and notprime
-      show "False" by simp
+    apply (rule primeideal.intro [OF is_ideal is_cring])
+    apply (rule primeideal_axioms.intro)
+     apply (rule InR)
+    apply (erule (2) I_prime)
+    done
+  with notprime show False by simp
 qed
 
 lemma (in ideal) primeidealCE:
@@ -721,47 +687,44 @@
 lemma (in cring) zeroprimeideal_domainI:
   assumes pi: "primeideal {\<zero>} R"
   shows "domain R"
-apply (rule domain.intro, rule is_cring)
-apply (rule domain_axioms.intro)
+  apply (rule domain.intro, rule is_cring)
+  apply (rule domain_axioms.intro)
 proof (rule ccontr, simp)
   interpret primeideal "{\<zero>}" "R" by (rule pi)
   assume "\<one> = \<zero>"
-  hence "carrier R = {\<zero>}" by (rule one_zeroD)
-  from this[symmetric] and I_notcarr
-      show "False" by simp
+  then have "carrier R = {\<zero>}" by (rule one_zeroD)
+  from this[symmetric] and I_notcarr show False
+    by simp
 next
   interpret primeideal "{\<zero>}" "R" by (rule pi)
   fix a b
-  assume ab: "a \<otimes> b = \<zero>"
-     and carr: "a \<in> carrier R" "b \<in> carrier R"
-  from ab
-      have abI: "a \<otimes> b \<in> {\<zero>}" by fast
-  from carr and this
-      have "a \<in> {\<zero>} \<or> b \<in> {\<zero>}" by (rule I_prime)
-  thus "a = \<zero> \<or> b = \<zero>" by simp
+  assume ab: "a \<otimes> b = \<zero>" and carr: "a \<in> carrier R" "b \<in> carrier R"
+  from ab have abI: "a \<otimes> b \<in> {\<zero>}"
+    by fast
+  with carr have "a \<in> {\<zero>} \<or> b \<in> {\<zero>}"
+    by (rule I_prime)
+  then show "a = \<zero> \<or> b = \<zero>" by simp
 qed
 
-corollary (in cring) domain_eq_zeroprimeideal:
-  shows "domain R = primeideal {\<zero>} R"
-apply rule
- apply (erule domain.zeroprimeideal)
-apply (erule zeroprimeideal_domainI)
-done
+corollary (in cring) domain_eq_zeroprimeideal: "domain R = primeideal {\<zero>} R"
+  apply rule
+   apply (erule domain.zeroprimeideal)
+  apply (erule zeroprimeideal_domainI)
+  done
 
 
 subsection {* Maximal Ideals *}
 
 lemma (in ideal) helper_I_closed:
   assumes carr: "a \<in> carrier R" "x \<in> carrier R" "y \<in> carrier R"
-      and axI: "a \<otimes> x \<in> I"
+    and axI: "a \<otimes> x \<in> I"
   shows "a \<otimes> (x \<otimes> y) \<in> I"
 proof -
-  from axI and carr
-     have "(a \<otimes> x) \<otimes> y \<in> I" by (simp add: I_r_closed)
-  also from carr
-     have "(a \<otimes> x) \<otimes> y = a \<otimes> (x \<otimes> y)" by (simp add: m_assoc)
-  finally
-     show "a \<otimes> (x \<otimes> y) \<in> I" .
+  from axI and carr have "(a \<otimes> x) \<otimes> y \<in> I"
+    by (simp add: I_r_closed)
+  also from carr have "(a \<otimes> x) \<otimes> y = a \<otimes> (x \<otimes> y)"
+    by (simp add: m_assoc)
+  finally show "a \<otimes> (x \<otimes> y) \<in> I" .
 qed
 
 lemma (in ideal) helper_max_prime:
@@ -787,19 +750,18 @@
     have "\<ominus>(a \<otimes> x) \<in> I" by simp
     also from acarr xcarr
     have "\<ominus>(a \<otimes> x) = a \<otimes> (\<ominus>x)" by algebra
-    finally
-    show "a \<otimes> (\<ominus>x) \<in> I" .
-    from acarr
-    have "a \<otimes> \<zero> = \<zero>" by simp
+    finally show "a \<otimes> (\<ominus>x) \<in> I" .
+    from acarr have "a \<otimes> \<zero> = \<zero>" by simp
   next
     fix x y
     assume xcarr: "x \<in> carrier R"
       and ycarr: "y \<in> carrier R"
       and ayI: "a \<otimes> y \<in> I"
-    from ayI and acarr xcarr ycarr
-    have "a \<otimes> (y \<otimes> x) \<in> I" by (simp add: helper_I_closed)
-    moreover from xcarr ycarr
-    have "y \<otimes> x = x \<otimes> y" by (simp add: m_comm)
+    from ayI and acarr xcarr ycarr have "a \<otimes> (y \<otimes> x) \<in> I"
+      by (simp add: helper_I_closed)
+    moreover
+    from xcarr ycarr have "y \<otimes> x = x \<otimes> y"
+      by (simp add: m_comm)
     ultimately
     show "a \<otimes> (x \<otimes> y) \<in> I" by simp
   qed
@@ -818,40 +780,36 @@
     apply (simp add: I_notcarr)
   proof -
     assume "\<exists>a b. a \<in> carrier R \<and> b \<in> carrier R \<and> a \<otimes> b \<in> I \<and> a \<notin> I \<and> b \<notin> I"
-    from this
-    obtain a b
-      where acarr: "a \<in> carrier R"
-      and bcarr: "b \<in> carrier R"
-      and abI: "a \<otimes> b \<in> I"
-      and anI: "a \<notin> I"
-      and bnI: "b \<notin> I"
-      by fast
+    then obtain a b where
+      acarr: "a \<in> carrier R" and
+      bcarr: "b \<in> carrier R" and
+      abI: "a \<otimes> b \<in> I" and
+      anI: "a \<notin> I" and
+      bnI: "b \<notin> I" by fast
     def J \<equiv> "{x\<in>carrier R. a \<otimes> x \<in> I}"
     
-    from is_cring and acarr
-    have idealJ: "ideal J R" unfolding J_def by (rule helper_max_prime)
+    from is_cring and acarr have idealJ: "ideal J R"
+      unfolding J_def by (rule helper_max_prime)
     
     have IsubJ: "I \<subseteq> J"
     proof
       fix x
       assume xI: "x \<in> I"
-      from this and acarr
-      have "a \<otimes> x \<in> I" by (intro I_l_closed)
-      from xI[THEN a_Hcarr] this
-      show "x \<in> J" unfolding J_def by fast
+      with acarr have "a \<otimes> x \<in> I"
+        by (intro I_l_closed)
+      with xI[THEN a_Hcarr] show "x \<in> J"
+        unfolding J_def by fast
     qed
     
-    from abI and acarr bcarr
-    have "b \<in> J" unfolding J_def by fast
-    from bnI and this
-    have JnI: "J \<noteq> I" by fast
+    from abI and acarr bcarr have "b \<in> J"
+      unfolding J_def by fast
+    with bnI have JnI: "J \<noteq> I" by fast
     from acarr
     have "a = a \<otimes> \<one>" by algebra
-    from this and anI
-    have "a \<otimes> \<one> \<notin> I" by simp
-    from one_closed and this
-    have "\<one> \<notin> J" unfolding J_def by fast
-    hence Jncarr: "J \<noteq> carrier R" by fast
+    with anI have "a \<otimes> \<one> \<notin> I" by simp
+    with one_closed have "\<one> \<notin> J"
+      unfolding J_def by fast
+    then have Jncarr: "J \<noteq> carrier R" by fast
     
     interpret ideal J R by (rule idealJ)
     
@@ -862,8 +820,7 @@
       apply (rule a_subset)
       done
     
-    from this and JnI and Jncarr
-    show "False" by simp
+    with JnI and Jncarr show False by simp
   qed
 qed
 
@@ -873,111 +830,93 @@
 --"A non-zero cring that has only the two trivial ideals is a field"
 lemma (in cring) trivialideals_fieldI:
   assumes carrnzero: "carrier R \<noteq> {\<zero>}"
-      and haveideals: "{I. ideal I R} = {{\<zero>}, carrier R}"
+    and haveideals: "{I. ideal I R} = {{\<zero>}, carrier R}"
   shows "field R"
-apply (rule cring_fieldI)
-apply (rule, rule, rule)
- apply (erule Units_closed)
-defer 1
-  apply rule
-defer 1
+  apply (rule cring_fieldI)
+  apply (rule, rule, rule)
+   apply (erule Units_closed)
+  defer 1
+    apply rule
+  defer 1
 proof (rule ccontr, simp)
   assume zUnit: "\<zero> \<in> Units R"
-  hence a: "\<zero> \<otimes> inv \<zero> = \<one>" by (rule Units_r_inv)
-  from zUnit
-      have "\<zero> \<otimes> inv \<zero> = \<zero>" by (intro l_null, rule Units_inv_closed)
-  from a[symmetric] and this
-      have "\<one> = \<zero>" by simp
-  hence "carrier R = {\<zero>}" by (rule one_zeroD)
-  from this and carrnzero
-      show "False" by simp
+  then have a: "\<zero> \<otimes> inv \<zero> = \<one>" by (rule Units_r_inv)
+  from zUnit have "\<zero> \<otimes> inv \<zero> = \<zero>"
+    by (intro l_null) (rule Units_inv_closed)
+  with a[symmetric] have "\<one> = \<zero>" by simp
+  then have "carrier R = {\<zero>}" by (rule one_zeroD)
+  with carrnzero show False by simp
 next
   fix x
   assume xcarr': "x \<in> carrier R - {\<zero>}"
-  hence xcarr: "x \<in> carrier R" by fast
-  from xcarr'
-      have xnZ: "x \<noteq> \<zero>" by fast
-  from xcarr
-      have xIdl: "ideal (PIdl x) R" by (intro cgenideal_ideal, fast)
+  then have xcarr: "x \<in> carrier R" by fast
+  from xcarr' have xnZ: "x \<noteq> \<zero>" by fast
+  from xcarr have xIdl: "ideal (PIdl x) R"
+    by (intro cgenideal_ideal) fast
 
-  from xcarr
-      have "x \<in> PIdl x" by (intro cgenideal_self, fast)
-  from this and xnZ
-      have "PIdl x \<noteq> {\<zero>}" by fast
-  from haveideals and this
-      have "PIdl x = carrier R"
-      by (blast intro!: xIdl)
-  hence "\<one> \<in> PIdl x" by simp
-  hence "\<exists>y. \<one> = y \<otimes> x \<and> y \<in> carrier R" unfolding cgenideal_def by blast
-  from this
-      obtain y
-        where ycarr: " y \<in> carrier R"
-        and ylinv: "\<one> = y \<otimes> x"
-      by fast+
-  from ylinv and xcarr ycarr
-      have yrinv: "\<one> = x \<otimes> y" by (simp add: m_comm)
+  from xcarr have "x \<in> PIdl x"
+    by (intro cgenideal_self) fast
+  with xnZ have "PIdl x \<noteq> {\<zero>}" by fast
+  with haveideals have "PIdl x = carrier R"
+    by (blast intro!: xIdl)
+  then have "\<one> \<in> PIdl x" by simp
+  then have "\<exists>y. \<one> = y \<otimes> x \<and> y \<in> carrier R"
+    unfolding cgenideal_def by blast
+  then obtain y where ycarr: " y \<in> carrier R" and ylinv: "\<one> = y \<otimes> x"
+    by fast+
+  from ylinv and xcarr ycarr have yrinv: "\<one> = x \<otimes> y"
+    by (simp add: m_comm)
   from ycarr and ylinv[symmetric] and yrinv[symmetric]
-      have "\<exists>y \<in> carrier R. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>" by fast
-  from this and xcarr
-      show "x \<in> Units R"
-      unfolding Units_def
-      by fast
+  have "\<exists>y \<in> carrier R. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>" by fast
+  with xcarr show "x \<in> Units R"
+    unfolding Units_def by fast
 qed
 
-lemma (in field) all_ideals:
-  shows "{I. ideal I R} = {{\<zero>}, carrier R}"
-apply (rule, rule)
+lemma (in field) all_ideals: "{I. ideal I R} = {{\<zero>}, carrier R}"
+  apply (rule, rule)
 proof -
   fix I
   assume a: "I \<in> {I. ideal I R}"
-  with this
-      interpret ideal I R by simp
+  then interpret ideal I R by simp
 
   show "I \<in> {{\<zero>}, carrier R}"
   proof (cases "\<exists>a. a \<in> I - {\<zero>}")
-    assume "\<exists>a. a \<in> I - {\<zero>}"
-    from this
-        obtain a
-          where aI: "a \<in> I"
-          and anZ: "a \<noteq> \<zero>"
-        by fast+
-    from aI[THEN a_Hcarr] anZ
-        have aUnit: "a \<in> Units R" by (simp add: field_Units)
-    hence a: "a \<otimes> inv a = \<one>" by (rule Units_r_inv)
-    from aI and aUnit
-        have "a \<otimes> inv a \<in> I" by (simp add: I_r_closed del: Units_r_inv)
-    hence oneI: "\<one> \<in> I" by (simp add: a[symmetric])
+    case True
+    then obtain a where aI: "a \<in> I" and anZ: "a \<noteq> \<zero>"
+      by fast+
+    from aI[THEN a_Hcarr] anZ have aUnit: "a \<in> Units R"
+      by (simp add: field_Units)
+    then have a: "a \<otimes> inv a = \<one>" by (rule Units_r_inv)
+    from aI and aUnit have "a \<otimes> inv a \<in> I"
+      by (simp add: I_r_closed del: Units_r_inv)
+    then have oneI: "\<one> \<in> I" by (simp add: a[symmetric])
 
     have "carrier R \<subseteq> I"
     proof
       fix x
       assume xcarr: "x \<in> carrier R"
-      from oneI and this
-          have "\<one> \<otimes> x \<in> I" by (rule I_r_closed)
-      from this and xcarr
-          show "x \<in> I" by simp
+      with oneI have "\<one> \<otimes> x \<in> I" by (rule I_r_closed)
+      with xcarr show "x \<in> I" by simp
     qed
-    from this and a_subset
-        have "I = carrier R" by fast
-    thus "I \<in> {{\<zero>}, carrier R}" by fast
+    with a_subset have "I = carrier R" by fast
+    then show "I \<in> {{\<zero>}, carrier R}" by fast
   next
-    assume "\<not> (\<exists>a. a \<in> I - {\<zero>})"
-    hence IZ: "\<And>a. a \<in> I \<Longrightarrow> a = \<zero>" by simp
+    case False
+    then have IZ: "\<And>a. a \<in> I \<Longrightarrow> a = \<zero>" by simp
 
     have a: "I \<subseteq> {\<zero>}"
     proof
       fix x
       assume "x \<in> I"
-      hence "x = \<zero>" by (rule IZ)
-      thus "x \<in> {\<zero>}" by fast
+      then have "x = \<zero>" by (rule IZ)
+      then show "x \<in> {\<zero>}" by fast
     qed
 
     have "\<zero> \<in> I" by simp
-    hence "{\<zero>} \<subseteq> I" by fast
+    then have "{\<zero>} \<subseteq> I" by fast
 
-    from this and a
-        have "I = {\<zero>}" by fast
-    thus "I \<in> {{\<zero>}, carrier R}" by fast
+    with a have "I = {\<zero>}" by fast
+    then show "I \<in> {{\<zero>}, carrier R}" by fast
   qed
 qed (simp add: zeroideal oneideal)
 
@@ -985,52 +924,47 @@
 lemma (in cring) trivialideals_eq_field:
   assumes carrnzero: "carrier R \<noteq> {\<zero>}"
   shows "({I. ideal I R} = {{\<zero>}, carrier R}) = field R"
-by (fast intro!: trivialideals_fieldI[OF carrnzero] field.all_ideals)
+  by (fast intro!: trivialideals_fieldI[OF carrnzero] field.all_ideals)
 
 
 text {* Like zeroprimeideal for domains *}
-lemma (in field) zeromaximalideal:
-  "maximalideal {\<zero>} R"
-apply (rule maximalidealI)
-  apply (rule zeroideal)
+lemma (in field) zeromaximalideal: "maximalideal {\<zero>} R"
+  apply (rule maximalidealI)
+    apply (rule zeroideal)
 proof-
-  from one_not_zero
-      have "\<one> \<notin> {\<zero>}" by simp
-  from this and one_closed
-      show "carrier R \<noteq> {\<zero>}" by fast
+  from one_not_zero have "\<one> \<notin> {\<zero>}" by simp
+  with one_closed show "carrier R \<noteq> {\<zero>}" by fast
 next
   fix J
   assume Jideal: "ideal J R"
-  hence "J \<in> {I. ideal I R}"
-      by fast
-
-  from this and all_ideals
-      show "J = {\<zero>} \<or> J = carrier R" by simp
+  then have "J \<in> {I. ideal I R}" by fast
+  with all_ideals show "J = {\<zero>} \<or> J = carrier R"
+    by simp
 qed
 
 lemma (in cring) zeromaximalideal_fieldI:
   assumes zeromax: "maximalideal {\<zero>} R"
   shows "field R"
-apply (rule trivialideals_fieldI, rule maximalideal.I_notcarr[OF zeromax])
-apply rule apply clarsimp defer 1
- apply (simp add: zeroideal oneideal)
+  apply (rule trivialideals_fieldI, rule maximalideal.I_notcarr[OF zeromax])
+  apply rule apply clarsimp defer 1
+   apply (simp add: zeroideal oneideal)
 proof -
   fix J
   assume Jn0: "J \<noteq> {\<zero>}"
-     and idealJ: "ideal J R"
+    and idealJ: "ideal J R"
   interpret ideal J R by (rule idealJ)
-  have "{\<zero>} \<subseteq> J" by (rule ccontr, simp)
+  have "{\<zero>} \<subseteq> J" by (rule ccontr) simp
   from zeromax and idealJ and this and a_subset
-      have "J = {\<zero>} \<or> J = carrier R" by (rule maximalideal.I_maximal)
-  from this and Jn0
-      show "J = carrier R" by simp
+  have "J = {\<zero>} \<or> J = carrier R"
+    by (rule maximalideal.I_maximal)
+  with Jn0 show "J = carrier R"
+    by simp
 qed
 
-lemma (in cring) zeromaximalideal_eq_field:
-  "maximalideal {\<zero>} R = field R"
-apply rule
- apply (erule zeromaximalideal_fieldI)
-apply (erule field.zeromaximalideal)
-done
+lemma (in cring) zeromaximalideal_eq_field: "maximalideal {\<zero>} R = field R"
+  apply rule
+   apply (erule zeromaximalideal_fieldI)
+  apply (erule field.zeromaximalideal)
+  done
 
 end
--- a/src/HOL/Algebra/Ring.thy	Sat Sep 03 11:10:38 2011 -0700
+++ b/src/HOL/Algebra/Ring.thy	Sat Sep 03 14:33:45 2011 -0700
@@ -367,7 +367,7 @@
 proof -
   assume R: "x \<in> carrier R" "y \<in> carrier R"
   then have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = (\<ominus> x \<oplus> x) \<otimes> y" by (simp add: l_distr)
-  also from R have "... = \<zero>" by (simp add: l_neg l_null)
+  also from R have "... = \<zero>" by (simp add: l_neg)
   finally have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = \<zero>" .
   with R have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) = \<zero> \<oplus> \<ominus> (x \<otimes> y)" by simp
   with R show ?thesis by (simp add: a_assoc r_neg)
@@ -378,7 +378,7 @@
 proof -
   assume R: "x \<in> carrier R" "y \<in> carrier R"
   then have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = x \<otimes> (\<ominus> y \<oplus> y)" by (simp add: r_distr)
-  also from R have "... = \<zero>" by (simp add: l_neg r_null)
+  also from R have "... = \<zero>" by (simp add: l_neg)
   finally have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = \<zero>" .
   with R have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) = \<zero> \<oplus> \<ominus> (x \<otimes> y)" by simp
   with R show ?thesis by (simp add: a_assoc r_neg )
@@ -464,7 +464,6 @@
 proof -
   interpret ring R by fact
   interpret cring S by fact
-ML_val {* Algebra.print_structures @{context} *}
   from RS show ?thesis by algebra
 qed
 
--- a/src/Pure/General/markup.ML	Sat Sep 03 11:10:38 2011 -0700
+++ b/src/Pure/General/markup.ML	Sat Sep 03 14:33:45 2011 -0700
@@ -116,6 +116,7 @@
   val badN: string val bad: T
   val functionN: string
   val assign_execs: Properties.T
+  val removed_versions: Properties.T
   val invoke_scala: string -> string -> Properties.T
   val cancel_scala: string -> Properties.T
   val no_output: Output.output * Output.output
@@ -367,6 +368,7 @@
 val functionN = "function"
 
 val assign_execs = [(functionN, "assign_execs")];
+val removed_versions = [(functionN, "removed_versions")];
 
 fun invoke_scala name id = [(functionN, "invoke_scala"), (nameN, name), (idN, id)];
 fun cancel_scala id = [(functionN, "cancel_scala"), (idN, id)];
--- a/src/Pure/General/markup.scala	Sat Sep 03 11:10:38 2011 -0700
+++ b/src/Pure/General/markup.scala	Sat Sep 03 14:33:45 2011 -0700
@@ -274,6 +274,7 @@
   val Function = new Properties.String(FUNCTION)
 
   val Assign_Execs: Properties.T = List((FUNCTION, "assign_execs"))
+  val Removed_Versions: Properties.T = List((FUNCTION, "removed_versions"))
 
   val INVOKE_SCALA = "invoke_scala"
   object Invoke_Scala
--- a/src/Pure/PIDE/document.ML	Sat Sep 03 11:10:38 2011 -0700
+++ b/src/Pure/PIDE/document.ML	Sat Sep 03 14:33:45 2011 -0700
@@ -30,6 +30,7 @@
   val update: version_id -> version_id -> edit list -> state ->
     ((command_id * exec_id option) list * (string * command_id option) list) * state
   val execute: version_id -> state -> state
+  val remove_versions: version_id list -> state -> state
   val state: unit -> state
   val change_state: (state -> state) -> unit
 end;
@@ -61,11 +62,15 @@
 type perspective = (command_id -> bool) * command_id option; (*visible commands, last*)
 structure Entries = Linear_Set(type key = command_id val ord = int_ord);
 
+type exec = exec_id * (Toplevel.state * unit lazy) lazy;  (*eval/print process*)
+val no_print = Lazy.value ();
+val no_exec = (no_id, Lazy.value (Toplevel.toplevel, no_print));
+
 abstype node = Node of
  {touched: bool,
   header: node_header,
   perspective: perspective,
-  entries: exec_id option Entries.T,  (*command entries with excecutions*)
+  entries: exec option Entries.T,  (*command entries with excecutions*)
   last_exec: command_id option,  (*last command with exec state assignment*)
   result: Toplevel.state lazy}
 and version = Version of node Graph.T  (*development graph wrt. static imports*)
@@ -83,7 +88,6 @@
 
 val no_header = Exn.Exn (ERROR "Bad theory header");
 val no_perspective = make_perspective [];
-val no_print = Lazy.value ();
 val no_result = Lazy.value Toplevel.toplevel;
 
 val empty_node = make_node (false, no_header, no_perspective, Entries.empty, NONE, no_result);
@@ -156,8 +160,8 @@
     NONE => err_undef "command entry" id
   | SOME (exec, _) => exec);
 
-fun the_default_entry node (SOME id) = the_default no_id (the_entry node id)
-  | the_default_entry _ NONE = no_id;
+fun the_default_entry node (SOME id) = (id, (the_default no_exec (the_entry node id)))
+  | the_default_entry _ NONE = (no_id, no_exec);
 
 fun update_entry id exec =
   map_entries (Entries.update (id, exec));
@@ -237,42 +241,40 @@
 abstype state = State of
  {versions: version Inttab.table,  (*version_id -> document content*)
   commands: (string * Toplevel.transition future) Inttab.table,  (*command_id -> name * transition*)
-  execs: (command_id * (Toplevel.state * unit lazy) lazy) Inttab.table,
-    (*exec_id -> command_id with eval/print process*)
-  execution: Future.group}  (*global execution process*)
+  execution: version_id * Future.group}  (*current execution process*)
 with
 
-fun make_state (versions, commands, execs, execution) =
-  State {versions = versions, commands = commands, execs = execs, execution = execution};
+fun make_state (versions, commands, execution) =
+  State {versions = versions, commands = commands, execution = execution};
 
-fun map_state f (State {versions, commands, execs, execution}) =
-  make_state (f (versions, commands, execs, execution));
+fun map_state f (State {versions, commands, execution}) =
+  make_state (f (versions, commands, execution));
 
 val init_state =
-  make_state (Inttab.make [(no_id, empty_version)],
-    Inttab.make [(no_id, (Toplevel.name_of Toplevel.empty, Future.value Toplevel.empty))],
-    Inttab.make [(no_id, (no_id, Lazy.value (Toplevel.toplevel, no_print)))],
-    Future.new_group NONE);
+  make_state (Inttab.make [(no_id, empty_version)], Inttab.empty, (no_id, Future.new_group NONE));
 
 
 (* document versions *)
 
 fun define_version (id: version_id) version =
-  map_state (fn (versions, commands, execs, execution) =>
+  map_state (fn (versions, commands, execution) =>
     let val versions' = Inttab.update_new (id, version) versions
       handle Inttab.DUP dup => err_dup "document version" dup
-    in (versions', commands, execs, execution) end);
+    in (versions', commands, execution) end);
 
 fun the_version (State {versions, ...}) (id: version_id) =
   (case Inttab.lookup versions id of
     NONE => err_undef "document version" id
   | SOME version => version);
 
+fun delete_version (id: version_id) versions = Inttab.delete id versions
+  handle Inttab.UNDEF _ => err_undef "document version" id;
+
 
 (* commands *)
 
 fun define_command (id: command_id) name text =
-  map_state (fn (versions, commands, execs, execution) =>
+  map_state (fn (versions, commands, execution) =>
     let
       val id_string = print_id id;
       val future =
@@ -285,7 +287,7 @@
       val commands' =
         Inttab.update_new (id, (name, future)) commands
           handle Inttab.DUP dup => err_dup "command" dup;
-    in (versions, commands', execs, execution) end);
+    in (versions, commands', execution) end);
 
 fun the_command (State {commands, ...}) (id: command_id) =
   (case Inttab.lookup commands id of
@@ -293,23 +295,9 @@
   | SOME command => command);
 
 
-(* command executions *)
-
-fun define_exec (exec_id, exec) =
-  map_state (fn (versions, commands, execs, execution) =>
-    let val execs' = Inttab.update_new (exec_id, exec) execs
-      handle Inttab.DUP dup => err_dup "command execution" dup
-    in (versions, commands, execs', execution) end);
-
-fun the_exec (State {execs, ...}) exec_id =
-  (case Inttab.lookup execs exec_id of
-    NONE => err_undef "command execution" exec_id
-  | SOME exec => exec);
-
-
 (* document execution *)
 
-fun cancel_execution (State {execution, ...}) = Future.cancel_group execution;
+fun cancel_execution (State {execution, ...}) = Future.cancel_group (#2 execution);
 
 end;
 
@@ -398,7 +386,8 @@
     fun get_common ((prev, id), exec) (found, (_, flags)) =
       if found then NONE
       else
-        let val found' = is_none exec orelse exec <> lookup_entry node0 id
+        let val found' =
+          is_none exec orelse op <> (pairself (Option.map #1) (exec, lookup_entry node0 id));
         in SOME (found', (prev, update_flags prev flags)) end;
     val (found, (common, flags)) =
       iterate_entries get_common node (false, (NONE, (true, true)));
@@ -411,7 +400,7 @@
 
 fun illegal_init () = error "Illegal theory header after end of theory";
 
-fun new_exec state bad_init command_id' (execs, exec, init) =
+fun new_exec state bad_init command_id' (execs, command_exec, init) =
   if bad_init andalso is_none init then NONE
   else
     let
@@ -424,8 +413,9 @@
       val tr = tr0
         |> modify_init
         |> Toplevel.put_id (print_id exec_id');
-      val exec' = (command_id', Lazy.map (fn (st, _) => run_command tr st) (snd exec));
-    in SOME ((exec_id', exec') :: execs, exec', init') end;
+      val exec' = snd (snd command_exec) |> Lazy.map (fn (st, _) => run_command tr st);
+      val command_exec' = (command_id', (exec_id', exec'));
+    in SOME (command_exec' :: execs, command_exec', init') end;
 
 fun make_required nodes =
   let
@@ -489,10 +479,10 @@
                     val required = is_required name;
                     val last_visible = #2 (get_perspective node);
                     val (common, (visible, initial)) = last_common state last_visible node0 node;
-                    val common_exec = the_exec state (the_default_entry node common);
+                    val common_command_exec = the_default_entry node common;
 
-                    val (execs, exec, _) =
-                      ([], common_exec, if initial then SOME init else NONE) |>
+                    val (execs, (command_id, (_, exec)), _) =
+                      ([], common_command_exec, if initial then SOME init else NONE) |>
                       (visible orelse required) ?
                         iterate_entries_after common
                           (fn ((prev, id), _) => fn res =>
@@ -506,14 +496,14 @@
                           else if exists (fn (_, (id, _)) => id0 = id) execs then SOME res
                           else SOME (id0 :: res)) node0 [];
 
-                    val last_exec = if #1 exec = no_id then NONE else SOME (#1 exec);
+                    val last_exec = if command_id = no_id then NONE else SOME command_id;
                     val result =
                       if is_some (after_entry node last_exec) then no_result
-                      else Lazy.map #1 (#2 exec);
+                      else Lazy.map #1 exec;
 
                     val node' = node
                       |> fold reset_entry no_execs
-                      |> fold (fn (exec_id, (id, _)) => update_entry id (SOME exec_id)) execs
+                      |> fold (fn (id, exec) => update_entry id (SOME exec)) execs
                       |> set_last_exec last_exec
                       |> set_result result
                       |> set_touched false;
@@ -523,12 +513,11 @@
 
     val command_execs =
       map (rpair NONE) (maps #1 updated) @
-      map (fn (exec_id, (command_id, _)) => (command_id, SOME exec_id)) (maps #2 updated);
+      map (fn (command_id, (exec_id, _)) => (command_id, SOME exec_id)) (maps #2 updated);
     val updated_nodes = maps #3 updated;
     val last_execs = map (fn (name, node) => (name, get_last_exec node)) updated_nodes;
 
     val state' = state
-      |> fold (fold define_exec o #2) updated
       |> define_version new_id (fold put_node updated_nodes new_version);
   in ((command_execs, last_execs), state') end;
 
@@ -538,14 +527,13 @@
 (* execute *)
 
 fun execute version_id state =
-  state |> map_state (fn (versions, commands, execs, _) =>
+  state |> map_state (fn (versions, commands, _) =>
     let
       val version = the_version state version_id;
 
-      fun force_exec _ NONE = ()
-        | force_exec node (SOME exec_id) =
+      fun force_exec _ _ NONE = ()
+        | force_exec node command_id (SOME (_, exec)) =
             let
-              val (command_id, exec) = the_exec state exec_id;
               val (_, print) = Lazy.force exec;
               val _ =
                 if #1 (get_perspective node) command_id
@@ -553,16 +541,34 @@
                 else ();
             in () end;
 
-      val execution = Future.new_group NONE;
+      val group = Future.new_group NONE;
       val _ =
         nodes_of version |> Graph.schedule
           (fn deps => fn (name, node) =>
             (singleton o Future.forks)
-              {name = "theory:" ^ name, group = SOME (Future.new_group (SOME execution)),
+              {name = "theory:" ^ name, group = SOME (Future.new_group (SOME group)),
                 deps = map (Future.task_of o #2) deps, pri = 1, interrupts = false}
-              (iterate_entries (fn (_, exec) => fn () => SOME (force_exec node exec)) node));
+              (iterate_entries (fn ((_, id), exec) => fn () =>
+                SOME (force_exec node id exec)) node));
+
+    in (versions, commands, (version_id, group)) end);
+
+
+(* remove versions *)
 
-    in (versions, commands, execs, execution) end);
+fun remove_versions ids state = state |> map_state (fn (versions, _, execution) =>
+  let
+    val _ = member (op =) ids (#1 execution) andalso
+      error ("Attempt to remove execution version " ^ print_id (#1 execution));
+
+    val versions' = fold delete_version ids versions;
+    val commands' =
+      (versions', Inttab.empty) |->
+        Inttab.fold (fn (_, version) => nodes_of version |>
+          Graph.fold (fn (_, (node, _)) => node |>
+            iterate_entries (fn ((_, id), _) =>
+              SOME o Inttab.insert (K true) (id, the_command state id))));
+  in (versions', commands', execution) end);
 
 
 
--- a/src/Pure/PIDE/document.scala	Sat Sep 03 11:10:38 2011 -0700
+++ b/src/Pure/PIDE/document.scala	Sat Sep 03 14:33:45 2011 -0700
@@ -168,15 +168,19 @@
 
   object Change
   {
-    val init: Change = Change(Future.value(Version.init), Nil, Future.value(Version.init))
+    val init: Change = Change(Some(Future.value(Version.init)), Nil, Future.value(Version.init))
   }
 
   sealed case class Change(
-    val previous: Future[Version],
+    val previous: Option[Future[Version]],
     val edits: List[Edit_Text],
     val version: Future[Version])
   {
-    def is_finished: Boolean = previous.is_finished && version.is_finished
+    def is_finished: Boolean =
+      (previous match { case None => true case Some(future) => future.is_finished }) &&
+      version.is_finished
+
+    def truncate: Change = copy(previous = None, edits = Nil)
   }
 
 
@@ -184,16 +188,16 @@
 
   object History
   {
-    val init: History = new History(List(Change.init))
+    val init: History = History(List(Change.init))
   }
 
   // FIXME pruning, purging of state
-  class History(val undo_list: List[Change])
+  sealed case class History(val undo_list: List[Change])
   {
     require(!undo_list.isEmpty)
 
     def tip: Change = undo_list.head
-    def +(change: Change): History = new History(change :: undo_list)
+    def +(change: Change): History = copy(undo_list = change :: undo_list)
   }
 
 
@@ -260,7 +264,6 @@
     val commands: Map[Command_ID, Command.State] = Map(),  // static markup from define_command
     val execs: Map[Exec_ID, Command.State] = Map(),  // dynamic markup from execution
     val assignments: Map[Version_ID, State.Assignment] = Map(),
-    val disposed: Set[ID] = Set(),  // FIXME unused!?
     val history: History = History.init)
   {
     private def fail[A]: A = throw new State.Fail(this)
@@ -268,7 +271,6 @@
     def define_version(version: Version, assignment: State.Assignment): State =
     {
       val id = version.id
-      if (versions.isDefinedAt(id) || disposed(id)) fail
       copy(versions = versions + (id -> version),
         assignments = assignments + (id -> assignment.unfinished))
     }
@@ -276,7 +278,6 @@
     def define_command(command: Command): State =
     {
       val id = command.id
-      if (commands.isDefinedAt(id) || disposed(id)) fail
       copy(commands = commands + (id -> command.empty_state))
     }
 
@@ -342,7 +343,7 @@
     def is_stable(change: Change): Boolean =
       change.is_finished && is_assigned(change.version.get_finished)
 
-    def recent_stable: Option[Change] = history.undo_list.find(is_stable)
+    def recent_stable: Change = history.undo_list.find(is_stable) getOrElse fail
     def tip_stable: Boolean = is_stable(history.tip)
     def tip_version: Version = history.tip.version.get_finished
 
@@ -366,10 +367,49 @@
         edits: List[Edit_Text],
         version: Future[Version]): (Change, State) =
     {
-      val change = Change(previous, edits, version)
+      val change = Change(Some(previous), edits, version)
       (change, copy(history = history + change))
     }
 
+    def prune_history(retain: Int = 0): (List[Version], State) =
+    {
+      val undo_list = history.undo_list
+      val n = undo_list.iterator.zipWithIndex.find(p => is_stable(p._1)).get._2 + 1
+      val (retained, dropped) = undo_list.splitAt(n max retain)
+
+      retained.splitAt(retained.length - 1) match {
+        case (prefix, List(last)) =>
+          val dropped_versions = dropped.map(change => change.version.get_finished)
+          val state1 = copy(history = History(prefix ::: List(last.truncate)))
+          (dropped_versions, state1)
+        case _ => fail
+      }
+    }
+
+    def removed_versions(removed: List[Version_ID]): State =
+    {
+      val versions1 = versions -- removed
+      val assignments1 = assignments -- removed
+      var commands1 = Map.empty[Command_ID, Command.State]
+      var execs1 = Map.empty[Exec_ID, Command.State]
+      for {
+        (version_id, version) <- versions1.iterator
+        val command_execs = assignments1(version_id).command_execs
+        (_, node) <- version.nodes.iterator
+        command <- node.commands.iterator
+      } {
+        val id = command.id
+        if (!commands1.isDefinedAt(id) && commands.isDefinedAt(id))
+          commands1 += (id -> commands(id))
+        if (command_execs.isDefinedAt(id)) {
+          val exec_id = command_execs(id)
+          if (!execs1.isDefinedAt(exec_id) && execs.isDefinedAt(exec_id))
+            execs1 += (exec_id -> execs(exec_id))
+        }
+      }
+      copy(versions = versions1, commands = commands1, execs = execs1, assignments = assignments1)
+    }
+
     def command_state(version: Version, command: Command): Command.State =
     {
       require(is_assigned(version))
@@ -384,7 +424,7 @@
     // persistent user-view
     def snapshot(name: Node.Name, pending_edits: List[Text.Edit]): Snapshot =
     {
-      val stable = recent_stable.get
+      val stable = recent_stable
       val latest = history.tip
 
       // FIXME proper treatment of removed nodes
--- a/src/Pure/PIDE/isar_document.ML	Sat Sep 03 11:10:38 2011 -0700
+++ b/src/Pure/PIDE/isar_document.ML	Sat Sep 03 14:33:45 2011 -0700
@@ -65,6 +65,17 @@
       in state2 end));
 
 val _ =
+  Isabelle_Process.add_command "Isar_Document.remove_versions"
+    (fn [versions_yxml] => Document.change_state (fn state =>
+      let
+        val versions =
+          YXML.parse_body versions_yxml |>
+            let open XML.Decode in list int end;
+        val state1 = Document.remove_versions versions state;
+        val _ = Output.raw_message Markup.removed_versions versions_yxml;
+      in state1 end));
+
+val _ =
   Isabelle_Process.add_command "Isar_Document.invoke_scala"
     (fn [id, tag, res] => Invoke_Scala.fulfill_method id tag res);
 
--- a/src/Pure/PIDE/isar_document.scala	Sat Sep 03 11:10:38 2011 -0700
+++ b/src/Pure/PIDE/isar_document.scala	Sat Sep 03 14:33:45 2011 -0700
@@ -26,6 +26,20 @@
       }
   }
 
+  object Removed
+  {
+    def unapply(text: String): Option[List[Document.Version_ID]] =
+      try {
+        import XML.Decode._
+        Some(list(long)(YXML.parse_body(text)))
+      }
+      catch {
+        case ERROR(_) => None
+        case _: XML.XML_Atom => None
+        case _: XML.XML_Body => None
+      }
+  }
+
 
   /* toplevel transactions */
 
@@ -188,6 +202,14 @@
     input("Isar_Document.update", Document.ID(old_id), Document.ID(new_id), edits_yxml)
   }
 
+  def remove_versions(versions: List[Document.Version])
+  {
+    val versions_yxml =
+      { import XML.Encode._
+        YXML.string_of_body(list(long)(versions.map(_.id))) }
+    input("Isar_Document.remove_versions", versions_yxml)
+  }
+
 
   /* method invocation service */
 
--- a/src/Pure/System/session.scala	Sat Sep 03 11:10:38 2011 -0700
+++ b/src/Pure/System/session.scala	Sat Sep 03 14:33:45 2011 -0700
@@ -42,6 +42,8 @@
   val output_delay = Time.seconds(0.1)  // prover output (markup, common messages)
   val update_delay = Time.seconds(0.5)  // GUI layout updates
   val load_delay = Time.seconds(0.5)  // file load operations (new buffers etc.)
+  val prune_delay = Time.seconds(60.0)  // prune history -- delete old versions
+  val prune_size = 0  // size of retained history
 
 
   /* pervasive event buses */
@@ -180,6 +182,8 @@
     val this_actor = self
     var prover: Option[Isabelle_Process with Isar_Document] = None
 
+    var prune_next = System.currentTimeMillis() + prune_delay.ms
+
 
     /* perspective */
 
@@ -239,6 +243,16 @@
     //}}}
 
 
+    /* removed versions */
+
+    def handle_removed(removed: List[Document.Version_ID])
+    //{{{
+    {
+      global_state.change(_.removed_versions(removed))
+    }
+    //}}}
+
+
     /* resulting changes */
 
     def handle_change(change: Change_Node)
@@ -295,6 +309,19 @@
               catch { case _: Document.State.Fail => bad_result(result) }
             case _ => bad_result(result)
           }
+          // FIXME separate timeout event/message!?
+          if (prover.isDefined && System.currentTimeMillis() > prune_next) {
+            val old_versions = global_state.change_yield(_.prune_history(prune_size))
+            if (!old_versions.isEmpty) prover.get.remove_versions(old_versions)
+            prune_next = System.currentTimeMillis() + prune_delay.ms
+          }
+        case Markup.Removed_Versions if result.is_raw =>
+          XML.content(result.body).mkString match {
+            case Isar_Document.Removed(removed) =>
+              try { handle_removed(removed) }
+              catch { case _: Document.State.Fail => bad_result(result) }
+            case _ => bad_result(result)
+          }
         case Markup.Invoke_Scala(name, id) if result.is_raw =>
           Future.fork {
             val arg = XML.content(result.body).mkString
--- a/src/Tools/jEdit/src/session_dockable.scala	Sat Sep 03 11:10:38 2011 -0700
+++ b/src/Tools/jEdit/src/session_dockable.scala	Sat Sep 03 14:33:45 2011 -0700
@@ -80,23 +80,21 @@
     Swing_Thread.now {
       // FIXME correlation to changed_nodes!?
       val state = Isabelle.session.current_state()
-      state.recent_stable.map(change => change.version.get_finished) match {
-        case None =>
-        case Some(version) =>
-          var nodes_status1 = nodes_status
-          for {
-            name <- changed_nodes
-            node <- version.nodes.get(name)
-            val status = Isar_Document.node_status(state, version, node)
-          } nodes_status1 += (name -> status.toString)
+      val version = state.recent_stable.version.get_finished
 
-          if (nodes_status != nodes_status1) {
-            nodes_status = nodes_status1
-            val order =
-              Library.sort_wrt((name: Document.Node.Name) => name.theory,
-                nodes_status.keySet.toList)
-            status.listData = order.map(name => name.theory + " " + nodes_status(name))
-          }
+      var nodes_status1 = nodes_status
+      for {
+        name <- changed_nodes
+        node <- version.nodes.get(name)
+        val status = Isar_Document.node_status(state, version, node)
+      } nodes_status1 += (name -> status.toString)
+
+      if (nodes_status != nodes_status1) {
+        nodes_status = nodes_status1
+        val order =
+          Library.sort_wrt((name: Document.Node.Name) => name.theory,
+            nodes_status.keySet.toList)
+        status.listData = order.map(name => name.theory + " " + nodes_status(name))
       }
     }
   }