src/ZF/upair.thy
author wenzelm
Sun, 24 Sep 2023 15:07:40 +0200
changeset 78688 ff7db9055002
parent 76216 9fc34f76b4e8
child 79340 3ef7606a0d11
permissions -rw-r--r--
clarified signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 124
diff changeset
     1
(*  Title:      ZF/upair.thy
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 124
diff changeset
     2
    Author:     Lawrence C Paulson and Martin D Coen, CU Computer Laboratory
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 124
diff changeset
     3
    Copyright   1993  University of Cambridge
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
     4
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
     5
Observe the order of dependence:
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
     6
    Upair is defined in terms of Replace
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
     7
    \<union> is defined in terms of Upair and \<Union>(similarly for Int)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
     8
    cons is defined in terms of Upair and Un
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
     9
    Ordered pairs and descriptions are defined using cons ("set notation")
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 124
diff changeset
    10
*)
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 124
diff changeset
    11
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    12
section\<open>Unordered Pairs\<close>
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
    13
46950
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 46821
diff changeset
    14
theory upair
65449
c82e63b11b8b clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 63901
diff changeset
    15
imports ZF_Base
46950
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 46821
diff changeset
    16
keywords "print_tcset" :: diag
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 46821
diff changeset
    17
begin
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 3924
diff changeset
    18
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    19
ML_file \<open>Tools/typechk.ML\<close>
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 3924
diff changeset
    20
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    21
lemma atomize_ball [symmetric, rulify]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    22
     "(\<And>x. x \<in> A \<Longrightarrow> P(x)) \<equiv> Trueprop (\<forall>x\<in>A. P(x))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    23
by (simp add: Ball_def atomize_all atomize_imp)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    24
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    25
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68233
diff changeset
    26
subsection\<open>Unordered Pairs: constant \<^term>\<open>Upair\<close>\<close>
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    27
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
    28
lemma Upair_iff [simp]: "c \<in> Upair(a,b) \<longleftrightarrow> (c=a | c=b)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    29
by (unfold Upair_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    30
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
    31
lemma UpairI1: "a \<in> Upair(a,b)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    32
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    33
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
    34
lemma UpairI2: "b \<in> Upair(a,b)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    35
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    36
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    37
lemma UpairE: "\<lbrakk>a \<in> Upair(b,c);  a=b \<Longrightarrow> P;  a=c \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    38
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    39
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68233
diff changeset
    40
subsection\<open>Rules for Binary Union, Defined via \<^term>\<open>Upair\<close>\<close>
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    41
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    42
lemma Un_iff [simp]: "c \<in> A \<union> B \<longleftrightarrow> (c \<in> A | c \<in> B)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    43
apply (simp add: Un_def)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    44
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    45
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    46
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    47
lemma UnI1: "c \<in> A \<Longrightarrow> c \<in> A \<union> B"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    48
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    49
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    50
lemma UnI2: "c \<in> B \<Longrightarrow> c \<in> A \<union> B"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    51
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    52
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13259
diff changeset
    53
declare UnI1 [elim?]  UnI2 [elim?]
c9cfe1638bf2 improved presentation markup
paulson
parents: 13259
diff changeset
    54
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    55
lemma UnE [elim!]: "\<lbrakk>c \<in> A \<union> B;  c \<in> A \<Longrightarrow> P;  c \<in> B \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    56
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    57
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    58
(*Stronger version of the rule above*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    59
lemma UnE': "\<lbrakk>c \<in> A \<union> B;  c \<in> A \<Longrightarrow> P;  \<lbrakk>c \<in> B;  c\<notin>A\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    60
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    61
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    62
(*Classical introduction rule: no commitment to A vs B*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    63
lemma UnCI [intro!]: "(c \<notin> B \<Longrightarrow> c \<in> A) \<Longrightarrow> c \<in> A \<union> B"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    64
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    65
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68233
diff changeset
    66
subsection\<open>Rules for Binary Intersection, Defined via \<^term>\<open>Upair\<close>\<close>
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    67
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
    68
lemma Int_iff [simp]: "c \<in> A \<inter> B \<longleftrightarrow> (c \<in> A \<and> c \<in> B)"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
    69
  unfolding Int_def
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    70
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    71
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    72
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    73
lemma IntI [intro!]: "\<lbrakk>c \<in> A;  c \<in> B\<rbrakk> \<Longrightarrow> c \<in> A \<inter> B"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    74
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    75
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    76
lemma IntD1: "c \<in> A \<inter> B \<Longrightarrow> c \<in> A"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    77
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    78
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    79
lemma IntD2: "c \<in> A \<inter> B \<Longrightarrow> c \<in> B"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    80
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    81
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    82
lemma IntE [elim!]: "\<lbrakk>c \<in> A \<inter> B;  \<lbrakk>c \<in> A; c \<in> B\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    83
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    84
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    85
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68233
diff changeset
    86
subsection\<open>Rules for Set Difference, Defined via \<^term>\<open>Upair\<close>\<close>
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    87
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
    88
lemma Diff_iff [simp]: "c \<in> A-B \<longleftrightarrow> (c \<in> A \<and> c\<notin>B)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    89
by (unfold Diff_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    90
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    91
lemma DiffI [intro!]: "\<lbrakk>c \<in> A;  c \<notin> B\<rbrakk> \<Longrightarrow> c \<in> A - B"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    92
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    93
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    94
lemma DiffD1: "c \<in> A - B \<Longrightarrow> c \<in> A"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    95
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    96
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
    97
lemma DiffD2: "c \<in> A - B \<Longrightarrow> c \<notin> B"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    98
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    99
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   100
lemma DiffE [elim!]: "\<lbrakk>c \<in> A - B;  \<lbrakk>c \<in> A; c\<notin>B\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   101
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   102
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   103
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68233
diff changeset
   104
subsection\<open>Rules for \<^term>\<open>cons\<close>\<close>
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   105
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   106
lemma cons_iff [simp]: "a \<in> cons(b,A) \<longleftrightarrow> (a=b | a \<in> A)"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   107
  unfolding cons_def
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   108
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   109
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   110
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   111
(*risky as a typechecking rule, but solves otherwise unconstrained goals of
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   112
the form x \<in> ?A*)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   113
lemma consI1 [simp,TC]: "a \<in> cons(a,B)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   114
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   115
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   116
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   117
lemma consI2: "a \<in> B \<Longrightarrow> a \<in> cons(b,B)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   118
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   119
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   120
lemma consE [elim!]: "\<lbrakk>a \<in> cons(b,A);  a=b \<Longrightarrow> P;  a \<in> A \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   121
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   122
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   123
(*Stronger version of the rule above*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   124
lemma consE':
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   125
    "\<lbrakk>a \<in> cons(b,A);  a=b \<Longrightarrow> P;  \<lbrakk>a \<in> A;  a\<noteq>b\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   126
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   127
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   128
(*Classical introduction rule*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   129
lemma consCI [intro!]: "(a\<notin>B \<Longrightarrow> a=b) \<Longrightarrow> a \<in> cons(b,B)"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   130
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   131
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   132
lemma cons_not_0 [simp]: "cons(a,B) \<noteq> 0"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   133
by (blast elim: equalityE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   134
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 32960
diff changeset
   135
lemmas cons_neq_0 = cons_not_0 [THEN notE]
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   136
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   137
declare cons_not_0 [THEN not_sym, simp]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   138
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   139
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   140
subsection\<open>Singletons\<close>
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   141
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   142
lemma singleton_iff: "a \<in> {b} \<longleftrightarrow> a=b"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   143
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   144
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   145
lemma singletonI [intro!]: "a \<in> {a}"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   146
by (rule consI1)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   147
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 32960
diff changeset
   148
lemmas singletonE = singleton_iff [THEN iffD1, elim_format, elim!]
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   149
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   150
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   151
subsection\<open>Descriptions\<close>
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   152
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   153
lemma the_equality [intro]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   154
    "\<lbrakk>P(a);  \<And>x. P(x) \<Longrightarrow> x=a\<rbrakk> \<Longrightarrow> (THE x. P(x)) = a"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   155
  unfolding the_def
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   156
apply (fast dest: subst)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   157
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   158
63901
4ce989e962e0 more symbols;
wenzelm
parents: 61798
diff changeset
   159
(* Only use this if you already know \<exists>!x. P(x) *)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   160
lemma the_equality2: "\<lbrakk>\<exists>!x. P(x);  P(a)\<rbrakk> \<Longrightarrow> (THE x. P(x)) = a"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   161
by blast
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   162
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   163
lemma theI: "\<exists>!x. P(x) \<Longrightarrow> P(THE x. P(x))"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   164
apply (erule ex1E)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   165
apply (subst the_equality)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   166
apply (blast+)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   167
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   168
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   169
(*No congruence rule is necessary: if @{term"\<forall>y.P(y)\<longleftrightarrow>Q(y)"} then
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   170
  @{term "THE x.P(x)"}  rewrites to @{term "THE x.Q(x)"} *)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   171
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   172
(*If it's "undefined", it's zero!*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   173
lemma the_0: "\<not> (\<exists>!x. P(x)) \<Longrightarrow> (THE x. P(x))=0"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   174
  unfolding the_def
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   175
apply (blast elim!: ReplaceE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   176
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   177
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   178
(*Easier to apply than theI: conclusion has only one occurrence of P*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   179
lemma theI2:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   180
    assumes p1: "\<not> Q(0) \<Longrightarrow> \<exists>!x. P(x)"
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   181
        and p2: "\<And>x. P(x) \<Longrightarrow> Q(x)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   182
    shows "Q(THE x. P(x))"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   183
apply (rule classical)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   184
apply (rule p2)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   185
apply (rule theI)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   186
apply (rule classical)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   187
apply (rule p1)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   188
apply (erule the_0 [THEN subst], assumption)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   189
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   190
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   191
lemma the_eq_trivial [simp]: "(THE x. x = a) = a"
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   192
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   193
13544
895994073bdf various new lemmas for Constructible
paulson
parents: 13357
diff changeset
   194
lemma the_eq_trivial2 [simp]: "(THE x. a = x) = a"
895994073bdf various new lemmas for Constructible
paulson
parents: 13357
diff changeset
   195
by blast
895994073bdf various new lemmas for Constructible
paulson
parents: 13357
diff changeset
   196
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   197
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   198
subsection\<open>Conditional Terms: \<open>if-then-else\<close>\<close>
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   199
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   200
lemma if_true [simp]: "(if True then a else b) = a"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   201
by (unfold if_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   202
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   203
lemma if_false [simp]: "(if False then a else b) = b"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   204
by (unfold if_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   205
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   206
(*Never use with case splitting, or if P is known to be true or false*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   207
lemma if_cong:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   208
    "\<lbrakk>P\<longleftrightarrow>Q;  Q \<Longrightarrow> a=c;  \<not>Q \<Longrightarrow> b=d\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   209
     \<Longrightarrow> (if P then a else b) = (if Q then c else d)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   210
by (simp add: if_def cong add: conj_cong)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   211
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   212
(*Prevents simplification of x and y \<in> faster and allows the execution
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   213
  of functional programs. NOW THE DEFAULT.*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   214
lemma if_weak_cong: "P\<longleftrightarrow>Q \<Longrightarrow> (if P then x else y) = (if Q then x else y)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   215
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   216
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   217
(*Not needed for rewriting, since P would rewrite to True anyway*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   218
lemma if_P: "P \<Longrightarrow> (if P then a else b) = a"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   219
by (unfold if_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   220
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   221
(*Not needed for rewriting, since P would rewrite to False anyway*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   222
lemma if_not_P: "\<not>P \<Longrightarrow> (if P then a else b) = b"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   223
by (unfold if_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   224
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   225
lemma split_if [split]:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   226
     "P(if Q then x else y) \<longleftrightarrow> ((Q \<longrightarrow> P(x)) \<and> (\<not>Q \<longrightarrow> P(y)))"
14153
76a6ba67bd15 new case_tac
paulson
parents: 13780
diff changeset
   227
by (case_tac Q, simp_all)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   228
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45602
diff changeset
   229
(** Rewrite rules for boolean case-splitting: faster than split_if [split]
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   230
**)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   231
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   232
lemmas split_if_eq1 = split_if [of "\<lambda>x. x = b"] for b
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   233
lemmas split_if_eq2 = split_if [of "\<lambda>x. a = x"] for a
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   234
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   235
lemmas split_if_mem1 = split_if [of "\<lambda>x. x \<in> b"] for b
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   236
lemmas split_if_mem2 = split_if [of "\<lambda>x. a \<in> x"] for a
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   237
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   238
lemmas split_ifs = split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   239
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   240
(*Logically equivalent to split_if_mem2*)
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   241
lemma if_iff: "a: (if P then x else y) \<longleftrightarrow> P \<and> a \<in> x | \<not>P \<and> a \<in> y"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   242
by simp
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   243
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   244
lemma if_type [TC]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   245
    "\<lbrakk>P \<Longrightarrow> a \<in> A;  \<not>P \<Longrightarrow> b \<in> A\<rbrakk> \<Longrightarrow> (if P then a else b): A"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   246
by simp
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   247
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   248
(** Splitting IFs in the assumptions **)
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   249
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   250
lemma split_if_asm: "P(if Q then x else y) \<longleftrightarrow> (\<not>((Q \<and> \<not>P(x)) | (\<not>Q \<and> \<not>P(y))))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   251
by simp
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   252
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   253
lemmas if_splits = split_if split_if_asm
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   254
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   255
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   256
subsection\<open>Consequences of Foundation\<close>
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   257
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   258
(*was called mem_anti_sym*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   259
lemma mem_asym: "\<lbrakk>a \<in> b;  \<not>P \<Longrightarrow> b \<in> a\<rbrakk> \<Longrightarrow> P"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   260
apply (rule classical)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   261
apply (rule_tac A1 = "{a,b}" in foundation [THEN disjE])
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   262
apply (blast elim!: equalityE)+
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   263
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   264
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   265
(*was called mem_anti_refl*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   266
lemma mem_irrefl: "a \<in> a \<Longrightarrow> P"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   267
by (blast intro: mem_asym)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   268
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   269
(*mem_irrefl should NOT be added to default databases:
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   270
      it would be tried on most goals, making proofs slower!*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   271
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   272
lemma mem_not_refl: "a \<notin> a"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   273
apply (rule notI)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   274
apply (erule mem_irrefl)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   275
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   276
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   277
(*Good for proving inequalities by rewriting*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   278
lemma mem_imp_not_eq: "a \<in> A \<Longrightarrow> a \<noteq> A"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   279
by (blast elim!: mem_irrefl)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   280
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   281
lemma eq_imp_not_mem: "a=A \<Longrightarrow> a \<notin> A"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   282
by (blast intro: elim: mem_irrefl)
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   283
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   284
subsection\<open>Rules for Successor\<close>
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   285
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   286
lemma succ_iff: "i \<in> succ(j) \<longleftrightarrow> i=j | i \<in> j"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   287
by (unfold succ_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   288
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   289
lemma succI1 [simp]: "i \<in> succ(i)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   290
by (simp add: succ_iff)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   291
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   292
lemma succI2: "i \<in> j \<Longrightarrow> i \<in> succ(j)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   293
by (simp add: succ_iff)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   294
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   295
lemma succE [elim!]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   296
    "\<lbrakk>i \<in> succ(j);  i=j \<Longrightarrow> P;  i \<in> j \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   297
apply (simp add: succ_iff, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   298
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   299
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   300
(*Classical introduction rule*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   301
lemma succCI [intro!]: "(i\<notin>j \<Longrightarrow> i=j) \<Longrightarrow> i \<in> succ(j)"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   302
by (simp add: succ_iff, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   303
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   304
lemma succ_not_0 [simp]: "succ(n) \<noteq> 0"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   305
by (blast elim!: equalityE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   306
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 32960
diff changeset
   307
lemmas succ_neq_0 = succ_not_0 [THEN notE, elim!]
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   308
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   309
declare succ_not_0 [THEN not_sym, simp]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   310
declare sym [THEN succ_neq_0, elim!]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   311
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   312
(* @{term"succ(c) \<subseteq> B \<Longrightarrow> c \<in> B"} *)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   313
lemmas succ_subsetD = succI1 [THEN [2] subsetD]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   314
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   315
(* @{term"succ(b) \<noteq> b"} *)
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 32960
diff changeset
   316
lemmas succ_neq_self = succI1 [THEN mem_imp_not_eq, THEN not_sym]
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   317
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   318
lemma succ_inject_iff [simp]: "succ(m) = succ(n) \<longleftrightarrow> m=n"
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   319
by (blast elim: mem_asym elim!: equalityE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   320
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 32960
diff changeset
   321
lemmas succ_inject = succ_inject_iff [THEN iffD1, dest!]
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   322
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   323
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   324
subsection\<open>Miniscoping of the Bounded Universal Quantifier\<close>
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   325
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   326
lemma ball_simps1:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   327
     "(\<forall>x\<in>A. P(x) \<and> Q)   \<longleftrightarrow> (\<forall>x\<in>A. P(x)) \<and> (A=0 | Q)"
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   328
     "(\<forall>x\<in>A. P(x) | Q)   \<longleftrightarrow> ((\<forall>x\<in>A. P(x)) | Q)"
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   329
     "(\<forall>x\<in>A. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<exists>x\<in>A. P(x)) \<longrightarrow> Q)"
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   330
     "(\<not>(\<forall>x\<in>A. P(x))) \<longleftrightarrow> (\<exists>x\<in>A. \<not>P(x))"
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   331
     "(\<forall>x\<in>0.P(x)) \<longleftrightarrow> True"
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   332
     "(\<forall>x\<in>succ(i).P(x)) \<longleftrightarrow> P(i) \<and> (\<forall>x\<in>i. P(x))"
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   333
     "(\<forall>x\<in>cons(a,B).P(x)) \<longleftrightarrow> P(a) \<and> (\<forall>x\<in>B. P(x))"
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   334
     "(\<forall>x\<in>RepFun(A,f). P(x)) \<longleftrightarrow> (\<forall>y\<in>A. P(f(y)))"
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   335
     "(\<forall>x\<in>\<Union>(A).P(x)) \<longleftrightarrow> (\<forall>y\<in>A. \<forall>x\<in>y. P(x))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   336
by blast+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   337
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   338
lemma ball_simps2:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   339
     "(\<forall>x\<in>A. P \<and> Q(x))   \<longleftrightarrow> (A=0 | P) \<and> (\<forall>x\<in>A. Q(x))"
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   340
     "(\<forall>x\<in>A. P | Q(x))   \<longleftrightarrow> (P | (\<forall>x\<in>A. Q(x)))"
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   341
     "(\<forall>x\<in>A. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<forall>x\<in>A. Q(x)))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   342
by blast+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   343
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   344
lemma ball_simps3:
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   345
     "(\<forall>x\<in>Collect(A,Q).P(x)) \<longleftrightarrow> (\<forall>x\<in>A. Q(x) \<longrightarrow> P(x))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   346
by blast+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   347
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   348
lemmas ball_simps [simp] = ball_simps1 ball_simps2 ball_simps3
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   349
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   350
lemma ball_conj_distrib:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   351
    "(\<forall>x\<in>A. P(x) \<and> Q(x)) \<longleftrightarrow> ((\<forall>x\<in>A. P(x)) \<and> (\<forall>x\<in>A. Q(x)))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   352
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   353
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   354
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   355
subsection\<open>Miniscoping of the Bounded Existential Quantifier\<close>
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   356
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   357
lemma bex_simps1:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   358
     "(\<exists>x\<in>A. P(x) \<and> Q) \<longleftrightarrow> ((\<exists>x\<in>A. P(x)) \<and> Q)"
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   359
     "(\<exists>x\<in>A. P(x) | Q) \<longleftrightarrow> (\<exists>x\<in>A. P(x)) | (A\<noteq>0 \<and> Q)"
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   360
     "(\<exists>x\<in>A. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<forall>x\<in>A. P(x)) \<longrightarrow> (A\<noteq>0 \<and> Q))"
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   361
     "(\<exists>x\<in>0.P(x)) \<longleftrightarrow> False"
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   362
     "(\<exists>x\<in>succ(i).P(x)) \<longleftrightarrow> P(i) | (\<exists>x\<in>i. P(x))"
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   363
     "(\<exists>x\<in>cons(a,B).P(x)) \<longleftrightarrow> P(a) | (\<exists>x\<in>B. P(x))"
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   364
     "(\<exists>x\<in>RepFun(A,f). P(x)) \<longleftrightarrow> (\<exists>y\<in>A. P(f(y)))"
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   365
     "(\<exists>x\<in>\<Union>(A).P(x)) \<longleftrightarrow> (\<exists>y\<in>A. \<exists>x\<in>y.  P(x))"
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69605
diff changeset
   366
     "(\<not>(\<exists>x\<in>A. P(x))) \<longleftrightarrow> (\<forall>x\<in>A. \<not>P(x))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   367
by blast+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   368
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   369
lemma bex_simps2:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   370
     "(\<exists>x\<in>A. P \<and> Q(x)) \<longleftrightarrow> (P \<and> (\<exists>x\<in>A. Q(x)))"
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   371
     "(\<exists>x\<in>A. P | Q(x)) \<longleftrightarrow> (A\<noteq>0 \<and> P) | (\<exists>x\<in>A. Q(x))"
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   372
     "(\<exists>x\<in>A. P \<longrightarrow> Q(x)) \<longleftrightarrow> ((A=0 | P) \<longrightarrow> (\<exists>x\<in>A. Q(x)))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   373
by blast+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   374
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   375
lemma bex_simps3:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   376
     "(\<exists>x\<in>Collect(A,Q).P(x)) \<longleftrightarrow> (\<exists>x\<in>A. Q(x) \<and> P(x))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   377
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   378
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   379
lemmas bex_simps [simp] = bex_simps1 bex_simps2 bex_simps3
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   380
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   381
lemma bex_disj_distrib:
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   382
    "(\<exists>x\<in>A. P(x) | Q(x)) \<longleftrightarrow> ((\<exists>x\<in>A. P(x)) | (\<exists>x\<in>A. Q(x)))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   383
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   384
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   385
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   386
(** One-point rule for bounded quantifiers: see HOL/Set.ML **)
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   387
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   388
lemma bex_triv_one_point1 [simp]: "(\<exists>x\<in>A. x=a) \<longleftrightarrow> (a \<in> A)"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   389
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   390
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   391
lemma bex_triv_one_point2 [simp]: "(\<exists>x\<in>A. a=x) \<longleftrightarrow> (a \<in> A)"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   392
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   393
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   394
lemma bex_one_point1 [simp]: "(\<exists>x\<in>A. x=a \<and> P(x)) \<longleftrightarrow> (a \<in> A \<and> P(a))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   395
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   396
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   397
lemma bex_one_point2 [simp]: "(\<exists>x\<in>A. a=x \<and> P(x)) \<longleftrightarrow> (a \<in> A \<and> P(a))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   398
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   399
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   400
lemma ball_one_point1 [simp]: "(\<forall>x\<in>A. x=a \<longrightarrow> P(x)) \<longleftrightarrow> (a \<in> A \<longrightarrow> P(a))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   401
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   402
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   403
lemma ball_one_point2 [simp]: "(\<forall>x\<in>A. a=x \<longrightarrow> P(x)) \<longleftrightarrow> (a \<in> A \<longrightarrow> P(a))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   404
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   405
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   406
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   407
subsection\<open>Miniscoping of the Replacement Operator\<close>
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   408
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68233
diff changeset
   409
text\<open>These cover both \<^term>\<open>Replace\<close> and \<^term>\<open>Collect\<close>\<close>
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   410
lemma Rep_simps [simp]:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   411
     "{x. y \<in> 0, R(x,y)} = 0"
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   412
     "{x \<in> 0. P(x)} = 0"
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   413
     "{x \<in> A. Q} = (if Q then A else 0)"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   414
     "RepFun(0,f) = 0"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   415
     "RepFun(succ(i),f) = cons(f(i), RepFun(i,f))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   416
     "RepFun(cons(a,B),f) = cons(f(a), RepFun(B,f))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   417
by (simp_all, blast+)
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   418
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   419
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   420
subsection\<open>Miniscoping of Unions\<close>
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   421
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   422
lemma UN_simps1:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   423
     "(\<Union>x\<in>C. cons(a, B(x))) = (if C=0 then 0 else cons(a, \<Union>x\<in>C. B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   424
     "(\<Union>x\<in>C. A(x) \<union> B')   = (if C=0 then 0 else (\<Union>x\<in>C. A(x)) \<union> B')"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   425
     "(\<Union>x\<in>C. A' \<union> B(x))   = (if C=0 then 0 else A' \<union> (\<Union>x\<in>C. B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   426
     "(\<Union>x\<in>C. A(x) \<inter> B')  = ((\<Union>x\<in>C. A(x)) \<inter> B')"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   427
     "(\<Union>x\<in>C. A' \<inter> B(x))  = (A' \<inter> (\<Union>x\<in>C. B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   428
     "(\<Union>x\<in>C. A(x) - B')    = ((\<Union>x\<in>C. A(x)) - B')"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   429
     "(\<Union>x\<in>C. A' - B(x))    = (if C=0 then 0 else A' - (\<Inter>x\<in>C. B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   430
apply (simp_all add: Inter_def)
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   431
apply (blast intro!: equalityI )+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   432
done
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   433
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   434
lemma UN_simps2:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   435
      "(\<Union>x\<in>\<Union>(A). B(x)) = (\<Union>y\<in>A. \<Union>x\<in>y. B(x))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   436
      "(\<Union>z\<in>(\<Union>x\<in>A. B(x)). C(z)) = (\<Union>x\<in>A. \<Union>z\<in>B(x). C(z))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   437
      "(\<Union>x\<in>RepFun(A,f). B(x))     = (\<Union>a\<in>A. B(f(a)))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   438
by blast+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   439
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   440
lemmas UN_simps [simp] = UN_simps1 UN_simps2
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   441
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   442
text\<open>Opposite of miniscoping: pull the operator out\<close>
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   443
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   444
lemma UN_extend_simps1:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   445
     "(\<Union>x\<in>C. A(x)) \<union> B   = (if C=0 then B else (\<Union>x\<in>C. A(x) \<union> B))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   446
     "((\<Union>x\<in>C. A(x)) \<inter> B) = (\<Union>x\<in>C. A(x) \<inter> B)"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   447
     "((\<Union>x\<in>C. A(x)) - B) = (\<Union>x\<in>C. A(x) - B)"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   448
apply simp_all
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   449
apply blast+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   450
done
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   451
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   452
lemma UN_extend_simps2:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   453
     "cons(a, \<Union>x\<in>C. B(x)) = (if C=0 then {a} else (\<Union>x\<in>C. cons(a, B(x))))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   454
     "A \<union> (\<Union>x\<in>C. B(x))   = (if C=0 then A else (\<Union>x\<in>C. A \<union> B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   455
     "(A \<inter> (\<Union>x\<in>C. B(x))) = (\<Union>x\<in>C. A \<inter> B(x))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   456
     "A - (\<Inter>x\<in>C. B(x))    = (if C=0 then A else (\<Union>x\<in>C. A - B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   457
     "(\<Union>y\<in>A. \<Union>x\<in>y. B(x)) = (\<Union>x\<in>\<Union>(A). B(x))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   458
     "(\<Union>a\<in>A. B(f(a))) = (\<Union>x\<in>RepFun(A,f). B(x))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   459
apply (simp_all add: Inter_def)
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   460
apply (blast intro!: equalityI)+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   461
done
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   462
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   463
lemma UN_UN_extend:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   464
     "(\<Union>x\<in>A. \<Union>z\<in>B(x). C(z)) = (\<Union>z\<in>(\<Union>x\<in>A. B(x)). C(z))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   465
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   466
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   467
lemmas UN_extend_simps = UN_extend_simps1 UN_extend_simps2 UN_UN_extend
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   468
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   469
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   470
subsection\<open>Miniscoping of Intersections\<close>
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   471
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   472
lemma INT_simps1:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   473
     "(\<Inter>x\<in>C. A(x) \<inter> B) = (\<Inter>x\<in>C. A(x)) \<inter> B"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   474
     "(\<Inter>x\<in>C. A(x) - B)   = (\<Inter>x\<in>C. A(x)) - B"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   475
     "(\<Inter>x\<in>C. A(x) \<union> B)  = (if C=0 then 0 else (\<Inter>x\<in>C. A(x)) \<union> B)"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   476
by (simp_all add: Inter_def, blast+)
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   477
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   478
lemma INT_simps2:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   479
     "(\<Inter>x\<in>C. A \<inter> B(x)) = A \<inter> (\<Inter>x\<in>C. B(x))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   480
     "(\<Inter>x\<in>C. A - B(x))   = (if C=0 then 0 else A - (\<Union>x\<in>C. B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   481
     "(\<Inter>x\<in>C. cons(a, B(x))) = (if C=0 then 0 else cons(a, \<Inter>x\<in>C. B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   482
     "(\<Inter>x\<in>C. A \<union> B(x))  = (if C=0 then 0 else A \<union> (\<Inter>x\<in>C. B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   483
apply (simp_all add: Inter_def)
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   484
apply (blast intro!: equalityI)+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   485
done
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   486
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   487
lemmas INT_simps [simp] = INT_simps1 INT_simps2
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   488
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   489
text\<open>Opposite of miniscoping: pull the operator out\<close>
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   490
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   491
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   492
lemma INT_extend_simps1:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   493
     "(\<Inter>x\<in>C. A(x)) \<inter> B = (\<Inter>x\<in>C. A(x) \<inter> B)"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   494
     "(\<Inter>x\<in>C. A(x)) - B = (\<Inter>x\<in>C. A(x) - B)"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   495
     "(\<Inter>x\<in>C. A(x)) \<union> B  = (if C=0 then B else (\<Inter>x\<in>C. A(x) \<union> B))"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   496
apply (simp_all add: Inter_def, blast+)
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   497
done
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   498
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   499
lemma INT_extend_simps2:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   500
     "A \<inter> (\<Inter>x\<in>C. B(x)) = (\<Inter>x\<in>C. A \<inter> B(x))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   501
     "A - (\<Union>x\<in>C. B(x))   = (if C=0 then A else (\<Inter>x\<in>C. A - B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   502
     "cons(a, \<Inter>x\<in>C. B(x)) = (if C=0 then {a} else (\<Inter>x\<in>C. cons(a, B(x))))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   503
     "A \<union> (\<Inter>x\<in>C. B(x))  = (if C=0 then A else (\<Inter>x\<in>C. A \<union> B(x)))"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   504
apply (simp_all add: Inter_def)
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   505
apply (blast intro!: equalityI)+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   506
done
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   507
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   508
lemmas INT_extend_simps = INT_extend_simps1 INT_extend_simps2
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   509
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   510
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   511
subsection\<open>Other simprules\<close>
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   512
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   513
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   514
(*** Miniscoping: pushing in big Unions, Intersections, quantifiers, etc. ***)
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   515
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   516
lemma misc_simps [simp]:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   517
     "0 \<union> A = A"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   518
     "A \<union> 0 = A"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   519
     "0 \<inter> A = 0"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   520
     "A \<inter> 0 = 0"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   521
     "0 - A = 0"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   522
     "A - 0 = A"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   523
     "\<Union>(0) = 0"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   524
     "\<Union>(cons(b,A)) = b \<union> \<Union>(A)"
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45620
diff changeset
   525
     "\<Inter>({b}) = b"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   526
by blast+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   527
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 3924
diff changeset
   528
end