src/HOL/Typedef.thy
author wenzelm
Thu, 15 Nov 2001 18:20:13 +0100
changeset 12207 4dff931b852f
parent 12023 d982f98e0f0d
child 13412 666137b488a4
permissions -rw-r--r--
added Induct/Binary_Trees.thy, Induct/Tree_Forest (converted from former ex/TF.ML ex/TF.thy ex/Term.ML ex/Term.thy);
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
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
     8
theory Typedef = Set
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11770
diff changeset
     9
files ("Tools/typedef_package.ML"):
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    10
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    11
constdefs
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    12
  type_definition :: "('a => 'b) => ('b => 'a) => 'b set => bool"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    13
  "type_definition Rep Abs A ==
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    14
    (\<forall>x. Rep x \<in> A) \<and>
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    15
    (\<forall>x. Abs (Rep x) = x) \<and>
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    16
    (\<forall>y \<in> A. Rep (Abs y) = y)"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    17
  -- {* This will be stated as an axiom for each typedef! *}
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    18
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    19
lemma type_definitionI [intro]:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    20
  "(!!x. Rep x \<in> A) ==>
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    21
    (!!x. Abs (Rep x) = x) ==>
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    22
    (!!y. y \<in> A ==> Rep (Abs y) = y) ==>
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    23
    type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    24
  by (unfold type_definition_def) blast
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    25
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    26
theorem Rep: "type_definition Rep Abs A ==> Rep x \<in> A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    27
  by (unfold type_definition_def) blast
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    28
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    29
theorem Rep_inverse: "type_definition Rep Abs A ==> Abs (Rep x) = x"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    30
  by (unfold type_definition_def) blast
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    31
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    32
theorem Abs_inverse: "type_definition Rep Abs A ==> y \<in> A ==> Rep (Abs y) = y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    33
  by (unfold type_definition_def) blast
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    34
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    35
theorem Rep_inject: "type_definition Rep Abs A ==> (Rep x = Rep y) = (x = y)"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    36
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    37
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    38
  show ?thesis
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    39
  proof
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    40
    assume "Rep x = Rep y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    41
    hence "Abs (Rep x) = Abs (Rep y)" by (simp only:)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    42
    thus "x = y" by (simp only: Rep_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    43
  next
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    44
    assume "x = y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    45
    thus "Rep x = Rep y" by simp
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    46
  qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    47
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    48
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    49
theorem Abs_inject:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    50
  "type_definition Rep Abs A ==> x \<in> A ==> y \<in> A ==> (Abs x = Abs y) = (x = y)"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    51
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    52
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    53
  assume x: "x \<in> A" and y: "y \<in> A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    54
  show ?thesis
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    55
  proof
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    56
    assume "Abs x = Abs y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    57
    hence "Rep (Abs x) = Rep (Abs y)" by simp
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    58
    moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    59
    moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    60
    ultimately show "x = y" by (simp only:)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    61
  next
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    62
    assume "x = y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    63
    thus "Abs x = Abs y" by simp
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    64
  qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    65
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    66
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    67
theorem Rep_cases:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    68
  "type_definition Rep Abs A ==> y \<in> A ==> (!!x. y = Rep x ==> P) ==> P"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    69
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    70
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    71
  assume y: "y \<in> A" and r: "(!!x. y = Rep x ==> P)"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    72
  show P
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    73
  proof (rule r)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    74
    from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    75
    thus "y = Rep (Abs y)" ..
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    76
  qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    77
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    78
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    79
theorem Abs_cases:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    80
  "type_definition Rep Abs A ==> (!!y. x = Abs y ==> y \<in> A ==> P) ==> P"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    81
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    82
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    83
  assume r: "!!y. x = Abs y ==> y \<in> A ==> P"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    84
  show P
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    85
  proof (rule r)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    86
    have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    87
    thus "x = Abs (Rep x)" ..
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    88
    show "Rep x \<in> A" by (rule Rep [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    89
  qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    90
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    91
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    92
theorem Rep_induct:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    93
  "type_definition Rep Abs A ==> y \<in> A ==> (!!x. P (Rep x)) ==> P y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    94
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    95
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    96
  assume "!!x. P (Rep x)" hence "P (Rep (Abs y))" .
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    97
  moreover assume "y \<in> A" hence "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    98
  ultimately show "P y" by (simp only:)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    99
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   100
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   101
theorem Abs_induct:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   102
  "type_definition Rep Abs A ==> (!!y. y \<in> A ==> P (Abs y)) ==> P x"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   103
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   104
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   105
  assume r: "!!y. y \<in> A ==> P (Abs y)"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   106
  have "Rep x \<in> A" by (rule Rep [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   107
  hence "P (Abs (Rep x))" by (rule r)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   108
  moreover have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   109
  ultimately show "P x" by (simp only:)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   110
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   111
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   112
use "Tools/typedef_package.ML"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   113
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   114
end