| author | wenzelm | 
| Mon, 02 Mar 2020 13:57:03 +0100 | |
| changeset 71506 | 4197e30040f3 | 
| parent 70326 | aa7c49651f4e | 
| child 74383 | 107941e8fa01 | 
| permissions | -rw-r--r-- | 
| 48124 | 1  | 
(* Title: HOL/Tools/set_comprehension_pointfree.ML  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
2  | 
Author: Felix Kuperjans, Lukas Bulwahn, TU Muenchen  | 
| 48124 | 3  | 
Author: Rafal Kolanski, NICTA  | 
| 
48049
 
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  | 
Simproc for rewriting set comprehensions to pointfree expressions.  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
8  | 
signature SET_COMPREHENSION_POINTFREE =  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
9  | 
sig  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
10  | 
val base_simproc : Proof.context -> cterm -> thm option  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
11  | 
val code_simproc : Proof.context -> cterm -> thm option  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
12  | 
val simproc : Proof.context -> cterm -> thm option  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
13  | 
end  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
structure Set_Comprehension_Pointfree : SET_COMPREHENSION_POINTFREE =  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
struct  | 
| 
 
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  | 
(* syntactic operations *)  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
fun mk_inf (t1, t2) =  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
21  | 
let  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
val T = fastype_of t1  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
in  | 
| 69593 | 24  | 
Const (\<^const_name>\<open>Lattices.inf_class.inf\<close>, T --> T --> T) $ t1 $ t2  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
end  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
|
| 
49768
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
27  | 
fun mk_sup (t1, t2) =  | 
| 
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
28  | 
let  | 
| 
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
29  | 
val T = fastype_of t1  | 
| 
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
30  | 
in  | 
| 69593 | 31  | 
Const (\<^const_name>\<open>Lattices.sup_class.sup\<close>, T --> T --> T) $ t1 $ t2  | 
| 
49768
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
32  | 
end  | 
| 
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
33  | 
|
| 
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
34  | 
fun mk_Compl t =  | 
| 
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
35  | 
let  | 
| 
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
36  | 
val T = fastype_of t  | 
| 
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
37  | 
in  | 
| 69593 | 38  | 
Const (\<^const_name>\<open>Groups.uminus_class.uminus\<close>, T --> T) $ t  | 
| 
49768
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
39  | 
end  | 
| 
 
3ecfba7e731d
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
 
bulwahn 
parents: 
49765 
diff
changeset
 | 
40  | 
|
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
41  | 
fun mk_image t1 t2 =  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
42  | 
let  | 
| 69593 | 43  | 
val T as Type (\<^type_name>\<open>fun\<close>, [_ , R]) = fastype_of t1  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
44  | 
in  | 
| 69593 | 45  | 
Const (\<^const_name>\<open>image\<close>,  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
46  | 
T --> fastype_of t2 --> HOLogic.mk_setT R) $ t1 $ t2  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
47  | 
end;  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
48  | 
|
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
49  | 
fun mk_sigma (t1, t2) =  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
50  | 
let  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
51  | 
val T1 = fastype_of t1  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
52  | 
val T2 = fastype_of t2  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
53  | 
val setT = HOLogic.dest_setT T1  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
54  | 
val resT = HOLogic.mk_setT (HOLogic.mk_prodT (setT, HOLogic.dest_setT T2))  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
55  | 
in  | 
| 69593 | 56  | 
Const (\<^const_name>\<open>Sigma\<close>,  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
57  | 
T1 --> (setT --> T2) --> resT) $ t1 $ absdummy setT t2  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
58  | 
end;  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
59  | 
|
| 
49874
 
72b6d5fb407f
term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
 
bulwahn 
parents: 
49873 
diff
changeset
 | 
60  | 
fun mk_vimage f s =  | 
| 
 
72b6d5fb407f
term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
 
bulwahn 
parents: 
49873 
diff
changeset
 | 
61  | 
let  | 
| 69593 | 62  | 
val T as Type (\<^type_name>\<open>fun\<close>, [T1, T2]) = fastype_of f  | 
| 
49874
 
72b6d5fb407f
term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
 
bulwahn 
parents: 
49873 
diff
changeset
 | 
63  | 
in  | 
| 69593 | 64  | 
Const (\<^const_name>\<open>vimage\<close>, T --> HOLogic.mk_setT T2 --> HOLogic.mk_setT T1) $ f $ s  | 
| 
49874
 
72b6d5fb407f
term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
 
bulwahn 
parents: 
49873 
diff
changeset
 | 
65  | 
end;  | 
| 
 
72b6d5fb407f
term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
 
bulwahn 
parents: 
49873 
diff
changeset
 | 
66  | 
|
| 69593 | 67  | 
fun dest_Collect (Const (\<^const_name>\<open>Collect\<close>, _) $ Abs (x, T, t)) = ((x, T), t)  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
68  | 
  | dest_Collect t = raise TERM ("dest_Collect", [t])
 | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
70  | 
(* Copied from predicate_compile_aux.ML *)  | 
| 69593 | 71  | 
fun strip_ex (Const (\<^const_name>\<open>Ex\<close>, _) $ Abs (x, T, t)) =  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
72  | 
let  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
73  | 
val (xTs, t') = strip_ex t  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
74  | 
in  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
75  | 
((x, T) :: xTs, t')  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
76  | 
end  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
77  | 
| strip_ex t = ([], t)  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
78  | 
|
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
79  | 
fun mk_prod1 Ts (t1, t2) =  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
80  | 
let  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
81  | 
val (T1, T2) = apply2 (curry fastype_of1 Ts) (t1, t2)  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
82  | 
in  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
83  | 
HOLogic.pair_const T1 T2 $ t1 $ t2  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
84  | 
end;  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
85  | 
|
| 
49874
 
72b6d5fb407f
term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
 
bulwahn 
parents: 
49873 
diff
changeset
 | 
86  | 
fun mk_split_abs vs (Bound i) t = let val (x, T) = nth vs i in Abs (x, T, t) end  | 
| 69593 | 87  | 
| mk_split_abs vs (Const (\<^const_name>\<open>Product_Type.Pair\<close>, _) $ u $ v) t =  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
88  | 
HOLogic.mk_case_prod (mk_split_abs vs u (mk_split_abs vs v t))  | 
| 
49874
 
72b6d5fb407f
term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
 
bulwahn 
parents: 
49873 
diff
changeset
 | 
89  | 
  | mk_split_abs _ t _ = raise TERM ("mk_split_abs: bad term", [t]);
 | 
| 
 
72b6d5fb407f
term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
 
bulwahn 
parents: 
49873 
diff
changeset
 | 
90  | 
|
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
91  | 
(* a variant of HOLogic.strip_ptupleabs *)  | 
| 
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
92  | 
val strip_ptupleabs =  | 
| 
49901
 
58cac1b3b535
comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
 
bulwahn 
parents: 
49900 
diff
changeset
 | 
93  | 
let  | 
| 
 
58cac1b3b535
comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
 
bulwahn 
parents: 
49900 
diff
changeset
 | 
94  | 
fun strip [] qs vs t = (t, rev vs, qs)  | 
| 69593 | 95  | 
| strip (p :: ps) qs vs (Const (\<^const_name>\<open>case_prod\<close>, _) $ t) =  | 
| 
49901
 
58cac1b3b535
comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
 
bulwahn 
parents: 
49900 
diff
changeset
 | 
96  | 
strip ((1 :: p) :: (2 :: p) :: ps) (p :: qs) vs t  | 
| 
49957
 
6250121bfffb
passing around the simpset instead of the context; rewriting tactics to avoid the 'renamed bound variable' warnings in nested simplifier calls
 
bulwahn 
parents: 
49946 
diff
changeset
 | 
97  | 
| strip (_ :: ps) qs vs (Abs (s, T, t)) = strip ps qs ((s, T) :: vs) t  | 
| 
 
6250121bfffb
passing around the simpset instead of the context; rewriting tactics to avoid the 'renamed bound variable' warnings in nested simplifier calls
 
bulwahn 
parents: 
49946 
diff
changeset
 | 
98  | 
| strip (_ :: ps) qs vs t = strip ps qs  | 
| 
49901
 
58cac1b3b535
comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
 
bulwahn 
parents: 
49900 
diff
changeset
 | 
99  | 
((Name.uu_, hd (binder_types (fastype_of1 (map snd vs, t)))) :: vs)  | 
| 
 
58cac1b3b535
comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
 
bulwahn 
parents: 
49900 
diff
changeset
 | 
100  | 
(incr_boundvars 1 t $ Bound 0)  | 
| 
 
58cac1b3b535
comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
 
bulwahn 
parents: 
49900 
diff
changeset
 | 
101  | 
in strip [[]] [] [] end;  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
102  | 
|
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
103  | 
(* patterns *)  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
104  | 
|
| 
50024
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
105  | 
datatype pattern = Pattern of int list  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
106  | 
|
| 
50024
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
107  | 
fun dest_Pattern (Pattern bs) = bs  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
108  | 
|
| 
50024
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
109  | 
fun dest_bound (Bound i) = i  | 
| 
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
110  | 
  | dest_bound t = raise TERM("dest_bound", [t]);
 | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
111  | 
|
| 
50024
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
112  | 
fun type_of_pattern Ts (Pattern bs) = HOLogic.mk_tupleT (map (nth Ts) bs)  | 
| 
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
113  | 
|
| 
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
114  | 
fun term_of_pattern Ts (Pattern bs) =  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
115  | 
let  | 
| 
50024
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
116  | 
fun mk [b] = Bound b  | 
| 
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
117  | 
| mk (b :: bs) = HOLogic.pair_const (nth Ts b) (type_of_pattern Ts (Pattern bs))  | 
| 
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
118  | 
$ Bound b $ mk bs  | 
| 
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
119  | 
in mk bs end;  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
120  | 
|
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
121  | 
(* formulas *)  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
122  | 
|
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
123  | 
datatype formula = Atom of (pattern * term) | Int of formula * formula | Un of formula * formula  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
124  | 
|
| 
49900
 
89b118c0c070
checking for bound variables in the set expression; handling negation more generally
 
bulwahn 
parents: 
49898 
diff
changeset
 | 
125  | 
fun map_atom f (Atom a) = Atom (f a)  | 
| 
 
89b118c0c070
checking for bound variables in the set expression; handling negation more generally
 
bulwahn 
parents: 
49898 
diff
changeset
 | 
126  | 
| map_atom _ x = x  | 
| 
 
89b118c0c070
checking for bound variables in the set expression; handling negation more generally
 
bulwahn 
parents: 
49898 
diff
changeset
 | 
127  | 
|
| 69593 | 128  | 
fun is_collect_atom (Atom (_, Const(\<^const_name>\<open>Collect\<close>, _) $ _)) = true  | 
129  | 
| is_collect_atom (Atom (_, Const (\<^const_name>\<open>Groups.uminus_class.uminus\<close>, _) $ (Const(\<^const_name>\<open>Collect\<close>, _) $ _))) = true  | 
|
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
130  | 
| is_collect_atom _ = false  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
131  | 
|
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
132  | 
fun mk_case_prod _ [(x, T)] t = (T, Abs (x, T, t))  | 
| 
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
133  | 
| mk_case_prod rT ((x, T) :: vs) t =  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
134  | 
let  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
135  | 
val (T', t') = mk_case_prod rT vs t  | 
| 
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
136  | 
val t'' = HOLogic.case_prod_const (T, T', rT) $ (Abs (x, T, t'))  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
137  | 
in (domain_type (fastype_of t''), t'') end  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
138  | 
|
| 
50030
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
139  | 
fun mk_term vs t =  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
140  | 
let  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
141  | 
val bs = loose_bnos t  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
142  | 
val vs' = map (nth (rev vs)) bs  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
143  | 
val subst = map_index (fn (i, j) => (j, Bound i)) (rev bs)  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
144  | 
|> sort (fn (p1, p2) => int_ord (fst p1, fst p2))  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
145  | 
|> (fn subst' => map (fn i => the_default (Bound i) (AList.lookup (op =) subst' i)) (0 upto (fst (snd (split_last subst')))))  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
146  | 
val t' = subst_bounds (subst, t)  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
147  | 
val tuple = Pattern bs  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
148  | 
in (tuple, (vs', t')) end  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
149  | 
|
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
150  | 
fun default_atom vs t =  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
151  | 
let  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
152  | 
val (tuple, (vs', t')) = mk_term vs t  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
153  | 
val T = HOLogic.mk_tupleT (map snd vs')  | 
| 69593 | 154  | 
val s = HOLogic.Collect_const T $ (snd (mk_case_prod \<^typ>\<open>bool\<close> vs' t'))  | 
| 
50030
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
155  | 
in  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
156  | 
(tuple, Atom (tuple, s))  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
157  | 
end  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
158  | 
|
| 69593 | 159  | 
fun mk_atom vs (t as Const (\<^const_name>\<open>Set.member\<close>, _) $ x $ s) =  | 
| 
49900
 
89b118c0c070
checking for bound variables in the set expression; handling negation more generally
 
bulwahn 
parents: 
49898 
diff
changeset
 | 
160  | 
if not (null (loose_bnos s)) then  | 
| 
50032
 
a439a9d14ba3
handling x : S y pattern with the default mechanism instead of raising an exception in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50031 
diff
changeset
 | 
161  | 
default_atom vs t  | 
| 
49900
 
89b118c0c070
checking for bound variables in the set expression; handling negation more generally
 
bulwahn 
parents: 
49898 
diff
changeset
 | 
162  | 
else  | 
| 
50030
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
163  | 
(case try ((map dest_bound) o HOLogic.strip_tuple) x of  | 
| 
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
164  | 
SOME pat => (Pattern pat, Atom (Pattern pat, s))  | 
| 
49874
 
72b6d5fb407f
term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
 
bulwahn 
parents: 
49873 
diff
changeset
 | 
165  | 
| NONE =>  | 
| 
49900
 
89b118c0c070
checking for bound variables in the set expression; handling negation more generally
 
bulwahn 
parents: 
49898 
diff
changeset
 | 
166  | 
let  | 
| 
50030
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
167  | 
val (tuple, (vs', x')) = mk_term vs x  | 
| 
49900
 
89b118c0c070
checking for bound variables in the set expression; handling negation more generally
 
bulwahn 
parents: 
49898 
diff
changeset
 | 
168  | 
val rT = HOLogic.dest_setT (fastype_of s)  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
169  | 
val s = mk_vimage (snd (mk_case_prod rT vs' x')) s  | 
| 
50030
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
170  | 
in (tuple, Atom (tuple, s)) end)  | 
| 69593 | 171  | 
| mk_atom vs (Const (\<^const_name>\<open>HOL.Not\<close>, _) $ t) = apsnd (map_atom (apsnd mk_Compl)) (mk_atom vs t)  | 
| 
50030
 
349f651ec203
syntactic tuning and restructuring of set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50029 
diff
changeset
 | 
172  | 
| mk_atom vs t = default_atom vs t  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
173  | 
|
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
174  | 
fun merge' [] (pats1, pats2) = ([], (pats1, pats2))  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
175  | 
| merge' pat (pats, []) = (pat, (pats, []))  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
176  | 
| merge' pat (pats1, pats) =  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
177  | 
let  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
178  | 
fun disjoint_to_pat p = null (inter (op =) pat p)  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
179  | 
val overlap_pats = filter_out disjoint_to_pat pats  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
180  | 
val rem_pats = filter disjoint_to_pat pats  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
181  | 
val (pat, (pats', pats1')) = merge' (distinct (op =) (flat overlap_pats @ pat)) (rem_pats, pats1)  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
182  | 
in  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
183  | 
(pat, (pats1', pats'))  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
184  | 
end  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
185  | 
|
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
186  | 
fun merge ([], pats) = pats  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
187  | 
| merge (pat :: pats', pats) =  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
188  | 
let val (pat', (pats1', pats2')) = merge' pat (pats', pats)  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
189  | 
in pat' :: merge (pats1', pats2') end;  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
190  | 
|
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
191  | 
fun restricted_merge ([], pats) = pats  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
192  | 
| restricted_merge (pat :: pats', pats) =  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
193  | 
let  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
194  | 
fun disjoint_to_pat p = null (inter (op =) pat p)  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
195  | 
val overlap_pats = filter_out disjoint_to_pat pats  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
196  | 
val rem_pats = filter disjoint_to_pat pats  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
197  | 
in  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
198  | 
case overlap_pats of  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
199  | 
[] => pat :: restricted_merge (pats', rem_pats)  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
200  | 
| [pat'] => if subset (op =) (pat, pat') then  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
201  | 
pat' :: restricted_merge (pats', rem_pats)  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
202  | 
else if subset (op =) (pat', pat) then  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
203  | 
pat :: restricted_merge (pats', rem_pats)  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
204  | 
else error "restricted merge: two patterns require relational join"  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
205  | 
| _ => error "restricted merge: multiple patterns overlap"  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
206  | 
end;  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
207  | 
|
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
208  | 
fun map_atoms f (Atom a) = Atom (f a)  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
209  | 
| map_atoms f (Un (fm1, fm2)) = Un (apply2 (map_atoms f) (fm1, fm2))  | 
| 
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
210  | 
| map_atoms f (Int (fm1, fm2)) = Int (apply2 (map_atoms f) (fm1, fm2))  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
211  | 
|
| 
50028
 
d05f859558a0
improving the extension of sets in case of more than one bound variable; rearranging the tactic to prefer simpler steps before more involved ones
 
bulwahn 
parents: 
50026 
diff
changeset
 | 
212  | 
fun extend Ts bs t = foldr1 mk_sigma (t :: map (fn b => HOLogic.mk_UNIV (nth Ts b)) bs)  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
213  | 
|
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
214  | 
fun rearrange vs (pat, pat') t =  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
215  | 
let  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
216  | 
val subst = map_index (fn (i, b) => (b, i)) (rev pat)  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
217  | 
val vs' = map (nth (rev vs)) pat  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
218  | 
val Ts' = map snd (rev vs')  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
219  | 
val bs = map (fn b => the (AList.lookup (op =) subst b)) pat'  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
220  | 
val rt = term_of_pattern Ts' (Pattern bs)  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
221  | 
val rT = type_of_pattern Ts' (Pattern bs)  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
222  | 
val (_, f) = mk_case_prod rT vs' rt  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
223  | 
in  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
224  | 
mk_image f t  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
225  | 
end;  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
226  | 
|
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
227  | 
fun adjust vs pats (Pattern pat, t) =  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
228  | 
let  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
229  | 
val SOME p = find_first (fn p => not (null (inter (op =) pat p))) pats  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
230  | 
val missing = subtract (op =) pat p  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
231  | 
val Ts = rev (map snd vs)  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
232  | 
val t' = extend Ts missing t  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
233  | 
in (Pattern p, rearrange vs (pat @ missing, p) t') end  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
234  | 
|
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
235  | 
fun adjust_atoms vs pats fm = map_atoms (adjust vs pats) fm  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
236  | 
|
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
237  | 
fun merge_inter vs (pats1, fm1) (pats2, fm2) =  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
238  | 
let  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
239  | 
val pats = restricted_merge (map dest_Pattern pats1, map dest_Pattern pats2)  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
240  | 
val (fm1', fm2') = apply2 (adjust_atoms vs pats) (fm1, fm2)  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
241  | 
in  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
242  | 
(map Pattern pats, Int (fm1', fm2'))  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
243  | 
end;  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
244  | 
|
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
245  | 
fun merge_union vs (pats1, fm1) (pats2, fm2) =  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
246  | 
let  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
247  | 
val pats = merge (map dest_Pattern pats1, map dest_Pattern pats2)  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
248  | 
val (fm1', fm2') = apply2 (adjust_atoms vs pats) (fm1, fm2)  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
249  | 
in  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
250  | 
(map Pattern pats, Un (fm1', fm2'))  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
251  | 
end;  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
252  | 
|
| 69597 | 253  | 
fun mk_formula vs (\<^const>\<open>HOL.conj\<close> $ t1 $ t2) = merge_inter vs (mk_formula vs t1) (mk_formula vs t2)  | 
254  | 
| mk_formula vs (\<^const>\<open>HOL.disj\<close> $ t1 $ t2) = merge_union vs (mk_formula vs t1) (mk_formula vs t2)  | 
|
| 
49874
 
72b6d5fb407f
term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
 
bulwahn 
parents: 
49873 
diff
changeset
 | 
255  | 
| mk_formula vs t = apfst single (mk_atom vs t)  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
256  | 
|
| 
49852
 
caaa1956f0da
refined tactic in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49850 
diff
changeset
 | 
257  | 
fun strip_Int (Int (fm1, fm2)) = fm1 :: (strip_Int fm2)  | 
| 
 
caaa1956f0da
refined tactic in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49850 
diff
changeset
 | 
258  | 
| strip_Int fm = [fm]  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
259  | 
|
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
260  | 
(* term construction *)  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
261  | 
|
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
262  | 
fun reorder_bounds pats t =  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
263  | 
let  | 
| 
50024
 
b7265db3a1dc
simplified structure of patterns in set_comprehension_simproc
 
bulwahn 
parents: 
49959 
diff
changeset
 | 
264  | 
val bounds = maps dest_Pattern pats  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
265  | 
val bperm = bounds ~~ ((length bounds - 1) downto 0)  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
266  | 
|> sort (fn (i,j) => int_ord (fst i, fst j)) |> map snd  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
267  | 
in  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
268  | 
subst_bounds (map Bound bperm, t)  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
269  | 
end;  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
270  | 
|
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
271  | 
fun is_reordering t =  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
272  | 
let val (t', _, _) = HOLogic.strip_ptupleabs t  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
273  | 
in forall (fn Bound _ => true) (HOLogic.strip_tuple t') end  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
274  | 
|
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
275  | 
fun mk_pointfree_expr t =  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
276  | 
let  | 
| 
49857
 
7bf407d77152
setcomprehension_pointfree simproc also works for set comprehension without an equation
 
bulwahn 
parents: 
49852 
diff
changeset
 | 
277  | 
val ((x, T), (vs, t'')) = apsnd strip_ex (dest_Collect t)  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
278  | 
val Ts = map snd (rev vs)  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
279  | 
fun mk_mem_UNIV n = HOLogic.mk_mem (Bound n, HOLogic.mk_UNIV (nth Ts n))  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
280  | 
fun lookup (pat', t) pat = if pat = pat' then t else HOLogic.mk_UNIV (type_of_pattern Ts pat)  | 
| 
49761
 
b7772f3b6c03
set_comprehension_pointfree also handles terms where the equation is not at the first position, which is a necessary generalisation to eventually handle bounded existentials; tuned
 
bulwahn 
parents: 
48128 
diff
changeset
 | 
281  | 
val conjs = HOLogic.dest_conj t''  | 
| 
49857
 
7bf407d77152
setcomprehension_pointfree simproc also works for set comprehension without an equation
 
bulwahn 
parents: 
49852 
diff
changeset
 | 
282  | 
val refl = HOLogic.eq_const T $ Bound (length vs) $ Bound (length vs)  | 
| 
49761
 
b7772f3b6c03
set_comprehension_pointfree also handles terms where the equation is not at the first position, which is a necessary generalisation to eventually handle bounded existentials; tuned
 
bulwahn 
parents: 
48128 
diff
changeset
 | 
283  | 
val is_the_eq =  | 
| 
 
b7772f3b6c03
set_comprehension_pointfree also handles terms where the equation is not at the first position, which is a necessary generalisation to eventually handle bounded existentials; tuned
 
bulwahn 
parents: 
48128 
diff
changeset
 | 
284  | 
the_default false o (try (fn eq => fst (HOLogic.dest_eq eq) = Bound (length vs)))  | 
| 
49857
 
7bf407d77152
setcomprehension_pointfree simproc also works for set comprehension without an equation
 
bulwahn 
parents: 
49852 
diff
changeset
 | 
285  | 
val eq = the_default refl (find_first is_the_eq conjs)  | 
| 
49761
 
b7772f3b6c03
set_comprehension_pointfree also handles terms where the equation is not at the first position, which is a necessary generalisation to eventually handle bounded existentials; tuned
 
bulwahn 
parents: 
48128 
diff
changeset
 | 
286  | 
val f = snd (HOLogic.dest_eq eq)  | 
| 
 
b7772f3b6c03
set_comprehension_pointfree also handles terms where the equation is not at the first position, which is a necessary generalisation to eventually handle bounded existentials; tuned
 
bulwahn 
parents: 
48128 
diff
changeset
 | 
287  | 
val conjs' = filter_out (fn t => eq = t) conjs  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
288  | 
val unused_bounds = subtract (op =) (distinct (op =) (maps loose_bnos conjs'))  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
289  | 
(0 upto (length vs - 1))  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
290  | 
val (pats, fm) =  | 
| 
49943
 
6a26fed71755
passing names and types of all bounds around in the simproc
 
bulwahn 
parents: 
49942 
diff
changeset
 | 
291  | 
mk_formula ((x, T) :: vs) (foldr1 HOLogic.mk_conj (conjs' @ map mk_mem_UNIV unused_bounds))  | 
| 50031 | 292  | 
fun mk_set (Atom pt) = foldr1 mk_sigma (map (lookup pt) pats)  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
293  | 
| mk_set (Un (f1, f2)) = mk_sup (mk_set f1, mk_set f2)  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
294  | 
| mk_set (Int (f1, f2)) = mk_inf (mk_set f1, mk_set f2)  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
295  | 
val pat = foldr1 (mk_prod1 Ts) (map (term_of_pattern Ts) pats)  | 
| 
49857
 
7bf407d77152
setcomprehension_pointfree simproc also works for set comprehension without an equation
 
bulwahn 
parents: 
49852 
diff
changeset
 | 
296  | 
val t = mk_split_abs (rev ((x, T) :: vs)) pat (reorder_bounds pats f)  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
297  | 
in  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
298  | 
if the_default false (try is_reordering t) andalso is_collect_atom fm then  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
299  | 
error "mk_pointfree_expr: trivial case"  | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
300  | 
else (fm, mk_image t (mk_set fm))  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
301  | 
end;  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
302  | 
|
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
303  | 
val rewrite_term = try mk_pointfree_expr  | 
| 
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
304  | 
|
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
305  | 
|
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
306  | 
(* proof tactic *)  | 
| 
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
307  | 
|
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
308  | 
val case_prod_beta = @{lemma "case_prod g x z = case_prod (\<lambda>x y. (g x y) z) x" by (simp add: case_prod_beta)}
 | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
309  | 
|
| 
49944
 
28cd3c9ca278
tuned tactic in set_comprehension_pointfree simproc to handle composition of negation and vimage
 
bulwahn 
parents: 
49943 
diff
changeset
 | 
310  | 
val vimageI2' = @{lemma "f a \<notin> A ==> a \<notin> f -` A" by simp}
 | 
| 
 
28cd3c9ca278
tuned tactic in set_comprehension_pointfree simproc to handle composition of negation and vimage
 
bulwahn 
parents: 
49943 
diff
changeset
 | 
311  | 
val vimageE' =  | 
| 
 
28cd3c9ca278
tuned tactic in set_comprehension_pointfree simproc to handle composition of negation and vimage
 
bulwahn 
parents: 
49943 
diff
changeset
 | 
312  | 
  @{lemma "a \<notin> f -` B ==> (\<And> x. f a = x ==> x \<notin> B ==> P) ==> P" by simp}
 | 
| 
 
28cd3c9ca278
tuned tactic in set_comprehension_pointfree simproc to handle composition of negation and vimage
 
bulwahn 
parents: 
49943 
diff
changeset
 | 
313  | 
|
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
314  | 
val collectI' = @{lemma "\<not> P a ==> a \<notin> {x. P x}" by auto}
 | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
315  | 
val collectE' = @{lemma "a \<notin> {x. P x} ==> (\<not> P a ==> Q) ==> Q" by auto}
 | 
| 
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
316  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
317  | 
fun elim_Collect_tac ctxt =  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
318  | 
  dresolve_tac ctxt @{thms iffD1 [OF mem_Collect_eq]}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
319  | 
  THEN' (REPEAT_DETERM o (eresolve_tac ctxt @{thms exE}))
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
320  | 
  THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms conjE}
 | 
| 51798 | 321  | 
THEN' TRY o hyp_subst_tac ctxt;  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
322  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
323  | 
fun intro_image_tac ctxt =  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
324  | 
  resolve_tac ctxt @{thms image_eqI}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
325  | 
THEN' (REPEAT_DETERM1 o  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
326  | 
      (resolve_tac ctxt @{thms refl}
 | 
| 67399 | 327  | 
      ORELSE' resolve_tac ctxt @{thms arg_cong2 [OF refl, where f = "(=)", OF prod.case, THEN iffD2]}
 | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
328  | 
ORELSE' CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1  | 
| 
51315
 
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
 
wenzelm 
parents: 
51314 
diff
changeset
 | 
329  | 
(HOLogic.Trueprop_conv  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
330  | 
(HOLogic.eq_conv Conv.all_conv (Conv.rewr_conv (mk_meta_eq case_prod_beta)))))) ctxt)))  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
331  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
332  | 
fun elim_image_tac ctxt =  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
333  | 
  eresolve_tac ctxt @{thms imageE}
 | 
| 
50028
 
d05f859558a0
improving the extension of sets in case of more than one bound variable; rearranging the tactic to prefer simpler steps before more involved ones
 
bulwahn 
parents: 
50026 
diff
changeset
 | 
334  | 
THEN' REPEAT_DETERM o CHANGED o  | 
| 
55642
 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 
blanchet 
parents: 
55414 
diff
changeset
 | 
335  | 
    (TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_paired_all prod.case})
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
336  | 
    THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject}
 | 
| 51798 | 337  | 
THEN' TRY o hyp_subst_tac ctxt)  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
338  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
339  | 
fun tac1_of_formula ctxt (Int (fm1, fm2)) =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
340  | 
    TRY o eresolve_tac ctxt @{thms conjE}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
341  | 
    THEN' resolve_tac ctxt @{thms IntI}
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
342  | 
THEN' (fn i => tac1_of_formula ctxt fm2 (i + 1))  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
343  | 
THEN' tac1_of_formula ctxt fm1  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
344  | 
| tac1_of_formula ctxt (Un (fm1, fm2)) =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
345  | 
    eresolve_tac ctxt @{thms disjE} THEN' resolve_tac ctxt @{thms UnI1}
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
346  | 
THEN' tac1_of_formula ctxt fm1  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
347  | 
    THEN' resolve_tac ctxt @{thms UnI2}
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
348  | 
THEN' tac1_of_formula ctxt fm2  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
349  | 
| tac1_of_formula ctxt (Atom _) =  | 
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
350  | 
REPEAT_DETERM1 o (assume_tac ctxt  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
351  | 
      ORELSE' resolve_tac ctxt @{thms SigmaI}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
352  | 
      ORELSE' ((resolve_tac ctxt @{thms CollectI} ORELSE' resolve_tac ctxt [collectI']) THEN'
 | 
| 
55642
 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 
blanchet 
parents: 
55414 
diff
changeset
 | 
353  | 
        TRY o simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}]))
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
354  | 
      ORELSE' ((resolve_tac ctxt @{thms vimageI2} ORELSE' resolve_tac ctxt [vimageI2']) THEN'
 | 
| 
55642
 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 
blanchet 
parents: 
55414 
diff
changeset
 | 
355  | 
        TRY o simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}]))
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
356  | 
      ORELSE' (resolve_tac ctxt @{thms image_eqI} THEN'
 | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
357  | 
(REPEAT_DETERM o  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
358  | 
      (resolve_tac ctxt @{thms refl}
 | 
| 67399 | 359  | 
      ORELSE' resolve_tac ctxt @{thms arg_cong2[OF refl, where f = "(=)", OF prod.case, THEN iffD2]})))
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
360  | 
      ORELSE' resolve_tac ctxt @{thms UNIV_I}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
361  | 
      ORELSE' resolve_tac ctxt @{thms iffD2[OF Compl_iff]}
 | 
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
362  | 
ORELSE' assume_tac ctxt)  | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
363  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
364  | 
fun tac2_of_formula ctxt (Int (fm1, fm2)) =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
365  | 
    TRY o eresolve_tac ctxt @{thms IntE}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
366  | 
    THEN' TRY o resolve_tac ctxt @{thms conjI}
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
367  | 
THEN' (fn i => tac2_of_formula ctxt fm2 (i + 1))  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
368  | 
THEN' tac2_of_formula ctxt fm1  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
369  | 
| tac2_of_formula ctxt (Un (fm1, fm2)) =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
370  | 
    eresolve_tac ctxt @{thms UnE} THEN' resolve_tac ctxt @{thms disjI1}
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
371  | 
THEN' tac2_of_formula ctxt fm1  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
372  | 
    THEN' resolve_tac ctxt @{thms disjI2}
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
373  | 
THEN' tac2_of_formula ctxt fm2  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
374  | 
| tac2_of_formula ctxt (Atom _) =  | 
| 
50025
 
19965e6a705e
handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50024 
diff
changeset
 | 
375  | 
REPEAT_DETERM o  | 
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
376  | 
(assume_tac ctxt  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
377  | 
       ORELSE' dresolve_tac ctxt @{thms iffD1[OF mem_Sigma_iff]}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
378  | 
       ORELSE' eresolve_tac ctxt @{thms conjE}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
379  | 
       ORELSE' ((eresolve_tac ctxt @{thms CollectE} ORELSE' eresolve_tac ctxt [collectE']) THEN'
 | 
| 
55642
 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 
blanchet 
parents: 
55414 
diff
changeset
 | 
380  | 
         TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}]) THEN'
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
381  | 
         REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject} THEN' TRY o hyp_subst_tac ctxt THEN'
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
382  | 
         TRY o resolve_tac ctxt @{thms refl})
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
383  | 
       ORELSE' (eresolve_tac ctxt @{thms imageE}
 | 
| 
50028
 
d05f859558a0
improving the extension of sets in case of more than one bound variable; rearranging the tactic to prefer simpler steps before more involved ones
 
bulwahn 
parents: 
50026 
diff
changeset
 | 
384  | 
THEN' (REPEAT_DETERM o CHANGED o  | 
| 
55642
 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 
blanchet 
parents: 
55414 
diff
changeset
 | 
385  | 
         (TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_paired_all prod.case})
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
386  | 
         THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
387  | 
         THEN' TRY o hyp_subst_tac ctxt THEN' TRY o resolve_tac ctxt @{thms refl})))
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
388  | 
       ORELSE' eresolve_tac ctxt @{thms ComplE}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
389  | 
       ORELSE' ((eresolve_tac ctxt @{thms vimageE} ORELSE' eresolve_tac ctxt [vimageE'])
 | 
| 
55642
 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 
blanchet 
parents: 
55414 
diff
changeset
 | 
390  | 
        THEN' TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}])
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
391  | 
        THEN' TRY o hyp_subst_tac ctxt THEN' TRY o resolve_tac ctxt @{thms refl}))
 | 
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
392  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
393  | 
fun tac ctxt fm =  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
394  | 
let  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
395  | 
    val subset_tac1 = resolve_tac ctxt @{thms subsetI}
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
396  | 
THEN' elim_Collect_tac ctxt  | 
| 
50029
 
31c9294eebe6
using more proper simpset in tactic of set_comprehension_pointfree simproc to avoid renamed bound variable warnings in recursive simplifier calls
 
bulwahn 
parents: 
50028 
diff
changeset
 | 
397  | 
THEN' intro_image_tac ctxt  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
398  | 
THEN' tac1_of_formula ctxt fm  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
399  | 
    val subset_tac2 = resolve_tac ctxt @{thms subsetI}
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
400  | 
THEN' elim_image_tac ctxt  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
401  | 
      THEN' resolve_tac ctxt @{thms iffD2[OF mem_Collect_eq]}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
402  | 
      THEN' REPEAT_DETERM o resolve_tac ctxt @{thms exI}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
403  | 
      THEN' (TRY o REPEAT_ALL_NEW (resolve_tac ctxt @{thms conjI}))
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
404  | 
      THEN' (K (TRY (FIRSTGOAL ((TRY o hyp_subst_tac ctxt) THEN' resolve_tac ctxt @{thms refl}))))
 | 
| 
49852
 
caaa1956f0da
refined tactic in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49850 
diff
changeset
 | 
405  | 
THEN' (fn i => EVERY (rev (map_index (fn (j, f) =>  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
406  | 
        REPEAT_DETERM (eresolve_tac ctxt @{thms IntE} (i + j)) THEN
 | 
| 58839 | 407  | 
tac2_of_formula ctxt f (i + j)) (strip_Int fm))))  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
408  | 
in  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
409  | 
    resolve_tac ctxt @{thms subset_antisym} THEN' subset_tac1 THEN' subset_tac2
 | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
410  | 
end;  | 
| 
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
411  | 
|
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
412  | 
|
| 
49896
 
a39deedd5c3f
employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49875 
diff
changeset
 | 
413  | 
(* preprocessing conversion:  | 
| 
 
a39deedd5c3f
employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49875 
diff
changeset
 | 
414  | 
  rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} *)
 | 
| 
 
a39deedd5c3f
employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49875 
diff
changeset
 | 
415  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
416  | 
fun comprehension_conv ctxt ct =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
417  | 
let  | 
| 69593 | 418  | 
fun dest_Collect (Const (\<^const_name>\<open>Collect\<close>, T) $ t) = (HOLogic.dest_setT (body_type T), t)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
419  | 
      | dest_Collect t = raise TERM ("dest_Collect", [t])
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
420  | 
fun list_ex vs t = fold_rev (fn (x, T) => fn t => HOLogic.exists_const T $ Abs (x, T, t)) vs t  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
421  | 
fun mk_term t =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
422  | 
let  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
423  | 
val (T, t') = dest_Collect t  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
424  | 
val (t'', vs, fp) = case strip_ptupleabs t' of  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
425  | 
            (_, [_], _) => raise TERM("mk_term", [t'])
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
426  | 
| (t'', vs, fp) => (t'', vs, fp)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
427  | 
val Ts = map snd vs  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
428  | 
val eq = HOLogic.eq_const T $ Bound (length Ts) $  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
429  | 
(HOLogic.mk_ptuple fp (HOLogic.mk_ptupleT fp Ts) (rev (map_index (fn (i, _) => Bound i) Ts)))  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
430  | 
in  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
431  | 
HOLogic.Collect_const T $ absdummy T (list_ex vs (HOLogic.mk_conj (eq, t'')))  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
432  | 
end;  | 
| 59582 | 433  | 
fun is_eq th = is_some (try (HOLogic.dest_eq o HOLogic.dest_Trueprop) (Thm.prop_of th))  | 
| 
55642
 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 
blanchet 
parents: 
55414 
diff
changeset
 | 
434  | 
    val unfold_thms = @{thms split_paired_all mem_Collect_eq prod.case}
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
435  | 
fun tac ctxt =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
436  | 
      resolve_tac ctxt @{thms set_eqI}
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
437  | 
THEN' simp_tac (put_simpset HOL_basic_ss ctxt addsimps unfold_thms)  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
438  | 
      THEN' resolve_tac ctxt @{thms iffI}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
439  | 
      THEN' REPEAT_DETERM o resolve_tac ctxt @{thms exI}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
440  | 
      THEN' resolve_tac ctxt @{thms conjI} THEN' resolve_tac ctxt @{thms refl} THEN' assume_tac ctxt
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
441  | 
      THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms exE}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
442  | 
      THEN' eresolve_tac ctxt @{thms conjE}
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
443  | 
      THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject}
 | 
| 60696 | 444  | 
      THEN' Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
 | 
445  | 
simp_tac (put_simpset HOL_basic_ss ctxt' addsimps (filter is_eq prems)) 1) ctxt  | 
|
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
446  | 
THEN' TRY o assume_tac ctxt  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
447  | 
in  | 
| 59582 | 448  | 
case try mk_term (Thm.term_of ct) of  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
449  | 
NONE => Thm.reflexive ct  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
450  | 
| SOME t' =>  | 
| 59582 | 451  | 
Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (Thm.term_of ct, t')))  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
452  | 
          (fn {context, ...} => tac context 1)
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
453  | 
        RS @{thm eq_reflection}
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
454  | 
end  | 
| 
49896
 
a39deedd5c3f
employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49875 
diff
changeset
 | 
455  | 
|
| 
 
a39deedd5c3f
employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49875 
diff
changeset
 | 
456  | 
|
| 
49849
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
457  | 
(* main simprocs *)  | 
| 
 
d9822ec4f434
extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
 
bulwahn 
parents: 
49831 
diff
changeset
 | 
458  | 
|
| 
49942
 
50e457bbc5fe
locally inverting previously applied simplifications with ex_simps in set_comprehension_pointfree
 
bulwahn 
parents: 
49901 
diff
changeset
 | 
459  | 
val prep_thms =  | 
| 
 
50e457bbc5fe
locally inverting previously applied simplifications with ex_simps in set_comprehension_pointfree
 
bulwahn 
parents: 
49901 
diff
changeset
 | 
460  | 
  map mk_meta_eq ([@{thm Bex_def}, @{thm Pow_iff[symmetric]}] @ @{thms ex_simps[symmetric]})
 | 
| 
49873
 
4b7c2e4991fc
extending preprocessing of simproc to rewrite subset inequality into membership of powerset
 
bulwahn 
parents: 
49857 
diff
changeset
 | 
461  | 
|
| 
49850
 
873fa7156468
adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49849 
diff
changeset
 | 
462  | 
val post_thms =  | 
| 
 
873fa7156468
adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49849 
diff
changeset
 | 
463  | 
  map mk_meta_eq [@{thm Times_Un_distrib1[symmetric]},
 | 
| 
 
873fa7156468
adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49849 
diff
changeset
 | 
464  | 
  @{lemma "A \<times> B \<union> A \<times> C = A \<times> (B \<union> C)" by auto},
 | 
| 
 
873fa7156468
adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49849 
diff
changeset
 | 
465  | 
  @{lemma "(A \<times> B \<inter> C \<times> D) = (A \<inter> C) \<times> (B \<inter> D)" by auto}]
 | 
| 
 
873fa7156468
adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49849 
diff
changeset
 | 
466  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
467  | 
fun conv ctxt t =  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
468  | 
let  | 
| 70326 | 469  | 
val (t', ctxt') = yield_singleton (Variable.import_terms true) t (Variable.declare_term t ctxt)  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59582 
diff
changeset
 | 
470  | 
val ct = Thm.cterm_of ctxt' t'  | 
| 
49957
 
6250121bfffb
passing around the simpset instead of the context; rewriting tactics to avoid the 'renamed bound variable' warnings in nested simplifier calls
 
bulwahn 
parents: 
49946 
diff
changeset
 | 
471  | 
fun unfold_conv thms =  | 
| 
 
6250121bfffb
passing around the simpset instead of the context; rewriting tactics to avoid the 'renamed bound variable' warnings in nested simplifier calls
 
bulwahn 
parents: 
49946 
diff
changeset
 | 
472  | 
Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE))  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
473  | 
(empty_simpset ctxt' addsimps thms)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
474  | 
val prep_eq = (comprehension_conv ctxt' then_conv unfold_conv prep_thms) ct  | 
| 59582 | 475  | 
val t'' = Thm.term_of (Thm.rhs_of prep_eq)  | 
| 
50026
 
d9871e5ea0e1
importing term with schematic type variables properly before passing it to the tactic in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50025 
diff
changeset
 | 
476  | 
fun mk_thm (fm, t''') = Goal.prove ctxt' [] []  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
477  | 
      (HOLogic.mk_Trueprop (HOLogic.mk_eq (t'', t'''))) (fn {context, ...} => tac context fm 1)
 | 
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67399 
diff
changeset
 | 
478  | 
    fun unfold th = th RS (HOLogic.mk_obj_eq prep_eq RS @{thm trans})
 | 
| 
51315
 
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
 
wenzelm 
parents: 
51314 
diff
changeset
 | 
479  | 
val post =  | 
| 
 
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
 
wenzelm 
parents: 
51314 
diff
changeset
 | 
480  | 
Conv.fconv_rule  | 
| 
 
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
 
wenzelm 
parents: 
51314 
diff
changeset
 | 
481  | 
(HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (unfold_conv post_thms)))  | 
| 
50026
 
d9871e5ea0e1
importing term with schematic type variables properly before passing it to the tactic in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50025 
diff
changeset
 | 
482  | 
val export = singleton (Variable.export ctxt' ctxt)  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
483  | 
in  | 
| 
50026
 
d9871e5ea0e1
importing term with schematic type variables properly before passing it to the tactic in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
50025 
diff
changeset
 | 
484  | 
Option.map (export o post o unfold o mk_thm) (rewrite_term t'')  | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
485  | 
end;  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
486  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
487  | 
fun base_simproc ctxt redex =  | 
| 
48122
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
488  | 
let  | 
| 59582 | 489  | 
val set_compr = Thm.term_of redex  | 
| 
48122
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
490  | 
in  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
491  | 
conv ctxt set_compr  | 
| 
48122
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
492  | 
    |> Option.map (fn thm => thm RS @{thm eq_reflection})
 | 
| 
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
493  | 
end;  | 
| 
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
494  | 
|
| 
49763
 
bed063d0c526
generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
 
bulwahn 
parents: 
49761 
diff
changeset
 | 
495  | 
fun instantiate_arg_cong ctxt pred =  | 
| 
 
bed063d0c526
generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
 
bulwahn 
parents: 
49761 
diff
changeset
 | 
496  | 
let  | 
| 
49831
 
b28dbb7a45d9
increading indexes to avoid clashes in the set_comprehension_pointfree simproc
 
bulwahn 
parents: 
49768 
diff
changeset
 | 
497  | 
    val arg_cong = Thm.incr_indexes (maxidx_of_term pred + 1) @{thm arg_cong}
 | 
| 60781 | 498  | 
val (Var (f, _) $ _, _) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.concl_of arg_cong))  | 
| 
49763
 
bed063d0c526
generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
 
bulwahn 
parents: 
49761 
diff
changeset
 | 
499  | 
in  | 
| 60781 | 500  | 
infer_instantiate ctxt [(f, Thm.cterm_of ctxt pred)] arg_cong  | 
| 
49763
 
bed063d0c526
generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
 
bulwahn 
parents: 
49761 
diff
changeset
 | 
501  | 
end;  | 
| 
 
bed063d0c526
generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
 
bulwahn 
parents: 
49761 
diff
changeset
 | 
502  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
503  | 
fun simproc ctxt redex =  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
504  | 
let  | 
| 59582 | 505  | 
val pred $ set_compr = Thm.term_of redex  | 
| 
49763
 
bed063d0c526
generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
 
bulwahn 
parents: 
49761 
diff
changeset
 | 
506  | 
val arg_cong' = instantiate_arg_cong ctxt pred  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
507  | 
in  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
508  | 
conv ctxt set_compr  | 
| 
49763
 
bed063d0c526
generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
 
bulwahn 
parents: 
49761 
diff
changeset
 | 
509  | 
    |> Option.map (fn thm => thm RS arg_cong' RS @{thm eq_reflection})
 | 
| 
48108
 
f93433b451b8
Improved tactic for rewriting set comprehensions into pointfree form.
 
Rafal Kolanski <rafal.kolanski@nicta.com.au> 
parents: 
48055 
diff
changeset
 | 
510  | 
end;  | 
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
511  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
512  | 
fun code_simproc ctxt redex =  | 
| 
48122
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
513  | 
let  | 
| 
50033
 
c78f9cddc907
rewriting with the simpset that is passed to the simproc
 
bulwahn 
parents: 
50032 
diff
changeset
 | 
514  | 
fun unfold_conv thms =  | 
| 
 
c78f9cddc907
rewriting with the simpset that is passed to the simproc
 
bulwahn 
parents: 
50032 
diff
changeset
 | 
515  | 
Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE))  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
516  | 
(empty_simpset ctxt addsimps thms)  | 
| 
50033
 
c78f9cddc907
rewriting with the simpset that is passed to the simproc
 
bulwahn 
parents: 
50032 
diff
changeset
 | 
517  | 
    val prep_thm = unfold_conv @{thms eq_equal[symmetric]} redex
 | 
| 
48122
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
518  | 
in  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51315 
diff
changeset
 | 
519  | 
case base_simproc ctxt (Thm.rhs_of prep_thm) of  | 
| 
48122
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
520  | 
SOME rewr_thm => SOME (transitive_thm OF [transitive_thm OF [prep_thm, rewr_thm],  | 
| 
50033
 
c78f9cddc907
rewriting with the simpset that is passed to the simproc
 
bulwahn 
parents: 
50032 
diff
changeset
 | 
521  | 
        unfold_conv @{thms eq_equal} (Thm.rhs_of rewr_thm)])
 | 
| 
48122
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
522  | 
| NONE => NONE  | 
| 
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
523  | 
end;  | 
| 
 
f479f36ed2ff
adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
 
bulwahn 
parents: 
48109 
diff
changeset
 | 
524  | 
|
| 
48049
 
d862b0d56c49
adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
 
bulwahn 
parents:  
diff
changeset
 | 
525  | 
end;  |