# HG changeset patch # User wenzelm # Date 1393013283 -3600 # Node ID 12448c1798517b0924465988c5f3b6b6baf83d20 # Parent b45af39fcdaee2875d6787ee7cd50a70db4ed526 more standard theory name; diff -r b45af39fcdae -r 12448c179851 src/HOL/ROOT --- a/src/HOL/ROOT Fri Feb 21 20:54:13 2014 +0100 +++ b/src/HOL/ROOT Fri Feb 21 21:08:03 2014 +0100 @@ -570,7 +570,7 @@ Simproc_Tests Executable_Relation FinFunPred - Set_Comprehension_Pointfree_Tests + Set_Comprehension_Pointfree_Examples Parallel_Example IArray_Examples SVC_Oracle diff -r b45af39fcdae -r 12448c179851 src/HOL/ex/Set_Comprehension_Pointfree_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Set_Comprehension_Pointfree_Examples.thy Fri Feb 21 21:08:03 2014 +0100 @@ -0,0 +1,140 @@ +(* Title: HOL/ex/Set_Comprehension_Pointfree_Examples.thy + Author: Lukas Bulwahn, Rafal Kolanski + Copyright 2012 TU Muenchen +*) + +header {* Examples for the set comprehension to pointfree simproc *} + +theory Set_Comprehension_Pointfree_Examples +imports Main +begin + +declare [[simproc add: finite_Collect]] + +lemma + "finite (UNIV::'a set) ==> finite {p. EX x::'a. p = (x, x)}" + by simp + +lemma + "finite A ==> finite B ==> finite {f a b| a b. a : A \ b : B}" + by simp + +lemma + "finite B ==> finite A' ==> finite {f a b| a b. a : A \ a : A' \ b : B}" + by simp + +lemma + "finite A ==> finite B ==> finite {f a b| a b. a : A \ b : B \ b : B'}" + by simp + +lemma + "finite A ==> finite B ==> finite C ==> finite {f a b c| a b c. a : A \ b : B \ c : C}" + by simp + +lemma + "finite A ==> finite B ==> finite C ==> finite D ==> + finite {f a b c d| a b c d. a : A \ b : B \ c : C \ d : D}" + by simp + +lemma + "finite A ==> finite B ==> finite C ==> finite D ==> finite E ==> + finite {f a b c d e | a b c d e. a : A \ b : B \ c : C \ d : D \ e : E}" + by simp + +lemma + "finite A ==> finite B ==> finite C ==> finite D ==> finite E \ + finite {f a d c b e | e b c d a. b : B \ a : A \ e : E' \ c : C \ d : D \ e : E \ b : B'}" + by simp + +lemma + "\ finite A ; finite B ; finite C ; finite D \ + \ finite ({f a b c d| a b c d. a : A \ b : B \ c : C \ d : D})" + by simp + +lemma + "finite ((\(a,b,c,d). f a b c d) ` (A \ B \ C \ D)) + \ finite ({f a b c d| a b c d. a : A \ b : B \ c : C \ d : D})" + by simp + +lemma + "finite S ==> finite {s'. EX s:S. s' = f a e s}" + by simp + +lemma + "finite A ==> finite B ==> finite {f a b| a b. a : A \ b : B \ a \ Z}" + by simp + +lemma + "finite A ==> finite B ==> finite R ==> finite {f a b x y| a b x y. a : A \ b : B \ (x,y) \ R}" +by simp + +lemma + "finite A ==> finite B ==> finite R ==> finite {f a b x y| a b x y. a : A \ (x,y) \ R \ b : B}" +by simp + +lemma + "finite A ==> finite B ==> finite R ==> finite {f a (x, b) y| y b x a. a : A \ (x,y) \ R \ b : B}" +by simp + +lemma + "finite A ==> finite AA ==> finite B ==> finite {f a b| a b. (a : A \ a : AA) \ b : B \ a \ Z}" +by simp + +lemma + "finite A1 ==> finite A2 ==> finite A3 ==> finite A4 ==> finite A5 ==> finite B ==> + finite {f a b c | a b c. ((a : A1 \ a : A2) \ (a : A3 \ (a : A4 \ a : A5))) \ b : B \ a \ Z}" +apply simp +oops + +lemma "finite B ==> finite {c. EX x. x : B & c = a * x}" +by simp + +lemma + "finite A ==> finite B ==> finite {f a * g b |a b. a : A & b : B}" +by simp + +lemma + "finite S ==> inj (%(x, y). g x y) ==> finite {f x y| x y. g x y : S}" + by (auto intro: finite_vimageI) + +lemma + "finite A ==> finite S ==> inj (%(x, y). g x y) ==> finite {f x y z | x y z. g x y : S & z : A}" + by (auto intro: finite_vimageI) + +lemma + "finite S ==> finite A ==> inj (%(x, y). g x y) ==> inj (%(x, y). h x y) + ==> finite {f a b c d | a b c d. g a c : S & h b d : A}" + by (auto intro: finite_vimageI) + +lemma + assumes "finite S" shows "finite {(a,b,c,d). ([a, b], [c, d]) : S}" +using assms by (auto intro!: finite_vimageI simp add: inj_on_def) + (* injectivity to be automated with further rules and automation *) + +schematic_lemma (* check interaction with schematics *) + "finite {x :: ?'A \ ?'B \ bool. \a b. x = Pair_Rep a b} + = finite ((\(b :: ?'B, a:: ?'A). Pair_Rep a b) ` (UNIV \ UNIV))" + by simp + +declare [[simproc del: finite_Collect]] + + +section {* Testing simproc in code generation *} + +definition union :: "nat set => nat set => nat set" +where + "union A B = {x. x : A \ x : B}" + +definition common_subsets :: "nat set => nat set => nat set set" +where + "common_subsets S1 S2 = {S. S \ S1 \ S \ S2}" + +definition products :: "nat set => nat set => nat set" +where + "products A B = {c. EX a b. a : A & b : B & c = a * b}" + +export_code products in Haskell + +export_code union common_subsets products in Haskell + +end diff -r b45af39fcdae -r 12448c179851 src/HOL/ex/Set_Comprehension_Pointfree_Tests.thy --- a/src/HOL/ex/Set_Comprehension_Pointfree_Tests.thy Fri Feb 21 20:54:13 2014 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,140 +0,0 @@ -(* Title: HOL/ex/Set_Comprehension_Pointfree_Tests.thy - Author: Lukas Bulwahn, Rafal Kolanski - Copyright 2012 TU Muenchen -*) - -header {* Tests for the set comprehension to pointfree simproc *} - -theory Set_Comprehension_Pointfree_Tests -imports Main -begin - -declare [[simproc add: finite_Collect]] - -lemma - "finite (UNIV::'a set) ==> finite {p. EX x::'a. p = (x, x)}" - by simp - -lemma - "finite A ==> finite B ==> finite {f a b| a b. a : A \ b : B}" - by simp - -lemma - "finite B ==> finite A' ==> finite {f a b| a b. a : A \ a : A' \ b : B}" - by simp - -lemma - "finite A ==> finite B ==> finite {f a b| a b. a : A \ b : B \ b : B'}" - by simp - -lemma - "finite A ==> finite B ==> finite C ==> finite {f a b c| a b c. a : A \ b : B \ c : C}" - by simp - -lemma - "finite A ==> finite B ==> finite C ==> finite D ==> - finite {f a b c d| a b c d. a : A \ b : B \ c : C \ d : D}" - by simp - -lemma - "finite A ==> finite B ==> finite C ==> finite D ==> finite E ==> - finite {f a b c d e | a b c d e. a : A \ b : B \ c : C \ d : D \ e : E}" - by simp - -lemma - "finite A ==> finite B ==> finite C ==> finite D ==> finite E \ - finite {f a d c b e | e b c d a. b : B \ a : A \ e : E' \ c : C \ d : D \ e : E \ b : B'}" - by simp - -lemma - "\ finite A ; finite B ; finite C ; finite D \ - \ finite ({f a b c d| a b c d. a : A \ b : B \ c : C \ d : D})" - by simp - -lemma - "finite ((\(a,b,c,d). f a b c d) ` (A \ B \ C \ D)) - \ finite ({f a b c d| a b c d. a : A \ b : B \ c : C \ d : D})" - by simp - -lemma - "finite S ==> finite {s'. EX s:S. s' = f a e s}" - by simp - -lemma - "finite A ==> finite B ==> finite {f a b| a b. a : A \ b : B \ a \ Z}" - by simp - -lemma - "finite A ==> finite B ==> finite R ==> finite {f a b x y| a b x y. a : A \ b : B \ (x,y) \ R}" -by simp - -lemma - "finite A ==> finite B ==> finite R ==> finite {f a b x y| a b x y. a : A \ (x,y) \ R \ b : B}" -by simp - -lemma - "finite A ==> finite B ==> finite R ==> finite {f a (x, b) y| y b x a. a : A \ (x,y) \ R \ b : B}" -by simp - -lemma - "finite A ==> finite AA ==> finite B ==> finite {f a b| a b. (a : A \ a : AA) \ b : B \ a \ Z}" -by simp - -lemma - "finite A1 ==> finite A2 ==> finite A3 ==> finite A4 ==> finite A5 ==> finite B ==> - finite {f a b c | a b c. ((a : A1 \ a : A2) \ (a : A3 \ (a : A4 \ a : A5))) \ b : B \ a \ Z}" -apply simp -oops - -lemma "finite B ==> finite {c. EX x. x : B & c = a * x}" -by simp - -lemma - "finite A ==> finite B ==> finite {f a * g b |a b. a : A & b : B}" -by simp - -lemma - "finite S ==> inj (%(x, y). g x y) ==> finite {f x y| x y. g x y : S}" - by (auto intro: finite_vimageI) - -lemma - "finite A ==> finite S ==> inj (%(x, y). g x y) ==> finite {f x y z | x y z. g x y : S & z : A}" - by (auto intro: finite_vimageI) - -lemma - "finite S ==> finite A ==> inj (%(x, y). g x y) ==> inj (%(x, y). h x y) - ==> finite {f a b c d | a b c d. g a c : S & h b d : A}" - by (auto intro: finite_vimageI) - -lemma - assumes "finite S" shows "finite {(a,b,c,d). ([a, b], [c, d]) : S}" -using assms by (auto intro!: finite_vimageI simp add: inj_on_def) - (* injectivity to be automated with further rules and automation *) - -schematic_lemma (* check interaction with schematics *) - "finite {x :: ?'A \ ?'B \ bool. \a b. x = Pair_Rep a b} - = finite ((\(b :: ?'B, a:: ?'A). Pair_Rep a b) ` (UNIV \ UNIV))" - by simp - -declare [[simproc del: finite_Collect]] - - -section {* Testing simproc in code generation *} - -definition union :: "nat set => nat set => nat set" -where - "union A B = {x. x : A \ x : B}" - -definition common_subsets :: "nat set => nat set => nat set set" -where - "common_subsets S1 S2 = {S. S \ S1 \ S \ S2}" - -definition products :: "nat set => nat set => nat set" -where - "products A B = {c. EX a b. a : A & b : B & c = a * b}" - -export_code products in Haskell - -export_code union common_subsets products in Haskell - -end