| author | Andreas Lochbihler | 
| Fri, 20 Sep 2013 10:09:16 +0200 | |
| changeset 53745 | 788730ab7da4 | 
| parent 51143 | 0a2371e7ced3 | 
| child 55437 | 3fd63b92ea3b | 
| permissions | -rw-r--r-- | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Tools/Predicate_Compile/predicate_compile_compilations.ML  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
2  | 
Author: Lukas Bulwahn, TU Muenchen  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
3  | 
|
| 41941 | 4  | 
Structures for different compilations of the predicate compiler.  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
6  | 
|
| 45450 | 7  | 
structure Predicate_Comp_Funs =  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
8  | 
struct  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
9  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
10  | 
fun mk_monadT T = Type (@{type_name Predicate.pred}, [T])
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
12  | 
fun dest_monadT (Type (@{type_name Predicate.pred}, [T])) = T
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
13  | 
  | dest_monadT T = raise TYPE ("dest_monadT", [T], []);
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
15  | 
fun mk_empty T = Const (@{const_name Orderings.bot}, mk_monadT T);
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
fun mk_single t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
let val T = fastype_of t  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
19  | 
  in Const(@{const_name Predicate.single}, T --> mk_monadT T) $ t end;
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
21  | 
fun mk_bind (x, f) =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
  let val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
in  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
24  | 
    Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
27  | 
val mk_plus = HOLogic.mk_binop @{const_name sup};
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
28  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
29  | 
fun mk_if cond = Const (@{const_name Predicate.if_pred},
 | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
30  | 
HOLogic.boolT --> mk_monadT HOLogic.unitT) $ cond;  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
31  | 
|
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
32  | 
fun mk_iterate_upto T (f, from, to) =  | 
| 
46664
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
33  | 
  list_comb (Const (@{const_name Predicate.iterate_upto},
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
34  | 
      [@{typ natural} --> T, @{typ natural}, @{typ natural}] ---> mk_monadT T),
 | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
35  | 
[f, from, to])  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
36  | 
|
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
37  | 
fun mk_not t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
38  | 
let  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
39  | 
val T = mk_monadT HOLogic.unitT  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
40  | 
  in Const (@{const_name Predicate.not_pred}, T --> T) $ t end
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
42  | 
fun mk_Enum f =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
43  | 
  let val T as Type ("fun", [T', _]) = fastype_of f
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
44  | 
in  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
45  | 
    Const (@{const_name Predicate.Pred}, T --> mk_monadT T') $ f    
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
46  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
47  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
48  | 
fun mk_Eval (f, x) =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
49  | 
let  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
50  | 
val T = dest_monadT (fastype_of f)  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
51  | 
in  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
52  | 
    Const (@{const_name Predicate.eval}, mk_monadT T --> T --> HOLogic.boolT) $ f $ x
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
53  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
55  | 
fun dest_Eval (Const (@{const_name Predicate.eval}, _) $ f $ x) = (f, x)
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
57  | 
fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map},
 | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
58  | 
(T1 --> T2) --> mk_monadT T1 --> mk_monadT T2) $ tf $ tp;  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
59  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
60  | 
val compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
61  | 
    {mk_monadT = mk_monadT, dest_monadT = dest_monadT, mk_empty = mk_empty,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
62  | 
mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
63  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map};  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
64  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
65  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
66  | 
|
| 45450 | 67  | 
structure CPS_Comp_Funs =  | 
68  | 
struct  | 
|
69  | 
||
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
70  | 
fun mk_monadT T = (T --> @{typ "Code_Evaluation.term list option"}) --> @{typ "Code_Evaluation.term list option"}
 | 
| 45450 | 71  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
72  | 
fun dest_monadT (Type ("fun", [Type ("fun", [T, @{typ "term list option"}]), @{typ "term list option"}])) = T
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
73  | 
  | dest_monadT T = raise TYPE ("dest_monadT", [T], []);
 | 
| 45450 | 74  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
75  | 
fun mk_empty T = Const (@{const_name Quickcheck_Exhaustive.cps_empty}, mk_monadT T);
 | 
| 45450 | 76  | 
|
77  | 
fun mk_single t =  | 
|
78  | 
let val T = fastype_of t  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
79  | 
  in Const(@{const_name Quickcheck_Exhaustive.cps_single}, T --> mk_monadT T) $ t end;
 | 
| 45450 | 80  | 
|
81  | 
fun mk_bind (x, f) =  | 
|
82  | 
  let val T as Type ("fun", [_, U]) = fastype_of f
 | 
|
83  | 
in  | 
|
84  | 
    Const (@{const_name Quickcheck_Exhaustive.cps_bind}, fastype_of x --> T --> U) $ x $ f
 | 
|
85  | 
end;  | 
|
86  | 
||
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
87  | 
val mk_plus = HOLogic.mk_binop @{const_name Quickcheck_Exhaustive.cps_plus};
 | 
| 45450 | 88  | 
|
89  | 
fun mk_if cond = Const (@{const_name Quickcheck_Exhaustive.cps_if},
 | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
90  | 
HOLogic.boolT --> mk_monadT HOLogic.unitT) $ cond;  | 
| 45450 | 91  | 
|
| 50056 | 92  | 
fun mk_iterate_upto _ _ = error "not implemented yet"  | 
| 45450 | 93  | 
|
94  | 
fun mk_not t =  | 
|
95  | 
let  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
96  | 
val T = mk_monadT HOLogic.unitT  | 
| 45450 | 97  | 
  in Const (@{const_name Quickcheck_Exhaustive.cps_not}, T --> T) $ t end
 | 
98  | 
||
| 50056 | 99  | 
fun mk_Enum _ = error "not implemented"  | 
| 45450 | 100  | 
|
| 50056 | 101  | 
fun mk_Eval _ = error "not implemented"  | 
| 45450 | 102  | 
|
| 50056 | 103  | 
fun dest_Eval _ = error "not implemented"  | 
| 45450 | 104  | 
|
| 50056 | 105  | 
fun mk_map _ _ _ _ = error "not implemented"  | 
| 45450 | 106  | 
|
107  | 
val compfuns = Predicate_Compile_Aux.CompilationFuns  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
108  | 
    {mk_monadT = mk_monadT, dest_monadT = dest_monadT, mk_empty = mk_empty,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
109  | 
mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 45450 | 110  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map};  | 
111  | 
||
112  | 
end;  | 
|
113  | 
||
114  | 
structure Pos_Bounded_CPS_Comp_Funs =  | 
|
115  | 
struct  | 
|
116  | 
||
| 
45750
 
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
 
bulwahn 
parents: 
45461 
diff
changeset
 | 
117  | 
val resultT = @{typ "(bool * Code_Evaluation.term list) option"}
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
118  | 
fun mk_monadT T = (T --> resultT) --> @{typ "natural"} --> resultT
 | 
| 45450 | 119  | 
|
| 
45750
 
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
 
bulwahn 
parents: 
45461 
diff
changeset
 | 
120  | 
fun dest_monadT (Type ("fun", [Type ("fun", [T, @{typ "(bool * term list) option"}]),
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
121  | 
  @{typ "natural => (bool * term list) option"}])) = T
 | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
122  | 
  | dest_monadT T = raise TYPE ("dest_monadT", [T], []);
 | 
| 45450 | 123  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
124  | 
fun mk_empty T = Const (@{const_name Quickcheck_Exhaustive.pos_bound_cps_empty}, mk_monadT T);
 | 
| 45450 | 125  | 
|
126  | 
fun mk_single t =  | 
|
127  | 
let val T = fastype_of t  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
128  | 
  in Const(@{const_name Quickcheck_Exhaustive.pos_bound_cps_single}, T --> mk_monadT T) $ t end;
 | 
| 45450 | 129  | 
|
130  | 
fun mk_bind (x, f) =  | 
|
131  | 
  let val T as Type ("fun", [_, U]) = fastype_of f
 | 
|
132  | 
in  | 
|
133  | 
    Const (@{const_name Quickcheck_Exhaustive.pos_bound_cps_bind}, fastype_of x --> T --> U) $ x $ f
 | 
|
134  | 
end;  | 
|
135  | 
||
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
136  | 
val mk_plus = HOLogic.mk_binop @{const_name Quickcheck_Exhaustive.pos_bound_cps_plus};
 | 
| 45450 | 137  | 
|
138  | 
fun mk_if cond = Const (@{const_name Quickcheck_Exhaustive.pos_bound_cps_if},
 | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
139  | 
HOLogic.boolT --> mk_monadT HOLogic.unitT) $ cond;  | 
| 45450 | 140  | 
|
| 50056 | 141  | 
fun mk_iterate_upto _ _ = error "not implemented yet"  | 
| 45450 | 142  | 
|
143  | 
fun mk_not t =  | 
|
144  | 
let  | 
|
145  | 
    val nT = @{typ "(unit Quickcheck_Exhaustive.unknown =>
 | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
146  | 
Code_Evaluation.term list Quickcheck_Exhaustive.three_valued) => natural =>  | 
| 45450 | 147  | 
Code_Evaluation.term list Quickcheck_Exhaustive.three_valued"}  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
148  | 
val T = mk_monadT HOLogic.unitT  | 
| 45450 | 149  | 
  in Const (@{const_name Quickcheck_Exhaustive.pos_bound_cps_not}, nT --> T) $ t end
 | 
150  | 
||
| 50056 | 151  | 
fun mk_Enum _ = error "not implemented"  | 
| 45450 | 152  | 
|
| 50056 | 153  | 
fun mk_Eval _ = error "not implemented"  | 
| 45450 | 154  | 
|
| 50056 | 155  | 
fun dest_Eval _ = error "not implemented"  | 
| 45450 | 156  | 
|
| 50056 | 157  | 
fun mk_map _ _ _ _ = error "not implemented"  | 
| 45450 | 158  | 
|
159  | 
val compfuns = Predicate_Compile_Aux.CompilationFuns  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
160  | 
    {mk_monadT = mk_monadT, dest_monadT = dest_monadT, mk_empty = mk_empty,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
161  | 
mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 45450 | 162  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map};  | 
163  | 
||
164  | 
end;  | 
|
165  | 
||
166  | 
structure Neg_Bounded_CPS_Comp_Funs =  | 
|
167  | 
struct  | 
|
168  | 
||
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
169  | 
fun mk_monadT T =  | 
| 45450 | 170  | 
  (Type (@{type_name "Quickcheck_Exhaustive.unknown"}, [T])
 | 
171  | 
    --> @{typ "Code_Evaluation.term list Quickcheck_Exhaustive.three_valued"})
 | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
172  | 
    --> @{typ "natural => Code_Evaluation.term list Quickcheck_Exhaustive.three_valued"}
 | 
| 45450 | 173  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
174  | 
fun dest_monadT (Type ("fun", [Type ("fun", [Type (@{type_name "Quickcheck_Exhaustive.unknown"}, [T]),
 | 
| 45450 | 175  | 
    @{typ "term list Quickcheck_Exhaustive.three_valued"}]),
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
176  | 
    @{typ "natural => term list Quickcheck_Exhaustive.three_valued"}])) = T
 | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
177  | 
  | dest_monadT T = raise TYPE ("dest_monadT", [T], []);
 | 
| 45450 | 178  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
179  | 
fun mk_empty T = Const (@{const_name Quickcheck_Exhaustive.neg_bound_cps_empty}, mk_monadT T);
 | 
| 45450 | 180  | 
|
181  | 
fun mk_single t =  | 
|
182  | 
let val T = fastype_of t  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
183  | 
  in Const(@{const_name Quickcheck_Exhaustive.neg_bound_cps_single}, T --> mk_monadT T) $ t end;
 | 
| 45450 | 184  | 
|
185  | 
fun mk_bind (x, f) =  | 
|
186  | 
  let val T as Type ("fun", [_, U]) = fastype_of f
 | 
|
187  | 
in  | 
|
188  | 
    Const (@{const_name Quickcheck_Exhaustive.neg_bound_cps_bind}, fastype_of x --> T --> U) $ x $ f
 | 
|
189  | 
end;  | 
|
190  | 
||
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
191  | 
val mk_plus = HOLogic.mk_binop @{const_name Quickcheck_Exhaustive.neg_bound_cps_plus};
 | 
| 45450 | 192  | 
|
193  | 
fun mk_if cond = Const (@{const_name Quickcheck_Exhaustive.neg_bound_cps_if},
 | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
194  | 
HOLogic.boolT --> mk_monadT HOLogic.unitT) $ cond;  | 
| 45450 | 195  | 
|
| 50056 | 196  | 
fun mk_iterate_upto _ _ = error "not implemented"  | 
| 45450 | 197  | 
|
198  | 
fun mk_not t =  | 
|
199  | 
let  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
200  | 
val T = mk_monadT HOLogic.unitT  | 
| 
45750
 
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
 
bulwahn 
parents: 
45461 
diff
changeset
 | 
201  | 
    val pT = @{typ "(unit => (bool * Code_Evaluation.term list) option)"}
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
202  | 
      --> @{typ "natural => (bool * Code_Evaluation.term list) option"}
 | 
| 45450 | 203  | 
  in Const (@{const_name Quickcheck_Exhaustive.neg_bound_cps_not}, pT --> T) $ t end
 | 
204  | 
||
| 50056 | 205  | 
fun mk_Enum _ = error "not implemented"  | 
| 45450 | 206  | 
|
| 50056 | 207  | 
fun mk_Eval _ = error "not implemented"  | 
| 45450 | 208  | 
|
| 50056 | 209  | 
fun dest_Eval _ = error "not implemented"  | 
| 45450 | 210  | 
|
| 50056 | 211  | 
fun mk_map _ _ _ _ = error "not implemented"  | 
| 45450 | 212  | 
|
213  | 
val compfuns = Predicate_Compile_Aux.CompilationFuns  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
214  | 
    {mk_monadT = mk_monadT, dest_monadT = dest_monadT, mk_empty = mk_empty,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
215  | 
mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 45450 | 216  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map};  | 
217  | 
||
218  | 
end;  | 
|
219  | 
||
220  | 
||
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
221  | 
structure RandomPredCompFuns =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
222  | 
struct  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
223  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
224  | 
fun mk_randompredT T =  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
225  | 
  @{typ Random.seed} --> HOLogic.mk_prodT (Predicate_Comp_Funs.mk_monadT T, @{typ Random.seed})
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
226  | 
|
| 
37678
 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 
haftmann 
parents: 
36049 
diff
changeset
 | 
227  | 
fun dest_randompredT (Type ("fun", [@{typ Random.seed}, Type (@{type_name Product_Type.prod},
 | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
228  | 
  [Type (@{type_name Predicate.pred}, [T]), @{typ Random.seed}])])) = T
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
229  | 
  | dest_randompredT T = raise TYPE ("dest_randompredT", [T], []);
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
230  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
231  | 
fun mk_empty T = Const(@{const_name Random_Pred.empty}, mk_randompredT T)
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
232  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
233  | 
fun mk_single t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
234  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
235  | 
val T = fastype_of t  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
236  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
237  | 
    Const (@{const_name Random_Pred.single}, T --> mk_randompredT T) $ t
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
238  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
239  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
240  | 
fun mk_bind (x, f) =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
241  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
242  | 
    val T as (Type ("fun", [_, U])) = fastype_of f
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
243  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
244  | 
    Const (@{const_name Random_Pred.bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
245  | 
end  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
246  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
247  | 
val mk_plus = HOLogic.mk_binop @{const_name Random_Pred.union}
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
248  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
249  | 
fun mk_if cond = Const (@{const_name Random_Pred.if_randompred},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
250  | 
HOLogic.boolT --> mk_randompredT HOLogic.unitT) $ cond;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
251  | 
|
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
252  | 
fun mk_iterate_upto T (f, from, to) =  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
253  | 
  list_comb (Const (@{const_name Random_Pred.iterate_upto},
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
254  | 
      [@{typ natural} --> T, @{typ natural}, @{typ natural}] ---> mk_randompredT T),
 | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
255  | 
[f, from, to])  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
256  | 
|
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
257  | 
fun mk_not t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
258  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
259  | 
val T = mk_randompredT HOLogic.unitT  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
260  | 
  in Const (@{const_name Random_Pred.not_randompred}, T --> T) $ t end
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
261  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
262  | 
fun mk_map T1 T2 tf tp = Const (@{const_name Random_Pred.map},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
263  | 
(T1 --> T2) --> mk_randompredT T1 --> mk_randompredT T2) $ tf $ tp  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
264  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
265  | 
val compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
266  | 
    {mk_monadT = mk_randompredT, dest_monadT = dest_randompredT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
267  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
268  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map};  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
269  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
270  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
271  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
272  | 
structure DSequence_CompFuns =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
273  | 
struct  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
274  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
275  | 
fun mk_dseqT T = Type ("fun", [@{typ natural}, Type ("fun", [@{typ bool},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
276  | 
  Type (@{type_name Option.option}, [Type  (@{type_name Lazy_Sequence.lazy_sequence}, [T])])])])
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
277  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
278  | 
fun dest_dseqT (Type ("fun", [@{typ natural}, Type ("fun", [@{typ bool},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
279  | 
  Type (@{type_name Option.option}, [Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])])])])) = T
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
280  | 
  | dest_dseqT T = raise TYPE ("dest_dseqT", [T], []);
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
281  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
282  | 
fun mk_empty T = Const (@{const_name Limited_Sequence.empty}, mk_dseqT T);
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
283  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
284  | 
fun mk_single t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
285  | 
let val T = fastype_of t  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
286  | 
  in Const(@{const_name Limited_Sequence.single}, T --> mk_dseqT T) $ t end;
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
287  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
288  | 
fun mk_bind (x, f) =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
289  | 
  let val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
290  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
291  | 
    Const (@{const_name Limited_Sequence.bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
292  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
293  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
294  | 
val mk_plus = HOLogic.mk_binop @{const_name Limited_Sequence.union};
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
295  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
296  | 
fun mk_if cond = Const (@{const_name Limited_Sequence.if_seq},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
297  | 
HOLogic.boolT --> mk_dseqT HOLogic.unitT) $ cond;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
298  | 
|
| 50056 | 299  | 
fun mk_iterate_upto _ _ = raise Fail "No iterate_upto compilation"  | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
300  | 
|
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
301  | 
fun mk_not t = let val T = mk_dseqT HOLogic.unitT  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
302  | 
  in Const (@{const_name Limited_Sequence.not_seq}, T --> T) $ t end
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
303  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
304  | 
fun mk_map T1 T2 tf tp = Const (@{const_name Limited_Sequence.map},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
305  | 
(T1 --> T2) --> mk_dseqT T1 --> mk_dseqT T2) $ tf $ tp  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
306  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
307  | 
val compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
308  | 
    {mk_monadT = mk_dseqT, dest_monadT = dest_dseqT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
309  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
310  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
312  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
313  | 
|
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
314  | 
structure New_Pos_DSequence_CompFuns =  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
315  | 
struct  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
316  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
317  | 
fun mk_pos_dseqT T =  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
318  | 
    @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
319  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
320  | 
fun dest_pos_dseqT (Type ("fun", [@{typ natural},
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
321  | 
    Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])])) = T
 | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
322  | 
  | dest_pos_dseqT T = raise TYPE ("dest_pos_dseqT", [T], []);
 | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
323  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
324  | 
fun mk_empty T = Const (@{const_name Limited_Sequence.pos_empty}, mk_pos_dseqT T);
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
325  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
326  | 
fun mk_single t =  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
327  | 
let  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
328  | 
val T = fastype_of t  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
329  | 
  in Const(@{const_name Limited_Sequence.pos_single}, T --> mk_pos_dseqT T) $ t end;
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
330  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
331  | 
fun mk_bind (x, f) =  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
332  | 
let  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
333  | 
    val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
334  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
335  | 
    Const (@{const_name Limited_Sequence.pos_bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
336  | 
end;  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
337  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
338  | 
fun mk_decr_bind (x, f) =  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
339  | 
let  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
340  | 
    val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
341  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
342  | 
    Const (@{const_name Limited_Sequence.pos_decr_bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
343  | 
end;  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
344  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
345  | 
val mk_plus = HOLogic.mk_binop @{const_name Limited_Sequence.pos_union};
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
346  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
347  | 
fun mk_if cond = Const (@{const_name Limited_Sequence.pos_if_seq},
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
348  | 
HOLogic.boolT --> mk_pos_dseqT HOLogic.unitT) $ cond;  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
349  | 
|
| 50056 | 350  | 
fun mk_iterate_upto _ _ = raise Fail "No iterate_upto compilation"  | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
351  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
352  | 
fun mk_not t =  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
353  | 
let  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
354  | 
val pT = mk_pos_dseqT HOLogic.unitT  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
355  | 
val nT =  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
356  | 
      @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
357  | 
        [Type (@{type_name Option.option}, [@{typ unit}])])
 | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
358  | 
  in Const (@{const_name Limited_Sequence.pos_not_seq}, nT --> pT) $ t end
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
359  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
360  | 
fun mk_map T1 T2 tf tp = Const (@{const_name Limited_Sequence.pos_map},
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
361  | 
(T1 --> T2) --> mk_pos_dseqT T1 --> mk_pos_dseqT T2) $ tf $ tp  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
362  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
363  | 
val depth_limited_compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
364  | 
    {mk_monadT = mk_pos_dseqT, dest_monadT = dest_pos_dseqT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
365  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_decr_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
366  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
367  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
368  | 
val depth_unlimited_compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
369  | 
    {mk_monadT = mk_pos_dseqT, dest_monadT = dest_pos_dseqT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
370  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
371  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
372  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
373  | 
end;  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
374  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
375  | 
structure New_Neg_DSequence_CompFuns =  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
376  | 
struct  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
377  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
378  | 
fun mk_neg_dseqT T = @{typ natural} -->
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
379  | 
  Type (@{type_name Lazy_Sequence.lazy_sequence}, [Type (@{type_name Option.option}, [T])])
 | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
380  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
381  | 
fun dest_neg_dseqT (Type ("fun", [@{typ natural},
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
382  | 
    Type (@{type_name Lazy_Sequence.lazy_sequence}, [Type (@{type_name Option.option}, [T])])])) = T
 | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
383  | 
  | dest_neg_dseqT T = raise TYPE ("dest_neg_dseqT", [T], []);
 | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
384  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
385  | 
fun mk_empty T = Const (@{const_name Limited_Sequence.neg_empty}, mk_neg_dseqT T);
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
386  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
387  | 
fun mk_single t =  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
388  | 
let  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
389  | 
val T = fastype_of t  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
390  | 
  in Const(@{const_name Limited_Sequence.neg_single}, T --> mk_neg_dseqT T) $ t end;
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
391  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
392  | 
fun mk_bind (x, f) =  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
393  | 
let  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
394  | 
    val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
395  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
396  | 
    Const (@{const_name Limited_Sequence.neg_bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
397  | 
end;  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
398  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
399  | 
fun mk_decr_bind (x, f) =  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
400  | 
let  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
401  | 
    val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
402  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
403  | 
    Const (@{const_name Limited_Sequence.neg_decr_bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
404  | 
end;  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
405  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
406  | 
val mk_plus = HOLogic.mk_binop @{const_name Limited_Sequence.neg_union};
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
407  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
408  | 
fun mk_if cond = Const (@{const_name Limited_Sequence.neg_if_seq},
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
409  | 
HOLogic.boolT --> mk_neg_dseqT HOLogic.unitT) $ cond;  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
410  | 
|
| 50056 | 411  | 
fun mk_iterate_upto _ _ = raise Fail "No iterate_upto compilation"  | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
412  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
413  | 
fun mk_not t =  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
414  | 
let  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
415  | 
val nT = mk_neg_dseqT HOLogic.unitT  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
416  | 
val pT =  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
417  | 
      @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
418  | 
        [@{typ unit}])
 | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
419  | 
  in Const (@{const_name Limited_Sequence.neg_not_seq}, pT --> nT) $ t end
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
420  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
421  | 
fun mk_map T1 T2 tf tp = Const (@{const_name Limited_Sequence.neg_map},
 | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
422  | 
(T1 --> T2) --> mk_neg_dseqT T1 --> mk_neg_dseqT T2) $ tf $ tp  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
423  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
424  | 
val depth_limited_compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
425  | 
    {mk_monadT = mk_neg_dseqT, dest_monadT = dest_neg_dseqT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
426  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_decr_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
427  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
428  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
429  | 
val depth_unlimited_compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
430  | 
    {mk_monadT = mk_neg_dseqT, dest_monadT = dest_neg_dseqT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
431  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
432  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
433  | 
|
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
434  | 
end;  | 
| 
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
435  | 
|
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
436  | 
structure New_Pos_Random_Sequence_CompFuns =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
437  | 
struct  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
438  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
439  | 
fun mk_pos_random_dseqT T =  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
440  | 
  @{typ natural} --> @{typ natural} --> @{typ Random.seed} -->
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
441  | 
    @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
442  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
443  | 
fun dest_pos_random_dseqT (Type ("fun", [@{typ natural}, Type ("fun", [@{typ natural},
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
444  | 
    Type ("fun", [@{typ Random.seed}, Type ("fun", [@{typ natural},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
445  | 
    Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])])])])])) = T
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
446  | 
  | dest_pos_random_dseqT T = raise TYPE ("dest_random_dseqT", [T], []);
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
447  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
448  | 
fun mk_empty T = Const (@{const_name Random_Sequence.pos_empty}, mk_pos_random_dseqT T);
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
449  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
450  | 
fun mk_single t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
451  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
452  | 
val T = fastype_of t  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
453  | 
  in Const(@{const_name Random_Sequence.pos_single}, T --> mk_pos_random_dseqT T) $ t end;
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
454  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
455  | 
fun mk_bind (x, f) =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
456  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
457  | 
    val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
458  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
459  | 
    Const (@{const_name Random_Sequence.pos_bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
460  | 
end;  | 
| 
40051
 
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
461  | 
|
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
462  | 
fun mk_decr_bind (x, f) =  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
463  | 
let  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
464  | 
    val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
465  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
466  | 
    Const (@{const_name Random_Sequence.pos_decr_bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
467  | 
end;  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
468  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
469  | 
val mk_plus = HOLogic.mk_binop @{const_name Random_Sequence.pos_union};
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
470  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
471  | 
fun mk_if cond = Const (@{const_name Random_Sequence.pos_if_random_dseq},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
472  | 
HOLogic.boolT --> mk_pos_random_dseqT HOLogic.unitT) $ cond;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
473  | 
|
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
474  | 
fun mk_iterate_upto T (f, from, to) =  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
475  | 
  list_comb (Const (@{const_name Random_Sequence.pos_iterate_upto},
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
476  | 
      [@{typ natural} --> T, @{typ natural}, @{typ natural}]
 | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
477  | 
---> mk_pos_random_dseqT T),  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
478  | 
[f, from, to])  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
479  | 
|
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
480  | 
fun mk_not t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
481  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
482  | 
val pT = mk_pos_random_dseqT HOLogic.unitT  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
483  | 
    val nT = @{typ natural} --> @{typ natural} --> @{typ Random.seed} -->
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
484  | 
      @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
485  | 
        [Type (@{type_name Option.option}, [@{typ unit}])])
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
486  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
487  | 
  in Const (@{const_name Random_Sequence.pos_not_random_dseq}, nT --> pT) $ t end
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
488  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
489  | 
fun mk_map T1 T2 tf tp = Const (@{const_name Random_Sequence.pos_map},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
490  | 
(T1 --> T2) --> mk_pos_random_dseqT T1 --> mk_pos_random_dseqT T2) $ tf $ tp  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
491  | 
|
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
492  | 
val depth_limited_compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
493  | 
    {mk_monadT = mk_pos_random_dseqT, dest_monadT = dest_pos_random_dseqT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
494  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_decr_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
495  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
496  | 
|
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
497  | 
val depth_unlimited_compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
498  | 
    {mk_monadT = mk_pos_random_dseqT, dest_monadT = dest_pos_random_dseqT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
499  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
500  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
501  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
502  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
503  | 
structure New_Neg_Random_Sequence_CompFuns =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
504  | 
struct  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
505  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
506  | 
fun mk_neg_random_dseqT T =  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
507  | 
   @{typ natural} --> @{typ natural} --> @{typ Random.seed} -->
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
508  | 
    @{typ natural} --> 
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
509  | 
    Type (@{type_name Lazy_Sequence.lazy_sequence}, [Type (@{type_name Option.option}, [T])])
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
510  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
511  | 
fun dest_neg_random_dseqT (Type ("fun", [@{typ natural}, Type ("fun", [@{typ natural},
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
512  | 
    Type ("fun", [@{typ Random.seed}, Type ("fun", [@{typ natural},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
513  | 
      Type (@{type_name Lazy_Sequence.lazy_sequence},
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
514  | 
        [Type (@{type_name Option.option}, [T])])])])])])) = T
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
515  | 
  | dest_neg_random_dseqT T = raise TYPE ("dest_random_dseqT", [T], []);
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
516  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
517  | 
fun mk_empty T = Const (@{const_name Random_Sequence.neg_empty}, mk_neg_random_dseqT T);
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
518  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
519  | 
fun mk_single t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
520  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
521  | 
val T = fastype_of t  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
522  | 
  in Const(@{const_name Random_Sequence.neg_single}, T --> mk_neg_random_dseqT T) $ t end;
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
523  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
524  | 
fun mk_bind (x, f) =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
525  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
526  | 
    val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
527  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
528  | 
    Const (@{const_name Random_Sequence.neg_bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
529  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
530  | 
|
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
531  | 
fun mk_decr_bind (x, f) =  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
532  | 
let  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
533  | 
    val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
534  | 
in  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
535  | 
    Const (@{const_name Random_Sequence.neg_decr_bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
536  | 
end;  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
537  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
538  | 
val mk_plus = HOLogic.mk_binop @{const_name Random_Sequence.neg_union};
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
539  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
540  | 
fun mk_if cond = Const (@{const_name Random_Sequence.neg_if_random_dseq},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
541  | 
HOLogic.boolT --> mk_neg_random_dseqT HOLogic.unitT) $ cond;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
542  | 
|
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
543  | 
fun mk_iterate_upto T (f, from, to) =  | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
544  | 
  list_comb (Const (@{const_name Random_Sequence.neg_iterate_upto},
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
545  | 
      [@{typ natural} --> T, @{typ natural}, @{typ natural}]
 | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
546  | 
---> mk_neg_random_dseqT T),  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
547  | 
[f, from, to])  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
548  | 
|
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
549  | 
fun mk_not t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
550  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
551  | 
val nT = mk_neg_random_dseqT HOLogic.unitT  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
552  | 
    val pT = @{typ natural} --> @{typ natural} --> @{typ Random.seed} -->
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
553  | 
    @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence}, [@{typ unit}])
 | 
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
554  | 
  in Const (@{const_name Random_Sequence.neg_not_random_dseq}, pT --> nT) $ t end
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
555  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
50056 
diff
changeset
 | 
556  | 
fun mk_map T1 T2 tf tp = Const (@{const_name Random_Sequence.neg_map},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
557  | 
(T1 --> T2) --> mk_neg_random_dseqT T1 --> mk_neg_random_dseqT T2) $ tf $ tp  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
558  | 
|
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
559  | 
val depth_limited_compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
560  | 
    {mk_monadT = mk_neg_random_dseqT, dest_monadT = dest_neg_random_dseqT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
561  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_decr_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
562  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
563  | 
|
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
39648 
diff
changeset
 | 
564  | 
val depth_unlimited_compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
565  | 
    {mk_monadT = mk_neg_random_dseqT, dest_monadT = dest_neg_random_dseqT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
566  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
567  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
568  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
569  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
570  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
571  | 
structure Random_Sequence_CompFuns =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
572  | 
struct  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
573  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
574  | 
fun mk_random_dseqT T =  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
575  | 
  @{typ natural} --> @{typ natural} --> @{typ Random.seed} -->
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
576  | 
    HOLogic.mk_prodT (DSequence_CompFuns.mk_dseqT T, @{typ Random.seed})
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
577  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
578  | 
fun dest_random_dseqT (Type ("fun", [@{typ natural}, Type ("fun", [@{typ natural},
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
579  | 
  Type ("fun", [@{typ Random.seed},
 | 
| 
37678
 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 
haftmann 
parents: 
36049 
diff
changeset
 | 
580  | 
  Type (@{type_name Product_Type.prod}, [T, @{typ Random.seed}])])])])) = DSequence_CompFuns.dest_dseqT T
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
581  | 
  | dest_random_dseqT T = raise TYPE ("dest_random_dseqT", [T], []);
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
582  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
583  | 
fun mk_empty T = Const (@{const_name Random_Sequence.empty}, mk_random_dseqT T);
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
584  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
585  | 
fun mk_single t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
586  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
587  | 
val T = fastype_of t  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
588  | 
  in Const(@{const_name Random_Sequence.single}, T --> mk_random_dseqT T) $ t end;
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
589  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
590  | 
fun mk_bind (x, f) =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
591  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
592  | 
    val T as Type ("fun", [_, U]) = fastype_of f
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
593  | 
in  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
594  | 
    Const (@{const_name Random_Sequence.bind}, fastype_of x --> T --> U) $ x $ f
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
595  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
596  | 
|
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
597  | 
val mk_plus = HOLogic.mk_binop @{const_name Random_Sequence.union};
 | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
598  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
599  | 
fun mk_if cond = Const (@{const_name Random_Sequence.if_random_dseq},
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
600  | 
HOLogic.boolT --> mk_random_dseqT HOLogic.unitT) $ cond;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
601  | 
|
| 50056 | 602  | 
fun mk_iterate_upto _ _ = raise Fail "No iterate_upto compilation"  | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
603  | 
|
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
604  | 
fun mk_not t =  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
605  | 
let  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
606  | 
val T = mk_random_dseqT HOLogic.unitT  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
607  | 
  in Const (@{const_name Random_Sequence.not_random_dseq}, T --> T) $ t end
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
608  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
609  | 
fun mk_map T1 T2 tf tp = Const (@{const_name Random_Sequence.map},
 | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
610  | 
(T1 --> T2) --> mk_random_dseqT T1 --> mk_random_dseqT T2) $ tf $ tp  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
611  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
612  | 
val compfuns = Predicate_Compile_Aux.CompilationFuns  | 
| 
45461
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
613  | 
    {mk_monadT = mk_random_dseqT, dest_monadT = dest_random_dseqT,
 | 
| 
 
130c90bb80b4
using more conventional names for monad plus operations
 
bulwahn 
parents: 
45450 
diff
changeset
 | 
614  | 
mk_empty = mk_empty, mk_single = mk_single, mk_bind = mk_bind, mk_plus = mk_plus, mk_if = mk_if,  | 
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36046 
diff
changeset
 | 
615  | 
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}  | 
| 
36046
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
616  | 
|
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
617  | 
end;  | 
| 
 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 
bulwahn 
parents:  
diff
changeset
 | 
618  |