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