| author | huffman | 
| Thu, 14 Oct 2010 09:34:00 -0700 | |
| changeset 40013 | 9db8fb58fddc | 
| parent 39183 | 512c10416590 | 
| child 40049 | 75d9f57123d6 | 
| 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 | header {* Depth-Limited Sequences with failure element *}
 | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 5 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 6 | theory New_DSequence | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 7 | imports Random_Sequence | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 8 | begin | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 9 | |
| 36902 
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
 huffman parents: 
36176diff
changeset | 10 | subsection {* Positive Depth-Limited Sequence *}
 | 
| 36017 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 11 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 12 | types 'a pos_dseq = "code_numeral => 'a Lazy_Sequence.lazy_sequence" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 13 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 14 | definition pos_empty :: "'a pos_dseq" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 15 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 16 | "pos_empty = (%i. Lazy_Sequence.empty)" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 17 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 18 | definition pos_single :: "'a => 'a pos_dseq" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 19 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 20 | "pos_single x = (%i. Lazy_Sequence.single x)" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 21 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 22 | definition pos_bind :: "'a pos_dseq => ('a => 'b pos_dseq) => 'b pos_dseq"
 | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 23 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 24 | "pos_bind x f = (%i. | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 25 | if i = 0 then | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 26 | Lazy_Sequence.empty | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 27 | else | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 28 | Lazy_Sequence.bind (x (i - 1)) (%a. f a i))" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 29 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 30 | definition pos_union :: "'a pos_dseq => 'a pos_dseq => 'a pos_dseq" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 31 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 32 | "pos_union xq yq = (%i. Lazy_Sequence.append (xq i) (yq i))" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 33 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 34 | definition pos_if_seq :: "bool => unit pos_dseq" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 35 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 36 | "pos_if_seq 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 | 37 | |
| 36049 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 bulwahn parents: 
36017diff
changeset | 38 | definition pos_iterate_upto :: "(code_numeral => 'a) => code_numeral => code_numeral => 'a pos_dseq" | 
| 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 bulwahn parents: 
36017diff
changeset | 39 | where | 
| 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 bulwahn parents: 
36017diff
changeset | 40 | "pos_iterate_upto f n m = (%i. Lazy_Sequence.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: 
36017diff
changeset | 41 | |
| 36017 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 42 | definition pos_map :: "('a => 'b) => 'a pos_dseq => 'b pos_dseq"
 | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 43 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 44 | "pos_map f xq = (%i. Lazy_Sequence.map f (xq i))" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 45 | |
| 36902 
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
 huffman parents: 
36176diff
changeset | 46 | subsection {* Negative Depth-Limited Sequence *}
 | 
| 36017 
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 | types 'a neg_dseq = "code_numeral => 'a Lazy_Sequence.hit_bound_lazy_sequence" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 49 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 50 | definition neg_empty :: "'a neg_dseq" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 51 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 52 | "neg_empty = (%i. Lazy_Sequence.empty)" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 53 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 54 | definition neg_single :: "'a => 'a neg_dseq" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 55 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 56 | "neg_single x = (%i. Lazy_Sequence.hb_single x)" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 57 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 58 | definition neg_bind :: "'a neg_dseq => ('a => 'b neg_dseq) => 'b neg_dseq"
 | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 59 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 60 | "neg_bind x f = (%i. | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 61 | if i = 0 then | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 62 | Lazy_Sequence.hit_bound | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 63 | else | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 64 | hb_bind (x (i - 1)) (%a. f a i))" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 65 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 66 | definition neg_union :: "'a neg_dseq => 'a neg_dseq => 'a neg_dseq" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 67 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 68 | "neg_union x y = (%i. Lazy_Sequence.append (x i) (y i))" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 69 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 70 | definition neg_if_seq :: "bool => unit neg_dseq" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 71 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 72 | "neg_if_seq 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 | 73 | |
| 36049 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 bulwahn parents: 
36017diff
changeset | 74 | definition neg_iterate_upto | 
| 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 bulwahn parents: 
36017diff
changeset | 75 | where | 
| 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 bulwahn parents: 
36017diff
changeset | 76 | "neg_iterate_upto f n m = (%i. Lazy_Sequence.iterate_upto (%i. Some (f i)) 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: 
36017diff
changeset | 77 | |
| 36017 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 78 | definition neg_map :: "('a => 'b) => 'a neg_dseq => 'b neg_dseq"
 | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 79 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 80 | "neg_map f xq = (%i. Lazy_Sequence.hb_map f (xq i))" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 81 | |
| 36902 
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
 huffman parents: 
36176diff
changeset | 82 | subsection {* Negation *}
 | 
| 36017 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 83 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 84 | definition pos_not_seq :: "unit neg_dseq => unit pos_dseq" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 85 | where | 
| 39183 
512c10416590
adding dependencies to IsaMakefile; increasing negative search limit for predicate_compile_quickcheck; adding tracing of introduction rules in code_prolog
 bulwahn parents: 
36902diff
changeset | 86 | "pos_not_seq xq = (%i. Lazy_Sequence.hb_not_seq (xq (3 * i)))" | 
| 36017 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 87 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 88 | definition neg_not_seq :: "unit pos_dseq => unit neg_dseq" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 89 | where | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 90 | "neg_not_seq x = (%i. case Lazy_Sequence.yield (x i) of | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 91 | None => Lazy_Sequence.hb_single () | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 92 | | Some ((), xq) => Lazy_Sequence.empty)" | 
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 93 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36049diff
changeset | 94 | hide_type (open) pos_dseq neg_dseq | 
| 36017 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 95 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36049diff
changeset | 96 | hide_const (open) | 
| 36049 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 bulwahn parents: 
36017diff
changeset | 97 | pos_empty pos_single pos_bind pos_union pos_if_seq pos_iterate_upto pos_not_seq pos_map | 
| 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 bulwahn parents: 
36017diff
changeset | 98 | neg_empty neg_single neg_bind neg_union neg_if_seq neg_iterate_upto neg_not_seq 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: 
36049diff
changeset | 99 | hide_fact (open) | 
| 36049 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 bulwahn parents: 
36017diff
changeset | 100 | pos_empty_def pos_single_def pos_bind_def pos_union_def pos_if_seq_def pos_iterate_upto_def pos_not_seq_def pos_map_def | 
| 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 bulwahn parents: 
36017diff
changeset | 101 | neg_empty_def neg_single_def neg_bind_def neg_union_def neg_if_seq_def neg_iterate_upto_def neg_not_seq_def neg_map_def | 
| 36017 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 102 | |
| 
7516568bebeb
adding new depth-limited and new random monad for the predicate compiler
 bulwahn parents: diff
changeset | 103 | end |