src/HOL/Typedef.thy
author wenzelm
Tue, 02 Sep 2008 14:10:45 +0200
changeset 28083 103d9282a946
parent 27295 cfe5244301dd
child 28084 a05ca48ef263
permissions -rw-r--r--
explicit type Name.binding for higher-specification elements;
     1 (*  Title:      HOL/Typedef.thy
     2     ID:         $Id$
     3     Author:     Markus Wenzel, TU Munich
     4 *)
     5 
     6 header {* HOL type definitions *}
     7 
     8 theory Typedef
     9 imports Set
    10 uses
    11   ("Tools/typedef_package.ML")
    12   ("Tools/typecopy_package.ML")
    13   ("Tools/typedef_codegen.ML")
    14 begin
    15 
    16 ML {*
    17 structure HOL = struct val thy = theory "HOL" end;
    18 *}  -- "belongs to theory HOL"
    19 
    20 locale type_definition =
    21   fixes Rep and Abs and A
    22   assumes Rep: "Rep x \<in> A"
    23     and Rep_inverse: "Abs (Rep x) = x"
    24     and Abs_inverse: "y \<in> A ==> Rep (Abs y) = y"
    25   -- {* This will be axiomatized for each typedef! *}
    26 begin
    27 
    28 lemma Rep_inject:
    29   "(Rep x = Rep y) = (x = y)"
    30 proof
    31   assume "Rep x = Rep y"
    32   then have "Abs (Rep x) = Abs (Rep y)" by (simp only:)
    33   moreover have "Abs (Rep x) = x" by (rule Rep_inverse)
    34   moreover have "Abs (Rep y) = y" by (rule Rep_inverse)
    35   ultimately show "x = y" by simp
    36 next
    37   assume "x = y"
    38   thus "Rep x = Rep y" by (simp only:)
    39 qed
    40 
    41 lemma Abs_inject:
    42   assumes x: "x \<in> A" and y: "y \<in> A"
    43   shows "(Abs x = Abs y) = (x = y)"
    44 proof
    45   assume "Abs x = Abs y"
    46   then have "Rep (Abs x) = Rep (Abs y)" by (simp only:)
    47   moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse)
    48   moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse)
    49   ultimately show "x = y" by simp
    50 next
    51   assume "x = y"
    52   thus "Abs x = Abs y" by (simp only:)
    53 qed
    54 
    55 lemma Rep_cases [cases set]:
    56   assumes y: "y \<in> A"
    57     and hyp: "!!x. y = Rep x ==> P"
    58   shows P
    59 proof (rule hyp)
    60   from y have "Rep (Abs y) = y" by (rule Abs_inverse)
    61   thus "y = Rep (Abs y)" ..
    62 qed
    63 
    64 lemma Abs_cases [cases type]:
    65   assumes r: "!!y. x = Abs y ==> y \<in> A ==> P"
    66   shows P
    67 proof (rule r)
    68   have "Abs (Rep x) = x" by (rule Rep_inverse)
    69   thus "x = Abs (Rep x)" ..
    70   show "Rep x \<in> A" by (rule Rep)
    71 qed
    72 
    73 lemma Rep_induct [induct set]:
    74   assumes y: "y \<in> A"
    75     and hyp: "!!x. P (Rep x)"
    76   shows "P y"
    77 proof -
    78   have "P (Rep (Abs y))" by (rule hyp)
    79   moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse)
    80   ultimately show "P y" by simp
    81 qed
    82 
    83 lemma Abs_induct [induct type]:
    84   assumes r: "!!y. y \<in> A ==> P (Abs y)"
    85   shows "P x"
    86 proof -
    87   have "Rep x \<in> A" by (rule Rep)
    88   then have "P (Abs (Rep x))" by (rule r)
    89   moreover have "Abs (Rep x) = x" by (rule Rep_inverse)
    90   ultimately show "P x" by simp
    91 qed
    92 
    93 lemma Rep_range: "range Rep = A"
    94 proof
    95   show "range Rep <= A" using Rep by (auto simp add: image_def)
    96   show "A <= range Rep"
    97   proof
    98     fix x assume "x : A"
    99     hence "x = Rep (Abs x)" by (rule Abs_inverse [symmetric])
   100     thus "x : range Rep" by (rule range_eqI)
   101   qed
   102 qed
   103 
   104 lemma Abs_image: "Abs ` A = UNIV"
   105 proof
   106   show "Abs ` A <= UNIV" by (rule subset_UNIV)
   107 next
   108   show "UNIV <= Abs ` A"
   109   proof
   110     fix x
   111     have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric])
   112     moreover have "Rep x : A" by (rule Rep)
   113     ultimately show "x : Abs ` A" by (rule image_eqI)
   114   qed
   115 qed
   116 
   117 end
   118 
   119 use "Tools/typedef_package.ML"
   120 use "Tools/typecopy_package.ML"
   121 use "Tools/typedef_codegen.ML"
   122 
   123 setup {*
   124   TypedefPackage.setup
   125   #> TypecopyPackage.setup
   126   #> TypedefCodegen.setup
   127 *}
   128 
   129 text {* This class is just a workaround for classes without parameters;
   130   it shall disappear as soon as possible. *}
   131 
   132 class itself = type + 
   133   fixes itself :: "'a itself"
   134 
   135 setup {*
   136 let fun add_itself tyco thy =
   137   let
   138     val vs = Name.names Name.context "'a"
   139       (replicate (Sign.arity_number thy tyco) @{sort type});
   140     val ty = Type (tyco, map TFree vs);
   141     val lhs = Const (@{const_name itself}, Term.itselfT ty);
   142     val rhs = Logic.mk_type ty;
   143     val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
   144   in
   145     thy
   146     |> TheoryTarget.instantiation ([tyco], vs, @{sort itself})
   147     |> `(fn lthy => Syntax.check_term lthy eq)
   148     |-> (fn eq => Specification.definition (NONE, ((Name.no_binding, []), eq)))
   149     |> snd
   150     |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
   151     |> LocalTheory.exit
   152     |> ProofContext.theory_of
   153   end
   154 in TypedefPackage.interpretation add_itself end
   155 *}
   156 
   157 instantiation bool :: itself
   158 begin
   159 
   160 definition "itself = TYPE(bool)"
   161 
   162 instance ..
   163 
   164 end
   165 
   166 instantiation "fun" :: ("type", "type") itself
   167 begin
   168 
   169 definition "itself = TYPE('a \<Rightarrow> 'b)"
   170 
   171 instance ..
   172 
   173 end
   174 
   175 hide (open) const itself
   176 
   177 end