src/Benchmarks/Datatype_Benchmark/Misc_N2M.thy
changeset 62286 705d4c4003ea
parent 61260 e6f03fae14d5
child 62688 a3cccaef566a
equal deleted inserted replaced
62285:747fc3692fca 62286:705d4c4003ea
       
     1 (*  Title:      Benchmarks/Datatype_Benchmark/Misc_N2M.thy
       
     2     Author:     Dmitriy Traytel, TU Muenchen
       
     3     Copyright   2014
       
     4 
       
     5 Miscellaneous tests for the nested-to-mutual reduction.
       
     6 *)
       
     7 
       
     8 section \<open>Miscellaneous Tests for the Nested-to-Mutual Reduction\<close>
       
     9 
       
    10 theory Misc_N2M
       
    11 imports "~~/src/HOL/Library/BNF_Axiomatization"
       
    12 begin
       
    13 
       
    14 declare [[typedef_overloaded]]
       
    15 
       
    16 
       
    17 locale misc
       
    18 begin
       
    19 
       
    20 datatype 'a li = Ni | Co 'a "'a li"
       
    21 datatype 'a tr = Tr "'a \<Rightarrow> 'a tr li"
       
    22 
       
    23 primrec (nonexhaustive)
       
    24   f_tl :: "'a \<Rightarrow> 'a tr li \<Rightarrow> nat" and
       
    25   f_t :: "'a \<Rightarrow> 'a tr \<Rightarrow> nat"
       
    26 where
       
    27   "f_tl _ Ni = 0" |
       
    28   "f_t a (Tr ts) = 1 + f_tl a (ts a)"
       
    29 
       
    30 datatype ('a, 'b) l = N | C 'a 'b "('a, 'b) l"
       
    31 datatype ('a, 'b) t = T "(('a, 'b) t, 'a) l" "(('a, 'b) t, 'b) l"
       
    32 
       
    33 primrec (nonexhaustive)
       
    34   f_tl2 :: "(('a, 'a) t, 'a) l \<Rightarrow> nat" and
       
    35   f_t2 :: "('a, 'a) t \<Rightarrow> nat"
       
    36 where
       
    37   "f_tl2 N = 0" |
       
    38   "f_t2 (T ts us) = f_tl2 ts + f_tl2 us"
       
    39 
       
    40 primrec  (nonexhaustive)
       
    41   g_tla :: "(('a, 'b) t, 'a) l \<Rightarrow> nat" and
       
    42   g_tlb :: "(('a, 'b) t, 'b) l \<Rightarrow> nat" and
       
    43   g_t :: "('a, 'b) t \<Rightarrow> nat"
       
    44 where
       
    45   "g_tla N = 0" |
       
    46   "g_tlb N = 1" |
       
    47   "g_t (T ts us) = g_tla ts + g_tlb us"
       
    48 
       
    49 
       
    50 datatype
       
    51   'a l1 = N1 | C1 'a "'a l1"
       
    52 
       
    53 datatype
       
    54   ('a, 'b) t1 = T1 'a 'b "('a, 'b) t1 l1" "(nat \<times> ('a, 'b) t1) l1" and
       
    55   ('a, 'b) t2 = T2 "('a, 'b) t1"
       
    56 
       
    57 primrec
       
    58   h1_tl1 :: "(nat, 'a) t1 l1 \<Rightarrow> nat" and
       
    59   h1_natl1 :: "(nat \<times> (nat, 'a) t1) l1 \<Rightarrow> nat" and
       
    60   h1_t1 :: "(nat, 'a) t1 \<Rightarrow> nat"
       
    61 where
       
    62   "h1_tl1 N1 = 0" |
       
    63   "h1_tl1 (C1 t ts) = h1_t1 t + h1_tl1 ts" |
       
    64   "h1_natl1 N1 = Suc 0" |
       
    65   "h1_natl1 (C1 n ts) = fst n + h1_natl1 ts" |
       
    66   "h1_t1 (T1 n _ ts _) = n + h1_tl1 ts"
       
    67 
       
    68 end
       
    69 
       
    70 
       
    71 bnf_axiomatization ('a, 'b) M0 [wits: "('a, 'b) M0"]
       
    72 bnf_axiomatization ('a, 'b) N0 [wits: "('a, 'b) N0"]
       
    73 bnf_axiomatization ('a, 'b) K0 [wits: "('a, 'b) K0"]
       
    74 bnf_axiomatization ('a, 'b) L0 [wits: "('a, 'b) L0"]
       
    75 bnf_axiomatization ('a, 'b, 'c) G0 [wits: "('a, 'b, 'c) G0"]
       
    76 
       
    77 locale (*co*)mplicated
       
    78 begin
       
    79 
       
    80 datatype 'a M = CM "('a, 'a M) M0"
       
    81 datatype 'a N = CN "('a N, 'a) N0"
       
    82 datatype ('a, 'b) K = CK "('a, ('a, 'b) L) K0"
       
    83          and ('a, 'b) L = CL "('b, ('a, 'b) K) L0"
       
    84 datatype 'a G = CG "('a, ('a G, 'a G N) K, ('a G M, 'a G) L) G0"
       
    85 
       
    86 primrec
       
    87     fG :: "'a G \<Rightarrow> 'a G"
       
    88 and fK :: "('a G, 'a G N) K \<Rightarrow> ('a G, 'a G N) K"
       
    89 and fL :: "('a G, 'a G N) L \<Rightarrow> ('a G, 'a G N) L"
       
    90 and fM :: "'a G M \<Rightarrow> 'a G M" where
       
    91   "fG (CG x) = CG (map_G0 id fK (map_L fM fG) x)"
       
    92 | "fK (CK y) = CK (map_K0 fG fL y)"
       
    93 | "fL (CL z) = CL (map_L0 (map_N fG) fK z)"
       
    94 | "fM (CM w) = CM (map_M0 fG fM w)"
       
    95 thm fG_def fK_def fL_def fM_def fG.simps fK.simps fL.simps fM.simps
       
    96 
       
    97 end
       
    98 
       
    99 locale complicated
       
   100 begin
       
   101 
       
   102 codatatype 'a M = CM "('a, 'a M) M0"
       
   103 codatatype 'a N = CN "('a N, 'a) N0"
       
   104 codatatype ('a, 'b) K = CK "('a, ('a, 'b) L) K0"
       
   105          and ('a, 'b) L = CL "('b, ('a, 'b) K) L0"
       
   106 codatatype 'a G = CG "('a, ('a G, 'a G N) K, ('a G M, 'a G) L) G0"
       
   107 
       
   108 primcorec
       
   109     fG :: "'a G \<Rightarrow> 'a G"
       
   110 and fK :: "('a G, 'a G N) K \<Rightarrow> ('a G, 'a G N) K"
       
   111 and fL :: "('a G, 'a G N) L \<Rightarrow> ('a G, 'a G N) L"
       
   112 and fM :: "'a G M \<Rightarrow> 'a G M" where
       
   113   "fG x = CG (map_G0 id fK (map_L fM fG) (un_CG x))"
       
   114 | "fK y = CK (map_K0 fG fL (un_CK y))"
       
   115 | "fL z = CL (map_L0 (map_N fG) fK (un_CL z))"
       
   116 | "fM w = CM (map_M0 fG fM (un_CM w))"
       
   117 thm fG_def fK_def fL_def fM_def fG.simps fK.simps fL.simps fM.simps
       
   118 
       
   119 end
       
   120 
       
   121 datatype ('a, 'b) F1 = F1 'a 'b
       
   122 datatype F2 = F2 "((unit, nat) F1, F2) F1" | F2'
       
   123 datatype 'a kk1 = K1 'a | K2 "'a kk2" and 'a kk2 = K3 "'a kk1"
       
   124 datatype 'a ll1 = L1 'a | L2 "'a ll2 kk2" and 'a ll2 = L3 "'a ll1"
       
   125 
       
   126 datatype_compat F1
       
   127 datatype_compat F2
       
   128 datatype_compat kk1 kk2
       
   129 datatype_compat ll1 ll2
       
   130 
       
   131 
       
   132 subsection \<open>Deep Nesting\<close>
       
   133 
       
   134 datatype 'a tree = Empty | Node 'a "'a tree list"
       
   135 datatype_compat tree
       
   136 
       
   137 datatype 'a ttree = TEmpty | TNode 'a "'a ttree list tree"
       
   138 datatype_compat ttree
       
   139 
       
   140 datatype 'a tttree = TEmpty | TNode 'a "'a tttree list ttree list tree"
       
   141 datatype_compat tttree
       
   142 (*
       
   143 datatype 'a ttttree = TEmpty | TNode 'a "'a ttttree list tttree list ttree list tree"
       
   144 datatype_compat ttttree
       
   145 *)
       
   146 
       
   147 datatype ('a,'b)complex = 
       
   148   C1 nat "'a ttree" 
       
   149 | C2 "('a,'b)complex list tree tree" 'b "('a,'b)complex" "('a,'b)complex2 ttree list"
       
   150 and ('a,'b)complex2 = D1 "('a,'b)complex ttree"
       
   151 datatype_compat complex complex2
       
   152 
       
   153 datatype 'a F = F1 'a | F2 "'a F"
       
   154 datatype 'a G = G1 'a | G2 "'a G F"
       
   155 datatype H = H1 | H2 "H G"
       
   156 
       
   157 datatype_compat F
       
   158 datatype_compat G
       
   159 datatype_compat H
       
   160 
       
   161 
       
   162 subsection \<open>Primrec cache\<close>
       
   163 
       
   164 datatype 'a l = N | C 'a "'a l"
       
   165 datatype ('a, 'b) t = T 'a 'b "('a, 'b) t l"
       
   166 
       
   167 context linorder
       
   168 begin
       
   169 
       
   170 (* slow *)
       
   171 primrec
       
   172   f1_tl :: "(nat, 'a) t l \<Rightarrow> nat" and
       
   173   f1_t :: "(nat, 'a) t \<Rightarrow> nat"
       
   174 where
       
   175   "f1_tl N = 0" |
       
   176   "f1_tl (C t ts) = f1_t t + f1_tl ts" |
       
   177   "f1_t (T n _ ts) = n + f1_tl ts"
       
   178 
       
   179 (* should be fast *)
       
   180 primrec
       
   181   f2_t :: "(nat, 'b) t \<Rightarrow> nat" and
       
   182   f2_tl :: "(nat, 'b) t l \<Rightarrow> nat"
       
   183 where
       
   184   "f2_tl N = 0" |
       
   185   "f2_tl (C t ts) = f2_t t + f2_tl ts" |
       
   186   "f2_t (T n _ ts) = n + f2_tl ts"
       
   187 
       
   188 end
       
   189 
       
   190 (* should be fast *)
       
   191 primrec
       
   192   g1_t :: "('a, int) t \<Rightarrow> nat" and
       
   193   g1_tl :: "('a, int) t l \<Rightarrow> nat"
       
   194 where
       
   195   "g1_t (T _ _ ts) = g1_tl ts" |
       
   196   "g1_tl N = 0" |
       
   197   "g1_tl (C _ ts) = g1_tl ts"
       
   198 
       
   199 (* should be fast *)
       
   200 primrec
       
   201   g2_t :: "(int, int) t \<Rightarrow> nat" and
       
   202   g2_tl :: "(int, int) t l \<Rightarrow> nat"
       
   203 where
       
   204   "g2_t (T _ _ ts) = g2_tl ts" |
       
   205   "g2_tl N = 0" |
       
   206   "g2_tl (C _ ts) = g2_tl ts"
       
   207 
       
   208 
       
   209 datatype
       
   210   'a l1 = N1 | C1 'a "'a l2" and
       
   211   'a l2 = N2 | C2 'a "'a l1"
       
   212 
       
   213 primrec
       
   214   sum_l1 :: "'a::{zero,plus} l1 \<Rightarrow> 'a" and
       
   215   sum_l2 :: "'a::{zero,plus} l2 \<Rightarrow> 'a"
       
   216 where
       
   217   "sum_l1 N1 = 0" |
       
   218   "sum_l1 (C1 n ns) = n + sum_l2 ns" |
       
   219   "sum_l2 N2 = 0" |
       
   220   "sum_l2 (C2 n ns) = n + sum_l1 ns"
       
   221 
       
   222 datatype
       
   223   ('a, 'b) t1 = T1 'a 'b "('a, 'b) t1 l1" and
       
   224   ('a, 'b) t2 = T2 "('a, 'b) t1"
       
   225 
       
   226 (* slow *)
       
   227 primrec
       
   228   h1_tl1 :: "(nat, 'a) t1 l1 \<Rightarrow> nat" and
       
   229   h1_tl2 :: "(nat, 'a) t1 l2 \<Rightarrow> nat" and
       
   230   h1_t1 :: "(nat, 'a) t1 \<Rightarrow> nat"
       
   231 where
       
   232   "h1_tl1 N1 = 0" |
       
   233   "h1_tl1 (C1 t ts) = h1_t1 t + h1_tl2 ts" |
       
   234   "h1_tl2 N2 = 0" |
       
   235   "h1_tl2 (C2 t ts) = h1_t1 t + h1_tl1 ts" |
       
   236   "h1_t1 (T1 n _ ts) = n + h1_tl1 ts"
       
   237 
       
   238 (* should be fast *)
       
   239 primrec
       
   240   h2_tl1 :: "(nat, 'a) t1 l1 \<Rightarrow> nat" and
       
   241   h2_tl2 :: "(nat, 'a) t1 l2 \<Rightarrow> nat" and
       
   242   h2_t1 :: "(nat, 'a) t1 \<Rightarrow> nat"
       
   243 where
       
   244   "h2_tl1 N1 = 0" |
       
   245   "h2_tl1 (C1 t ts) = h2_t1 t + h2_tl2 ts" |
       
   246   "h2_tl2 N2 = 0" |
       
   247   "h2_tl2 (C2 t ts) = h2_t1 t + h2_tl1 ts" |
       
   248   "h2_t1 (T1 n _ ts) = n + h2_tl1 ts"
       
   249 
       
   250 (* should be fast *)
       
   251 primrec
       
   252   h3_tl2 :: "(nat, 'a) t1 l2 \<Rightarrow> nat" and
       
   253   h3_tl1 :: "(nat, 'a) t1 l1 \<Rightarrow> nat" and
       
   254   h3_t1 :: "(nat, 'a) t1 \<Rightarrow> nat"
       
   255 where
       
   256   "h3_tl1 N1 = 0" |
       
   257   "h3_tl1 (C1 t ts) = h3_t1 t + h3_tl2 ts" |
       
   258   "h3_tl2 N2 = 0" |
       
   259   "h3_tl2 (C2 t ts) = h3_t1 t + h3_tl1 ts" |
       
   260   "h3_t1 (T1 n _ ts) = n + h3_tl1 ts"
       
   261 
       
   262 (* should be fast *)
       
   263 primrec
       
   264   i1_tl2 :: "(nat, 'a) t1 l2 \<Rightarrow> nat" and
       
   265   i1_tl1 :: "(nat, 'a) t1 l1 \<Rightarrow> nat" and
       
   266   i1_t1 :: "(nat, 'a) t1 \<Rightarrow> nat" and
       
   267   i1_t2 :: "(nat, 'a) t2 \<Rightarrow> nat"
       
   268 where
       
   269   "i1_tl1 N1 = 0" |
       
   270   "i1_tl1 (C1 t ts) = i1_t1 t + i1_tl2 ts" |
       
   271   "i1_tl2 N2 = 0" |
       
   272   "i1_tl2 (C2 t ts) = i1_t1 t + i1_tl1 ts" |
       
   273   "i1_t1 (T1 n _ ts) = n + i1_tl1 ts" |
       
   274   "i1_t2 (T2 t) = i1_t1 t"
       
   275 
       
   276 (* should be fast *)
       
   277 primrec
       
   278   j1_t2 :: "(nat, 'a) t2 \<Rightarrow> nat" and
       
   279   j1_t1 :: "(nat, 'a) t1 \<Rightarrow> nat" and
       
   280   j1_tl1 :: "(nat, 'a) t1 l1 \<Rightarrow> nat" and
       
   281   j1_tl2 :: "(nat, 'a) t1 l2 \<Rightarrow> nat"
       
   282 where
       
   283   "j1_tl1 N1 = 0" |
       
   284   "j1_tl1 (C1 t ts) = j1_t1 t + j1_tl2 ts" |
       
   285   "j1_tl2 N2 = 0" |
       
   286   "j1_tl2 (C2 t ts) = j1_t1 t + j1_tl1 ts" |
       
   287   "j1_t1 (T1 n _ ts) = n + j1_tl1 ts" |
       
   288   "j1_t2 (T2 t) = j1_t1 t"
       
   289 
       
   290 
       
   291 datatype 'a l3 = N3 | C3 'a "'a l3"
       
   292 datatype 'a l4 = N4 | C4 'a "'a l4"
       
   293 datatype ('a, 'b) t3 = T3 'a 'b "('a, 'b) t3 l3" "('a, 'b) t3 l4"
       
   294 
       
   295 (* slow *)
       
   296 primrec
       
   297   k1_tl3 :: "(nat, 'a) t3 l3 \<Rightarrow> nat" and
       
   298   k1_tl4 :: "(nat, 'a) t3 l4 \<Rightarrow> nat" and
       
   299   k1_t3 :: "(nat, 'a) t3 \<Rightarrow> nat"
       
   300 where
       
   301   "k1_tl3 N3 = 0" |
       
   302   "k1_tl3 (C3 t ts) = k1_t3 t + k1_tl3 ts" |
       
   303   "k1_tl4 N4 = 0" |
       
   304   "k1_tl4 (C4 t ts) = k1_t3 t + k1_tl4 ts" |
       
   305   "k1_t3 (T3 n _ ts us) = n + k1_tl3 ts + k1_tl4 us"
       
   306 
       
   307 (* should be fast *)
       
   308 primrec
       
   309   k2_tl4 :: "(nat, int) t3 l4 \<Rightarrow> nat" and
       
   310   k2_tl3 :: "(nat, int) t3 l3 \<Rightarrow> nat" and
       
   311   k2_t3 :: "(nat, int) t3 \<Rightarrow> nat"
       
   312 where
       
   313   "k2_tl4 N4 = 0" |
       
   314   "k2_tl4 (C4 t ts) = k2_t3 t + k2_tl4 ts" |
       
   315   "k2_tl3 N3 = 0" |
       
   316   "k2_tl3 (C3 t ts) = k2_t3 t + k2_tl3 ts" |
       
   317   "k2_t3 (T3 n _ ts us) = n + k2_tl3 ts + k2_tl4 us"
       
   318 
       
   319 
       
   320 datatype ('a, 'b) l5 = N5 | C5 'a 'b "('a, 'b) l5"
       
   321 datatype ('a, 'b) l6 = N6 | C6 'a 'b "('a, 'b) l6"
       
   322 datatype ('a, 'b, 'c) t4 = T4 'a 'b "(('a, 'b, 'c) t4, 'b) l5" "(('a, 'b, 'c) t4, 'c) l6"
       
   323 
       
   324 (* slow *)
       
   325 primrec
       
   326   l1_tl5 :: "((nat, 'a, 'b) t4, 'a) l5 \<Rightarrow> nat" and
       
   327   l1_tl6 :: "((nat, 'a, 'b) t4, 'b) l6 \<Rightarrow> nat" and
       
   328   l1_t4 :: "(nat, 'a, 'b) t4 \<Rightarrow> nat"
       
   329 where
       
   330   "l1_tl5 N5 = 0" |
       
   331   "l1_tl5 (C5 t _ ts) = l1_t4 t + l1_tl5 ts" |
       
   332   "l1_tl6 N6 = 0" |
       
   333   "l1_tl6 (C6 t _ ts) = l1_t4 t + l1_tl6 ts" |
       
   334   "l1_t4 (T4 n _ ts us) = n + l1_tl5 ts + l1_tl6 us"
       
   335 
       
   336 
       
   337 subsection \<open>Primcorec Cache\<close>
       
   338 
       
   339 codatatype 'a col = N | C 'a "'a col"
       
   340 codatatype ('a, 'b) cot = T 'a 'b "('a, 'b) cot col"
       
   341 
       
   342 context linorder
       
   343 begin
       
   344 
       
   345 (* slow *)
       
   346 primcorec
       
   347   f1_cotcol :: "nat \<Rightarrow> (nat, 'a) cot col" and
       
   348   f1_cot :: "nat \<Rightarrow> (nat, 'a) cot"
       
   349 where
       
   350   "n = 0 \<Longrightarrow> f1_cotcol n = N" |
       
   351   "_ \<Longrightarrow> f1_cotcol n = C (f1_cot n) (f1_cotcol n)" |
       
   352   "f1_cot n = T n undefined (f1_cotcol n)"
       
   353 
       
   354 (* should be fast *)
       
   355 primcorec
       
   356   f2_cotcol :: "nat \<Rightarrow> (nat, 'b) cot col" and
       
   357   f2_cot :: "nat \<Rightarrow> (nat, 'b) cot"
       
   358 where
       
   359   "n = 0 \<Longrightarrow> f2_cotcol n = N" |
       
   360   "_ \<Longrightarrow> f2_cotcol n = C (f2_cot n) (f2_cotcol n)" |
       
   361   "f2_cot n = T n undefined (f2_cotcol n)"
       
   362 
       
   363 end
       
   364 
       
   365 (* should be fast *)
       
   366 primcorec
       
   367   g1_cot :: "int \<Rightarrow> (int, 'a) cot" and
       
   368   g1_cotcol :: "int \<Rightarrow> (int, 'a) cot col"
       
   369 where
       
   370   "g1_cot n = T n undefined (g1_cotcol n)" |
       
   371   "n = 0 \<Longrightarrow> g1_cotcol n = N" |
       
   372   "_ \<Longrightarrow> g1_cotcol n = C (g1_cot n) (g1_cotcol n)"
       
   373 
       
   374 (* should be fast *)
       
   375 primcorec
       
   376   g2_cot :: "int \<Rightarrow> (int, int) cot" and
       
   377   g2_cotcol :: "int \<Rightarrow> (int, int) cot col"
       
   378 where
       
   379   "g2_cot n = T n undefined (g2_cotcol n)" |
       
   380   "n = 0 \<Longrightarrow> g2_cotcol n = N" |
       
   381   "_ \<Longrightarrow> g2_cotcol n = C (g2_cot n) (g2_cotcol n)"
       
   382 
       
   383 
       
   384 codatatype
       
   385   'a col1 = N1 | C1 'a "'a col2" and
       
   386   'a col2 = N2 | C2 'a "'a col1"
       
   387 
       
   388 codatatype
       
   389   ('a, 'b) cot1 = T1 'a 'b "('a, 'b) cot1 col1" and
       
   390   ('a, 'b) cot2 = T2 "('a, 'b) cot1"
       
   391 
       
   392 (* slow *)
       
   393 primcorec
       
   394   h1_cotcol1 :: "nat \<Rightarrow> (nat, 'a) cot1 col1" and
       
   395   h1_cotcol2 :: "nat \<Rightarrow> (nat, 'a) cot1 col2" and
       
   396   h1_cot1 :: "nat \<Rightarrow> (nat, 'a) cot1"
       
   397 where
       
   398   "h1_cotcol1 n = C1 (h1_cot1 n) (h1_cotcol2 n)" |
       
   399   "h1_cotcol2 n = C2 (h1_cot1 n) (h1_cotcol1 n)" |
       
   400   "h1_cot1 n = T1 n undefined (h1_cotcol1 n)"
       
   401 
       
   402 (* should be fast *)
       
   403 primcorec
       
   404   h2_cotcol1 :: "nat \<Rightarrow> (nat, 'a) cot1 col1" and
       
   405   h2_cotcol2 :: "nat \<Rightarrow> (nat, 'a) cot1 col2" and
       
   406   h2_cot1 :: "nat \<Rightarrow> (nat, 'a) cot1"
       
   407 where
       
   408   "h2_cotcol1 n = C1 (h2_cot1 n) (h2_cotcol2 n)" |
       
   409   "h2_cotcol2 n = C2 (h2_cot1 n) (h2_cotcol1 n)" |
       
   410   "h2_cot1 n = T1 n undefined (h2_cotcol1 n)"
       
   411 
       
   412 (* should be fast *)
       
   413 primcorec
       
   414   h3_cotcol2 :: "nat \<Rightarrow> (nat, 'a) cot1 col2" and
       
   415   h3_cotcol1 :: "nat \<Rightarrow> (nat, 'a) cot1 col1" and
       
   416   h3_cot1 :: "nat \<Rightarrow> (nat, 'a) cot1"
       
   417 where
       
   418   "h3_cotcol1 n = C1 (h3_cot1 n) (h3_cotcol2 n)" |
       
   419   "h3_cotcol2 n = C2 (h3_cot1 n) (h3_cotcol1 n)" |
       
   420   "h3_cot1 n = T1 n undefined (h3_cotcol1 n)"
       
   421 
       
   422 (* should be fast *)
       
   423 primcorec
       
   424   i1_cotcol2 :: "nat \<Rightarrow> (nat, 'a) cot1 col2" and
       
   425   i1_cotcol1 :: "nat \<Rightarrow> (nat, 'a) cot1 col1" and
       
   426   i1_cot1 :: "nat \<Rightarrow> (nat, 'a) cot1" and
       
   427   i1_cot2 :: "nat \<Rightarrow> (nat, 'a) cot2"
       
   428 where
       
   429   "i1_cotcol1 n = C1 (i1_cot1 n) (i1_cotcol2 n)" |
       
   430   "i1_cotcol2 n = C2 (i1_cot1 n) (i1_cotcol1 n)" |
       
   431   "i1_cot1 n = T1 n undefined (i1_cotcol1 n)" |
       
   432   "i1_cot2 n = T2 (i1_cot1 n)"
       
   433 
       
   434 (* should be fast *)
       
   435 primcorec
       
   436   j1_cot2 :: "nat \<Rightarrow> (nat, 'a) cot2" and
       
   437   j1_cot1 :: "nat \<Rightarrow> (nat, 'a) cot1" and
       
   438   j1_cotcol1 :: "nat \<Rightarrow> (nat, 'a) cot1 col1" and
       
   439   j1_cotcol2 :: "nat \<Rightarrow> (nat, 'a) cot1 col2"
       
   440 where
       
   441   "j1_cotcol1 n = C1 (j1_cot1 n) (j1_cotcol2 n)" |
       
   442   "j1_cotcol2 n = C2 (j1_cot1 n) (j1_cotcol1 n)" |
       
   443   "j1_cot1 n = T1 n undefined (j1_cotcol1 n)" |
       
   444   "j1_cot2 n = T2 (j1_cot1 n)"
       
   445 
       
   446 
       
   447 codatatype 'a col3 = N3 | C3 'a "'a col3"
       
   448 codatatype 'a col4 = N4 | C4 'a "'a col4"
       
   449 codatatype ('a, 'b) cot3 = T3 'a 'b "('a, 'b) cot3 col3" "('a, 'b) cot3 col4"
       
   450 
       
   451 (* slow *)
       
   452 primcorec
       
   453   k1_cotcol3 :: "nat \<Rightarrow> (nat, 'a) cot3 col3" and
       
   454   k1_cotcol4 :: "nat \<Rightarrow> (nat, 'a) cot3 col4" and
       
   455   k1_cot3 :: "nat \<Rightarrow> (nat, 'a) cot3"
       
   456 where
       
   457   "k1_cotcol3 n = C3 (k1_cot3 n) (k1_cotcol3 n)" |
       
   458   "k1_cotcol4 n = C4 (k1_cot3 n) (k1_cotcol4 n)" |
       
   459   "k1_cot3 n = T3 n undefined (k1_cotcol3 n) (k1_cotcol4 n)"
       
   460 
       
   461 (* should be fast *)
       
   462 primcorec
       
   463   k2_cotcol4 :: "nat \<Rightarrow> (nat, 'a) cot3 col4" and
       
   464   k2_cotcol3 :: "nat \<Rightarrow> (nat, 'a) cot3 col3" and
       
   465   k2_cot3 :: "nat \<Rightarrow> (nat, 'a) cot3"
       
   466 where
       
   467   "k2_cotcol4 n = C4 (k2_cot3 n) (k2_cotcol4 n)" |
       
   468   "k2_cotcol3 n = C3 (k2_cot3 n) (k2_cotcol3 n)" |
       
   469   "k2_cot3 n = T3 n undefined (k2_cotcol3 n) (k2_cotcol4 n)"
       
   470 
       
   471 end