removed dead internal constants/theorems
authortraytel
Tue, 07 May 2013 18:40:23 +0200
changeset 51909 eb3169abcbd5
parent 51908 5aaa9e2f7dd1
child 51910 31bb70ddee7e
removed dead internal constants/theorems
src/HOL/BNF/BNF.thy
src/HOL/BNF/BNF_Def.thy
src/HOL/BNF/BNF_GFP.thy
--- a/src/HOL/BNF/BNF.thy	Tue May 07 17:35:29 2013 +0200
+++ b/src/HOL/BNF/BNF.thy	Tue May 07 18:40:23 2013 +0200
@@ -14,7 +14,7 @@
 begin
 
 hide_const (open) Gr Grp collect fsts snds setl setr 
-  convol thePull pick_middle pick_middlep fstO sndO fstOp sndOp csquare inver
+  convol thePull pick_middlep fstOp sndOp csquare inver
   image2 relImage relInvImage prefCl PrefCl Succ Shift shift
 
 end
--- a/src/HOL/BNF/BNF_Def.thy	Tue May 07 17:35:29 2013 +0200
+++ b/src/HOL/BNF/BNF_Def.thy	Tue May 07 18:40:23 2013 +0200
@@ -46,10 +46,6 @@
 apply(rule ext)
 unfolding convol_def by simp
 
-lemma convol_memI:
-"\<lbrakk>f x = f' x; g x = g' x; P x\<rbrakk> \<Longrightarrow> <f , g> x \<in> {(f' a, g' a) |a. P a}"
-unfolding convol_def by auto
-
 lemma convol_mem_GrpI:
 "\<lbrakk>g x = g' x; x \<in> A\<rbrakk> \<Longrightarrow> <id , g> x \<in> (Collect (split (Grp A g)))"
 unfolding convol_def Grp_def by auto
@@ -100,9 +96,6 @@
    wppull UNIV UNIV UNIV f1 f2 e1 e2 p1 p2"
 by (erule wpull_wppull) auto
 
-lemma Id_alt: "Id = Gr UNIV id"
-unfolding Gr_def by auto
-
 lemma eq_alt: "op = = Grp UNIV id"
 unfolding Grp_def by auto
 
@@ -112,18 +105,12 @@
 lemma eq_OOI: "R = op = \<Longrightarrow> R = R OO R"
   by auto
 
-lemma Gr_UNIV_id: "f = id \<Longrightarrow> (Gr UNIV f)^-1 O Gr UNIV f = Gr UNIV f"
-unfolding Gr_def by auto
-
 lemma Grp_UNIV_id: "f = id \<Longrightarrow> (Grp UNIV f)^--1 OO Grp UNIV f = Grp UNIV f"
 unfolding Grp_def by auto
 
 lemma Grp_UNIV_idI: "x = y \<Longrightarrow> Grp UNIV id x y"
 unfolding Grp_def by auto
 
-lemma Gr_mono: "A \<subseteq> B \<Longrightarrow> Gr A f \<subseteq> Gr B f"
-unfolding Gr_def by auto
-
 lemma Grp_mono: "A \<le> B \<Longrightarrow> Grp A f \<le> Grp B f"
 unfolding Grp_def by auto
 
@@ -139,73 +126,40 @@
 lemma Collect_split_Grp_inD: "z \<in> Collect (split (Grp A f)) \<Longrightarrow> fst z \<in> A"
 unfolding Grp_def o_def by auto
 
-lemma wpull_Gr:
-"wpull (Gr A f) A (f ` A) f id fst snd"
-unfolding wpull_def Gr_def by auto
-
 lemma wpull_Grp:
 "wpull (Collect (split (Grp A f))) A (f ` A) f id fst snd"
 unfolding wpull_def Grp_def by auto
 
-definition "pick_middle P Q a c = (SOME b. (a,b) \<in> P \<and> (b,c) \<in> Q)"
-
 definition "pick_middlep P Q a c = (SOME b. P a b \<and> Q b c)"
 
-lemma pick_middle:
-"(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) by auto
-
 lemma pick_middlep:
 "(P OO Q) a c \<Longrightarrow> P a (pick_middlep P Q a c) \<and> Q (pick_middlep P Q a c) c"
 unfolding pick_middlep_def apply(rule someI_ex) by auto
 
-definition fstO where "fstO P Q ac = (fst ac, pick_middle P Q (fst ac) (snd ac))"
-definition sndO where "sndO P Q ac = (pick_middle P Q (fst ac) (snd ac), snd ac)"
-
 definition fstOp where "fstOp P Q ac = (fst ac, pick_middlep P Q (fst ac) (snd ac))"
 definition sndOp where "sndOp P Q ac = (pick_middlep P Q (fst ac) (snd ac), (snd ac))"
 
-lemma fstO_in: "ac \<in> P O Q \<Longrightarrow> fstO P Q ac \<in> P"
-unfolding fstO_def by (subst (asm) surjective_pairing) (rule pick_middle[THEN conjunct1])
-
 lemma fstOp_in: "ac \<in> Collect (split (P OO Q)) \<Longrightarrow> fstOp P Q ac \<in> Collect (split P)"
 unfolding fstOp_def mem_Collect_eq
 by (subst (asm) surjective_pairing, unfold prod.cases) (erule pick_middlep[THEN conjunct1])
 
-lemma fst_fstO: "fst bc = (fst \<circ> fstO P Q) bc"
-unfolding comp_def fstO_def by simp
-
 lemma fst_fstOp: "fst bc = (fst \<circ> fstOp P Q) bc"
 unfolding comp_def fstOp_def by simp
 
-lemma snd_sndO: "snd bc = (snd \<circ> sndO P Q) bc"
-unfolding comp_def sndO_def by simp
-
 lemma snd_sndOp: "snd bc = (snd \<circ> sndOp P Q) bc"
 unfolding comp_def sndOp_def by simp
 
-lemma sndO_in: "ac \<in> P O Q \<Longrightarrow> sndO P Q ac \<in> Q"
-unfolding sndO_def
-by (subst (asm) surjective_pairing) (rule pick_middle[THEN conjunct2])
-
 lemma sndOp_in: "ac \<in> Collect (split (P OO Q)) \<Longrightarrow> sndOp P Q ac \<in> Collect (split Q)"
 unfolding sndOp_def mem_Collect_eq
 by (subst (asm) surjective_pairing, unfold prod.cases) (erule pick_middlep[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 simp
-
 lemma csquare_fstOp_sndOp:
 "csquare (Collect (split (P OO Q))) snd fst (fstOp P Q) (sndOp P Q)"
 unfolding csquare_def fstOp_def sndOp_def using pick_middlep by simp
 
-lemma wppull_fstO_sndO:
-shows "wppull (P O Q) P Q snd fst fst snd (fstO P Q) (sndO P Q)"
-using pick_middle unfolding wppull_def fstO_def sndO_def relcomp_def by auto
-
 lemma wppull_fstOp_sndOp:
-shows "wppull (Collect (split (P OO Q))) (Collect (split P)) (Collect (split Q)) snd fst fst snd (fstOp P Q) (sndOp P Q)"
+shows "wppull (Collect (split (P OO Q))) (Collect (split P)) (Collect (split Q))
+  snd fst fst snd (fstOp P Q) (sndOp P Q)"
 using pick_middlep unfolding wppull_def fstOp_def sndOp_def relcompp.simps by auto
 
 lemma snd_fst_flip: "snd xy = (fst o (%(x, y). (y, x))) xy"
@@ -214,9 +168,6 @@
 lemma fst_snd_flip: "fst xy = (snd o (%(x, y). (y, x))) xy"
 by (simp split: prod.split)
 
-lemma flip_rel: "A \<subseteq> (R ^-1) \<Longrightarrow> (%(x, y). (y, x)) ` A \<subseteq> R"
-by auto
-
 lemma flip_pred: "A \<subseteq> Collect (split (R ^--1)) \<Longrightarrow> (%(x, y). (y, x)) ` A \<subseteq> Collect (split R)"
 by auto
 
--- a/src/HOL/BNF/BNF_GFP.thy	Tue May 07 17:35:29 2013 +0200
+++ b/src/HOL/BNF/BNF_GFP.thy	Tue May 07 18:40:23 2013 +0200
@@ -76,9 +76,6 @@
 lemma image2_eqI: "\<lbrakk>b = f x; c = g x; x \<in> A\<rbrakk> \<Longrightarrow> (b, c) \<in> image2 A f g"
 unfolding image2_def by auto
 
-lemma Id_subset: "Id \<subseteq> {(a, b). P a b \<or> a = b}"
-by auto
-
 lemma eq_subset: "op = \<le> (\<lambda>a b. P a b \<or> a = b)"
 by auto
 
@@ -88,12 +85,6 @@
 lemma image2_Gr: "image2 A f g = (Gr A f)^-1 O (Gr A g)"
 unfolding image2_def Gr_def by auto
 
-lemma GrI: "\<lbrakk>x \<in> A; f x = fx\<rbrakk> \<Longrightarrow> (x, fx) \<in> Gr A f"
-unfolding Gr_def by simp
-
-lemma GrE: "(x, fx) \<in> Gr A f \<Longrightarrow> (x \<in> A \<Longrightarrow> f x = fx \<Longrightarrow> P) \<Longrightarrow> P"
-unfolding Gr_def by simp
-
 lemma GrD1: "(x, fx) \<in> Gr A f \<Longrightarrow> x \<in> A"
 unfolding Gr_def by simp