| author | blanchet | 
| Fri, 18 Nov 2011 11:47:12 +0100 | |
| changeset 45572 | 08970468f99b | 
| parent 42163 | 392fd6c4669c | 
| child 50055 | 94041d602ecb | 
| permissions | -rw-r--r-- | 
| 
36017
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
1  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
2  | 
(* Author: Lukas Bulwahn, TU Muenchen *)  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
4  | 
theory New_Random_Sequence  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
5  | 
imports Quickcheck New_DSequence  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
6  | 
begin  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
7  | 
|
| 
42163
 
392fd6c4669c
renewing specifications in HOL: replacing types by type_synonym
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
8  | 
type_synonym 'a pos_random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> 'a New_DSequence.pos_dseq"  | 
| 
 
392fd6c4669c
renewing specifications in HOL: replacing types by type_synonym
 
bulwahn 
parents: 
40049 
diff
changeset
 | 
9  | 
type_synonym 'a neg_random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> 'a New_DSequence.neg_dseq"  | 
| 
36017
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
10  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
definition pos_empty :: "'a pos_random_dseq"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
12  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
13  | 
"pos_empty = (%nrandom size seed. New_DSequence.pos_empty)"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
definition pos_single :: "'a => 'a pos_random_dseq"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
"pos_single x = (%nrandom size seed. New_DSequence.pos_single x)"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
19  | 
definition pos_bind :: "'a pos_random_dseq => ('a \<Rightarrow> 'b pos_random_dseq) \<Rightarrow> 'b pos_random_dseq"
 | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
21  | 
"pos_bind R f = (\<lambda>nrandom size seed. New_DSequence.pos_bind (R nrandom size seed) (%a. f a nrandom size seed))"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
|
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
23  | 
definition pos_decr_bind :: "'a pos_random_dseq => ('a \<Rightarrow> 'b pos_random_dseq) \<Rightarrow> 'b pos_random_dseq"
 | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
24  | 
where  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
25  | 
"pos_decr_bind R f = (\<lambda>nrandom size seed. New_DSequence.pos_decr_bind (R nrandom size seed) (%a. f a nrandom size seed))"  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
26  | 
|
| 
36017
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
27  | 
definition pos_union :: "'a pos_random_dseq => 'a pos_random_dseq => 'a pos_random_dseq"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
28  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
29  | 
"pos_union R1 R2 = (\<lambda>nrandom size seed. New_DSequence.pos_union (R1 nrandom size seed) (R2 nrandom size seed))"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
31  | 
definition pos_if_random_dseq :: "bool => unit pos_random_dseq"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
32  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
33  | 
"pos_if_random_dseq b = (if b then pos_single () else pos_empty)"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
34  | 
|
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36017 
diff
changeset
 | 
35  | 
definition pos_iterate_upto :: "(code_numeral => 'a) => code_numeral => code_numeral => 'a pos_random_dseq"  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36017 
diff
changeset
 | 
36  | 
where  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36017 
diff
changeset
 | 
37  | 
"pos_iterate_upto f n m = (\<lambda>nrandom size seed. New_DSequence.pos_iterate_upto f n m)"  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36017 
diff
changeset
 | 
38  | 
|
| 
36017
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
39  | 
definition pos_not_random_dseq :: "unit neg_random_dseq => unit pos_random_dseq"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
40  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
41  | 
"pos_not_random_dseq R = (\<lambda>nrandom size seed. New_DSequence.pos_not_seq (R nrandom size seed))"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
42  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
43  | 
fun iter :: "(code_numeral * code_numeral => ('a * (unit => term)) * code_numeral * code_numeral) => code_numeral => code_numeral * code_numeral => 'a Lazy_Sequence.lazy_sequence"
 | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
44  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
45  | 
"iter random nrandom seed =  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
46  | 
(if nrandom = 0 then Lazy_Sequence.empty else Lazy_Sequence.Lazy_Sequence (%u. let ((x, _), seed') = random seed in Some (x, iter random (nrandom - 1) seed')))"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
47  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
48  | 
definition Random :: "(code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed)) \<Rightarrow> 'a pos_random_dseq"
 | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
49  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
50  | 
"Random g = (%nrandom size seed depth. iter (g size) nrandom seed)"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
52  | 
definition pos_map :: "('a => 'b) => 'a pos_random_dseq => 'b pos_random_dseq"
 | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
53  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
54  | 
"pos_map f P = pos_bind P (pos_single o f)"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
55  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
57  | 
definition neg_empty :: "'a neg_random_dseq"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
58  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
59  | 
"neg_empty = (%nrandom size seed. New_DSequence.neg_empty)"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
61  | 
definition neg_single :: "'a => 'a neg_random_dseq"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
62  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
63  | 
"neg_single x = (%nrandom size seed. New_DSequence.neg_single x)"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
64  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
65  | 
definition neg_bind :: "'a neg_random_dseq => ('a \<Rightarrow> 'b neg_random_dseq) \<Rightarrow> 'b neg_random_dseq"
 | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
66  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
67  | 
"neg_bind R f = (\<lambda>nrandom size seed. New_DSequence.neg_bind (R nrandom size seed) (%a. f a nrandom size seed))"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
68  | 
|
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
69  | 
definition neg_decr_bind :: "'a neg_random_dseq => ('a \<Rightarrow> 'b neg_random_dseq) \<Rightarrow> 'b neg_random_dseq"
 | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
70  | 
where  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
71  | 
"neg_decr_bind R f = (\<lambda>nrandom size seed. New_DSequence.neg_decr_bind (R nrandom size seed) (%a. f a nrandom size seed))"  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
72  | 
|
| 
36017
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
73  | 
definition neg_union :: "'a neg_random_dseq => 'a neg_random_dseq => 'a neg_random_dseq"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
74  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
75  | 
"neg_union R1 R2 = (\<lambda>nrandom size seed. New_DSequence.neg_union (R1 nrandom size seed) (R2 nrandom size seed))"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
76  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
77  | 
definition neg_if_random_dseq :: "bool => unit neg_random_dseq"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
78  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
79  | 
"neg_if_random_dseq b = (if b then neg_single () else neg_empty)"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
80  | 
|
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36017 
diff
changeset
 | 
81  | 
definition neg_iterate_upto :: "(code_numeral => 'a) => code_numeral => code_numeral => 'a neg_random_dseq"  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36017 
diff
changeset
 | 
82  | 
where  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36017 
diff
changeset
 | 
83  | 
"neg_iterate_upto f n m = (\<lambda>nrandom size seed. New_DSequence.neg_iterate_upto f n m)"  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
36017 
diff
changeset
 | 
84  | 
|
| 
36017
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
85  | 
definition neg_not_random_dseq :: "unit pos_random_dseq => unit neg_random_dseq"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
86  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
87  | 
"neg_not_random_dseq R = (\<lambda>nrandom size seed. New_DSequence.neg_not_seq (R nrandom size seed))"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
88  | 
(*  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
89  | 
fun iter :: "(code_numeral * code_numeral => ('a * (unit => term)) * code_numeral * code_numeral) => code_numeral => code_numeral * code_numeral => 'a Lazy_Sequence.lazy_sequence"
 | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
90  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
91  | 
"iter random nrandom seed =  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
92  | 
(if nrandom = 0 then Lazy_Sequence.empty else Lazy_Sequence.Lazy_Sequence (%u. let ((x, _), seed') = random seed in Some (x, iter random (nrandom - 1) seed')))"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
94  | 
definition Random :: "(code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed)) \<Rightarrow> 'a pos_random_dseq"
 | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
95  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
96  | 
"Random g = (%nrandom size seed depth. iter (g size) nrandom seed)"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
97  | 
*)  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
98  | 
definition neg_map :: "('a => 'b) => 'a neg_random_dseq => 'b neg_random_dseq"
 | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
99  | 
where  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
100  | 
"neg_map f P = neg_bind P (neg_single o f)"  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
101  | 
(*  | 
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36049 
diff
changeset
 | 
102  | 
hide_const DSequence.empty DSequence.single DSequence.eval  | 
| 
36017
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
103  | 
DSequence.map_seq DSequence.bind DSequence.union DSequence.if_seq DSequence.not_seq  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
104  | 
DSequence.map  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
105  | 
*)  | 
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
106  | 
|
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36049 
diff
changeset
 | 
107  | 
hide_const (open)  | 
| 
40049
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
108  | 
pos_empty pos_single pos_bind pos_decr_bind pos_union pos_if_random_dseq pos_iterate_upto pos_not_random_dseq iter Random pos_map  | 
| 
 
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
109  | 
neg_empty neg_single neg_bind neg_decr_bind neg_union neg_if_random_dseq neg_iterate_upto neg_not_random_dseq neg_map  | 
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36049 
diff
changeset
 | 
110  | 
hide_type New_DSequence.pos_dseq New_DSequence.neg_dseq pos_random_dseq neg_random_dseq  | 
| 
36017
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
111  | 
(* FIXME: hide facts *)  | 
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36049 
diff
changeset
 | 
112  | 
(*hide_fact (open) empty_def single_def bind_def union_def if_random_dseq_def not_random_dseq_def Random.simps map_def*)  | 
| 
36017
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
113  | 
|
| 
 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
114  | 
end  |