src/HOL/Typedef.thy
author wenzelm
Thu, 07 Jul 2011 23:55:15 +0200
changeset 44572 91c4d7397f0e
parent 42603 996b0c14a430
child 47821 b8c7eb0c2f89
permissions -rw-r--r--
simplified make_option/dest_option;
added make_variant/dest_variant -- usual representation of datatypes;
     1 (*  Title:      HOL/Typedef.thy
     2     Author:     Markus Wenzel, TU Munich
     3 *)
     4 
     5 header {* HOL type definitions *}
     6 
     7 theory Typedef
     8 imports Set
     9 uses ("Tools/typedef.ML")
    10 begin
    11 
    12 locale type_definition =
    13   fixes Rep and Abs and A
    14   assumes Rep: "Rep x \<in> A"
    15     and Rep_inverse: "Abs (Rep x) = x"
    16     and Abs_inverse: "y \<in> A ==> Rep (Abs y) = y"
    17   -- {* This will be axiomatized for each typedef! *}
    18 begin
    19 
    20 lemma Rep_inject:
    21   "(Rep x = Rep y) = (x = y)"
    22 proof
    23   assume "Rep x = Rep y"
    24   then have "Abs (Rep x) = Abs (Rep y)" by (simp only:)
    25   moreover have "Abs (Rep x) = x" by (rule Rep_inverse)
    26   moreover have "Abs (Rep y) = y" by (rule Rep_inverse)
    27   ultimately show "x = y" by simp
    28 next
    29   assume "x = y"
    30   thus "Rep x = Rep y" by (simp only:)
    31 qed
    32 
    33 lemma Abs_inject:
    34   assumes x: "x \<in> A" and y: "y \<in> A"
    35   shows "(Abs x = Abs y) = (x = y)"
    36 proof
    37   assume "Abs x = Abs y"
    38   then have "Rep (Abs x) = Rep (Abs y)" by (simp only:)
    39   moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse)
    40   moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse)
    41   ultimately show "x = y" by simp
    42 next
    43   assume "x = y"
    44   thus "Abs x = Abs y" by (simp only:)
    45 qed
    46 
    47 lemma Rep_cases [cases set]:
    48   assumes y: "y \<in> A"
    49     and hyp: "!!x. y = Rep x ==> P"
    50   shows P
    51 proof (rule hyp)
    52   from y have "Rep (Abs y) = y" by (rule Abs_inverse)
    53   thus "y = Rep (Abs y)" ..
    54 qed
    55 
    56 lemma Abs_cases [cases type]:
    57   assumes r: "!!y. x = Abs y ==> y \<in> A ==> P"
    58   shows P
    59 proof (rule r)
    60   have "Abs (Rep x) = x" by (rule Rep_inverse)
    61   thus "x = Abs (Rep x)" ..
    62   show "Rep x \<in> A" by (rule Rep)
    63 qed
    64 
    65 lemma Rep_induct [induct set]:
    66   assumes y: "y \<in> A"
    67     and hyp: "!!x. P (Rep x)"
    68   shows "P y"
    69 proof -
    70   have "P (Rep (Abs y))" by (rule hyp)
    71   moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse)
    72   ultimately show "P y" by simp
    73 qed
    74 
    75 lemma Abs_induct [induct type]:
    76   assumes r: "!!y. y \<in> A ==> P (Abs y)"
    77   shows "P x"
    78 proof -
    79   have "Rep x \<in> A" by (rule Rep)
    80   then have "P (Abs (Rep x))" by (rule r)
    81   moreover have "Abs (Rep x) = x" by (rule Rep_inverse)
    82   ultimately show "P x" by simp
    83 qed
    84 
    85 lemma Rep_range: "range Rep = A"
    86 proof
    87   show "range Rep <= A" using Rep by (auto simp add: image_def)
    88   show "A <= range Rep"
    89   proof
    90     fix x assume "x : A"
    91     hence "x = Rep (Abs x)" by (rule Abs_inverse [symmetric])
    92     thus "x : range Rep" by (rule range_eqI)
    93   qed
    94 qed
    95 
    96 lemma Abs_image: "Abs ` A = UNIV"
    97 proof
    98   show "Abs ` A <= UNIV" by (rule subset_UNIV)
    99 next
   100   show "UNIV <= Abs ` A"
   101   proof
   102     fix x
   103     have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric])
   104     moreover have "Rep x : A" by (rule Rep)
   105     ultimately show "x : Abs ` A" by (rule image_eqI)
   106   qed
   107 qed
   108 
   109 end
   110 
   111 use "Tools/typedef.ML" setup Typedef.setup
   112 
   113 end