src/HOL/New_DSequence.thy
changeset 36049 0ce5b7a5c2fd
parent 36017 7516568bebeb
child 36176 3fe7e97ccca8
equal deleted inserted replaced
36048:1d2faa488166 36049:0ce5b7a5c2fd
    33 
    33 
    34 definition pos_if_seq :: "bool => unit pos_dseq"
    34 definition pos_if_seq :: "bool => unit pos_dseq"
    35 where
    35 where
    36   "pos_if_seq b = (if b then pos_single () else pos_empty)"
    36   "pos_if_seq b = (if b then pos_single () else pos_empty)"
    37 
    37 
       
    38 definition pos_iterate_upto :: "(code_numeral => 'a) => code_numeral => code_numeral => 'a pos_dseq"
       
    39 where
       
    40   "pos_iterate_upto f n m = (%i. Lazy_Sequence.iterate_upto f n m)"
       
    41  
    38 definition pos_map :: "('a => 'b) => 'a pos_dseq => 'b pos_dseq"
    42 definition pos_map :: "('a => 'b) => 'a pos_dseq => 'b pos_dseq"
    39 where
    43 where
    40   "pos_map f xq = (%i. Lazy_Sequence.map f (xq i))"
    44   "pos_map f xq = (%i. Lazy_Sequence.map f (xq i))"
    41 
    45 
    42 section {* Negative Depth-Limited Sequence *}
    46 section {* Negative Depth-Limited Sequence *}
    65 
    69 
    66 definition neg_if_seq :: "bool => unit neg_dseq"
    70 definition neg_if_seq :: "bool => unit neg_dseq"
    67 where
    71 where
    68   "neg_if_seq b = (if b then neg_single () else neg_empty)"
    72   "neg_if_seq b = (if b then neg_single () else neg_empty)"
    69 
    73 
       
    74 definition neg_iterate_upto 
       
    75 where
       
    76   "neg_iterate_upto f n m = (%i. Lazy_Sequence.iterate_upto (%i. Some (f i)) n m)"
       
    77 
    70 definition neg_map :: "('a => 'b) => 'a neg_dseq => 'b neg_dseq"
    78 definition neg_map :: "('a => 'b) => 'a neg_dseq => 'b neg_dseq"
    71 where
    79 where
    72   "neg_map f xq = (%i. Lazy_Sequence.hb_map f (xq i))"
    80   "neg_map f xq = (%i. Lazy_Sequence.hb_map f (xq i))"
    73 
    81 
    74 section {* Negation *}
    82 section {* Negation *}
    84   | Some ((), xq) => Lazy_Sequence.empty)"
    92   | Some ((), xq) => Lazy_Sequence.empty)"
    85 
    93 
    86 hide (open) type pos_dseq neg_dseq
    94 hide (open) type pos_dseq neg_dseq
    87 
    95 
    88 hide (open) const 
    96 hide (open) const 
    89   pos_empty pos_single pos_bind pos_union pos_if_seq pos_not_seq
    97   pos_empty pos_single pos_bind pos_union pos_if_seq pos_iterate_upto pos_not_seq pos_map
    90   neg_empty neg_single neg_bind neg_union neg_if_seq neg_not_seq
    98   neg_empty neg_single neg_bind neg_union neg_if_seq neg_iterate_upto neg_not_seq neg_map
    91 hide (open) fact
    99 hide (open) fact
    92   pos_empty_def pos_single_def pos_bind_def pos_union_def pos_if_seq_def pos_not_seq_def pos_map_def
   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
    93   neg_empty_def neg_single_def neg_bind_def neg_union_def neg_if_seq_def neg_not_seq_def neg_map_def
   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
    94 
   102 
    95 end
   103 end