src/HOL/Library/Quotient_Sum.thy
changeset 53012 cb82606b8215
parent 53010 ec5e6f69bd65
child 53026 e1a548c11845
--- a/src/HOL/Library/Quotient_Sum.thy	Tue Aug 13 15:59:22 2013 +0200
+++ b/src/HOL/Library/Quotient_Sum.thy	Tue Aug 13 15:59:22 2013 +0200
@@ -1,5 +1,5 @@
 (*  Title:      HOL/Library/Quotient_Sum.thy
-    Author:     Cezary Kaliszyk, Christian Urban and Brian Huffman
+    Author:     Cezary Kaliszyk and Christian Urban
 *)
 
 header {* Quotient infrastructure for the sum type *}
@@ -8,31 +8,7 @@
 imports Main Quotient_Syntax
 begin
 
-subsection {* Relator for sum type *}
-
-fun
-  sum_rel :: "('a \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a + 'b \<Rightarrow> 'c + 'd \<Rightarrow> bool"
-where
-  "sum_rel R1 R2 (Inl a1) (Inl b1) = R1 a1 b1"
-| "sum_rel R1 R2 (Inl a1) (Inr b2) = False"
-| "sum_rel R1 R2 (Inr a2) (Inl b1) = False"
-| "sum_rel R1 R2 (Inr a2) (Inr b2) = R2 a2 b2"
-
-lemma sum_rel_unfold:
-  "sum_rel R1 R2 x y = (case (x, y) of (Inl x, Inl y) \<Rightarrow> R1 x y
-    | (Inr x, Inr y) \<Rightarrow> R2 x y
-    | _ \<Rightarrow> False)"
-  by (cases x) (cases y, simp_all)+
-
-fun sum_pred :: "('a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> 'a + 'b \<Rightarrow> bool"
-where
-  "sum_pred P1 P2 (Inl a) = P1 a"
-| "sum_pred P1 P2 (Inr a) = P2 a"
-
-lemma sum_pred_unfold:
-  "sum_pred P1 P2 x = (case x of Inl x \<Rightarrow> P1 x
-    | Inr x \<Rightarrow> P2 x)"
-by (cases x) simp_all
+subsection {* Rules for the Quotient package *}
 
 lemma sum_rel_map1:
   "sum_rel R1 R2 (sum_map f1 f2 x) y \<longleftrightarrow> sum_rel (\<lambda>x. R1 (f1 x)) (\<lambda>x. R2 (f2 x)) x y"
@@ -46,39 +22,10 @@
   "sum_map id id = id"
   by (simp add: id_def sum_map.identity fun_eq_iff)
 
-lemma sum_rel_eq [id_simps, relator_eq]:
+lemma sum_rel_eq [id_simps]:
   "sum_rel (op =) (op =) = (op =)"
   by (simp add: sum_rel_unfold fun_eq_iff split: sum.split)
 
-lemma sum_rel_mono[relator_mono]:
-  assumes "A \<le> C"
-  assumes "B \<le> D"
-  shows "(sum_rel A B) \<le> (sum_rel C D)"
-using assms by (auto simp: sum_rel_unfold split: sum.splits)
-
-lemma sum_rel_OO[relator_distr]:
-  "(sum_rel A B) OO (sum_rel C D) = sum_rel (A OO C) (B OO D)"
-by (rule ext)+ (auto simp add: sum_rel_unfold OO_def split_sum_ex split: sum.split)
-
-lemma Domainp_sum[relator_domain]:
-  assumes "Domainp R1 = P1"
-  assumes "Domainp R2 = P2"
-  shows "Domainp (sum_rel R1 R2) = (sum_pred P1 P2)"
-using assms
-by (auto simp add: Domainp_iff split_sum_ex sum_pred_unfold iff: fun_eq_iff split: sum.split)
-
-lemma reflp_sum_rel[reflexivity_rule]:
-  "reflp R1 \<Longrightarrow> reflp R2 \<Longrightarrow> reflp (sum_rel R1 R2)"
-  unfolding reflp_def split_sum_all sum_rel.simps by fast
-
-lemma left_total_sum_rel[reflexivity_rule]:
-  "left_total R1 \<Longrightarrow> left_total R2 \<Longrightarrow> left_total (sum_rel R1 R2)"
-  using assms unfolding left_total_def split_sum_all split_sum_ex by simp
-
-lemma left_unique_sum_rel [reflexivity_rule]:
-  "left_unique R1 \<Longrightarrow> left_unique R2 \<Longrightarrow> left_unique (sum_rel R1 R2)"
-  using assms unfolding left_unique_def split_sum_all by simp
-
 lemma sum_symp:
   "symp R1 \<Longrightarrow> symp R2 \<Longrightarrow> symp (sum_rel R1 R2)"
   unfolding symp_def split_sum_all sum_rel.simps by fast
@@ -91,50 +38,6 @@
   "equivp R1 \<Longrightarrow> equivp R2 \<Longrightarrow> equivp (sum_rel R1 R2)"
   by (blast intro: equivpI reflp_sum_rel sum_symp sum_transp elim: equivpE)
 
-lemma right_total_sum_rel [transfer_rule]:
-  "right_total R1 \<Longrightarrow> right_total R2 \<Longrightarrow> right_total (sum_rel R1 R2)"
-  unfolding right_total_def split_sum_all split_sum_ex by simp
-
-lemma right_unique_sum_rel [transfer_rule]:
-  "right_unique R1 \<Longrightarrow> right_unique R2 \<Longrightarrow> right_unique (sum_rel R1 R2)"
-  unfolding right_unique_def split_sum_all by simp
-
-lemma bi_total_sum_rel [transfer_rule]:
-  "bi_total R1 \<Longrightarrow> bi_total R2 \<Longrightarrow> bi_total (sum_rel R1 R2)"
-  using assms unfolding bi_total_def split_sum_all split_sum_ex by simp
-
-lemma bi_unique_sum_rel [transfer_rule]:
-  "bi_unique R1 \<Longrightarrow> bi_unique R2 \<Longrightarrow> bi_unique (sum_rel R1 R2)"
-  using assms unfolding bi_unique_def split_sum_all by simp
-
-subsection {* Transfer rules for transfer package *}
-
-lemma Inl_transfer [transfer_rule]: "(A ===> sum_rel A B) Inl Inl"
-  unfolding fun_rel_def by simp
-
-lemma Inr_transfer [transfer_rule]: "(B ===> sum_rel A B) Inr Inr"
-  unfolding fun_rel_def by simp
-
-lemma sum_case_transfer [transfer_rule]:
-  "((A ===> C) ===> (B ===> C) ===> sum_rel A B ===> C) sum_case sum_case"
-  unfolding fun_rel_def sum_rel_unfold by (simp split: sum.split)
-
-subsection {* Setup for lifting package *}
-
-lemma Quotient_sum[quot_map]:
-  assumes "Quotient R1 Abs1 Rep1 T1"
-  assumes "Quotient R2 Abs2 Rep2 T2"
-  shows "Quotient (sum_rel R1 R2) (sum_map Abs1 Abs2)
-    (sum_map Rep1 Rep2) (sum_rel T1 T2)"
-  using assms unfolding Quotient_alt_def
-  by (simp add: split_sum_all)
-
-lemma sum_invariant_commute [invariant_commute]: 
-  "sum_rel (Lifting.invariant P1) (Lifting.invariant P2) = Lifting.invariant (sum_pred P1 P2)"
-  by (auto simp add: fun_eq_iff Lifting.invariant_def sum_rel_unfold sum_pred_unfold split: sum.split)
-
-subsection {* Rules for quotient package *}
-
 lemma sum_quotient [quot_thm]:
   assumes q1: "Quotient3 R1 Abs1 Rep1"
   assumes q2: "Quotient3 R2 Abs2 Rep2"