| author | paulson <lp15@cam.ac.uk> | 
| Sun, 25 Feb 2018 12:54:55 +0000 | |
| changeset 67719 | bffb7482faaa | 
| parent 67613 | ce654b0e6d69 | 
| child 69624 | e02bdf853a4c | 
| permissions | -rw-r--r-- | 
| 55663 | 1  | 
(* Title: HOL/ex/Set_Comprehension_Pointfree_Examples.thy  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
2  | 
Author: Lukas Bulwahn, Rafal Kolanski  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
3  | 
Copyright 2012 TU Muenchen  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
5  | 
|
| 61343 | 6  | 
section \<open>Examples for the set comprehension to pointfree simproc\<close>  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
7  | 
|
| 55663 | 8  | 
theory Set_Comprehension_Pointfree_Examples  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
9  | 
imports Main  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
10  | 
begin  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
|
| 
54611
 
31afce809794
set_comprehension_pointfree simproc causes to many surprises if enabled by default
 
traytel 
parents: 
49947 
diff
changeset
 | 
12  | 
declare [[simproc add: finite_Collect]]  | 
| 
 
31afce809794
set_comprehension_pointfree simproc causes to many surprises if enabled by default
 
traytel 
parents: 
49947 
diff
changeset
 | 
13  | 
|
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
14  | 
lemma  | 
| 67613 | 15  | 
  "finite (UNIV::'a set) \<Longrightarrow> finite {p. \<exists>x::'a. p = (x, x)}"
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
16  | 
by simp  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
lemma  | 
| 67613 | 19  | 
  "finite A \<Longrightarrow> finite B \<Longrightarrow> finite {f a b| a b. a \<in> A \<and> b \<in> B}"
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
20  | 
by simp  | 
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
21  | 
|
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
22  | 
lemma  | 
| 67613 | 23  | 
  "finite B \<Longrightarrow> finite A' \<Longrightarrow> finite {f a b| a b. a \<in> A \<and> a \<in> A' \<and> b \<in> B}"
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
24  | 
by simp  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
lemma  | 
| 67613 | 27  | 
  "finite A \<Longrightarrow> finite B \<Longrightarrow> finite {f a b| a b. a \<in> A \<and> b \<in> B \<and> b \<in> B'}"
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
28  | 
by simp  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
30  | 
lemma  | 
| 67613 | 31  | 
  "finite A \<Longrightarrow> finite B \<Longrightarrow> finite C \<Longrightarrow> finite {f a b c| a b c. a \<in> A \<and> b \<in> B \<and> c \<in> C}"
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
32  | 
by simp  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
33  | 
|
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
34  | 
lemma  | 
| 67613 | 35  | 
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite C \<Longrightarrow> finite D \<Longrightarrow>  | 
36  | 
     finite {f a b c d| a b c d. a \<in> A \<and> b \<in> B \<and> c \<in> C \<and> d \<in> D}"
 | 
|
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
37  | 
by simp  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
38  | 
|
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
39  | 
lemma  | 
| 67613 | 40  | 
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite C \<Longrightarrow> finite D \<Longrightarrow> finite E \<Longrightarrow>  | 
41  | 
    finite {f a b c d e | a b c d e. a \<in> A \<and> b \<in> B \<and> c \<in> C \<and> d \<in> D \<and> e \<in> E}"
 | 
|
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
42  | 
by simp  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
43  | 
|
| 
49899
 
1f0b7d5bea4e
set_comprehension_pointfree simproc now handles the complicated test case; tuned
 
bulwahn 
parents: 
49876 
diff
changeset
 | 
44  | 
lemma  | 
| 67613 | 45  | 
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite C \<Longrightarrow> finite D \<Longrightarrow> finite E \<Longrightarrow>  | 
46  | 
    finite {f a d c b e | e b c d a. b \<in> B \<and> a \<in> A \<and> e \<in> E' \<and> c \<in> C \<and> d \<in> D \<and> e \<in> E \<and> b \<in> B'}"
 | 
|
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
47  | 
by simp  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
48  | 
|
| 
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
49  | 
lemma  | 
| 
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
50  | 
"\<lbrakk> finite A ; finite B ; finite C ; finite D \<rbrakk>  | 
| 67613 | 51  | 
  \<Longrightarrow> finite ({f a b c d| a b c d. a \<in> A \<and> b \<in> B \<and> c \<in> C \<and> d \<in> D})"
 | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
52  | 
by simp  | 
| 
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
53  | 
|
| 
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
54  | 
lemma  | 
| 
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
55  | 
"finite ((\<lambda>(a,b,c,d). f a b c d) ` (A \<times> B \<times> C \<times> D))  | 
| 67613 | 56  | 
  \<Longrightarrow> finite ({f a b c d| a b c d. a \<in> A \<and> b \<in> B \<and> c \<in> C \<and> d \<in> D})"
 | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
57  | 
by simp  | 
| 
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
58  | 
|
| 
49762
 
b5e355c41de3
adding some example that motivates some of the current changes in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
48174 
diff
changeset
 | 
59  | 
lemma  | 
| 67613 | 60  | 
  "finite S \<Longrightarrow> finite {s'. \<exists>s\<in>S. s' = f a e s}"
 | 
| 
49766
 
65318db3087b
test case for set_comprehension_pointfree simproc succeeds now
 
bulwahn 
parents: 
49762 
diff
changeset
 | 
61  | 
by simp  | 
| 
49762
 
b5e355c41de3
adding some example that motivates some of the current changes in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
48174 
diff
changeset
 | 
62  | 
|
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
63  | 
lemma  | 
| 67613 | 64  | 
  "finite A \<Longrightarrow> finite B \<Longrightarrow> finite {f a b| a b. a \<in> A \<and> b \<in> B \<and> a \<notin> Z}"
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
65  | 
by simp  | 
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
66  | 
|
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
67  | 
lemma  | 
| 67613 | 68  | 
  "finite A \<Longrightarrow> finite B \<Longrightarrow> finite R \<Longrightarrow> finite {f a b x y| a b x y. a \<in> A \<and> b \<in> B \<and> (x,y) \<in> R}"
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
69  | 
by simp  | 
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
70  | 
|
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
71  | 
lemma  | 
| 67613 | 72  | 
  "finite A \<Longrightarrow> finite B \<Longrightarrow> finite R \<Longrightarrow> finite {f a b x y| a b x y. a \<in> A \<and> (x,y) \<in> R \<and> b \<in> B}"
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
73  | 
by simp  | 
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
74  | 
|
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
75  | 
lemma  | 
| 67613 | 76  | 
  "finite A \<Longrightarrow> finite B \<Longrightarrow> finite R \<Longrightarrow> finite {f a (x, b) y| y b x a. a \<in> A \<and> (x,y) \<in> R \<and> b \<in> B}"
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
77  | 
by simp  | 
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
78  | 
|
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
79  | 
lemma  | 
| 67613 | 80  | 
  "finite A \<Longrightarrow> finite AA \<Longrightarrow> finite B \<Longrightarrow> finite {f a b| a b. (a \<in> A \<or> a \<in> AA) \<and> b \<in> B \<and> a \<notin> Z}"
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
81  | 
by simp  | 
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
82  | 
|
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
83  | 
lemma  | 
| 67613 | 84  | 
"finite A1 \<Longrightarrow> finite A2 \<Longrightarrow> finite A3 \<Longrightarrow> finite A4 \<Longrightarrow> finite A5 \<Longrightarrow> finite B \<Longrightarrow>  | 
85  | 
     finite {f a b c | a b c. ((a \<in> A1 \<and> a \<in> A2) \<or> (a \<in> A3 \<and> (a \<in> A4 \<or> a \<in> A5))) \<and> b \<in> B \<and> a \<notin> Z}"
 | 
|
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
86  | 
apply simp  | 
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
87  | 
oops  | 
| 
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
88  | 
|
| 67613 | 89  | 
lemma "finite B \<Longrightarrow> finite {c. \<exists>x. x \<in> B \<and> c = a * x}"
 | 
| 
49853
 
875ed58b3b65
adding further test cases for the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49851 
diff
changeset
 | 
90  | 
by simp  | 
| 
 
875ed58b3b65
adding further test cases for the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49851 
diff
changeset
 | 
91  | 
|
| 
 
875ed58b3b65
adding further test cases for the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49851 
diff
changeset
 | 
92  | 
lemma  | 
| 67613 | 93  | 
  "finite A \<Longrightarrow> finite B \<Longrightarrow> finite {f a * g b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 
49853
 
875ed58b3b65
adding further test cases for the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49851 
diff
changeset
 | 
94  | 
by simp  | 
| 
 
875ed58b3b65
adding further test cases for the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49851 
diff
changeset
 | 
95  | 
|
| 
49876
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
96  | 
lemma  | 
| 67613 | 97  | 
  "finite S \<Longrightarrow> inj (\<lambda>(x, y). g x y) \<Longrightarrow> finite {f x y| x y. g x y \<in> S}"
 | 
| 
49876
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
98  | 
by (auto intro: finite_vimageI)  | 
| 
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
99  | 
|
| 
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
100  | 
lemma  | 
| 67613 | 101  | 
  "finite A \<Longrightarrow> finite S \<Longrightarrow> inj (\<lambda>(x, y). g x y) \<Longrightarrow> finite {f x y z | x y z. g x y \<in> S & z \<in> A}"
 | 
| 
49876
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
102  | 
by (auto intro: finite_vimageI)  | 
| 
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
103  | 
|
| 
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
104  | 
lemma  | 
| 67613 | 105  | 
"finite S \<Longrightarrow> finite A \<Longrightarrow> inj (\<lambda>(x, y). g x y) \<Longrightarrow> inj (\<lambda>(x, y). h x y)  | 
106  | 
    \<Longrightarrow> finite {f a b c d | a b c d. g a c \<in> S \<and> h b d \<in> A}"
 | 
|
| 
49876
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
107  | 
by (auto intro: finite_vimageI)  | 
| 
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
108  | 
|
| 
49899
 
1f0b7d5bea4e
set_comprehension_pointfree simproc now handles the complicated test case; tuned
 
bulwahn 
parents: 
49876 
diff
changeset
 | 
109  | 
lemma  | 
| 67613 | 110  | 
  assumes "finite S" shows "finite {(a,b,c,d). ([a, b], [c, d]) \<in> S}"
 | 
| 
49899
 
1f0b7d5bea4e
set_comprehension_pointfree simproc now handles the complicated test case; tuned
 
bulwahn 
parents: 
49876 
diff
changeset
 | 
111  | 
using assms by (auto intro!: finite_vimageI simp add: inj_on_def)  | 
| 
 
1f0b7d5bea4e
set_comprehension_pointfree simproc now handles the complicated test case; tuned
 
bulwahn 
parents: 
49876 
diff
changeset
 | 
112  | 
(* injectivity to be automated with further rules and automation *)  | 
| 
49876
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
113  | 
|
| 61337 | 114  | 
schematic_goal (* check interaction with schematics *)  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
115  | 
  "finite {x :: ?'A \<Rightarrow> ?'B \<Rightarrow> bool. \<exists>a b. x = Pair_Rep a b}
 | 
| 
49851
 
4d33963962fa
adding further test cases to check new functionality of the simproc; strengthened test cases to check the success of the simproc more faithfully
 
bulwahn 
parents: 
49766 
diff
changeset
 | 
116  | 
= finite ((\<lambda>(b :: ?'B, a:: ?'A). Pair_Rep a b) ` (UNIV \<times> UNIV))"  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
117  | 
by simp  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
118  | 
|
| 
54611
 
31afce809794
set_comprehension_pointfree simproc causes to many surprises if enabled by default
 
traytel 
parents: 
49947 
diff
changeset
 | 
119  | 
declare [[simproc del: finite_Collect]]  | 
| 
 
31afce809794
set_comprehension_pointfree simproc causes to many surprises if enabled by default
 
traytel 
parents: 
49947 
diff
changeset
 | 
120  | 
|
| 
48174
 
eb72f99737be
adding a challenging example in the examples file
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
121  | 
|
| 61343 | 122  | 
section \<open>Testing simproc in code generation\<close>  | 
| 
49876
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
123  | 
|
| 
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
124  | 
definition union :: "nat set => nat set => nat set"  | 
| 
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
125  | 
where  | 
| 67613 | 126  | 
  "union A B = {x. x \<in> A \<or> x \<in> B}"
 | 
| 
49876
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
127  | 
|
| 67613 | 128  | 
definition common_subsets :: "nat set \<Rightarrow> nat set \<Rightarrow> nat set set"  | 
| 
49876
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
129  | 
where  | 
| 
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
130  | 
  "common_subsets S1 S2 = {S. S \<subseteq> S1 \<and> S \<subseteq> S2}"
 | 
| 
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
131  | 
|
| 
49947
 
29cd291bfea6
adding another test case for the set_comprehension_simproc to the theory in HOL/ex
 
bulwahn 
parents: 
49899 
diff
changeset
 | 
132  | 
definition products :: "nat set => nat set => nat set"  | 
| 
 
29cd291bfea6
adding another test case for the set_comprehension_simproc to the theory in HOL/ex
 
bulwahn 
parents: 
49899 
diff
changeset
 | 
133  | 
where  | 
| 67613 | 134  | 
  "products A B = {c. \<exists>a b. a \<in> A \<and> b \<in> B \<and> c = a * b}"
 | 
| 
49947
 
29cd291bfea6
adding another test case for the set_comprehension_simproc to the theory in HOL/ex
 
bulwahn 
parents: 
49899 
diff
changeset
 | 
135  | 
|
| 54934 | 136  | 
export_code products in Haskell  | 
| 
49947
 
29cd291bfea6
adding another test case for the set_comprehension_simproc to the theory in HOL/ex
 
bulwahn 
parents: 
49899 
diff
changeset
 | 
137  | 
|
| 54934 | 138  | 
export_code union common_subsets products in Haskell  | 
| 
49876
 
8cbd8340a21e
adding test cases for f x y : S patterns in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49853 
diff
changeset
 | 
139  | 
|
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
140  | 
end  |