1(*  Title:      HOL/Corec_Examples/Paper_Examples.thy
2    Author:     Andreas Lochbihler, ETH Zuerich
3    Author:     Andrei Popescu, TU Muenchen
4    Copyright   2016
5
6Small examples from the paper "Friends with Benefits".
7*)
8
9section \<open>Small Examples from the Paper ``Friends with Benefits''\<close>
10
11theory Paper_Examples
12imports "HOL-Library.BNF_Corec" "HOL-Library.FSet" Complex_Main
13begin
14
15section \<open>Examples from the introduction\<close>
16
17codatatype 'a stream = SCons (shd: 'a) (stl: "'a stream") (infixr "\<lhd>" 65)
18
19corec "natsFrom" :: "nat \<Rightarrow> nat stream" where
20  "natsFrom n = n \<lhd> natsFrom (n + 1)"
21
22corec (friend) add1 :: "nat stream \<Rightarrow> nat stream"
23where "add1 ns = (shd ns + 1) \<lhd> add1 (stl ns)"
24
25corec natsFrom' :: "nat \<Rightarrow> nat stream" where
26  "natsFrom' n = n \<lhd> add1 (natsFrom' n)"
27
28section \<open>Examples from section 3\<close>
29
30text \<open>We curry the example functions in this section because infix syntax works only for curried functions.\<close>
31
32corec (friend) Plus :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infix "\<oplus>" 67) where
33  "x\<^sub>1 \<oplus> x\<^sub>2 = (shd x\<^sub>1 + shd x\<^sub>2) \<lhd> (stl x\<^sub>1 \<oplus> stl x\<^sub>2)"
34
35section \<open>Examples from section 4\<close>
36
37codatatype 'a llist = LNil | LCons 'a "'a llist"
38
39corec collatz :: "nat \<Rightarrow> nat llist" where
40  "collatz n = (if n \<le> 1 then LNil
41     else if even n then collatz (n div 2)
42     else LCons n (collatz (3 * n + 1)))"
43
44datatype 'a nelist = NEList (hd:'a) (tl:"'a list")
45
46primrec (transfer) snoc :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a nelist" (infix "\<rhd>" 64) where
47 "[] \<rhd> a = NEList a []"
48|"(b # bs) \<rhd> a = NEList b (bs @ [a])"
49
50corec (friend) inter :: "nat stream nelist \<Rightarrow> nat stream" where
51"inter xss = shd (hd xss) \<lhd> inter (tl xss \<rhd> stl (hd xss))"
52
53corec (friend) inter' :: "nat stream nelist \<Rightarrow> nat stream" where
54"inter' xss = (case hd xss of x \<lhd> xs \<Rightarrow> x \<lhd> inter' (tl xss \<rhd> xs))"
55
56corec zero :: "nat stream" where "zero = 0 \<lhd> zero"
57
58section \<open>Examples from Blanchette et al. (ICFP 2015)\<close>
59
60corec oneTwos :: "nat stream" where "oneTwos = 1 \<lhd> 2 \<lhd> oneTwos"
61
62corec everyOther :: "'a stream \<Rightarrow> 'a stream"
63where "everyOther xs = shd xs \<lhd> everyOther (stl (stl xs))"
64
65corec fibA :: "nat stream"
66where "fibA = 0 \<lhd> (1 \<lhd> fibA \<oplus> fibA)"
67
68corec fibB :: "nat stream"
69where "fibB = (0 \<lhd> 1 \<lhd> fibB) \<oplus> (0 \<lhd> fibB)"
70
71corec (friend) times :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infix "\<otimes>" 69)
72where "xs \<otimes> ys = (shd xs * shd ys) \<lhd> xs \<otimes> stl ys \<oplus> stl xs \<otimes> ys"
73
74corec (friend) exp :: "nat stream \<Rightarrow> nat stream"
75where "exp xs = 2 ^ shd xs \<lhd> (stl xs \<otimes> exp xs)"
76
77corec facA :: "nat stream"
78where "facA = (1 \<lhd> facA) \<otimes> (1 \<lhd> facA)"
79
80corec facB :: "nat stream"
81where "facB = exp (0 \<lhd> facB)"
82
83corec (friend) sfsup :: "nat stream fset \<Rightarrow> nat stream"
84where "sfsup X = Sup (fset (fimage shd X)) \<lhd> sfsup (fimage stl X)"
85
86codatatype tree = Node (val: nat) (sub: "tree list")
87
88corec (friend) tplus :: "tree \<Rightarrow> tree \<Rightarrow> tree"
89where "tplus t u = Node (val t + val u) (map (\<lambda>(t', u'). tplus t' u') (zip (sub t) (sub u)))"
90
91corec (friend) ttimes :: "tree \<Rightarrow> tree \<Rightarrow> tree"
92where "ttimes t u = Node (val t * val u)
93  (map (\<lambda>(t, u). tplus (ttimes t u) (ttimes t u)) (zip (sub t) (sub u)))"
94
95corecursive primes :: "nat \<Rightarrow> nat \<Rightarrow> nat stream"
96where "primes m n =
97  (if (m = 0 \<and> n > 1) \<or> coprime m n then n \<lhd> primes (m * n) (n + 1) else primes m (n + 1))"
98apply (relation "measure (\<lambda>(m, n). if n = 0 then 1 else if coprime m n then 0 else m - n mod m)")
99   apply (auto simp: mod_Suc diff_less_mono2 intro: Suc_lessI elim!: not_coprimeE)
100   apply (metis dvd_1_iff_1 dvd_eq_mod_eq_0 mod_0 mod_Suc mod_Suc_eq mod_mod_cancel)
101  done
102
103corec facC :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat stream"
104where "facC n a i = (if i = 0 then a \<lhd> facC (n + 1) 1 (n + 1) else facC n (a * i) (i - 1))"
105
106corec catalan :: "nat \<Rightarrow> nat stream"
107where "catalan n = (if n > 0 then catalan (n - 1) \<oplus> (0 \<lhd> catalan (n + 1)) else 1 \<lhd> catalan 1)"
108
109corec (friend) heart :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infix "\<heartsuit>" 65)
110where "xs \<heartsuit> ys = SCons (shd xs * shd ys) ((((xs \<heartsuit> stl ys) \<oplus> (stl xs \<otimes> ys)) \<heartsuit> ys) \<otimes> ys)"
111
112corec (friend) g :: "'a stream \<Rightarrow> 'a stream"
113where "g xs = shd xs \<lhd> g (g (stl xs))"
114
115end
116