--- a/src/HOL/Codatatype/BNF_Def.thy Wed Sep 12 10:18:31 2012 +0200
+++ b/src/HOL/Codatatype/BNF_Def.thy Wed Sep 12 10:35:56 2012 +0200
@@ -51,7 +51,7 @@
moreover have "f1 (fst a') = f2 (snd a')"
using a' unfolding csquare_def thePull_def by auto
ultimately show "\<exists> ja'. ?phi a' ja'"
- using assms unfolding wppull_def by auto
+ using assms unfolding wppull_def by blast
qed
lemma wpull_wppull:
@@ -64,7 +64,7 @@
then obtain a' where a': "a' \<in> A'" and b1: "b1 = p1' a'" and b2: "b2 = p2' a'"
using wp unfolding wpull_def by blast
show "\<exists>a\<in>A. e1 (p1 a) = e1 b1 \<and> e2 (p2 a) = e2 b2"
- apply(rule bexI[of _ "j a'"]) unfolding b1 b2 using a' 1 by auto
+ apply (rule bexI[of _ "j a'"]) unfolding b1 b2 using a' 1 by auto
qed
lemma wppull_id: "\<lbrakk>wpull UNIV UNIV UNIV f1 f2 p1 p2; e1 = id; e2 = id\<rbrakk> \<Longrightarrow>
@@ -87,8 +87,7 @@
definition "pick_middle P Q a c = (SOME b. (a,b) \<in> P \<and> (b,c) \<in> Q)"
lemma pick_middle:
-assumes "(a,c) \<in> P O Q"
-shows "(a, pick_middle P Q a c) \<in> P \<and> (pick_middle P Q a c, c) \<in> Q"
+"(a,c) \<in> P O Q \<Longrightarrow> (a, pick_middle P Q a c) \<in> P \<and> (pick_middle P Q a c, c) \<in> Q"
unfolding pick_middle_def apply(rule someI_ex)
using assms unfolding relcomp_def by auto
@@ -96,7 +95,8 @@
definition sndO where "sndO P Q ac = (pick_middle P Q (fst ac) (snd ac), snd ac)"
lemma fstO_in: "ac \<in> P O Q \<Longrightarrow> fstO P Q ac \<in> P"
-by (metis assms fstO_def pick_middle surjective_pairing)
+unfolding fstO_def
+by (subst (asm) surjective_pairing) (rule pick_middle[THEN conjunct1])
lemma fst_fstO: "fst bc = (fst \<circ> fstO P Q) bc"
unfolding comp_def fstO_def by simp
@@ -105,11 +105,12 @@
unfolding comp_def sndO_def by simp
lemma sndO_in: "ac \<in> P O Q \<Longrightarrow> sndO P Q ac \<in> Q"
-by (metis assms sndO_def pick_middle surjective_pairing)
+unfolding sndO_def
+by (subst (asm) surjective_pairing) (rule pick_middle[THEN conjunct2])
lemma csquare_fstO_sndO:
"csquare (P O Q) snd fst (fstO P Q) (sndO P Q)"
-unfolding csquare_def fstO_def sndO_def using pick_middle by auto
+unfolding csquare_def fstO_def sndO_def using pick_middle by simp
lemma wppull_fstO_sndO:
shows "wppull (P O Q) P Q snd fst fst snd (fstO P Q) (sndO P Q)"
--- a/src/HOL/Codatatype/BNF_FP.thy Wed Sep 12 10:18:31 2012 +0200
+++ b/src/HOL/Codatatype/BNF_FP.thy Wed Sep 12 10:35:56 2012 +0200
@@ -80,16 +80,24 @@
lemma obj_one_pointE: "\<forall>x. s = x \<longrightarrow> P \<Longrightarrow> P"
by blast
+lemma obj_sumE_f':
+"\<lbrakk>\<forall>x. s = f (Inl x) \<longrightarrow> P; \<forall>x. s = f (Inr x) \<longrightarrow> P\<rbrakk> \<Longrightarrow> s = f x \<longrightarrow> P"
+by (cases x) blast+
+
lemma obj_sumE_f:
"\<lbrakk>\<forall>x. s = f (Inl x) \<longrightarrow> P; \<forall>x. s = f (Inr x) \<longrightarrow> P\<rbrakk> \<Longrightarrow> \<forall>x. s = f x \<longrightarrow> P"
-by (metis sum.exhaust)
+by (rule allI) (rule obj_sumE_f')
lemma obj_sumE: "\<lbrakk>\<forall>x. s = Inl x \<longrightarrow> P; \<forall>x. s = Inr x \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (cases s) auto
+lemma obj_sum_step':
+"\<lbrakk>\<forall>x. s = f (Inr (Inl x)) \<longrightarrow> P; \<forall>x. s = f (Inr (Inr x)) \<longrightarrow> P\<rbrakk> \<Longrightarrow> s = f (Inr x) \<longrightarrow> P"
+by (cases x) blast+
+
lemma obj_sum_step:
- "\<lbrakk>\<forall>x. s = f (Inr (Inl x)) \<longrightarrow> P; \<forall>x. s = f (Inr (Inr x)) \<longrightarrow> P\<rbrakk> \<Longrightarrow> \<forall>x. s = f (Inr x) \<longrightarrow> P"
-by (metis obj_sumE)
+"\<lbrakk>\<forall>x. s = f (Inr (Inl x)) \<longrightarrow> P; \<forall>x. s = f (Inr (Inr x)) \<longrightarrow> P\<rbrakk> \<Longrightarrow> \<forall>x. s = f (Inr x) \<longrightarrow> P"
+by (rule allI) (rule obj_sum_step')
lemma sum_case_if:
"sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)"
--- a/src/HOL/Codatatype/BNF_GFP.thy Wed Sep 12 10:18:31 2012 +0200
+++ b/src/HOL/Codatatype/BNF_GFP.thy Wed Sep 12 10:35:56 2012 +0200
@@ -278,12 +278,10 @@
(* pick according to the weak pullback *)
definition pickWP_pred where
-"pickWP_pred A p1 p2 b1 b2 a \<equiv>
- a \<in> A \<and> p1 a = b1 \<and> p2 a = b2"
+"pickWP_pred A p1 p2 b1 b2 a \<equiv> a \<in> A \<and> p1 a = b1 \<and> p2 a = b2"
definition pickWP where
-"pickWP A p1 p2 b1 b2 \<equiv>
- SOME a. pickWP_pred A p1 p2 b1 b2 a"
+"pickWP A p1 p2 b1 b2 \<equiv> SOME a. pickWP_pred A p1 p2 b1 b2 a"
lemma pickWP_pred:
assumes "wpull A B1 B2 f1 f2 p1 p2" and
--- a/src/HOL/Codatatype/Tools/bnf_def.ML Wed Sep 12 10:18:31 2012 +0200
+++ b/src/HOL/Codatatype/Tools/bnf_def.ML Wed Sep 12 10:35:56 2012 +0200
@@ -482,7 +482,6 @@
val bd_CnotzeroN = "bd_Cnotzero";
val collect_set_naturalN = "collect_set_natural";
val in_bdN = "in_bd";
-val in_congN = "in_cong";
val in_monoN = "in_mono";
val in_relN = "in_rel";
val map_idN = "map_id";
@@ -490,9 +489,7 @@
val map_compN = "map_comp";
val map_comp'N = "map_comp'";
val map_congN = "map_cong";
-val map_wppullN = "map_wppull";
val map_wpullN = "map_wpull";
-val rel_congN = "rel_cong";
val rel_IdN = "rel_Id";
val rel_GrN = "rel_Gr";
val rel_converseN = "rel_converse";