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