src/HOL/Typedef.thy
author blanchet
Fri, 17 Dec 2010 23:09:53 +0100
changeset 41260 ff38ea43aada
parent 38536 7e57a0dcbd4f
child 41732 996b0c14a430
permissions -rw-r--r--
remove option that doesn't work in Mirabelle anyway (Mirabelle uses Sledgehammer_Provers, not Sledgehammer_Run)
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
    Author:     Markus Wenzel, TU Munich
11743
wenzelm
parents: 11659
diff changeset
     3
*)
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
     4
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11770
diff changeset
     5
header {* HOL type definitions *}
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
     6
15131
c69542757a4d New theory header syntax.
nipkow
parents: 13421
diff changeset
     7
theory Typedef
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
     8
imports Set
38536
7e57a0dcbd4f dropped SML typedef_codegen: does not fit to code equations for record operations any longer
haftmann
parents: 38393
diff changeset
     9
uses ("Tools/typedef.ML")
15131
c69542757a4d New theory header syntax.
nipkow
parents: 13421
diff changeset
    10
begin
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    11
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    12
ML {*
37863
7f113caabcf4 discontinued pervasive val theory = Thy_Info.get_theory -- prefer antiquotations in most situations;
wenzelm
parents: 31723
diff changeset
    13
structure HOL = struct val thy = @{theory HOL} end;
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    14
*}  -- "belongs to theory HOL"
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    15
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    16
locale type_definition =
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    17
  fixes Rep and Abs and A
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    18
  assumes Rep: "Rep x \<in> A"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    19
    and Rep_inverse: "Abs (Rep x) = x"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    20
    and Abs_inverse: "y \<in> A ==> Rep (Abs y) = y"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    21
  -- {* This will be axiomatized for each typedef! *}
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    22
begin
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    23
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    24
lemma Rep_inject:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    25
  "(Rep x = Rep y) = (x = y)"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    26
proof
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    27
  assume "Rep x = Rep y"
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    28
  then have "Abs (Rep x) = Abs (Rep y)" by (simp only:)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    29
  moreover have "Abs (Rep x) = x" by (rule Rep_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    30
  moreover have "Abs (Rep y) = y" by (rule Rep_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    31
  ultimately show "x = y" by simp
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    32
next
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    33
  assume "x = y"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    34
  thus "Rep x = Rep y" by (simp only:)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    35
qed
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    36
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    37
lemma Abs_inject:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    38
  assumes x: "x \<in> A" and y: "y \<in> A"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    39
  shows "(Abs x = Abs y) = (x = y)"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    40
proof
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    41
  assume "Abs x = Abs y"
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    42
  then have "Rep (Abs x) = Rep (Abs y)" by (simp only:)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    43
  moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    44
  moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    45
  ultimately show "x = y" by simp
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    46
next
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    47
  assume "x = y"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    48
  thus "Abs x = Abs y" by (simp only:)
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    49
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    50
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    51
lemma Rep_cases [cases set]:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    52
  assumes y: "y \<in> A"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    53
    and hyp: "!!x. y = Rep x ==> P"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    54
  shows P
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    55
proof (rule hyp)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    56
  from y have "Rep (Abs y) = y" by (rule Abs_inverse)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    57
  thus "y = Rep (Abs y)" ..
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    58
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    59
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    60
lemma Abs_cases [cases type]:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    61
  assumes r: "!!y. x = Abs y ==> y \<in> A ==> P"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    62
  shows P
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    63
proof (rule r)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    64
  have "Abs (Rep x) = x" by (rule Rep_inverse)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    65
  thus "x = Abs (Rep x)" ..
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    66
  show "Rep x \<in> A" by (rule Rep)
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    67
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    68
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    69
lemma Rep_induct [induct set]:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    70
  assumes y: "y \<in> A"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    71
    and hyp: "!!x. P (Rep x)"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    72
  shows "P y"
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    73
proof -
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    74
  have "P (Rep (Abs y))" by (rule hyp)
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    75
  moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    76
  ultimately show "P y" by simp
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    77
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    78
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    79
lemma Abs_induct [induct type]:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    80
  assumes r: "!!y. y \<in> A ==> P (Abs y)"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    81
  shows "P x"
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    82
proof -
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    83
  have "Rep x \<in> A" by (rule Rep)
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    84
  then have "P (Abs (Rep x))" by (rule r)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    85
  moreover have "Abs (Rep x) = x" by (rule Rep_inverse)
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    86
  ultimately show "P x" by simp
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    87
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    88
27295
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
    89
lemma Rep_range: "range Rep = A"
24269
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
    90
proof
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
    91
  show "range Rep <= A" using Rep by (auto simp add: image_def)
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
    92
  show "A <= range Rep"
23433
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    93
  proof
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    94
    fix x assume "x : A"
24269
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
    95
    hence "x = Rep (Abs x)" by (rule Abs_inverse [symmetric])
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
    96
    thus "x : range Rep" by (rule range_eqI)
23433
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    97
  qed
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    98
qed
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    99
27295
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   100
lemma Abs_image: "Abs ` A = UNIV"
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   101
proof
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   102
  show "Abs ` A <= UNIV" by (rule subset_UNIV)
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   103
next
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   104
  show "UNIV <= Abs ` A"
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   105
  proof
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   106
    fix x
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   107
    have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric])
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   108
    moreover have "Rep x : A" by (rule Rep)
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   109
    ultimately show "x : Abs ` A" by (rule image_eqI)
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   110
  qed
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   111
qed
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   112
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
   113
end
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
   114
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 29797
diff changeset
   115
use "Tools/typedef.ML" setup Typedef.setup
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   116
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   117
end