merge
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Mon, 14 Jun 2010 14:45:40 +0200
changeset 2250 c3fd4e42bb4a
parent 2249 1476c26d4310 (diff)
parent 2248 a04040474800 (current diff)
child 2251 1a4fc8d3873f
merge
Nominal/FSet.thy
Quotient-Paper/Paper.thy
--- a/Nominal/FSet.thy	Mon Jun 14 12:30:08 2010 +0200
+++ b/Nominal/FSet.thy	Mon Jun 14 14:45:40 2010 +0200
@@ -1672,5 +1672,10 @@
 apply(auto simp add: mem_def)
 done
 
+lemma ball_reg_right_unfolded: "(\<forall>x. R x \<longrightarrow> P x \<longrightarrow> Q x) \<longrightarrow> (All P \<longrightarrow> Ball R Q)"
+apply rule
+apply (rule ball_reg_right)
+apply auto
+done
 
 end
--- a/Quotient-Paper/Paper.thy	Mon Jun 14 12:30:08 2010 +0200
+++ b/Quotient-Paper/Paper.thy	Mon Jun 14 14:45:40 2010 +0200
@@ -816,13 +816,13 @@
 
   Isabelle provides a set of \emph{mono} rules, that are used to split implications
   of similar statements into simpler implication subgoals. These are enhanced
-  with special quotient theorem in the regularization goal. Below we only show
+  with special quotient theorem in the regularization proof. Below we only show
   the versions for the universal quantifier. For the existential quantifier
   and abstraction they are analogous.
 
   First, bounded universal quantifiers can be removed on the right:
 
-  @{thm [display, indent=10] ball_reg_right[no_vars]}
+  @{thm [display, indent=10] ball_reg_right_unfolded[no_vars]}
 
   They can be removed anywhere if the relation is an equivalence relation: