src/HOL/Typedef.thy
author nipkow
Tue, 17 Jun 2025 14:11:40 +0200
changeset 82733 8b537e1af2ec
parent 81586 257f93d40d7c
permissions -rw-r--r--
reinstated intersection of lists as inter_list_set
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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     5
section \<open>HOL type definitions\<close>
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
63434
c956d995bec6 more indentation for quasi_command keywords;
wenzelm
parents: 61799
diff changeset
     9
keywords
69913
ca515cf61651 more specific keyword kinds;
wenzelm
parents: 69605
diff changeset
    10
  "typedef" :: thy_goal_defn and
63434
c956d995bec6 more indentation for quasi_command keywords;
wenzelm
parents: 61799
diff changeset
    11
  "morphisms" :: quasi_command
15131
c69542757a4d New theory header syntax.
nipkow
parents: 13421
diff changeset
    12
begin
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    13
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    14
locale type_definition =
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    15
  fixes Rep and Abs and A
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    16
  assumes Rep: "Rep x \<in> A"
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    17
    and Rep_inverse: "Abs (Rep x) = x"
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    18
    and Abs_inverse: "y \<in> A \<Longrightarrow> Rep (Abs y) = y"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61102
diff changeset
    19
  \<comment> \<open>This will be axiomatized for each typedef!\<close>
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    20
begin
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    21
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    22
lemma Rep_inject: "Rep x = Rep y \<longleftrightarrow> x = y"
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    23
proof
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    24
  assume "Rep x = Rep y"
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    25
  then have "Abs (Rep x) = Abs (Rep y)" by (simp only:)
81586
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    26
  also have "Abs (Rep x) = x" by (rule Rep_inverse)
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    27
  also have "Abs (Rep y) = y" by (rule Rep_inverse)
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    28
  finally show "x = y" .
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    29
next
81586
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    30
  show "x = y \<Longrightarrow> Rep x = Rep y" by (simp only:)
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    31
qed
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    32
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    33
lemma Abs_inject:
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    34
  assumes "x \<in> A" and "y \<in> A"
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    35
  shows "Abs x = Abs y \<longleftrightarrow> x = y"
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    36
proof
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    37
  assume "Abs x = Abs y"
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    38
  then have "Rep (Abs x) = Rep (Abs y)" by (simp only:)
81586
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    39
  also from \<open>x \<in> A\<close> have "Rep (Abs x) = x" by (rule Abs_inverse)
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    40
  also from \<open>y \<in> A\<close> have "Rep (Abs y) = y" by (rule Abs_inverse)
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    41
  finally show "x = y" .
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    42
next
81586
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    43
  show "x = y \<Longrightarrow> Abs x = Abs y" by (simp only:)
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    44
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    45
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    46
lemma Rep_cases [cases set]:
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    47
  assumes "y \<in> A"
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    48
    and hyp: "\<And>x. y = Rep x \<Longrightarrow> P"
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    49
  shows P
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    50
proof (rule hyp)
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    51
  from \<open>y \<in> A\<close> have "Rep (Abs y) = y" by (rule Abs_inverse)
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    52
  then show "y = Rep (Abs y)" ..
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    53
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    54
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    55
lemma Abs_cases [cases type]:
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    56
  assumes r: "\<And>y. x = Abs y \<Longrightarrow> y \<in> A \<Longrightarrow> P"
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    57
  shows P
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    58
proof (rule r)
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    59
  have "Abs (Rep x) = x" by (rule Rep_inverse)
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    60
  then show "x = Abs (Rep x)" ..
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    61
  show "Rep x \<in> A" by (rule Rep)
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    62
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    63
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    64
lemma Rep_induct [induct set]:
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    65
  assumes y: "y \<in> A"
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    66
    and hyp: "\<And>x. P (Rep x)"
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    67
  shows "P y"
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    68
proof -
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    69
  have "P (Rep (Abs y))" by (rule hyp)
81586
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    70
  also from y have "Rep (Abs y) = y" by (rule Abs_inverse)
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    71
  finally show "P y" .
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    72
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    73
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
    74
lemma Abs_induct [induct type]:
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    75
  assumes r: "\<And>y. y \<in> A \<Longrightarrow> P (Abs y)"
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    76
  shows "P x"
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    77
proof -
13412
666137b488a4 predicate defs via locales;
wenzelm
parents: 12023
diff changeset
    78
  have "Rep x \<in> A" by (rule Rep)
23710
a8ac2305eaf2 removed proof dependency on transitivity theorems
haftmann
parents: 23433
diff changeset
    79
  then have "P (Abs (Rep x))" by (rule r)
81586
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    80
  also have "Abs (Rep x) = x" by (rule Rep_inverse)
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
    81
  finally show "P x" .
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    82
qed
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
    83
27295
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
    84
lemma Rep_range: "range Rep = A"
24269
4b2aac7669b3 remove redundant assumption from Rep_range lemma
huffman
parents: 23710
diff changeset
    85
proof
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    86
  show "range Rep \<subseteq> A" using Rep by (auto simp add: image_def)
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    87
  show "A \<subseteq> range Rep"
23433
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    88
  proof
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    89
    fix x assume "x \<in> A"
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    90
    then have "x = Rep (Abs x)" by (rule Abs_inverse [symmetric])
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    91
    then show "x \<in> range Rep" by (rule range_eqI)
23433
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    92
  qed
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    93
qed
c2c10abd2a1e added lemmas
nipkow
parents: 23247
diff changeset
    94
27295
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
    95
lemma Abs_image: "Abs ` A = UNIV"
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
    96
proof
61102
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    97
  show "Abs ` A \<subseteq> UNIV" by (rule subset_UNIV)
0ec9fd8d8119 misc tuning and modernization;
wenzelm
parents: 60758
diff changeset
    98
  show "UNIV \<subseteq> Abs ` A"
27295
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
    99
  proof
81586
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
   100
    show "x \<in> Abs ` A" for x
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
   101
    proof (rule image_eqI)
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
   102
      show "x = Abs (Rep x)" by (rule Rep_inverse [symmetric])
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
   103
      show "Rep x \<in> A" by (rule Rep)
257f93d40d7c tuned proofs;
wenzelm
parents: 69913
diff changeset
   104
    qed
27295
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   105
  qed
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   106
qed
cfe5244301dd add lemma Abs_image
huffman
parents: 26802
diff changeset
   107
23247
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
   108
end
b99dce43d252 merged Code_Generator.thy into HOL.thy
haftmann
parents: 22846
diff changeset
   109
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 63434
diff changeset
   110
ML_file \<open>Tools/typedef.ML\<close>
11608
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   111
c760ea8154ee renamed theory "subset" to "Typedef";
wenzelm
parents:
diff changeset
   112
end