more eq_reflection & other cleaning.
--- a/Quot/QuotBase.thy Tue Jan 26 07:14:10 2010 +0100
+++ b/Quot/QuotBase.thy Tue Jan 26 07:42:52 2010 +0100
@@ -395,10 +395,9 @@
lemma babs_prs:
assumes q1: "Quotient R1 Abs1 Rep1"
and q2: "Quotient R2 Abs2 Rep2"
- shows "(Rep1 ---> Abs2) (Babs (Respects R1) ((Abs1 ---> Rep2) f)) \<equiv> f"
- apply(rule eq_reflection)
- apply(rule ext)
- apply simp
+ shows "((Rep1 ---> Abs2) (Babs (Respects R1) ((Abs1 ---> Rep2) f))) = f"
+ apply (rule ext)
+ apply (simp)
apply (subgoal_tac "Rep1 x \<in> Respects R1")
apply (simp add: Babs_def Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
apply (simp add: in_respects Quotient_rel_rep[OF q1])
--- a/Quot/QuotList.thy Tue Jan 26 07:14:10 2010 +0100
+++ b/Quot/QuotList.thy Tue Jan 26 07:42:52 2010 +0100
@@ -59,8 +59,7 @@
apply(rule list_rel_rel[OF q])
done
-lemma map_id[id_simps]: "map id \<equiv> id"
- apply (rule eq_reflection)
+lemma map_id[id_simps]: "map id = id"
apply (rule ext)
apply (rule_tac list="x" in list.induct)
apply (simp_all)
@@ -69,132 +68,128 @@
lemma cons_prs_aux:
assumes q: "Quotient R Abs Rep"
shows "(map Abs) ((Rep h) # (map Rep t)) = h # t"
-by (induct t) (simp_all add: Quotient_abs_rep[OF q])
+ by (induct t) (simp_all add: Quotient_abs_rep[OF q])
lemma cons_prs[quot_preserve]:
assumes q: "Quotient R Abs Rep"
shows "(Rep ---> (map Rep) ---> (map Abs)) (op #) = (op #)"
-by (simp only: expand_fun_eq fun_map_def cons_prs_aux[OF q])
- (simp)
+ by (simp only: expand_fun_eq fun_map_def cons_prs_aux[OF q]) (simp)
lemma cons_rsp[quot_respect]:
assumes q: "Quotient R Abs Rep"
shows "(R ===> list_rel R ===> list_rel R) op # op #"
-by (auto)
+ by auto
lemma nil_prs[quot_preserve]:
assumes q: "Quotient R Abs Rep"
- shows "map Abs [] \<equiv> []"
-by (simp)
+ shows "map Abs [] = []"
+ by simp
lemma nil_rsp[quot_respect]:
assumes q: "Quotient R Abs Rep"
shows "list_rel R [] []"
-by simp
+ by simp
lemma map_prs_aux:
assumes a: "Quotient R1 abs1 rep1"
and b: "Quotient R2 abs2 rep2"
shows "(map abs2) (map ((abs1 ---> rep2) f) (map rep1 l)) = map f l"
-by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
+ by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
lemma map_prs[quot_preserve]:
assumes a: "Quotient R1 abs1 rep1"
and b: "Quotient R2 abs2 rep2"
shows "((abs1 ---> rep2) ---> (map rep1) ---> (map abs2)) map = map"
-by (simp only: expand_fun_eq fun_map_def map_prs_aux[OF a b])
- (simp)
+ by (simp only: expand_fun_eq fun_map_def map_prs_aux[OF a b]) (simp)
lemma map_rsp[quot_respect]:
assumes q1: "Quotient R1 Abs1 Rep1"
and q2: "Quotient R2 Abs2 Rep2"
shows "((R1 ===> R2) ===> (list_rel R1) ===> list_rel R2) map map"
-apply(simp)
-apply(rule allI)+
-apply(rule impI)
-apply(rule allI)+
-apply (induct_tac xa ya rule: list_induct2')
-apply simp_all
-done
+ apply(simp)
+ apply(rule allI)+
+ apply(rule impI)
+ apply(rule allI)+
+ apply (induct_tac xa ya rule: list_induct2')
+ apply simp_all
+ done
lemma foldr_prs_aux:
assumes a: "Quotient R1 abs1 rep1"
and b: "Quotient R2 abs2 rep2"
shows "abs2 (foldr ((abs1 ---> abs2 ---> rep2) f) (map rep1 l) (rep2 e)) = foldr f l e"
-by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
+ by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
lemma foldr_prs[quot_preserve]:
assumes a: "Quotient R1 abs1 rep1"
and b: "Quotient R2 abs2 rep2"
shows "((abs1 ---> abs2 ---> rep2) ---> (map rep1) ---> rep2 ---> abs2) foldr = foldr"
-by (simp only: expand_fun_eq fun_map_def foldr_prs_aux[OF a b])
- (simp)
+ by (simp only: expand_fun_eq fun_map_def foldr_prs_aux[OF a b]) (simp)
lemma foldl_prs_aux:
assumes a: "Quotient R1 abs1 rep1"
and b: "Quotient R2 abs2 rep2"
shows "abs1 (foldl ((abs1 ---> abs2 ---> rep1) f) (rep1 e) (map rep2 l)) = foldl f e l"
-by (induct l arbitrary:e) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
-
+ by (induct l arbitrary:e) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
lemma foldl_prs[quot_preserve]:
assumes a: "Quotient R1 abs1 rep1"
and b: "Quotient R2 abs2 rep2"
shows "((abs1 ---> abs2 ---> rep1) ---> rep1 ---> (map rep2) ---> abs1) foldl = foldl"
-by (simp only: expand_fun_eq fun_map_def foldl_prs_aux[OF a b])
- (simp)
+ by (simp only: expand_fun_eq fun_map_def foldl_prs_aux[OF a b]) (simp)
-lemma list_rel_empty: "list_rel R [] b \<Longrightarrow> length b = 0"
-by (induct b) (simp_all)
+lemma list_rel_empty:
+ "list_rel R [] b \<Longrightarrow> length b = 0"
+ by (induct b) (simp_all)
-lemma list_rel_len: "list_rel R a b \<Longrightarrow> length a = length b"
-apply (induct a arbitrary: b)
-apply (simp add: list_rel_empty)
-apply (case_tac b)
-apply simp_all
-done
+lemma list_rel_len:
+ "list_rel R a b \<Longrightarrow> length a = length b"
+ apply (induct a arbitrary: b)
+ apply (simp add: list_rel_empty)
+ apply (case_tac b)
+ apply simp_all
+ done
(* induct_tac doesn't accept 'arbitrary', so we manually 'spec' *)
lemma foldl_rsp[quot_respect]:
assumes q1: "Quotient R1 Abs1 Rep1"
and q2: "Quotient R2 Abs2 Rep2"
shows "((R1 ===> R2 ===> R1) ===> R1 ===> list_rel R2 ===> R1) foldl foldl"
-apply(auto)
-apply (subgoal_tac "R1 xa ya \<longrightarrow> list_rel R2 xb yb \<longrightarrow> R1 (foldl x xa xb) (foldl y ya yb)")
-apply simp
-apply (rule_tac x="xa" in spec)
-apply (rule_tac x="ya" in spec)
-apply (rule_tac xs="xb" and ys="yb" in list_induct2)
-apply (rule list_rel_len)
-apply (simp_all)
-done
+ apply(auto)
+ apply (subgoal_tac "R1 xa ya \<longrightarrow> list_rel R2 xb yb \<longrightarrow> R1 (foldl x xa xb) (foldl y ya yb)")
+ apply simp
+ apply (rule_tac x="xa" in spec)
+ apply (rule_tac x="ya" in spec)
+ apply (rule_tac xs="xb" and ys="yb" in list_induct2)
+ apply (rule list_rel_len)
+ apply (simp_all)
+ done
lemma foldr_rsp[quot_respect]:
assumes q1: "Quotient R1 Abs1 Rep1"
and q2: "Quotient R2 Abs2 Rep2"
shows "((R1 ===> R2 ===> R2) ===> list_rel R1 ===> R2 ===> R2) foldr foldr"
-apply auto
-apply(subgoal_tac "R2 xb yb \<longrightarrow> list_rel R1 xa ya \<longrightarrow> R2 (foldr x xa xb) (foldr y ya yb)")
-apply simp
-apply (rule_tac xs="xa" and ys="ya" in list_induct2)
-apply (rule list_rel_len)
-apply (simp_all)
-done
+ apply auto
+ apply(subgoal_tac "R2 xb yb \<longrightarrow> list_rel R1 xa ya \<longrightarrow> R2 (foldr x xa xb) (foldr y ya yb)")
+ apply simp
+ apply (rule_tac xs="xa" and ys="ya" in list_induct2)
+ apply (rule list_rel_len)
+ apply (simp_all)
+ done
lemma list_rel_eq[id_simps]:
- shows "list_rel (op =) \<equiv> (op =)"
-apply(rule eq_reflection)
-unfolding expand_fun_eq
-apply(rule allI)+
-apply(induct_tac x xa rule: list_induct2')
-apply(simp_all)
-done
+ shows "(list_rel (op =)) = (op =)"
+ unfolding expand_fun_eq
+ apply(rule allI)+
+ apply(induct_tac x xa rule: list_induct2')
+ apply(simp_all)
+ done
lemma list_rel_refl:
assumes a: "\<And>x y. R x y = (R x = R y)"
shows "list_rel R x x"
-by (induct x) (auto simp add: a)
+ by (induct x) (auto simp add: a)
end
--- a/Quot/QuotOption.thy Tue Jan 26 07:14:10 2010 +0100
+++ b/Quot/QuotOption.thy Tue Jan 26 07:42:52 2010 +0100
@@ -73,17 +73,15 @@
apply(simp add: Quotient_abs_rep[OF q])
done
-lemma option_map_id[id_simps]:
- shows "option_map id \<equiv> id"
- apply (rule eq_reflection)
+lemma option_map_id[id_simps]:
+ shows "option_map id = id"
apply (auto simp add: expand_fun_eq)
apply (case_tac x)
apply (auto)
done
-lemma option_rel_eq[id_simps]:
- shows "option_rel (op =) \<equiv> (op =)"
- apply(rule eq_reflection)
+lemma option_rel_eq[id_simps]:
+ shows "option_rel (op =) = (op =)"
apply(auto simp add: expand_fun_eq)
apply(case_tac x)
apply(case_tac [!] xa)