| author | wenzelm | 
| Tue, 16 Aug 2011 21:54:06 +0200 | |
| changeset 44224 | 4040d0ffac7b | 
| 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  | 
header {* Depth-Limited Sequences with failure element *}
 | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
6  | 
theory DSequence  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
7  | 
imports Lazy_Sequence Code_Numeral  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
8  | 
begin  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
9  | 
|
| 
42163
 
392fd6c4669c
renewing specifications in HOL: replacing types by type_synonym
 
bulwahn 
parents: 
36176 
diff
changeset
 | 
10  | 
type_synonym 'a dseq = "code_numeral => bool => 'a Lazy_Sequence.lazy_sequence option"  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
12  | 
definition empty :: "'a dseq"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
13  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
"empty = (%i pol. Some Lazy_Sequence.empty)"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
definition single :: "'a => 'a dseq"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
"single x = (%i pol. Some (Lazy_Sequence.single x))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
fun eval :: "'a dseq => code_numeral => bool => 'a Lazy_Sequence.lazy_sequence option"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
21  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
"eval f i pol = f i pol"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
24  | 
definition yield :: "'a dseq => code_numeral => bool => ('a * 'a dseq) option" 
 | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
"yield dseq i pol = (case eval dseq i pol of  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
27  | 
None => None  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
28  | 
| Some s => Option.map (apsnd (%r i pol. Some r)) (Lazy_Sequence.yield s))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
30  | 
fun map_seq :: "('a => 'b dseq) => 'a Lazy_Sequence.lazy_sequence => 'b dseq"
 | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
31  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
32  | 
"map_seq f xq i pol = (case Lazy_Sequence.yield xq of  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
33  | 
None => Some Lazy_Sequence.empty  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
34  | 
| Some (x, xq') => (case eval (f x) i pol of  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
35  | 
None => None  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
36  | 
| Some yq => (case map_seq f xq' i pol of  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
37  | 
None => None  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
38  | 
| Some zq => Some (Lazy_Sequence.append yq zq))))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
39  | 
|
| 34953 | 40  | 
definition bind :: "'a dseq => ('a => 'b dseq) => 'b dseq"
 | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
41  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
42  | 
"bind x f = (%i pol.  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
43  | 
if i = 0 then  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
44  | 
(if pol then Some Lazy_Sequence.empty else None)  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
45  | 
else  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
46  | 
(case x (i - 1) pol of  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
47  | 
None => None  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
48  | 
| Some xq => map_seq f xq i pol))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
49  | 
|
| 34953 | 50  | 
definition union :: "'a dseq => 'a dseq => 'a dseq"  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
51  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
52  | 
"union x y = (%i pol. case (x i pol, y i pol) of  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
53  | 
(Some xq, Some yq) => Some (Lazy_Sequence.append xq yq)  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
54  | 
| _ => None)"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
55  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
56  | 
definition if_seq :: "bool => unit dseq"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
57  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
58  | 
"if_seq 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
 | 
59  | 
|
| 34953 | 60  | 
definition not_seq :: "unit dseq => unit dseq"  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
61  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
62  | 
"not_seq x = (%i pol. case x i (\<not>pol) of  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
63  | 
None => Some Lazy_Sequence.empty  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
64  | 
| Some xq => (case Lazy_Sequence.yield xq of  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
65  | 
None => Some (Lazy_Sequence.single ())  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
66  | 
| Some _ => Some (Lazy_Sequence.empty)))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
67  | 
|
| 34953 | 68  | 
definition map :: "('a => 'b) => 'a dseq => 'b dseq"
 | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
69  | 
where  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
70  | 
"map f g = (%i pol. case g i pol of  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
71  | 
None => None  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
72  | 
| Some xq => Some (Lazy_Sequence.map f xq))"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
74  | 
ML {*
 | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
75  | 
signature DSEQUENCE =  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
76  | 
sig  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
77  | 
type 'a dseq = int -> bool -> 'a Lazy_Sequence.lazy_sequence option  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
78  | 
  val yield : 'a dseq -> int -> bool -> ('a * 'a dseq) option
 | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
79  | 
val yieldn : int -> 'a dseq -> int -> bool -> 'a list * 'a dseq  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
80  | 
  val map : ('a -> 'b) -> 'a dseq -> 'b dseq
 | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
81  | 
end;  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
82  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
83  | 
structure DSequence : DSEQUENCE =  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
84  | 
struct  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
85  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
86  | 
type 'a dseq = int -> bool -> 'a Lazy_Sequence.lazy_sequence option  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
87  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
88  | 
val yield = @{code yield}
 | 
| 
36024
 
c1ce2f60b0f2
removed yieldn in Lazy_Sequence and put in the ML structure; corrects behaviour of values command
 
bulwahn 
parents: 
34953 
diff
changeset
 | 
89  | 
fun yieldn n s d pol = case s d pol of  | 
| 
 
c1ce2f60b0f2
removed yieldn in Lazy_Sequence and put in the ML structure; corrects behaviour of values command
 
bulwahn 
parents: 
34953 
diff
changeset
 | 
90  | 
NONE => ([], fn d => fn pol => NONE)  | 
| 
 
c1ce2f60b0f2
removed yieldn in Lazy_Sequence and put in the ML structure; corrects behaviour of values command
 
bulwahn 
parents: 
34953 
diff
changeset
 | 
91  | 
| SOME s => let val (xs, s') = Lazy_Sequence.yieldn n s in (xs, fn d => fn pol => SOME s') end  | 
| 
 
c1ce2f60b0f2
removed yieldn in Lazy_Sequence and put in the ML structure; corrects behaviour of values command
 
bulwahn 
parents: 
34953 
diff
changeset
 | 
92  | 
|
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
93  | 
val map = @{code map}
 | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
94  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
95  | 
end;  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
96  | 
*}  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
97  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
98  | 
code_reserved Eval DSequence  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
99  | 
(*  | 
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36024 
diff
changeset
 | 
100  | 
hide_type Sequence.seq  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
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: 
36024 
diff
changeset
 | 
102  | 
hide_const Sequence.Seq Sequence.yield Sequence.yieldn Sequence.empty Sequence.single  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
103  | 
Sequence.append Sequence.flat Sequence.map Sequence.bind Sequence.ifpred Sequence.not_seq  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
104  | 
*)  | 
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36024 
diff
changeset
 | 
105  | 
hide_const (open) empty single eval map_seq bind union if_seq not_seq map  | 
| 
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36024 
diff
changeset
 | 
106  | 
hide_fact (open) empty_def single_def eval.simps map_seq.simps bind_def union_def  | 
| 34953 | 107  | 
if_seq_def not_seq_def 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
 | 
108  | 
|
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents:  
diff
changeset
 | 
109  | 
end  |