| 68155 |      1 | (* Author: Andreas Lochbihler, Digital Asset *)
 | 
|  |      2 | 
 | 
|  |      3 | section \<open>Laziness tests\<close>
 | 
|  |      4 | 
 | 
|  |      5 | theory Code_Lazy_Test imports
 | 
|  |      6 |   "HOL-Library.Code_Lazy"
 | 
|  |      7 |   "HOL-Library.Stream" 
 | 
|  |      8 |   "HOL-Library.Code_Test"
 | 
|  |      9 |   "HOL-Library.BNF_Corec"
 | 
|  |     10 | begin
 | 
|  |     11 | 
 | 
|  |     12 | subsection \<open>Linear codatatype\<close>
 | 
|  |     13 | 
 | 
|  |     14 | code_lazy_type stream
 | 
|  |     15 | 
 | 
|  |     16 | value [code] "cycle ''ab''"
 | 
|  |     17 | value [code] "let x = cycle ''ab''; y = snth x 10 in x"
 | 
|  |     18 | 
 | 
|  |     19 | datatype 'a llist = LNil ("\<^bold>[\<^bold>]") | LCons (lhd: 'a) (ltl: "'a llist") (infixr "\<^bold>#" 65)
 | 
|  |     20 | 
 | 
|  |     21 | subsection \<open>Finite lazy lists\<close>
 | 
|  |     22 | 
 | 
|  |     23 | code_lazy_type llist
 | 
|  |     24 | 
 | 
|  |     25 | no_notation lazy_llist ("_")
 | 
|  |     26 | syntax "_llist" :: "args => 'a list"    ("\<^bold>[(_)\<^bold>]")
 | 
|  |     27 | translations
 | 
|  |     28 |   "\<^bold>[x, xs\<^bold>]" == "x\<^bold>#\<^bold>[xs\<^bold>]"
 | 
|  |     29 |   "\<^bold>[x\<^bold>]" == "x\<^bold>#\<^bold>[\<^bold>]"
 | 
|  |     30 |   "\<^bold>[x\<^bold>]" == "CONST lazy_llist x"
 | 
|  |     31 | 
 | 
|  |     32 | definition llist :: "nat llist" where
 | 
|  |     33 |   "llist = \<^bold>[1, 2, 3, hd [], 4\<^bold>]"
 | 
|  |     34 | 
 | 
|  |     35 | fun lnth :: "nat \<Rightarrow> 'a llist \<Rightarrow> 'a" where
 | 
|  |     36 |   "lnth 0 (x \<^bold># xs) = x"
 | 
|  |     37 | | "lnth (Suc n) (x \<^bold># xs) = lnth n xs"
 | 
|  |     38 | 
 | 
|  |     39 | value [code] "llist"
 | 
|  |     40 | value [code] "let x = lnth 2 llist in (x, llist)"
 | 
|  |     41 | value [code] "llist"
 | 
|  |     42 | 
 | 
|  |     43 | fun lfilter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
 | 
|  |     44 |   "lfilter P \<^bold>[\<^bold>] = \<^bold>[\<^bold>]"
 | 
|  |     45 | | "lfilter P (x \<^bold># xs) = (if P x then x \<^bold># lfilter P xs else lfilter P xs)"
 | 
|  |     46 | 
 | 
|  |     47 | value [code] "lhd (lfilter odd llist)"
 | 
|  |     48 | 
 | 
|  |     49 | definition lfilter_test :: "nat llist \<Rightarrow> _" where "lfilter_test xs = lhd (lfilter even xs)"
 | 
| 69597 |     50 |   \<comment> \<open>Filtering \<^term>\<open>llist\<close> for \<^term>\<open>even\<close> fails because only the datatype is lazy, not the
 | 
| 68155 |     51 |   filter function itself.\<close>
 | 
|  |     52 | ML_val \<open> (@{code lfilter_test} @{code llist}; raise Fail "Failure expected") handle Match => () \<close>
 | 
|  |     53 | 
 | 
|  |     54 | subsection \<open>Records as free type\<close>
 | 
|  |     55 | 
 | 
|  |     56 | record ('a, 'b) rec = 
 | 
|  |     57 |   rec1 :: 'a
 | 
|  |     58 |   rec2 :: 'b
 | 
|  |     59 | 
 | 
|  |     60 | free_constructors rec_ext for rec.rec_ext
 | 
|  |     61 |   subgoal by(rule rec.cases_scheme)
 | 
|  |     62 |   subgoal by(rule rec.ext_inject)
 | 
|  |     63 |   done
 | 
|  |     64 | 
 | 
|  |     65 | code_lazy_type rec_ext
 | 
|  |     66 | 
 | 
|  |     67 | definition rec_test1 where "rec_test1 = rec1 (\<lparr>rec1 = Suc 5, rec2 = True\<rparr>\<lparr>rec1 := 0\<rparr>)"
 | 
|  |     68 | definition rec_test2 :: "(bool, bool) rec" where "rec_test2 = \<lparr>rec1 = hd [], rec2 = True\<rparr>"
 | 
|  |     69 | test_code "rec_test1 = 0" in PolyML Scala
 | 
|  |     70 | value [code] "rec_test2"
 | 
|  |     71 | 
 | 
|  |     72 | subsection \<open>Branching codatatypes\<close>
 | 
|  |     73 | 
 | 
|  |     74 | codatatype tree = L | Node tree tree (infix "\<triangle>" 900)
 | 
|  |     75 | 
 | 
|  |     76 | code_lazy_type tree
 | 
|  |     77 | 
 | 
|  |     78 | fun mk_tree :: "nat \<Rightarrow> tree" where
 | 
|  |     79 |   mk_tree_0: "mk_tree 0 = L"
 | 
|  |     80 | |            "mk_tree (Suc n) = (let t = mk_tree n in t \<triangle> t)"
 | 
|  |     81 | 
 | 
|  |     82 | function subtree :: "bool list \<Rightarrow> tree \<Rightarrow> tree" where
 | 
|  |     83 |   "subtree [] t = t"
 | 
|  |     84 | | "subtree (True # p) (l \<triangle> r) = subtree p l"
 | 
|  |     85 | | "subtree (False # p) (l \<triangle> r) = subtree p r"
 | 
|  |     86 | | "subtree _ L = L"
 | 
|  |     87 |   by pat_completeness auto
 | 
|  |     88 | termination by lexicographic_order
 | 
|  |     89 | 
 | 
|  |     90 | value [code] "mk_tree 10"
 | 
|  |     91 | value [code] "let t = mk_tree 10; _ = subtree [True, True, False, False] t in t"
 | 
|  |     92 | 
 | 
|  |     93 | lemma mk_tree_Suc: "mk_tree (Suc n) = mk_tree n \<triangle> mk_tree n"
 | 
|  |     94 |   by(simp add: Let_def)
 | 
|  |     95 | lemmas [code] = mk_tree_0 mk_tree_Suc
 | 
|  |     96 | value [code] "let t = mk_tree 10; _ = subtree [True, True, False, False] t in t"
 | 
|  |     97 | value [code] "let t = mk_tree 4; _ = subtree [True, True, False, False] t in t"
 | 
|  |     98 | 
 | 
|  |     99 | subsection \<open>Corecursion\<close>
 | 
|  |    100 | 
 | 
|  |    101 | corec (friend) plus :: "'a :: plus stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
 | 
|  |    102 |   "plus xs ys = (shd xs + shd ys) ## plus (stl xs) (stl ys)"
 | 
|  |    103 | 
 | 
|  |    104 | corec (friend) times :: "'a :: {plus, times} stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
 | 
|  |    105 |   "times xs ys = (shd xs * shd ys) ## plus (times (stl xs) ys) (times xs (stl ys))"
 | 
|  |    106 | 
 | 
|  |    107 | subsection \<open>Pattern-matching tests\<close>
 | 
|  |    108 | 
 | 
|  |    109 | definition f1 :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> nat llist \<Rightarrow> unit" where
 | 
|  |    110 |   "f1 _ _ _ _ = ()"
 | 
|  |    111 | 
 | 
|  |    112 | declare [[code drop: f1]]
 | 
|  |    113 | lemma f1_code1 [code]: 
 | 
|  |    114 |   "f1 b c d    ns     = Code.abort (STR ''4'') (\<lambda>_. ())" 
 | 
|  |    115 |   "f1 b c True \<^bold>[n, m\<^bold>] = Code.abort (STR ''3'') (\<lambda>_. ())" 
 | 
|  |    116 |   "f1 b True d \<^bold>[n\<^bold>]    = Code.abort (STR ''2'') (\<lambda>_. ())" 
 | 
|  |    117 |   "f1 True c d \<^bold>[\<^bold>]     = ()"
 | 
|  |    118 |   by(simp_all add: f1_def)
 | 
|  |    119 | 
 | 
|  |    120 | value [code] "f1 True False False \<^bold>[\<^bold>]"
 | 
|  |    121 | deactivate_lazy_type llist
 | 
|  |    122 | value [code] "f1 True False False \<^bold>[\<^bold>]"
 | 
|  |    123 | declare f1_code1(1) [code del]
 | 
|  |    124 | value [code] "f1 True False False \<^bold>[\<^bold>]"
 | 
|  |    125 | activate_lazy_type llist
 | 
|  |    126 | value [code] "f1 True False False \<^bold>[\<^bold>]"
 | 
|  |    127 | 
 | 
|  |    128 | declare [[code drop: f1]]
 | 
|  |    129 | lemma f1_code2 [code]: 
 | 
|  |    130 |   "f1 b c d    ns     = Code.abort (STR ''4'') (\<lambda>_. ())" 
 | 
|  |    131 |   "f1 b c True \<^bold>[n, m\<^bold>] = Code.abort (STR ''3'') (\<lambda>_. ())" 
 | 
|  |    132 |   "f1 b True d \<^bold>[n\<^bold>]    = ()"
 | 
|  |    133 |   "f1 True c d \<^bold>[\<^bold>]     = Code.abort (STR ''1'') (\<lambda>_. ())"
 | 
|  |    134 |   by(simp_all add: f1_def)
 | 
|  |    135 | 
 | 
|  |    136 | value [code] "f1 True True True \<^bold>[0\<^bold>]"
 | 
|  |    137 | declare f1_code2(1)[code del]
 | 
|  |    138 | value [code] "f1 True True True \<^bold>[0\<^bold>]"
 | 
|  |    139 | 
 | 
|  |    140 | declare [[code drop: f1]]
 | 
|  |    141 | lemma f1_code3 [code]:
 | 
|  |    142 |   "f1 b c d    ns     = Code.abort (STR ''4'') (\<lambda>_. ())"
 | 
|  |    143 |   "f1 b c True \<^bold>[n, m\<^bold>] = ()" 
 | 
|  |    144 |   "f1 b True d \<^bold>[n\<^bold>]    = Code.abort (STR ''2'') (\<lambda>_. ())"
 | 
|  |    145 |   "f1 True c d \<^bold>[\<^bold>]     = Code.abort (STR ''1'') (\<lambda>_. ())"
 | 
|  |    146 |   by(simp_all add: f1_def)
 | 
|  |    147 | 
 | 
|  |    148 | value [code] "f1 True True True \<^bold>[0, 1\<^bold>]"
 | 
|  |    149 | declare f1_code3(1)[code del]
 | 
|  |    150 | value [code] "f1 True True True \<^bold>[0, 1\<^bold>]"
 | 
|  |    151 | 
 | 
|  |    152 | declare [[code drop: f1]]
 | 
|  |    153 | lemma f1_code4 [code]:
 | 
|  |    154 |   "f1 b c d    ns     = ()" 
 | 
|  |    155 |   "f1 b c True \<^bold>[n, m\<^bold>] = Code.abort (STR ''3'') (\<lambda>_. ())"
 | 
|  |    156 |   "f1 b True d \<^bold>[n\<^bold>]    = Code.abort (STR ''2'') (\<lambda>_. ())" 
 | 
|  |    157 |   "f1 True c d \<^bold>[\<^bold>]     = Code.abort (STR ''1'') (\<lambda>_. ())"
 | 
|  |    158 |   by(simp_all add: f1_def)
 | 
|  |    159 | 
 | 
|  |    160 | value [code] "f1 True True True \<^bold>[0, 1, 2\<^bold>]"
 | 
|  |    161 | value [code] "f1 True True False \<^bold>[0, 1\<^bold>]"
 | 
|  |    162 | value [code] "f1 True False True \<^bold>[0\<^bold>]"
 | 
|  |    163 | value [code] "f1 False True True \<^bold>[\<^bold>]"
 | 
|  |    164 | 
 | 
|  |    165 | definition f2 :: "nat llist llist list \<Rightarrow> unit" where "f2 _ = ()"
 | 
|  |    166 | 
 | 
|  |    167 | declare [[code drop: f2]]
 | 
|  |    168 | lemma f2_code1 [code]:
 | 
|  |    169 |   "f2 xs = Code.abort (STR ''a'') (\<lambda>_. ())"
 | 
|  |    170 |   "f2 [\<^bold>[\<^bold>[\<^bold>]\<^bold>]] = ()"
 | 
|  |    171 |   "f2 [\<^bold>[\<^bold>[Suc n\<^bold>]\<^bold>]] = ()"
 | 
|  |    172 |   "f2 [\<^bold>[\<^bold>[0, Suc n\<^bold>]\<^bold>]] = ()"
 | 
|  |    173 |   by(simp_all add: f2_def)
 | 
|  |    174 | 
 | 
|  |    175 | value [code] "f2 [\<^bold>[\<^bold>[\<^bold>]\<^bold>]]"
 | 
|  |    176 | value [code] "f2 [\<^bold>[\<^bold>[4\<^bold>]\<^bold>]]"
 | 
|  |    177 | value [code] "f2 [\<^bold>[\<^bold>[0, 1\<^bold>]\<^bold>]]"
 | 
|  |    178 | ML_val \<open> (@{code f2} []; error "Fail expected") handle Fail _ => () \<close>
 | 
|  |    179 | 
 | 
|  |    180 | definition f3 :: "nat set llist \<Rightarrow> unit" where "f3 _ = ()"
 | 
|  |    181 | 
 | 
|  |    182 | declare [[code drop: f3]]
 | 
|  |    183 | lemma f3_code1 [code]:
 | 
|  |    184 |   "f3 \<^bold>[\<^bold>] = ()"
 | 
|  |    185 |   "f3 \<^bold>[A\<^bold>] = ()"
 | 
|  |    186 |   by(simp_all add: f3_def)
 | 
|  |    187 | 
 | 
|  |    188 | value [code] "f3 \<^bold>[\<^bold>]"
 | 
|  |    189 | value [code] "f3 \<^bold>[{}\<^bold>]"
 | 
|  |    190 | value [code] "f3 \<^bold>[UNIV\<^bold>]"
 | 
|  |    191 | 
 | 
|  |    192 | end |