| author | huffman | 
| Mon, 12 Sep 2011 11:39:29 -0700 | |
| changeset 44906 | 8f3625167c76 | 
| parent 42163 | 392fd6c4669c | 
| child 50055 | 94041d602ecb | 
| permissions | -rw-r--r-- | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
1  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
2  | 
(* Author: Lukas Bulwahn, TU Muenchen *)  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
4  | 
theory Random_Sequence  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
5  | 
imports Quickcheck DSequence  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
6  | 
begin  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
7  | 
|
| 
42163
 
392fd6c4669c
renewing specifications in HOL: replacing types by type_synonym
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
8  | 
type_synonym 'a random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a DSequence.dseq \<times> Random.seed)"
 | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
10  | 
definition empty :: "'a random_dseq"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
12  | 
"empty = (%nrandom size. Pair (DSequence.empty))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
13  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
definition single :: "'a => 'a random_dseq"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
"single x = (%nrandom size. Pair (DSequence.single x))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
definition bind :: "'a random_dseq => ('a \<Rightarrow> 'b random_dseq) \<Rightarrow> 'b random_dseq"
 | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
19  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
"bind R f = (\<lambda>nrandom size s. let  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
21  | 
(P, s') = R nrandom size s;  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
(s1, s2) = Random.split_seed s'  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
in (DSequence.bind P (%a. fst (f a nrandom size s1)), s2))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
definition union :: "'a random_dseq => 'a random_dseq => 'a random_dseq"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
27  | 
"union R1 R2 = (\<lambda>nrandom size s. let  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
28  | 
(S1, s') = R1 nrandom size s; (S2, s'') = R2 nrandom size s'  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
29  | 
in (DSequence.union S1 S2, s''))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
31  | 
definition if_random_dseq :: "bool => unit random_dseq"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
32  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
33  | 
"if_random_dseq b = (if b then single () else empty)"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
34  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
35  | 
definition not_random_dseq :: "unit random_dseq => unit random_dseq"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
36  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
37  | 
"not_random_dseq R = (\<lambda>nrandom size s. let  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
38  | 
(S, s') = R nrandom size s  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
39  | 
in (DSequence.not_seq S, s'))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
40  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
41  | 
fun Random :: "(code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed)) \<Rightarrow> 'a random_dseq"
 | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
42  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
43  | 
"Random g nrandom = (%size. if nrandom <= 0 then (Pair DSequence.empty) else  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
44  | 
(scomp (g size) (%r. scomp (Random g (nrandom - 1) size) (%rs. Pair (DSequence.union (DSequence.single (fst r)) rs)))))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
46  | 
definition map :: "('a => 'b) => 'a random_dseq => 'b random_dseq"
 | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
47  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
48  | 
"map f P = bind P (single o f)"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
50  | 
(*  | 
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
34953 
diff
changeset
 | 
51  | 
hide_const DSequence.empty DSequence.single DSequence.eval  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
52  | 
DSequence.map_seq DSequence.bind DSequence.union DSequence.if_seq DSequence.not_seq  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
53  | 
DSequence.map  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
54  | 
*)  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
55  | 
|
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
34953 
diff
changeset
 | 
56  | 
hide_const (open) empty single bind union if_random_dseq not_random_dseq Random map  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
57  | 
|
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
34953 
diff
changeset
 | 
58  | 
hide_type DSequence.dseq random_dseq  | 
| 
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
34953 
diff
changeset
 | 
59  | 
hide_fact (open) empty_def single_def bind_def union_def if_random_dseq_def not_random_dseq_def Random.simps map_def  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
61  | 
end  |