| author | wenzelm | 
| Sat, 01 Jun 2024 16:26:35 +0200 | |
| changeset 80234 | cce5670be9f9 | 
| parent 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 62696 | 1 | (* Title: HOL/Corec_Examples/Tests/Merge_Poly.thy | 
| 2 | Author: Aymeric Bouzy, Ecole polytechnique | |
| 3 | Author: Jasmin Blanchette, Inria, LORIA, MPII | |
| 4 | Copyright 2015, 2016 | |
| 5 | ||
| 6 | Tests polymorphic merges. | |
| 7 | *) | |
| 8 | ||
| 62726 | 9 | section \<open>Tests Polymorphic Merges\<close> | 
| 62696 | 10 | |
| 11 | theory Merge_Poly | |
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
62726diff
changeset | 12 | imports "HOL-Library.BNF_Corec" | 
| 62696 | 13 | begin | 
| 14 | ||
| 62726 | 15 | subsection \<open>A Monomorphic Example\<close> | 
| 62696 | 16 | |
| 17 | codatatype r = R (rhd: nat) (rtl: r) | |
| 18 | ||
| 19 | corec id_r :: "r \<Rightarrow> r" where | |
| 20 | "id_r r = R (rhd r) (id_r (rtl r))" | |
| 21 | ||
| 22 | corec id_r' :: "r \<Rightarrow> r" where | |
| 23 | "id_r' r = R (rhd r) (id_r' (rtl r))" | |
| 24 | ||
| 25 | corec id_r'' :: "r \<Rightarrow> r" where | |
| 26 | "id_r'' r = R (rhd r) (id_r'' (rtl r))" | |
| 27 | ||
| 28 | consts | |
| 29 | f1 :: "r \<Rightarrow> r" | |
| 30 | f2 :: "r \<Rightarrow> r" | |
| 31 | f3 :: "r \<Rightarrow> r" | |
| 32 | ||
| 33 | friend_of_corec f1 :: "r \<Rightarrow> r" where | |
| 34 | "f1 r = R (rhd r) (f1 (R (rhd r) (rtl r)))" | |
| 35 | sorry | |
| 36 | ||
| 37 | friend_of_corec f2 :: "r \<Rightarrow> r" where | |
| 38 | "f2 r = R (rhd r) (f1 (f2 (rtl r)))" | |
| 39 | sorry | |
| 40 | ||
| 41 | friend_of_corec f3 :: "r \<Rightarrow> r" where | |
| 42 | "f3 r = R (rhd r) (f3 (rtl r))" | |
| 43 | sorry | |
| 44 | ||
| 45 | corec id_rx :: "r \<Rightarrow> r" where | |
| 46 | "id_rx r = f1 (f2 (f3 (R (rhd r) (id_rx (rtl r)))))" | |
| 47 | ||
| 48 | ||
| 62726 | 49 | subsection \<open>The Polymorphic Version\<close> | 
| 62696 | 50 | |
| 51 | codatatype 'a s = S (shd: 'a) (stl: "'a s") | |
| 52 | ||
| 53 | corec id_s :: "'a s \<Rightarrow> 'a s" where | |
| 54 | "id_s s = S (shd s) (id_s (stl s))" | |
| 55 | ||
| 56 | corec id_s' :: "'b s \<Rightarrow> 'b s" where | |
| 57 | "id_s' s = S (shd s) (id_s' (stl s))" | |
| 58 | ||
| 59 | corec id_s'' :: "'w s \<Rightarrow> 'w s" where | |
| 60 | "id_s'' s = S (shd s) (id_s'' (stl s))" | |
| 61 | ||
| 62 | (* Registers polymorphic and nonpolymorphic friends in an alternating fashion. *) | |
| 63 | ||
| 64 | consts | |
| 65 | g1 :: "'a \<Rightarrow> 'a s \<Rightarrow> 'a s" | |
| 66 | g2 :: "nat \<Rightarrow> nat s \<Rightarrow> nat s" | |
| 67 | g3 :: "'a \<Rightarrow> 'a s \<Rightarrow> 'a s" | |
| 68 | g4 :: "'a \<Rightarrow> 'a s \<Rightarrow> 'a s" | |
| 69 | g5 :: "nat \<Rightarrow> nat s \<Rightarrow> nat s" | |
| 70 | ||
| 71 | friend_of_corec g3 :: "'b \<Rightarrow> 'b s \<Rightarrow> 'b s" where | |
| 72 | "g3 x s = S (shd s) (g3 x (stl s))" | |
| 73 | sorry | |
| 74 | ||
| 75 | friend_of_corec g1 :: "'w \<Rightarrow> 'w s \<Rightarrow> 'w s" where | |
| 76 | "g1 x s = S (shd s) (g1 x (stl s))" | |
| 77 | sorry | |
| 78 | ||
| 79 | friend_of_corec g2 :: "nat \<Rightarrow> nat s \<Rightarrow> nat s" where | |
| 80 | "g2 x s = S (shd s) (g1 x (stl s))" | |
| 81 | sorry | |
| 82 | ||
| 83 | friend_of_corec g4 :: "'a \<Rightarrow> 'a s \<Rightarrow> 'a s" where | |
| 84 | "g4 x s = S (shd s) (g1 x (g4 x (stl s)))" | |
| 85 | sorry | |
| 86 | ||
| 87 | friend_of_corec g5 :: "nat \<Rightarrow> nat s \<Rightarrow> nat s" where | |
| 88 | "g5 x s = S (shd s) (g2 x (g5 x (stl s)))" | |
| 89 | sorry | |
| 90 | ||
| 91 | corec id_nat_s :: "nat s \<Rightarrow> nat s" where | |
| 92 | "id_nat_s s = g1 undefined (g2 undefined (g3 undefined (S (shd s) (id_nat_s (stl s)))))" | |
| 93 | ||
| 94 | codatatype ('a, 'b) ABstream = ACons 'a (ABtl: "('a, 'b) ABstream") | BCons 'b (ABtl: "('a, 'b) ABstream")
 | |
| 95 | ||
| 96 | consts | |
| 97 |   h1 :: "('a, 'b) ABstream \<Rightarrow> ('a, 'b) ABstream"
 | |
| 98 | h2 :: "(nat, 'b) ABstream \<Rightarrow> (nat, 'b) ABstream" | |
| 99 |   h3 :: "('a, nat) ABstream \<Rightarrow> ('a, nat) ABstream"
 | |
| 100 |   h4 :: "('a :: zero, 'a) ABstream \<Rightarrow> ('a :: zero, 'a) ABstream"
 | |
| 101 | ||
| 102 | friend_of_corec h1 where | |
| 103 | "h1 x = ACons undefined undefined" sorry | |
| 104 | ||
| 62698 | 105 | friend_of_corec h2 where | 
| 106 | "h2 x = (case x of | |
| 107 | ACons a t \<Rightarrow> ACons a (h1 (h2 t)) | |
| 62696 | 108 | | BCons b t \<Rightarrow> BCons b (h1 (h2 t)))" | 
| 109 | sorry | |
| 110 | ||
| 111 | friend_of_corec h3 where | |
| 62698 | 112 | "h3 x = (case x of | 
| 113 | ACons a t \<Rightarrow> ACons a (h1 (h3 t)) | |
| 62696 | 114 | | BCons b t \<Rightarrow> BCons b (h1 (h3 t)))" | 
| 115 | sorry | |
| 116 | ||
| 117 | friend_of_corec h4 where | |
| 62698 | 118 | "h4 x = (case x of | 
| 119 | ACons a t \<Rightarrow> ACons a (h1 (h4 t)) | |
| 62696 | 120 | | BCons b t \<Rightarrow> BCons b (h1 (h4 t)))" | 
| 121 | sorry | |
| 122 | ||
| 123 | corec (friend) h5 where | |
| 124 | "h5 x = (case x of | |
| 125 | ACons a t \<Rightarrow> ACons a (h1 (h2 (h3 (h4 (h5 t))))) | |
| 126 | | BCons b t \<Rightarrow> BCons b (h1 (h2 (h3 (h4 (h5 t))))))" | |
| 127 | ||
| 128 | corec (friend) h6 :: "(nat, nat) ABstream \<Rightarrow> (nat, nat) ABstream" where | |
| 129 | "h6 x = (case x of | |
| 130 | ACons a t \<Rightarrow> ACons a (h6 (h1 (h2 (h3 (h4 (h5 (h6 t))))))) | |
| 131 | | BCons b t \<Rightarrow> BCons b (h1 (h2 (h3 (h4 (h5 (h6 t)))))))" | |
| 132 | ||
| 133 | end |