src/HOL/Lazy_Sequence.thy
author wenzelm
Sat, 02 Aug 2014 19:29:02 +0200
changeset 57843 d8966c09025c
parent 56846 9df717fef2bb
child 58152 6fe60a9a5bad
permissions -rw-r--r--
proper priority for error over warning also for node_status (see 9c5220e05e04);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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 {* Lazy sequences *}
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 Lazy_Sequence
50055
94041d602ecb tuned import order
haftmann
parents: 45214
diff changeset
     7
imports Predicate
34948
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
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    10
subsection {* Type of lazy sequences *}
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    11
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    12
datatype 'a lazy_sequence = lazy_sequence_of_list "'a list"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    13
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    14
primrec list_of_lazy_sequence :: "'a lazy_sequence \<Rightarrow> 'a list"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    15
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    16
  "list_of_lazy_sequence (lazy_sequence_of_list xs) = xs"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    17
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    18
lemma lazy_sequence_of_list_of_lazy_sequence [simp]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    19
  "lazy_sequence_of_list (list_of_lazy_sequence xq) = xq"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    20
  by (cases xq) simp_all
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    21
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    22
lemma lazy_sequence_eqI:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    23
  "list_of_lazy_sequence xq = list_of_lazy_sequence yq \<Longrightarrow> xq = yq"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    24
  by (cases xq, cases yq) simp
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    25
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    26
lemma lazy_sequence_eq_iff:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    27
  "xq = yq \<longleftrightarrow> list_of_lazy_sequence xq = list_of_lazy_sequence yq"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    28
  by (auto intro: lazy_sequence_eqI)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    29
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    30
lemma size_lazy_sequence_eq [code]:
56846
9df717fef2bb renamed 'xxx_size' to 'size_xxx' for old datatype package
blanchet
parents: 56643
diff changeset
    31
  "size_lazy_sequence f xq = Suc (size_list f (list_of_lazy_sequence xq))"
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    32
  "size (xq :: 'a lazy_sequence) = 0"
56846
9df717fef2bb renamed 'xxx_size' to 'size_xxx' for old datatype package
blanchet
parents: 56643
diff changeset
    33
  by (cases xq, simp)+
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    34
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55413
diff changeset
    35
lemma case_lazy_sequence [simp]:
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55413
diff changeset
    36
  "case_lazy_sequence f xq = f (list_of_lazy_sequence xq)"
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    37
  by (cases xq) auto
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    38
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55413
diff changeset
    39
lemma rec_lazy_sequence [simp]:
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55413
diff changeset
    40
  "rec_lazy_sequence f xq = f (list_of_lazy_sequence xq)"
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    41
  by (cases xq) auto
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    42
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    43
definition Lazy_Sequence :: "(unit \<Rightarrow> ('a \<times> 'a lazy_sequence) option) \<Rightarrow> 'a lazy_sequence"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    44
where
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    45
  "Lazy_Sequence f = lazy_sequence_of_list (case f () of
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    46
    None \<Rightarrow> []
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    47
  | Some (x, xq) \<Rightarrow> x # list_of_lazy_sequence xq)"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    48
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    49
code_datatype Lazy_Sequence
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    50
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    51
declare list_of_lazy_sequence.simps [code del]
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55466
diff changeset
    52
declare lazy_sequence.case [code del]
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55466
diff changeset
    53
declare lazy_sequence.rec [code del]
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    54
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    55
lemma list_of_Lazy_Sequence [simp]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    56
  "list_of_lazy_sequence (Lazy_Sequence f) = (case f () of
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    57
    None \<Rightarrow> []
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    58
  | Some (x, xq) \<Rightarrow> x # list_of_lazy_sequence xq)"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    59
  by (simp add: Lazy_Sequence_def)
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    60
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    61
definition yield :: "'a lazy_sequence \<Rightarrow> ('a \<times> 'a lazy_sequence) option"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    62
where
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    63
  "yield xq = (case list_of_lazy_sequence xq of
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    64
    [] \<Rightarrow> None
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    65
  | x # xs \<Rightarrow> Some (x, lazy_sequence_of_list xs))" 
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    66
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    67
lemma yield_Seq [simp, code]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    68
  "yield (Lazy_Sequence f) = f ()"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    69
  by (cases "f ()") (simp_all add: yield_def split_def)
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    70
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 51143
diff changeset
    71
lemma case_yield_eq [simp]: "case_option g h (yield xq) =
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 51143
diff changeset
    72
  case_list g (\<lambda>x. curry h x \<circ> lazy_sequence_of_list) (list_of_lazy_sequence xq)"
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    73
  by (cases "list_of_lazy_sequence xq") (simp_all add: yield_def)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    74
56846
9df717fef2bb renamed 'xxx_size' to 'size_xxx' for old datatype package
blanchet
parents: 56643
diff changeset
    75
lemma size_lazy_sequence_code [code]:
9df717fef2bb renamed 'xxx_size' to 'size_xxx' for old datatype package
blanchet
parents: 56643
diff changeset
    76
  "size_lazy_sequence s xq = (case yield xq of
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    77
    None \<Rightarrow> 1
56846
9df717fef2bb renamed 'xxx_size' to 'size_xxx' for old datatype package
blanchet
parents: 56643
diff changeset
    78
  | Some (x, xq') \<Rightarrow> Suc (s x + size_lazy_sequence s xq'))"
9df717fef2bb renamed 'xxx_size' to 'size_xxx' for old datatype package
blanchet
parents: 56643
diff changeset
    79
  by (cases "list_of_lazy_sequence xq") (simp_all add: size_lazy_sequence_eq)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    80
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    81
lemma equal_lazy_sequence_code [code]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    82
  "HOL.equal xq yq = (case (yield xq, yield yq) of
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    83
    (None, None) \<Rightarrow> True
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    84
  | (Some (x, xq'), Some (y, yq')) \<Rightarrow> HOL.equal x y \<and> HOL.equal xq yq
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    85
  | _ \<Rightarrow> False)"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    86
  by (simp_all add: lazy_sequence_eq_iff equal_eq split: list.splits)
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 36902
diff changeset
    87
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 36902
diff changeset
    88
lemma [code nbe]:
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 36902
diff changeset
    89
  "HOL.equal (x :: 'a lazy_sequence) x \<longleftrightarrow> True"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 36902
diff changeset
    90
  by (fact equal_refl)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    91
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    92
definition empty :: "'a lazy_sequence"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    93
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    94
  "empty = lazy_sequence_of_list []"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    95
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    96
lemma list_of_lazy_sequence_empty [simp]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    97
  "list_of_lazy_sequence empty = []"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
    98
  by (simp add: empty_def)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
    99
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   100
lemma empty_code [code]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   101
  "empty = Lazy_Sequence (\<lambda>_. None)"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   102
  by (simp add: lazy_sequence_eq_iff)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   103
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   104
definition single :: "'a \<Rightarrow> 'a lazy_sequence"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   105
where
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   106
  "single x = lazy_sequence_of_list [x]"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   107
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   108
lemma list_of_lazy_sequence_single [simp]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   109
  "list_of_lazy_sequence (single x) = [x]"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   110
  by (simp add: single_def)
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   111
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   112
lemma single_code [code]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   113
  "single x = Lazy_Sequence (\<lambda>_. Some (x, empty))"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   114
  by (simp add: lazy_sequence_eq_iff)
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   115
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   116
definition append :: "'a lazy_sequence \<Rightarrow> 'a lazy_sequence \<Rightarrow> 'a lazy_sequence"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   117
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   118
  "append xq yq = lazy_sequence_of_list (list_of_lazy_sequence xq @ list_of_lazy_sequence yq)"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   119
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   120
lemma list_of_lazy_sequence_append [simp]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   121
  "list_of_lazy_sequence (append xq yq) = list_of_lazy_sequence xq @ list_of_lazy_sequence yq"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   122
  by (simp add: append_def)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   123
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   124
lemma append_code [code]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   125
  "append xq yq = Lazy_Sequence (\<lambda>_. case yield xq of
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   126
    None \<Rightarrow> yield yq
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   127
  | Some (x, xq') \<Rightarrow> Some (x, append xq' yq))"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   128
  by (simp_all add: lazy_sequence_eq_iff split: list.splits)
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   129
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   130
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a lazy_sequence \<Rightarrow> 'b lazy_sequence"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   131
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   132
  "map f xq = lazy_sequence_of_list (List.map f (list_of_lazy_sequence xq))"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   133
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   134
lemma list_of_lazy_sequence_map [simp]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   135
  "list_of_lazy_sequence (map f xq) = List.map f (list_of_lazy_sequence xq)"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   136
  by (simp add: map_def)
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   137
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   138
lemma map_code [code]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   139
  "map f xq =
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55416
diff changeset
   140
    Lazy_Sequence (\<lambda>_. map_option (\<lambda>(x, xq'). (f x, map f xq')) (yield xq))"
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   141
  by (simp_all add: lazy_sequence_eq_iff split: list.splits)
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   142
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   143
definition flat :: "'a lazy_sequence lazy_sequence \<Rightarrow> 'a lazy_sequence"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   144
where
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   145
  "flat xqq = lazy_sequence_of_list (concat (List.map list_of_lazy_sequence (list_of_lazy_sequence xqq)))"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   146
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   147
lemma list_of_lazy_sequence_flat [simp]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   148
  "list_of_lazy_sequence (flat xqq) = concat (List.map list_of_lazy_sequence (list_of_lazy_sequence xqq))"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   149
  by (simp add: flat_def)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   150
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   151
lemma flat_code [code]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   152
  "flat xqq = Lazy_Sequence (\<lambda>_. case yield xqq of
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   153
    None \<Rightarrow> None
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   154
  | Some (xq, xqq') \<Rightarrow> yield (append xq (flat xqq')))"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   155
  by (simp add: lazy_sequence_eq_iff split: list.splits)
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   156
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   157
definition bind :: "'a lazy_sequence \<Rightarrow> ('a \<Rightarrow> 'b lazy_sequence) \<Rightarrow> 'b lazy_sequence"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   158
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   159
  "bind xq f = flat (map f xq)"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   160
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   161
definition if_seq :: "bool \<Rightarrow> unit lazy_sequence"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   162
where
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   163
  "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
   164
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   165
definition those :: "'a option lazy_sequence \<Rightarrow> 'a lazy_sequence option"
36049
0ce5b7a5c2fd adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
bulwahn
parents: 36030
diff changeset
   166
where
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55416
diff changeset
   167
  "those xq = map_option lazy_sequence_of_list (List.those (list_of_lazy_sequence xq))"
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   168
  
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51126
diff changeset
   169
function iterate_upto :: "(natural \<Rightarrow> 'a) \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> 'a lazy_sequence"
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   170
where
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   171
  "iterate_upto f n m =
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   172
    Lazy_Sequence (\<lambda>_. if n > m then None else Some (f n, iterate_upto f (n + 1) m))"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   173
  by pat_completeness auto
36049
0ce5b7a5c2fd adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
bulwahn
parents: 36030
diff changeset
   174
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51126
diff changeset
   175
termination by (relation "measure (\<lambda>(f, n, m). nat_of_natural (m + 1 - n))")
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51126
diff changeset
   176
  (auto simp add: less_natural_def)
36049
0ce5b7a5c2fd adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
bulwahn
parents: 36030
diff changeset
   177
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   178
definition not_seq :: "unit lazy_sequence \<Rightarrow> unit lazy_sequence"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   179
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   180
  "not_seq xq = (case yield xq of
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   181
    None \<Rightarrow> single ()
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   182
  | Some ((), xq) \<Rightarrow> empty)"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   183
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   184
  
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   185
subsection {* Code setup *}
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   186
36533
f8df589ca2a5 dropped unnecessary ML code
haftmann
parents: 36516
diff changeset
   187
code_reflect Lazy_Sequence
f8df589ca2a5 dropped unnecessary ML code
haftmann
parents: 36516
diff changeset
   188
  datatypes lazy_sequence = Lazy_Sequence
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   189
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   190
ML {*
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   191
signature LAZY_SEQUENCE =
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   192
sig
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   193
  datatype 'a lazy_sequence = Lazy_Sequence of (unit -> ('a * 'a Lazy_Sequence.lazy_sequence) option)
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   194
  val map: ('a -> 'b) -> 'a lazy_sequence -> 'b lazy_sequence
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   195
  val yield: 'a lazy_sequence -> ('a * 'a lazy_sequence) option
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   196
  val yieldn: int -> 'a lazy_sequence -> 'a list * 'a lazy_sequence
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   197
end;
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   198
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   199
structure Lazy_Sequence : LAZY_SEQUENCE =
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   200
struct
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   201
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   202
datatype lazy_sequence = datatype Lazy_Sequence.lazy_sequence;
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   203
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   204
fun map f = @{code Lazy_Sequence.map} f;
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   205
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   206
fun yield P = @{code Lazy_Sequence.yield} P;
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   207
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   208
fun yieldn k = Predicate.anamorph yield k;
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   209
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   210
end;
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   211
*}
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   212
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   213
40051
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   214
subsection {* Generator Sequences *}
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   215
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   216
subsubsection {* General lazy sequence operation *}
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   217
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   218
definition product :: "'a lazy_sequence \<Rightarrow> 'b lazy_sequence \<Rightarrow> ('a \<times> 'b) lazy_sequence"
40051
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   219
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   220
  "product s1 s2 = bind s1 (\<lambda>a. bind s2 (\<lambda>b. single (a, b)))"
40051
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   221
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   222
40056
0bee30e3a4ad Changed section title to please LaTeX.
hoelzl
parents: 40051
diff changeset
   223
subsubsection {* Small lazy typeclasses *}
40051
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   224
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   225
class small_lazy =
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51126
diff changeset
   226
  fixes small_lazy :: "natural \<Rightarrow> 'a lazy_sequence"
40051
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   227
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   228
instantiation unit :: small_lazy
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   229
begin
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   230
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   231
definition "small_lazy d = single ()"
40051
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   232
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   233
instance ..
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   234
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   235
end
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   236
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   237
instantiation int :: small_lazy
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   238
begin
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   239
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   240
text {* maybe optimise this expression -> append (single x) xs == cons x xs 
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   241
Performance difference? *}
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   242
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   243
function small_lazy' :: "int \<Rightarrow> int \<Rightarrow> int lazy_sequence"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   244
where
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   245
  "small_lazy' d i = (if d < i then empty
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   246
    else append (single i) (small_lazy' d (i + 1)))"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   247
    by pat_completeness auto
40051
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   248
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   249
termination 
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   250
  by (relation "measure (%(d, i). nat (d + 1 - i))") auto
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   251
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   252
definition
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51126
diff changeset
   253
  "small_lazy d = small_lazy' (int (nat_of_natural d)) (- (int (nat_of_natural d)))"
40051
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   254
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   255
instance ..
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   256
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   257
end
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   258
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   259
instantiation prod :: (small_lazy, small_lazy) small_lazy
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   260
begin
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   261
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   262
definition
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   263
  "small_lazy d = product (small_lazy d) (small_lazy d)"
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   264
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   265
instance ..
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   266
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   267
end
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   268
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   269
instantiation list :: (small_lazy) small_lazy
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   270
begin
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   271
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51126
diff changeset
   272
fun small_lazy_list :: "natural \<Rightarrow> 'a list lazy_sequence"
40051
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   273
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   274
  "small_lazy_list d = append (single [])
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   275
    (if d > 0 then bind (product (small_lazy (d - 1))
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   276
      (small_lazy (d - 1))) (\<lambda>(x, xs). single (x # xs)) else empty)"
40051
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   277
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   278
instance ..
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   279
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   280
end
b6acda4d1c29 added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents: 38857
diff changeset
   281
36902
c6bae4456741 use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents: 36533
diff changeset
   282
subsection {* With Hit Bound Value *}
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   283
text {* assuming in negative context *}
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   284
42163
392fd6c4669c renewing specifications in HOL: replacing types by type_synonym
bulwahn
parents: 40056
diff changeset
   285
type_synonym 'a hit_bound_lazy_sequence = "'a option lazy_sequence"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   286
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   287
definition hit_bound :: "'a hit_bound_lazy_sequence"
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   288
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   289
  "hit_bound = Lazy_Sequence (\<lambda>_. Some (None, empty))"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   290
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   291
lemma list_of_lazy_sequence_hit_bound [simp]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   292
  "list_of_lazy_sequence hit_bound = [None]"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   293
  by (simp add: hit_bound_def)
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   294
  
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   295
definition hb_single :: "'a \<Rightarrow> 'a hit_bound_lazy_sequence"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   296
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   297
  "hb_single x = Lazy_Sequence (\<lambda>_. Some (Some x, empty))"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   298
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   299
definition hb_map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a hit_bound_lazy_sequence \<Rightarrow> 'b hit_bound_lazy_sequence"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   300
where
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55416
diff changeset
   301
  "hb_map f xq = map (map_option f) xq"
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   302
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   303
lemma hb_map_code [code]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   304
  "hb_map f xq =
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55416
diff changeset
   305
    Lazy_Sequence (\<lambda>_. map_option (\<lambda>(x, xq'). (map_option f x, hb_map f xq')) (yield xq))"
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55416
diff changeset
   306
  using map_code [of "map_option f" xq] by (simp add: hb_map_def)
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   307
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   308
definition hb_flat :: "'a hit_bound_lazy_sequence hit_bound_lazy_sequence \<Rightarrow> 'a hit_bound_lazy_sequence"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   309
where
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   310
  "hb_flat xqq = lazy_sequence_of_list (concat
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55416
diff changeset
   311
    (List.map ((\<lambda>x. case x of None \<Rightarrow> [None] | Some xs \<Rightarrow> xs) \<circ> map_option list_of_lazy_sequence) (list_of_lazy_sequence xqq)))"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   312
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   313
lemma list_of_lazy_sequence_hb_flat [simp]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   314
  "list_of_lazy_sequence (hb_flat xqq) =
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55416
diff changeset
   315
    concat (List.map ((\<lambda>x. case x of None \<Rightarrow> [None] | Some xs \<Rightarrow> xs) \<circ> map_option list_of_lazy_sequence) (list_of_lazy_sequence xqq))"
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   316
  by (simp add: hb_flat_def)
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   317
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   318
lemma hb_flat_code [code]:
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   319
  "hb_flat xqq = Lazy_Sequence (\<lambda>_. case yield xqq of
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   320
    None \<Rightarrow> None
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   321
  | Some (xq, xqq') \<Rightarrow> yield
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   322
     (append (case xq of None \<Rightarrow> hit_bound | Some xq \<Rightarrow> xq) (hb_flat xqq')))"
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   323
  by (simp add: lazy_sequence_eq_iff split: list.splits option.splits)
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   324
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   325
definition hb_bind :: "'a hit_bound_lazy_sequence \<Rightarrow> ('a \<Rightarrow> 'b hit_bound_lazy_sequence) \<Rightarrow> 'b hit_bound_lazy_sequence"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   326
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   327
  "hb_bind xq f = hb_flat (hb_map f xq)"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   328
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   329
definition hb_if_seq :: "bool \<Rightarrow> unit hit_bound_lazy_sequence"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   330
where
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   331
  "hb_if_seq b = (if b then hb_single () else empty)"
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   332
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   333
definition hb_not_seq :: "unit hit_bound_lazy_sequence \<Rightarrow> unit lazy_sequence"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   334
where
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   335
  "hb_not_seq xq = (case yield xq of
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   336
    None \<Rightarrow> single ()
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   337
  | Some (x, xq) \<Rightarrow> empty)"
36030
1cd962a0b1a6 adding Lazy_Sequences with explicit depth-bound
bulwahn
parents: 36024
diff changeset
   338
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   339
hide_const (open) yield empty single append flat map bind
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   340
  if_seq those iterate_upto not_seq product
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   341
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   342
hide_fact (open) yield_def empty_def single_def append_def flat_def map_def bind_def
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50055
diff changeset
   343
  if_seq_def those_def not_seq_def product_def 
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   344
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
diff changeset
   345
end