src/HOL/Typedef.thy
author wenzelm
Sun, 14 Oct 2001 22:08:29 +0200
changeset 11770 b6bb7a853dd2
parent 11743 b9739c85dd44
child 11979 0a3dace545c5
permissions -rw-r--r--
moved rulify to ObjectLogic;
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
11743
wenzelm
parents: 11659
diff changeset
     6
header {* Set-theory lemmas and 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
11659
a68f930bafb2 Tools/induct_attrib.ML now part of Pure;
wenzelm
parents: 11654
diff changeset
     9
files "subset.ML" "equalities.ML" "mono.ML" ("Tools/typedef_package.ML"):
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    10
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11743
diff changeset
    11
(*belongs to theory Set*)
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11743
diff changeset
    12
declare atomize_ball [symmetric, rulify]
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11743
diff changeset
    13
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    14
(* Courtesy of Stephan Merz *)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    15
lemma Least_mono: 
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    16
  "mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    17
    ==> (LEAST y. y : f ` S) = f (LEAST x. x : S)"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    18
  apply clarify
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    19
  apply (erule_tac P = "%x. x : S" in LeastI2)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    20
   apply fast
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    21
  apply (rule LeastI2)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    22
  apply (auto elim: monoD intro!: order_antisym)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    23
  done
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    24
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    25
11743
wenzelm
parents: 11659
diff changeset
    26
subsection {* HOL type definitions *}
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    27
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    28
constdefs
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    29
  type_definition :: "('a => 'b) => ('b => 'a) => 'b set => bool"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    30
  "type_definition Rep Abs A ==
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    31
    (\<forall>x. Rep x \<in> A) \<and>
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    32
    (\<forall>x. Abs (Rep x) = x) \<and>
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    33
    (\<forall>y \<in> A. Rep (Abs y) = y)"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    34
  -- {* This will be stated as an axiom for each typedef! *}
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    35
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    36
lemma type_definitionI [intro]:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    37
  "(!!x. Rep x \<in> A) ==>
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    38
    (!!x. Abs (Rep x) = x) ==>
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    39
    (!!y. y \<in> A ==> Rep (Abs y) = y) ==>
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    40
    type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    41
  by (unfold type_definition_def) blast
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    42
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    43
theorem Rep: "type_definition Rep Abs A ==> Rep x \<in> A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    44
  by (unfold type_definition_def) blast
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    45
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    46
theorem Rep_inverse: "type_definition Rep Abs A ==> Abs (Rep x) = x"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    47
  by (unfold type_definition_def) blast
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    48
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    49
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
    50
  by (unfold type_definition_def) blast
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    51
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    52
theorem Rep_inject: "type_definition Rep Abs A ==> (Rep x = Rep y) = (x = y)"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    53
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    54
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    55
  show ?thesis
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    56
  proof
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    57
    assume "Rep x = Rep y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    58
    hence "Abs (Rep x) = Abs (Rep y)" by (simp only:)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    59
    thus "x = y" by (simp only: Rep_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    60
  next
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    61
    assume "x = y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    62
    thus "Rep x = Rep y" by simp
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    63
  qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    64
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    65
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    66
theorem Abs_inject:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    67
  "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
    68
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    69
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    70
  assume x: "x \<in> A" and y: "y \<in> A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    71
  show ?thesis
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    72
  proof
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    73
    assume "Abs x = Abs y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    74
    hence "Rep (Abs x) = Rep (Abs y)" by simp
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    75
    moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    76
    moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    77
    ultimately show "x = y" by (simp only:)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    78
  next
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    79
    assume "x = y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    80
    thus "Abs x = Abs y" by simp
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    81
  qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    82
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    83
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    84
theorem Rep_cases:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    85
  "type_definition Rep Abs A ==> y \<in> A ==> (!!x. y = Rep x ==> P) ==> P"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    86
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    87
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    88
  assume y: "y \<in> A" and r: "(!!x. y = Rep x ==> P)"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    89
  show P
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    90
  proof (rule r)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    91
    from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    92
    thus "y = Rep (Abs y)" ..
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    93
  qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    94
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    95
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    96
theorem Abs_cases:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    97
  "type_definition Rep Abs A ==> (!!y. x = Abs y ==> y \<in> A ==> P) ==> P"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    98
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    99
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   100
  assume r: "!!y. x = Abs y ==> y \<in> A ==> P"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   101
  show P
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   102
  proof (rule r)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   103
    have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   104
    thus "x = Abs (Rep x)" ..
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   105
    show "Rep x \<in> A" by (rule Rep [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   106
  qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   107
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   108
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   109
theorem Rep_induct:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   110
  "type_definition Rep Abs A ==> y \<in> A ==> (!!x. P (Rep x)) ==> P y"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   111
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   112
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   113
  assume "!!x. P (Rep x)" hence "P (Rep (Abs y))" .
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   114
  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
   115
  ultimately show "P y" by (simp only:)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   116
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   117
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   118
theorem Abs_induct:
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   119
  "type_definition Rep Abs A ==> (!!y. y \<in> A ==> P (Abs y)) ==> P x"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   120
proof -
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   121
  assume tydef: "type_definition Rep Abs A"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   122
  assume r: "!!y. y \<in> A ==> P (Abs y)"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   123
  have "Rep x \<in> A" by (rule Rep [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   124
  hence "P (Abs (Rep x))" by (rule r)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   125
  moreover have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef])
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   126
  ultimately show "P x" by (simp only:)
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   127
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   128
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   129
use "Tools/typedef_package.ML"
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   130
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   131
end