src/HOL/New_DSequence.thy
changeset 36017 7516568bebeb
child 36049 0ce5b7a5c2fd
equal deleted inserted replaced
36010:a5e7574d8214 36017:7516568bebeb
       
     1 
       
     2 (* Author: Lukas Bulwahn, TU Muenchen *)
       
     3 
       
     4 header {* Depth-Limited Sequences with failure element *}
       
     5 
       
     6 theory New_DSequence
       
     7 imports Random_Sequence
       
     8 begin
       
     9 
       
    10 section {* Positive Depth-Limited Sequence *}
       
    11 
       
    12 types 'a pos_dseq = "code_numeral => 'a Lazy_Sequence.lazy_sequence"
       
    13 
       
    14 definition pos_empty :: "'a pos_dseq"
       
    15 where
       
    16   "pos_empty = (%i. Lazy_Sequence.empty)"
       
    17 
       
    18 definition pos_single :: "'a => 'a pos_dseq"
       
    19 where
       
    20   "pos_single x = (%i. Lazy_Sequence.single x)"
       
    21 
       
    22 definition pos_bind :: "'a pos_dseq => ('a => 'b pos_dseq) => 'b pos_dseq"
       
    23 where
       
    24   "pos_bind x f = (%i. 
       
    25      if i = 0 then
       
    26        Lazy_Sequence.empty
       
    27      else
       
    28        Lazy_Sequence.bind (x (i - 1)) (%a. f a i))"
       
    29 
       
    30 definition pos_union :: "'a pos_dseq => 'a pos_dseq => 'a pos_dseq"
       
    31 where
       
    32   "pos_union xq yq = (%i. Lazy_Sequence.append (xq i) (yq i))"
       
    33 
       
    34 definition pos_if_seq :: "bool => unit pos_dseq"
       
    35 where
       
    36   "pos_if_seq b = (if b then pos_single () else pos_empty)"
       
    37 
       
    38 definition pos_map :: "('a => 'b) => 'a pos_dseq => 'b pos_dseq"
       
    39 where
       
    40   "pos_map f xq = (%i. Lazy_Sequence.map f (xq i))"
       
    41 
       
    42 section {* Negative Depth-Limited Sequence *}
       
    43 
       
    44 types 'a neg_dseq = "code_numeral => 'a Lazy_Sequence.hit_bound_lazy_sequence"
       
    45 
       
    46 definition neg_empty :: "'a neg_dseq"
       
    47 where
       
    48   "neg_empty = (%i. Lazy_Sequence.empty)"
       
    49 
       
    50 definition neg_single :: "'a => 'a neg_dseq"
       
    51 where
       
    52   "neg_single x = (%i. Lazy_Sequence.hb_single x)"
       
    53 
       
    54 definition neg_bind :: "'a neg_dseq => ('a => 'b neg_dseq) => 'b neg_dseq"
       
    55 where
       
    56   "neg_bind x f = (%i. 
       
    57      if i = 0 then
       
    58        Lazy_Sequence.hit_bound
       
    59      else
       
    60        hb_bind (x (i - 1)) (%a. f a i))"
       
    61 
       
    62 definition neg_union :: "'a neg_dseq => 'a neg_dseq => 'a neg_dseq"
       
    63 where
       
    64   "neg_union x y = (%i. Lazy_Sequence.append (x i) (y i))"
       
    65 
       
    66 definition neg_if_seq :: "bool => unit neg_dseq"
       
    67 where
       
    68   "neg_if_seq b = (if b then neg_single () else neg_empty)"
       
    69 
       
    70 definition neg_map :: "('a => 'b) => 'a neg_dseq => 'b neg_dseq"
       
    71 where
       
    72   "neg_map f xq = (%i. Lazy_Sequence.hb_map f (xq i))"
       
    73 
       
    74 section {* Negation *}
       
    75 
       
    76 definition pos_not_seq :: "unit neg_dseq => unit pos_dseq"
       
    77 where
       
    78   "pos_not_seq xq = (%i. Lazy_Sequence.hb_not_seq (xq i))"
       
    79 
       
    80 definition neg_not_seq :: "unit pos_dseq => unit neg_dseq"
       
    81 where
       
    82   "neg_not_seq x = (%i. case Lazy_Sequence.yield (x i) of
       
    83     None => Lazy_Sequence.hb_single ()
       
    84   | Some ((), xq) => Lazy_Sequence.empty)"
       
    85 
       
    86 hide (open) type pos_dseq neg_dseq
       
    87 
       
    88 hide (open) const 
       
    89   pos_empty pos_single pos_bind pos_union pos_if_seq pos_not_seq
       
    90   neg_empty neg_single neg_bind neg_union neg_if_seq neg_not_seq
       
    91 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
       
    93   neg_empty_def neg_single_def neg_bind_def neg_union_def neg_if_seq_def neg_not_seq_def neg_map_def
       
    94 
       
    95 end