src/HOL/Datatype_Examples/Misc_Datatype.thy
author blanchet
Fri, 02 Oct 2015 21:06:32 +0200
changeset 61310 9a50ea544fd3
parent 58966 0297e1277ed2
child 62536 656e9653c645
permissions -rw-r--r--
better compliance with TPTP SZS standard
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58309
a09ec6daaa19 renamed 'BNF_Examples' to 'Datatype_Examples' (cf. 'datatypes.pdf')
blanchet
parents: 58232
diff changeset
     1
(*  Title:      HOL/Datatype_Examples/Misc_Datatype.thy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Author:     Andrei Popescu, TU Muenchen
51744
0468af6546ff more examples
blanchet
parents: 49601
diff changeset
     4
    Author:     Jasmin Blanchette, TU Muenchen
0468af6546ff more examples
blanchet
parents: 49601
diff changeset
     5
    Copyright   2012, 2013
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
53122
bc87b7af4767 renamed theory files to be closer to (new) command names
blanchet
parents: 52844
diff changeset
     7
Miscellaneous datatype definitions.
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58455
diff changeset
    10
section {* Miscellaneous Datatype Definitions *}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
53122
bc87b7af4767 renamed theory files to be closer to (new) command names
blanchet
parents: 52844
diff changeset
    12
theory Misc_Datatype
58163
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
    13
imports "~~/src/HOL/Library/Countable" "~~/src/HOL/Library/FSet"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    16
datatype (discs_sels) simple = X1 | X2 | X3 | X4
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    17
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    18
datatype (discs_sels) simple' = X1' unit | X2' unit | X3' unit | X4' unit
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    20
datatype (discs_sels) simple'' = X1'' nat int | X2''
52323
a11bbb5fef56 fixed failure in coinduction rule tactic
blanchet
parents: 51804
diff changeset
    21
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    22
datatype (discs_sels) 'a mylist = MyNil | MyCons (myhd: 'a) (mytl: "'a mylist")
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    24
datatype (discs_sels) ('b, 'c :: ord, 'd, 'e) some_passive =
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    25
  SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
58417
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58332
diff changeset
    27
datatype (discs_sels) 'a multi_live_direct1 = MultiLiveDirect1 'a
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58332
diff changeset
    28
datatype (discs_sels) 'a multi_live_direct2 = MultiLiveDirect2 'a 'a
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58332
diff changeset
    29
datatype (discs_sels) 'a multi_live_direct3 = MultiLiveDirect3 'a 'a 'a
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58332
diff changeset
    30
datatype (discs_sels) 'a multi_live_direct4 = MultiLiveDirect4 'a 'a 'a 'a
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58332
diff changeset
    31
datatype (discs_sels) 'a multi_live_direct5 = MultiLiveDirect5 'a 'a 'a 'a 'a
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58332
diff changeset
    32
datatype (discs_sels) 'a multi_live_direct6 = MultiLiveDirect6 'a 'a 'a 'a 'a 'a
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58332
diff changeset
    33
datatype (discs_sels) 'a multi_live_direct7 = MultiLiveDirect7 'a 'a 'a 'a 'a 'a 'a
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58332
diff changeset
    34
datatype (discs_sels) 'a multi_live_direct8 = MultiLiveDirect8 'a 'a 'a 'a 'a 'a 'a 'a
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58332
diff changeset
    35
datatype (discs_sels) 'a multi_live_direct9 = MultiLiveDirect9 'a 'a 'a 'a 'a 'a 'a 'a 'a
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58332
diff changeset
    36
58966
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58889
diff changeset
    37
datatype ('a, 'b) ite = ITE "'a \<Rightarrow> bool" "'a \<Rightarrow> 'b" "'a \<Rightarrow> 'b"
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58889
diff changeset
    38
58455
126c353540fc make 'case_transfer' tactic more robust
desharna
parents: 58417
diff changeset
    39
datatype 'a live_and_fun = LiveAndFun nat "nat \<Rightarrow> 'a"
126c353540fc make 'case_transfer' tactic more robust
desharna
parents: 58417
diff changeset
    40
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    41
datatype (discs_sels) hfset = HFset "hfset fset"
49601
ba31032887db modernized examples;
blanchet
parents: 49510
diff changeset
    42
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    43
datatype (discs_sels) lambda =
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    44
  Var string |
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    45
  App lambda lambda |
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    46
  Abs string lambda |
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    47
  Let "(string \<times> lambda) fset" lambda
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    49
datatype (discs_sels) 'a par_lambda =
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    50
  PVar 'a |
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    51
  PApp "'a par_lambda" "'a par_lambda" |
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    52
  PAbs 'a "'a par_lambda" |
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    53
  PLet "('a \<times> 'a par_lambda) fset" "'a par_lambda"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    54
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
(*
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    56
  ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    57
  ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    58
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    59
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    60
datatype (discs_sels) 'a I1 = I11 'a "'a I1" | I12 'a "'a I2"
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
    61
and 'a I2 = I21 | I22 "'a I1" "'a I2"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    63
datatype (discs_sels) 'a tree = TEmpty | TNode 'a "'a forest"
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
    64
and 'a forest = FNil | FCons "'a tree" "'a forest"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    65
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    66
datatype (discs_sels) 'a tree' = TEmpty' | TNode' "'a branch" "'a branch"
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
    67
and 'a branch = Branch 'a "'a tree'"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    69
datatype (discs_sels) 'a bin_rose_tree = BRTree 'a "'a bin_rose_tree mylist" "'a bin_rose_tree mylist"
55484
9deb5066508f added examples/tests
blanchet
parents: 55129
diff changeset
    70
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    71
datatype (discs_sels) ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
    72
and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm"
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
    73
and ('a, 'b) factor = C 'a | V 'b | Paren "('a, 'b) exp"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    74
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    75
datatype (discs_sels) 'a ftree = FTLeaf 'a | FTNode "'a \<Rightarrow> 'a ftree"
53134
4f8e156d2f19 added datatype example
blanchet
parents: 53123
diff changeset
    76
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    77
datatype (discs_sels) ('a, 'b, 'c) some_killing =
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    78
  SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here"
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
    79
and ('a, 'b, 'c) in_here =
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
    80
  IH1 'b 'a | IH2 'c
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    82
datatype (discs_sels) 'b nofail1 = NF11 "'b nofail1" 'b | NF12 'b
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    83
datatype (discs_sels) 'b nofail2 = NF2 "('b nofail2 \<times> 'b \<times> 'b nofail2 \<times> 'b) list"
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    84
datatype (discs_sels) 'b nofail3 = NF3 'b "('b nofail3 \<times> 'b \<times> 'b nofail3 \<times> 'b) fset"
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    85
datatype (discs_sels) 'b nofail4 = NF4 "('b nofail4 \<times> ('b nofail4 \<times> 'b \<times> 'b nofail4 \<times> 'b) fset) list"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    87
(*
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    88
datatype (discs_sels) 'b fail = F "'b fail" 'b "'b fail" "'b list"
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    89
datatype (discs_sels) 'b fail = F "'b fail" 'b "'b fail" 'b
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    90
datatype (discs_sels) 'b fail = F1 "'b fail" 'b | F2 "'b fail"
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    91
datatype (discs_sels) 'b fail = F "'b fail" 'b
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    92
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    94
datatype (discs_sels) l1 = L1 "l2 list"
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
    95
and l2 = L21 "l1 fset" | L22 l2
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    97
datatype (discs_sels) kk1 = KK1 kk2
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
    98
and kk2 = KK2 kk3
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
    99
and kk3 = KK3 "kk1 list"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   100
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   101
datatype (discs_sels) t1 = T11 t3 | T12 t2
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   102
and t2 = T2 t1
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   103
and t3 = T3
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   105
datatype (discs_sels) t1' = T11' t2' | T12' t3'
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   106
and t2' = T2' t1'
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   107
and t3' = T3'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   108
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   109
(*
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   110
datatype (discs_sels) fail1 = F1 fail2
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   111
and fail2 = F2 fail3
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   112
and fail3 = F3 fail1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   113
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   114
datatype (discs_sels) fail1 = F1 "fail2 list" fail2
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   115
and fail2 = F2 "fail2 fset" fail3
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   116
and fail3 = F3 fail1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   118
datatype (discs_sels) fail1 = F1 "fail2 list" fail2
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   119
and fail2 = F2 "fail1 fset" fail1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   120
*)
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
   121
49162
bd6a18a1a5af commented out slow examples again
blanchet
parents: 49158
diff changeset
   122
(* SLOW
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   123
datatype (discs_sels) ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list"
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   124
and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list"
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   125
and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5"
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   126
and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list"
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   127
and ('a, 'c) D5 = A5 "('a, 'c) D6"
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   128
and ('a, 'c) D6 = A6 "('a, 'c) D7"
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   129
and ('a, 'c) D7 = A7 "('a, 'c) D8"
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   130
and ('a, 'c) D8 = A8 "('a, 'c) D1 list"
49162
bd6a18a1a5af commented out slow examples again
blanchet
parents: 49158
diff changeset
   131
*)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   132
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   133
(* fail:
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   134
datatype (discs_sels) tt1 = TT11 tt2 tt3 | TT12 tt2 tt4
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   135
and tt2 = TT2
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   136
and tt3 = TT3 tt4
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   137
and tt4 = TT4 tt1
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   139
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   140
datatype (discs_sels) k1 = K11 k2 k3 | K12 k2 k4
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   141
and k2 = K2
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   142
and k3 = K3 k4
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   143
and k4 = K4
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   144
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   145
datatype (discs_sels) tt1 = TT11 tt3 tt2 | TT12 tt2 tt4
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   146
and tt2 = TT2
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   147
and tt3 = TT3 tt1
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   148
and tt4 = TT4
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49115
diff changeset
   149
49162
bd6a18a1a5af commented out slow examples again
blanchet
parents: 49158
diff changeset
   150
(* SLOW
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   151
datatype (discs_sels) s1 = S11 s2 s3 s4 | S12 s3 | S13 s2 s6 | S14 s4 s2 | S15 s2 s2
51804
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   152
and s2 = S21 s7 s5 | S22 s5 s4 s6
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   153
and s3 = S31 s1 s7 s2 | S32 s3 s3 | S33 s4 s5
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   154
and s4 = S4 s5
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   155
and s5 = S5
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   156
and s6 = S61 s6 | S62 s1 s2 | S63 s6
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   157
and s7 = S71 s8 | S72 s5
be6e703908f4 renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents: 51744
diff changeset
   158
and s8 = S8 nat
49162
bd6a18a1a5af commented out slow examples again
blanchet
parents: 49158
diff changeset
   159
*)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   160
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   161
datatype (discs_sels) 'a deadbar = DeadBar "'a \<Rightarrow> 'a"
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   162
datatype (discs_sels) 'a deadbar_option = DeadBarOption "'a option \<Rightarrow> 'a option"
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   163
datatype (discs_sels) ('a, 'b) bar = Bar "'b \<Rightarrow> 'a"
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   164
datatype (discs_sels) ('a, 'b, 'c, 'd) foo = Foo "'d + 'b \<Rightarrow> 'c + 'a"
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   165
datatype (discs_sels) 'a deadfoo = DeadFoo "'a \<Rightarrow> 'a + 'a"
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49166
diff changeset
   166
58332
be0f5d8d511b imported patch phantoms
blanchet
parents: 58327
diff changeset
   167
datatype (discs_sels) (dead 'a) dead_foo = A
be0f5d8d511b imported patch phantoms
blanchet
parents: 58327
diff changeset
   168
datatype (discs_sels) ('a, 'b) use_dead_foo = Y 'a "'b dead_foo"
be0f5d8d511b imported patch phantoms
blanchet
parents: 58327
diff changeset
   169
be0f5d8d511b imported patch phantoms
blanchet
parents: 58327
diff changeset
   170
datatype (discs_sels) 'a phantom = A
be0f5d8d511b imported patch phantoms
blanchet
parents: 58327
diff changeset
   171
datatype (discs_sels) 'a use_phantom = Y 'a "'a use_phantom phantom"
49186
4b5fa9d5e330 handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents: 49185
diff changeset
   172
58327
a147bd03cad0 make 'ctr_transfer' tactic more robust
desharna
parents: 58310
diff changeset
   173
datatype ('t, 'id) dead_sum_fun = Dead_sum_fun "('t list \<Rightarrow> 't) + 't" | Bar (bar: 'id)
a147bd03cad0 make 'ctr_transfer' tactic more robust
desharna
parents: 58310
diff changeset
   174
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   175
datatype (discs_sels) d1 = D
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   176
datatype (discs_sels) d1' = is_D: D
51744
0468af6546ff more examples
blanchet
parents: 49601
diff changeset
   177
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   178
datatype (discs_sels) d2 = D nat
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   179
datatype (discs_sels) d2' = is_D: D nat
51744
0468af6546ff more examples
blanchet
parents: 49601
diff changeset
   180
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   181
datatype (discs_sels) d3 = D | E
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   182
datatype (discs_sels) d3' = D | is_E: E
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   183
datatype (discs_sels) d3'' = is_D: D | E
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   184
datatype (discs_sels) d3''' = is_D: D | is_E: E
51744
0468af6546ff more examples
blanchet
parents: 49601
diff changeset
   185
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   186
datatype (discs_sels) d4 = D nat | E
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   187
datatype (discs_sels) d4' = D nat | is_E: E
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   188
datatype (discs_sels) d4'' = is_D: D nat | E
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   189
datatype (discs_sels) d4''' = is_D: D nat | is_E: E
51744
0468af6546ff more examples
blanchet
parents: 49601
diff changeset
   190
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   191
datatype (discs_sels) d5 = D nat | E int
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   192
datatype (discs_sels) d5' = D nat | is_E: E int
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   193
datatype (discs_sels) d5'' = is_D: D nat | E int
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   194
datatype (discs_sels) d5''' = is_D: D nat | is_E: E int
51744
0468af6546ff more examples
blanchet
parents: 49601
diff changeset
   195
58163
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   196
instance simple :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   197
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   198
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   199
instance simple'' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   200
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   201
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   202
instance mylist :: (countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   203
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   204
58232
7b70a2b4ec9b made new countable tactic work with sorts other than 'type'
blanchet
parents: 58228
diff changeset
   205
instance some_passive :: (countable, "{countable,ord}", countable, countable) countable
58163
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   206
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   207
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   208
(* TODO: Enable once "fset" is registered as countable:
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   209
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   210
instance hfset :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   211
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   212
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   213
instance lambda :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   214
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   215
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   216
instance par_lambda :: (countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   217
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   218
*)
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   219
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   220
instance I1 and I2 :: (countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   221
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   222
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   223
instance tree and forest :: (countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   224
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   225
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   226
instance tree' and branch :: (countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   227
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   228
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   229
instance bin_rose_tree :: (countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   230
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   231
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   232
instance exp and trm and factor :: (countable, countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   233
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   234
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   235
instance nofail1 :: (countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   236
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   237
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   238
instance nofail2 :: (countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   239
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   240
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   241
(* TODO: Enable once "fset" is registered as countable:
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   242
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   243
instance nofail3 :: (countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   244
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   245
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   246
instance nofail4 :: (countable) countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   247
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   248
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   249
instance l1 and l2 :: countable
58173
7a259137a0ba reenabled yet another example
blanchet
parents: 58171
diff changeset
   250
  by countable_datatype
58163
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   251
*)
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   252
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   253
instance kk1 and kk2 :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   254
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   255
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   256
instance t1 and t2 and t3 :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   257
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   258
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   259
instance t1' and t2' and t3' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   260
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   261
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   262
instance k1 and k2 and k3 and k4 :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   263
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   264
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   265
instance tt1 and tt2 and tt3 and tt4 :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   266
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   267
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   268
instance d1 :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   269
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   270
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   271
instance d1' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   272
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   273
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   274
instance d2 :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   275
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   276
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   277
instance d2' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   278
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   279
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   280
instance d3 :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   281
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   282
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   283
instance d3' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   284
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   285
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   286
instance d3'' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   287
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   288
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   289
instance d3''' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   290
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   291
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   292
instance d4 :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   293
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   294
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   295
instance d4' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   296
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   297
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   298
instance d4'' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   299
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   300
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   301
instance d4''' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   302
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   303
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   304
instance d5 :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   305
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   306
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   307
instance d5' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   308
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   309
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   310
instance d5'' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   311
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   312
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   313
instance d5''' :: countable
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   314
  by countable_datatype
c1e32fe387f4 added tests for new 'countable_datatype' proof method
blanchet
parents: 58139
diff changeset
   315
57795
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   316
datatype_compat simple
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   317
datatype_compat simple'
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   318
datatype_compat simple''
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   319
datatype_compat mylist
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   320
datatype_compat some_passive
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   321
datatype_compat I1 I2
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   322
datatype_compat tree forest
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   323
datatype_compat tree' branch
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   324
datatype_compat bin_rose_tree
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   325
datatype_compat exp trm factor
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   326
datatype_compat ftree
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   327
datatype_compat nofail1
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   328
datatype_compat kk1 kk2 kk3
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   329
datatype_compat t1 t2 t3
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   330
datatype_compat t1' t2' t3'
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   331
datatype_compat k1 k2 k3 k4
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   332
datatype_compat tt1 tt2 tt3 tt4
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   333
datatype_compat deadbar
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   334
datatype_compat deadbar_option
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   335
datatype_compat bar
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   336
datatype_compat foo
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   337
datatype_compat deadfoo
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   338
datatype_compat dead_foo
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   339
datatype_compat use_dead_foo
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   340
datatype_compat d1
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   341
datatype_compat d1'
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   342
datatype_compat d2
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   343
datatype_compat d2'
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   344
datatype_compat d3
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   345
datatype_compat d3'
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   346
datatype_compat d3''
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   347
datatype_compat d3'''
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   348
datatype_compat d4
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   349
datatype_compat d4'
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   350
datatype_compat d4''
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   351
datatype_compat d4'''
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   352
datatype_compat d5
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   353
datatype_compat d5'
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   354
datatype_compat d5''
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   355
datatype_compat d5'''
385d49c83943 added 'datatype_compat' tests
blanchet
parents: 55484
diff changeset
   356
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   357
end