src/HOL/Typedef.thy
author nipkow
Thu, 11 Dec 2008 08:56:02 +0100
changeset 29108 12ca66b887a0
parent 28965 1de908189869
child 29056 dc08e3990c77
permissions -rw-r--r--
codegen
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Typedef.thy
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Munich
11743
wenzelm
parents: 11659
diff changeset
     4
*)
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
     5
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11770
diff changeset
     6
header {* HOL type definitions *}
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 13421
diff changeset
     8
theory Typedef
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
     9
imports Set
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 19459
diff changeset
    10
uses
9ffea7a8b31c added typecopy_package
haftmann
parents: 19459
diff changeset
    11
  ("Tools/typedef_package.ML")
9ffea7a8b31c added typecopy_package
haftmann
parents: 19459
diff changeset
    12
  ("Tools/typecopy_package.ML")
9ffea7a8b31c added typecopy_package
haftmann
parents: 19459
diff changeset
    13
  ("Tools/typedef_codegen.ML")
15131
c69542757a4d New theory header syntax.
nipkow
parents: 13421
diff changeset
    14
begin
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    15
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    16
ML {*
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    17
structure HOL = struct val thy = theory "HOL" end;
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    18
*}  -- "belongs to theory HOL"
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    19
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    20
locale type_definition =
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    21
  fixes Rep and Abs and A
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    22
  assumes Rep: "Rep x \<in> A"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    23
    and Rep_inverse: "Abs (Rep x) = x"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    24
    and Abs_inverse: "y \<in> A ==> Rep (Abs y) = y"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    25
  -- {* This will be axiomatized for each typedef! *}
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    26
begin
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    27
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    28
lemma Rep_inject:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    29
  "(Rep x = Rep y) = (x = y)"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    30
proof
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    31
  assume "Rep x = Rep y"
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    32
  then have "Abs (Rep x) = Abs (Rep y)" by (simp only:)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    33
  moreover have "Abs (Rep x) = x" by (rule Rep_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    34
  moreover have "Abs (Rep y) = y" by (rule Rep_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    35
  ultimately show "x = y" by simp
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    36
next
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    37
  assume "x = y"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    38
  thus "Rep x = Rep y" by (simp only:)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    39
qed
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    40
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    41
lemma Abs_inject:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    42
  assumes x: "x \<in> A" and y: "y \<in> A"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    43
  shows "(Abs x = Abs y) = (x = y)"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    44
proof
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    45
  assume "Abs x = Abs y"
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    46
  then have "Rep (Abs x) = Rep (Abs y)" by (simp only:)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    47
  moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    48
  moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    49
  ultimately show "x = y" by simp
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    50
next
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    51
  assume "x = y"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    52
  thus "Abs x = Abs y" by (simp only:)
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    53
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    54
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    55
lemma Rep_cases [cases set]:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    56
  assumes y: "y \<in> A"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    57
    and hyp: "!!x. y = Rep x ==> P"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    58
  shows P
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    59
proof (rule hyp)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    60
  from y have "Rep (Abs y) = y" by (rule Abs_inverse)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    61
  thus "y = Rep (Abs y)" ..
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    62
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    63
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    64
lemma Abs_cases [cases type]:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    65
  assumes r: "!!y. x = Abs y ==> y \<in> A ==> P"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    66
  shows P
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    67
proof (rule r)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    68
  have "Abs (Rep x) = x" by (rule Rep_inverse)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    69
  thus "x = Abs (Rep x)" ..
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    70
  show "Rep x \<in> A" by (rule Rep)
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    71
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    72
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    73
lemma Rep_induct [induct set]:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    74
  assumes y: "y \<in> A"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    75
    and hyp: "!!x. P (Rep x)"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    76
  shows "P y"
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    77
proof -
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    78
  have "P (Rep (Abs y))" by (rule hyp)
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    79
  moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    80
  ultimately show "P y" by simp
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    81
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    82
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    83
lemma Abs_induct [induct type]:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    84
  assumes r: "!!y. y \<in> A ==> P (Abs y)"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    85
  shows "P x"
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    86
proof -
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    87
  have "Rep x \<in> A" by (rule Rep)
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    88
  then have "P (Abs (Rep x))" by (rule r)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    89
  moreover have "Abs (Rep x) = x" by (rule Rep_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    90
  ultimately show "P x" by simp
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    91
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    92
27295
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
    93
lemma Rep_range: "range Rep = A"
24269
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
    94
proof
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
    95
  show "range Rep <= A" using Rep by (auto simp add: image_def)
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
    96
  show "A <= range Rep"
23433
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    97
  proof
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    98
    fix x assume "x : A"
24269
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
    99
    hence "x = Rep (Abs x)" by (rule Abs_inverse [symmetric])
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
   100
    thus "x : range Rep" by (rule range_eqI)
23433
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
   101
  qed
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
   102
qed
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
   103
27295
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   104
lemma Abs_image: "Abs ` A = UNIV"
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   105
proof
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   106
  show "Abs ` A <= UNIV" by (rule subset_UNIV)
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   107
next
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   108
  show "UNIV <= Abs ` A"
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   109
  proof
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   110
    fix x
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   111
    have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric])
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   112
    moreover have "Rep x : A" by (rule Rep)
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   113
    ultimately show "x : Abs ` A" by (rule image_eqI)
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   114
  qed
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   115
qed
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   116
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
   117
end
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
   118
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   119
use "Tools/typedef_package.ML"
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 19459
diff changeset
   120
use "Tools/typecopy_package.ML"
19459
2041d472fc17 seperated typedef codegen from main code
haftmann
parents: 16417
diff changeset
   121
use "Tools/typedef_codegen.ML"
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   122
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 19459
diff changeset
   123
setup {*
25535
4975b7529a14 interpretation of typedefs
haftmann
parents: 24269
diff changeset
   124
  TypedefPackage.setup
4975b7529a14 interpretation of typedefs
haftmann
parents: 24269
diff changeset
   125
  #> TypecopyPackage.setup
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 19459
diff changeset
   126
  #> TypedefCodegen.setup
9ffea7a8b31c added typecopy_package
haftmann
parents: 19459
diff changeset
   127
*}
15260
a12e999a0113 Added setup for code generator.
berghofe
parents: 15140
diff changeset
   128
26151
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   129
text {* This class is just a workaround for classes without parameters;
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   130
  it shall disappear as soon as possible. *}
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   131
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   132
class itself = type + 
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   133
  fixes itself :: "'a itself"
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   134
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   135
setup {*
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   136
let fun add_itself tyco thy =
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   137
  let
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   138
    val vs = Name.names Name.context "'a"
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   139
      (replicate (Sign.arity_number thy tyco) @{sort type});
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   140
    val ty = Type (tyco, map TFree vs);
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   141
    val lhs = Const (@{const_name itself}, Term.itselfT ty);
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   142
    val rhs = Logic.mk_type ty;
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   143
    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   144
  in
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   145
    thy
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   146
    |> TheoryTarget.instantiation ([tyco], vs, @{sort itself})
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   147
    |> `(fn lthy => Syntax.check_term lthy eq)
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28394
diff changeset
   148
    |-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
26151
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   149
    |> snd
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   150
    |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
28394
b9c8e3a12a98 LocalTheory.exit_global;
wenzelm
parents: 28084
diff changeset
   151
    |> LocalTheory.exit_global
26151
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   152
  end
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   153
in TypedefPackage.interpretation add_itself end
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   154
*}
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   155
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   156
instantiation bool :: itself
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   157
begin
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   158
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   159
definition "itself = TYPE(bool)"
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   160
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   161
instance ..
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   162
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   163
end
26151
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   164
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   165
instantiation "fun" :: ("type", "type") itself
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   166
begin
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   167
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   168
definition "itself = TYPE('a \<Rightarrow> 'b)"
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   169
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   170
instance ..
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   171
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   172
end
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   173
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   174
hide (open) const itself
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   175
4a9b8f15ce7f class itself works around a problem with class interpretation in class finite
haftmann
parents: 25535
diff changeset
   176
end