--- a/src/HOL/Lifting_Sum.thy Wed Aug 14 00:15:03 2013 +0200
+++ b/src/HOL/Lifting_Sum.thy Tue Aug 13 18:22:55 2013 +0200
@@ -5,59 +5,52 @@
header {* Setup for Lifting/Transfer for the sum type *}
theory Lifting_Sum
-imports Lifting FunDef
+imports Lifting
begin
subsection {* Relator and predicator properties *}
-fun
- sum_rel :: "('a \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a + 'b \<Rightarrow> 'c + 'd \<Rightarrow> bool"
+definition
+ 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"
+ "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)"
-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)+
+lemma sum_rel_simps[simp]:
+ "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"
+ unfolding sum_rel_def by 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
+abbreviation (input) "sum_pred \<equiv> sum_case"
lemma sum_rel_eq [relator_eq]:
"sum_rel (op =) (op =) = (op =)"
- by (simp add: sum_rel_unfold fun_eq_iff split: sum.split)
+ by (simp add: sum_rel_def 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)
+using assms by (auto simp: sum_rel_def 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)
+by (rule ext)+ (auto simp add: sum_rel_def 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)
+by (auto simp add: Domainp_iff split_sum_ex 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
+ 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)"
@@ -85,7 +78,7 @@
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)
+ by (auto simp add: fun_eq_iff Lifting.invariant_def sum_rel_def split: sum.split)
subsection {* Quotient theorem for the Lifting package *}
@@ -111,7 +104,7 @@
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)
+ unfolding fun_rel_def sum_rel_def by (simp split: sum.split)
end