# HG changeset patch # User haftmann # Date 1289307733 -3600 # Node ID 2989f9f3aa10209de09a30d2b8b4cfbe68c83a7c # Parent e1db06cf625400b8e7ff67a133b6c4cfbb318119 more appropriate specification packages; fun_rel_def is no simp rule by default diff -r e1db06cf6254 -r 2989f9f3aa10 src/HOL/Library/Quotient_Product.thy --- a/src/HOL/Library/Quotient_Product.thy Tue Nov 09 14:02:12 2010 +0100 +++ b/src/HOL/Library/Quotient_Product.thy Tue Nov 09 14:02:13 2010 +0100 @@ -8,13 +8,16 @@ imports Main Quotient_Syntax begin -fun - prod_rel +definition + prod_rel :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ 'a \ 'b \ 'a \ 'b \ bool" where "prod_rel R1 R2 = (\(a, b) (c, d). R1 a c \ R2 b d)" declare [[map prod = (prod_fun, prod_rel)]] +lemma prod_rel_apply [simp]: + "prod_rel R1 R2 (a, b) (c, d) \ R1 a c \ R2 b d" + by (simp add: prod_rel_def) lemma prod_equivp[quot_equiv]: assumes a: "equivp R1" @@ -22,7 +25,7 @@ shows "equivp (prod_rel R1 R2)" apply(rule equivpI) unfolding reflp_def symp_def transp_def - apply(simp_all add: split_paired_all) + apply(simp_all add: split_paired_all prod_rel_def) apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b]) apply(blast intro: equivp_symp[OF a] equivp_symp[OF b]) apply(blast intro: equivp_transp[OF a] equivp_transp[OF b]) @@ -45,7 +48,7 @@ assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" shows "(R1 ===> R2 ===> prod_rel R1 R2) Pair Pair" - by simp + by (auto simp add: prod_rel_def) lemma Pair_prs[quot_preserve]: assumes q1: "Quotient R1 Abs1 Rep1" @@ -59,33 +62,29 @@ assumes "Quotient R1 Abs1 Rep1" assumes "Quotient R2 Abs2 Rep2" shows "(prod_rel R1 R2 ===> R1) fst fst" - by simp + by auto lemma fst_prs[quot_preserve]: assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" shows "(prod_fun Rep1 Rep2 ---> Abs1) fst = fst" - apply(simp add: fun_eq_iff) - apply(simp add: Quotient_abs_rep[OF q1]) - done + by (simp add: fun_eq_iff Quotient_abs_rep[OF q1]) lemma snd_rsp[quot_respect]: assumes "Quotient R1 Abs1 Rep1" assumes "Quotient R2 Abs2 Rep2" shows "(prod_rel R1 R2 ===> R2) snd snd" - by simp + by auto lemma snd_prs[quot_preserve]: assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" shows "(prod_fun Rep1 Rep2 ---> Abs2) snd = snd" - apply(simp add: fun_eq_iff) - apply(simp add: Quotient_abs_rep[OF q2]) - done + by (simp add: fun_eq_iff Quotient_abs_rep[OF q2]) lemma split_rsp[quot_respect]: shows "((R1 ===> R2 ===> (op =)) ===> (prod_rel R1 R2) ===> (op =)) split split" - by auto + by (auto intro!: fun_relI elim!: fun_relE) lemma split_prs[quot_preserve]: assumes q1: "Quotient R1 Abs1 Rep1" @@ -96,7 +95,7 @@ lemma [quot_respect]: shows "((R2 ===> R2 ===> op =) ===> (R1 ===> R1 ===> op =) ===> prod_rel R2 R1 ===> prod_rel R2 R1 ===> op =) prod_rel prod_rel" - by auto + by (auto simp add: fun_rel_def) lemma [quot_preserve]: assumes q1: "Quotient R1 abs1 rep1" @@ -114,7 +113,7 @@ lemma prod_fun_id[id_simps]: shows "prod_fun id id = id" - by (simp add: prod_fun_def) + by (simp add: fun_eq_iff) lemma prod_rel_eq[id_simps]: shows "prod_rel (op =) (op =) = (op =)" diff -r e1db06cf6254 -r 2989f9f3aa10 src/HOL/Library/Quotient_Sum.thy --- a/src/HOL/Library/Quotient_Sum.thy Tue Nov 09 14:02:12 2010 +0100 +++ b/src/HOL/Library/Quotient_Sum.thy Tue Nov 09 14:02:13 2010 +0100 @@ -9,15 +9,15 @@ begin fun - sum_rel + sum_rel :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ 'a + 'b \ 'a + 'b \ 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" -fun - sum_map +primrec + sum_map :: "('a \ 'c) \ ('b \ 'd) \ 'a + 'b \ 'c + 'd" where "sum_map f1 f2 (Inl a) = Inl (f1 a)" | "sum_map f1 f2 (Inr a) = Inr (f2 a)" @@ -62,13 +62,13 @@ assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" shows "(R1 ===> sum_rel R1 R2) Inl Inl" - by simp + by auto lemma sum_Inr_rsp[quot_respect]: assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" shows "(R2 ===> sum_rel R1 R2) Inr Inr" - by simp + by auto lemma sum_Inl_prs[quot_preserve]: assumes q1: "Quotient R1 Abs1 Rep1"