src/HOL/subset.thy
author paulson
Wed, 14 Feb 2001 12:16:32 +0100
changeset 11115 285b31e9e026
parent 11083 d8fda557e476
child 11454 7514e5e21cb8
permissions -rw-r--r--
a new theorem from Bryan Ford
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 923
diff changeset
     1
(*  Title:      HOL/subset.thy
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 923
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
     5
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
     6
Subset lemmas and HOL type definitions.
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
7705
222b715b5d24 Tools/typedef_package.ML;
wenzelm
parents: 5853
diff changeset
     9
theory subset = Set
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    10
files "Tools/induct_attrib.ML" ("Tools/typedef_package.ML"):
7705
222b715b5d24 Tools/typedef_package.ML;
wenzelm
parents: 5853
diff changeset
    11
9895
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    12
(*belongs to theory Ord*)
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    13
theorems linorder_cases [case_names less equal greater] =
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    14
  linorder_less_split
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    15
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    16
(*belongs to theory Set*)
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    17
setup Rulify.setup
7717
e7ecfa617443 Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents: 7705
diff changeset
    18
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    19
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    20
section {* HOL type definitions *}
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    21
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    22
constdefs
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    23
  type_definition :: "('a => 'b) => ('b => 'a) => 'b set => bool"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    24
  "type_definition Rep Abs A ==
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    25
    (\<forall>x. Rep x \<in> A) \<and>
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    26
    (\<forall>x. Abs (Rep x) = x) \<and>
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    27
    (\<forall>y \<in> A. Rep (Abs y) = y)"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    28
  -- {* This will be stated as an axiom for each typedef! *}
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    29
10290
8018d1743beb added type_definitionI;
wenzelm
parents: 10284
diff changeset
    30
lemma type_definitionI [intro]:
10291
a88d347db404 isatool unsymbolize;
wenzelm
parents: 10290
diff changeset
    31
  "(!!x. Rep x \<in> A) ==>
a88d347db404 isatool unsymbolize;
wenzelm
parents: 10290
diff changeset
    32
    (!!x. Abs (Rep x) = x) ==>
a88d347db404 isatool unsymbolize;
wenzelm
parents: 10290
diff changeset
    33
    (!!y. y \<in> A ==> Rep (Abs y) = y) ==>
10290
8018d1743beb added type_definitionI;
wenzelm
parents: 10284
diff changeset
    34
    type_definition Rep Abs A"
8018d1743beb added type_definitionI;
wenzelm
parents: 10284
diff changeset
    35
  by (unfold type_definition_def) blast
8018d1743beb added type_definitionI;
wenzelm
parents: 10284
diff changeset
    36
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    37
theorem Rep: "type_definition Rep Abs A ==> Rep x \<in> A"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    38
  by (unfold type_definition_def) blast
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    39
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    40
theorem Rep_inverse: "type_definition Rep Abs A ==> Abs (Rep x) = x"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    41
  by (unfold type_definition_def) blast
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    42
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    43
theorem Abs_inverse: "type_definition Rep Abs A ==> y \<in> A ==> Rep (Abs y) = y"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    44
  by (unfold type_definition_def) blast
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    45
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    46
theorem Rep_inject: "type_definition Rep Abs A ==> (Rep x = Rep y) = (x = y)"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    47
proof -
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    48
  assume tydef: "type_definition Rep Abs A"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    49
  show ?thesis
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    50
  proof
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    51
    assume "Rep x = Rep y"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    52
    hence "Abs (Rep x) = Abs (Rep y)" by (simp only:)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    53
    thus "x = y" by (simp only: Rep_inverse [OF tydef])
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    54
  next
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    55
    assume "x = y"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    56
    thus "Rep x = Rep y" by simp
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    57
  qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    58
qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    59
10284
wenzelm
parents: 10276
diff changeset
    60
theorem Abs_inject:
wenzelm
parents: 10276
diff changeset
    61
  "type_definition Rep Abs A ==> x \<in> A ==> y \<in> A ==> (Abs x = Abs y) = (x = y)"
wenzelm
parents: 10276
diff changeset
    62
proof -
wenzelm
parents: 10276
diff changeset
    63
  assume tydef: "type_definition Rep Abs A"
wenzelm
parents: 10276
diff changeset
    64
  assume x: "x \<in> A" and y: "y \<in> A"
wenzelm
parents: 10276
diff changeset
    65
  show ?thesis
wenzelm
parents: 10276
diff changeset
    66
  proof
wenzelm
parents: 10276
diff changeset
    67
    assume "Abs x = Abs y"
wenzelm
parents: 10276
diff changeset
    68
    hence "Rep (Abs x) = Rep (Abs y)" by simp
11083
wenzelm
parents: 10291
diff changeset
    69
    moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse [OF tydef])
wenzelm
parents: 10291
diff changeset
    70
    moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef])
10284
wenzelm
parents: 10276
diff changeset
    71
    ultimately show "x = y" by (simp only:)
wenzelm
parents: 10276
diff changeset
    72
  next
wenzelm
parents: 10276
diff changeset
    73
    assume "x = y"
wenzelm
parents: 10276
diff changeset
    74
    thus "Abs x = Abs y" by simp
wenzelm
parents: 10276
diff changeset
    75
  qed
wenzelm
parents: 10276
diff changeset
    76
qed
wenzelm
parents: 10276
diff changeset
    77
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    78
theorem Rep_cases:
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    79
  "type_definition Rep Abs A ==> y \<in> A ==> (!!x. y = Rep x ==> P) ==> P"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    80
proof -
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    81
  assume tydef: "type_definition Rep Abs A"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    82
  assume y: "y \<in> A" and r: "(!!x. y = Rep x ==> P)"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    83
  show P
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    84
  proof (rule r)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    85
    from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef])
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    86
    thus "y = Rep (Abs y)" ..
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    87
  qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    88
qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    89
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    90
theorem Abs_cases:
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    91
  "type_definition Rep Abs A ==> (!!y. x = Abs y ==> y \<in> A ==> P) ==> P"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    92
proof -
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    93
  assume tydef: "type_definition Rep Abs A"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    94
  assume r: "!!y. x = Abs y ==> y \<in> A ==> P"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    95
  show P
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    96
  proof (rule r)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    97
    have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef])
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    98
    thus "x = Abs (Rep x)" ..
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    99
    show "Rep x \<in> A" by (rule Rep [OF tydef])
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   100
  qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   101
qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   102
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   103
theorem Rep_induct:
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   104
  "type_definition Rep Abs A ==> y \<in> A ==> (!!x. P (Rep x)) ==> P y"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   105
proof -
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   106
  assume tydef: "type_definition Rep Abs A"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   107
  assume "!!x. P (Rep x)" hence "P (Rep (Abs y))" .
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   108
  moreover assume "y \<in> A" hence "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef])
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   109
  ultimately show "P y" by (simp only:)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   110
qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   111
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   112
theorem Abs_induct:
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   113
  "type_definition Rep Abs A ==> (!!y. y \<in> A ==> P (Abs y)) ==> P x"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   114
proof -
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   115
  assume tydef: "type_definition Rep Abs A"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   116
  assume r: "!!y. y \<in> A ==> P (Abs y)"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   117
  have "Rep x \<in> A" by (rule Rep [OF tydef])
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   118
  hence "P (Abs (Rep x))" by (rule r)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   119
  moreover have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef])
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   120
  ultimately show "P x" by (simp only:)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   121
qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   122
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   123
setup InductAttrib.setup
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   124
use "Tools/typedef_package.ML"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   125
7705
222b715b5d24 Tools/typedef_package.ML;
wenzelm
parents: 5853
diff changeset
   126
end