src/HOL/subset.thy
author paulson
Wed, 25 Jul 2001 17:58:26 +0200
changeset 11454 7514e5e21cb8
parent 11083 d8fda557e476
permissions -rw-r--r--
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
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
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    12
(** belongs to theory Ord **)
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    13
  
9895
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    14
theorems linorder_cases [case_names less equal greater] =
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    15
  linorder_less_split
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    16
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    17
(* Courtesy of Stephan Merz *)
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    18
lemma Least_mono: 
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    19
"[| mono (f::'a::order => 'b::order); EX x:S. ALL y:S. x <= y |]  
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    20
   ==> (LEAST y. y : f`S) = f(LEAST x. x : S)"
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    21
apply clarify
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    22
apply (erule_tac P = "%x. x : S" in LeastI2)
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    23
apply  fast
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    24
apply (rule LeastI2)
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    25
apply (auto elim: monoD intro!: order_antisym)
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    26
done
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    27
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    28
9895
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    29
(*belongs to theory Set*)
75e55370b1ae added linorder_cases;
wenzelm
parents: 7717
diff changeset
    30
setup Rulify.setup
7717
e7ecfa617443 Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents: 7705
diff changeset
    31
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    32
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    33
section {* HOL type definitions *}
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    34
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    35
constdefs
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    36
  type_definition :: "('a => 'b) => ('b => 'a) => 'b set => bool"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    37
  "type_definition Rep Abs A ==
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    38
    (\\<forall>x. Rep x \\<in> A) \\<and>
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    39
    (\\<forall>x. Abs (Rep x) = x) \\<and>
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    40
    (\\<forall>y \\<in> A. Rep (Abs y) = y)"
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    41
  -- {* This will be stated as an axiom for each typedef! *}
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    42
10290
8018d1743beb added type_definitionI;
wenzelm
parents: 10284
diff changeset
    43
lemma type_definitionI [intro]:
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    44
  "(!!x. Rep x \\<in> A) ==>
10291
a88d347db404 isatool unsymbolize;
wenzelm
parents: 10290
diff changeset
    45
    (!!x. Abs (Rep x) = x) ==>
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    46
    (!!y. y \\<in> A ==> Rep (Abs y) = y) ==>
10290
8018d1743beb added type_definitionI;
wenzelm
parents: 10284
diff changeset
    47
    type_definition Rep Abs A"
8018d1743beb added type_definitionI;
wenzelm
parents: 10284
diff changeset
    48
  by (unfold type_definition_def) blast
8018d1743beb added type_definitionI;
wenzelm
parents: 10284
diff changeset
    49
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    50
theorem Rep: "type_definition Rep Abs A ==> Rep x \\<in> A"
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    51
  by (unfold type_definition_def) blast
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    52
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    53
theorem Rep_inverse: "type_definition Rep Abs A ==> Abs (Rep x) = x"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    54
  by (unfold type_definition_def) blast
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    55
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    56
theorem Abs_inverse: "type_definition Rep Abs A ==> y \\<in> A ==> Rep (Abs y) = y"
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    57
  by (unfold type_definition_def) blast
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    58
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    59
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
    60
proof -
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    61
  assume tydef: "type_definition Rep Abs A"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    62
  show ?thesis
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    63
  proof
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    64
    assume "Rep x = Rep y"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    65
    hence "Abs (Rep x) = Abs (Rep y)" by (simp only:)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    66
    thus "x = y" by (simp only: Rep_inverse [OF tydef])
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    67
  next
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    68
    assume "x = y"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    69
    thus "Rep x = Rep y" by simp
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    70
  qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    71
qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    72
10284
wenzelm
parents: 10276
diff changeset
    73
theorem Abs_inject:
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    74
  "type_definition Rep Abs A ==> x \\<in> A ==> y \\<in> A ==> (Abs x = Abs y) = (x = y)"
10284
wenzelm
parents: 10276
diff changeset
    75
proof -
wenzelm
parents: 10276
diff changeset
    76
  assume tydef: "type_definition Rep Abs A"
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    77
  assume x: "x \\<in> A" and y: "y \\<in> A"
10284
wenzelm
parents: 10276
diff changeset
    78
  show ?thesis
wenzelm
parents: 10276
diff changeset
    79
  proof
wenzelm
parents: 10276
diff changeset
    80
    assume "Abs x = Abs y"
wenzelm
parents: 10276
diff changeset
    81
    hence "Rep (Abs x) = Rep (Abs y)" by simp
11083
wenzelm
parents: 10291
diff changeset
    82
    moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse [OF tydef])
wenzelm
parents: 10291
diff changeset
    83
    moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef])
10284
wenzelm
parents: 10276
diff changeset
    84
    ultimately show "x = y" by (simp only:)
wenzelm
parents: 10276
diff changeset
    85
  next
wenzelm
parents: 10276
diff changeset
    86
    assume "x = y"
wenzelm
parents: 10276
diff changeset
    87
    thus "Abs x = Abs y" by simp
wenzelm
parents: 10276
diff changeset
    88
  qed
wenzelm
parents: 10276
diff changeset
    89
qed
wenzelm
parents: 10276
diff changeset
    90
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    91
theorem Rep_cases:
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    92
  "type_definition Rep Abs A ==> y \\<in> A ==> (!!x. y = Rep x ==> P) ==> P"
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    93
proof -
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    94
  assume tydef: "type_definition Rep Abs A"
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
    95
  assume y: "y \\<in> A" and r: "(!!x. y = Rep x ==> P)"
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    96
  show P
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    97
  proof (rule r)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
    98
    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
    99
    thus "y = Rep (Abs y)" ..
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 Abs_cases:
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
   104
  "type_definition Rep Abs A ==> (!!y. x = Abs y ==> y \\<in> A ==> P) ==> P"
10276
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"
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
   107
  assume r: "!!y. x = Abs y ==> y \\<in> A ==> P"
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   108
  show P
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   109
  proof (rule r)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   110
    have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef])
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   111
    thus "x = Abs (Rep x)" ..
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
   112
    show "Rep x \\<in> A" by (rule Rep [OF tydef])
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   113
  qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   114
qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   115
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   116
theorem Rep_induct:
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
   117
  "type_definition Rep Abs A ==> y \\<in> A ==> (!!x. P (Rep x)) ==> P y"
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   118
proof -
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   119
  assume tydef: "type_definition Rep Abs A"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   120
  assume "!!x. P (Rep x)" hence "P (Rep (Abs y))" .
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
   121
  moreover assume "y \\<in> A" hence "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef])
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   122
  ultimately show "P y" by (simp only:)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   123
qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   124
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   125
theorem Abs_induct:
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
   126
  "type_definition Rep Abs A ==> (!!y. y \\<in> A ==> P (Abs y)) ==> P x"
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   127
proof -
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   128
  assume tydef: "type_definition Rep Abs A"
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
   129
  assume r: "!!y. y \\<in> A ==> P (Abs y)"
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11083
diff changeset
   130
  have "Rep x \\<in> A" by (rule Rep [OF tydef])
10276
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   131
  hence "P (Abs (Rep x))" by (rule r)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   132
  moreover have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef])
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   133
  ultimately show "P x" by (simp only:)
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   134
qed
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   135
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   136
setup InductAttrib.setup
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   137
use "Tools/typedef_package.ML"
75e2c6cb4153 added theory for HOL type definitions;
wenzelm
parents: 9895
diff changeset
   138
7705
222b715b5d24 Tools/typedef_package.ML;
wenzelm
parents: 5853
diff changeset
   139
end