src/ZF/upair.thy
author wenzelm
Sun, 12 Feb 2006 21:34:25 +0100
changeset 19031 0059b5b195a2
parent 16417 9bc16273c2d4
child 24893 b8ef7afe3a6b
permissions -rw-r--r--
export exception SAME (for join); join: use internal modify, no option type (handle SAME/DUP instead); defined: simplified copy of lookup code, avoids allocation of option constructor; added replace, which does not change equal entries;
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
    ID:         $Id$
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 124
diff changeset
     3
    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
     4
    Copyright   1993  University of Cambridge
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
     5
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
     6
Observe the order of dependence:
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
     7
    Upair is defined in terms of Replace
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
     8
    Un is defined in terms of Upair and Union (similarly for Int)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
     9
    cons is defined in terms of Upair and Un
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    10
    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
    11
*)
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 124
diff changeset
    12
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
    13
header{*Unordered Pairs*}
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
    14
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15091
diff changeset
    15
theory upair imports ZF
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15091
diff changeset
    16
uses "Tools/typechk.ML" begin
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 3924
diff changeset
    17
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9570
diff changeset
    18
setup TypeCheck.setup
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 3924
diff changeset
    19
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    20
lemma atomize_ball [symmetric, rulify]:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    21
     "(!!x. x:A ==> P(x)) == Trueprop (ALL x:A. P(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    22
by (simp add: Ball_def atomize_all atomize_imp)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    23
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    24
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
    25
subsection{*Unordered Pairs: constant @{term Upair}*}
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    26
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    27
lemma Upair_iff [simp]: "c : Upair(a,b) <-> (c=a | c=b)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    28
by (unfold Upair_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    29
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    30
lemma UpairI1: "a : Upair(a,b)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    31
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    32
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    33
lemma UpairI2: "b : Upair(a,b)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    34
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    35
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    36
lemma UpairE: "[| a : Upair(b,c);  a=b ==> P;  a=c ==> P |] ==> P"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    37
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    38
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
    39
subsection{*Rules for Binary Union, Defined via @{term Upair}*}
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    40
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    41
lemma Un_iff [simp]: "c : A Un B <-> (c:A | c:B)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    42
apply (simp add: Un_def)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    43
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    44
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    45
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    46
lemma UnI1: "c : A ==> c : A Un B"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    47
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    48
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    49
lemma UnI2: "c : B ==> c : A Un B"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    50
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    51
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13259
diff changeset
    52
declare UnI1 [elim?]  UnI2 [elim?]
c9cfe1638bf2 improved presentation markup
paulson
parents: 13259
diff changeset
    53
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    54
lemma UnE [elim!]: "[| c : A Un B;  c:A ==> P;  c:B ==> P |] ==> P"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    55
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    56
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    57
(*Stronger version of the rule above*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    58
lemma UnE': "[| c : A Un B;  c:A ==> P;  [| c:B;  c~:A |] ==> P |] ==> P"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    59
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    60
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    61
(*Classical introduction rule: no commitment to A vs B*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    62
lemma UnCI [intro!]: "(c ~: B ==> c : A) ==> c : A Un B"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
    63
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    64
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
    65
subsection{*Rules for Binary Intersection, Defined via @{term Upair}*}
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    66
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    67
lemma Int_iff [simp]: "c : A Int B <-> (c:A & c:B)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    68
apply (unfold Int_def)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    69
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    70
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    71
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    72
lemma IntI [intro!]: "[| c : A;  c : B |] ==> c : A Int B"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    73
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    74
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    75
lemma IntD1: "c : A Int B ==> c : A"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    76
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    77
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    78
lemma IntD2: "c : A Int B ==> c : B"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    79
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    80
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    81
lemma IntE [elim!]: "[| c : A Int B;  [| c:A; c:B |] ==> P |] ==> P"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    82
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    83
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    84
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
    85
subsection{*Rules for Set Difference, Defined via @{term Upair}*}
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    86
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    87
lemma Diff_iff [simp]: "c : A-B <-> (c:A & c~:B)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    88
by (unfold Diff_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    89
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    90
lemma DiffI [intro!]: "[| c : A;  c ~: B |] ==> c : A - B"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    91
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    92
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    93
lemma DiffD1: "c : A - B ==> c : A"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    94
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    95
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    96
lemma DiffD2: "c : A - B ==> c ~: B"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    97
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    98
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
    99
lemma DiffE [elim!]: "[| c : A - B;  [| c:A; c~:B |] ==> P |] ==> P"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   100
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   101
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   102
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   103
subsection{*Rules for @{term cons}*}
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   104
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   105
lemma cons_iff [simp]: "a : cons(b,A) <-> (a=b | a:A)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   106
apply (unfold cons_def)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   107
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   108
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   109
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   110
(*risky as a typechecking rule, but solves otherwise unconstrained goals of
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   111
the form x : ?A*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   112
lemma consI1 [simp,TC]: "a : cons(a,B)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   113
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   114
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
lemma consI2: "a : B ==> a : cons(b,B)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   117
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   118
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   119
lemma consE [elim!]: "[| a : cons(b,A);  a=b ==> P;  a:A ==> P |] ==> P"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   120
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   121
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   122
(*Stronger version of the rule above*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   123
lemma consE':
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   124
    "[| a : cons(b,A);  a=b ==> P;  [| a:A;  a~=b |] ==> P |] ==> P"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   125
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   126
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   127
(*Classical introduction rule*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   128
lemma consCI [intro!]: "(a~:B ==> a=b) ==> a: cons(b,B)"
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   129
by (simp, blast)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   130
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   131
lemma cons_not_0 [simp]: "cons(a,B) ~= 0"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   132
by (blast elim: equalityE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   133
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   134
lemmas cons_neq_0 = cons_not_0 [THEN notE, standard]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   135
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   136
declare cons_not_0 [THEN not_sym, simp]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   137
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   138
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   139
subsection{*Singletons*}
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   140
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   141
lemma singleton_iff: "a : {b} <-> a=b"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   142
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   143
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   144
lemma singletonI [intro!]: "a : {a}"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   145
by (rule consI1)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   146
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   147
lemmas singletonE = singleton_iff [THEN iffD1, elim_format, standard, elim!]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   148
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   149
14883
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14153
diff changeset
   150
subsection{*Descriptions*}
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   151
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   152
lemma the_equality [intro]:
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   153
    "[| P(a);  !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   154
apply (unfold the_def) 
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   155
apply (fast dest: subst)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   156
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   157
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   158
(* Only use this if you already know EX!x. P(x) *)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   159
lemma the_equality2: "[| EX! x. P(x);  P(a) |] ==> (THE x. P(x)) = a"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   160
by blast
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   161
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   162
lemma theI: "EX! x. P(x) ==> P(THE x. P(x))"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   163
apply (erule ex1E)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   164
apply (subst the_equality)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   165
apply (blast+)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   166
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   167
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   168
(*the_cong is no longer necessary: if (ALL y.P(y)<->Q(y)) then 
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   169
  (THE x.P(x))  rewrites to  (THE x. Q(x))  *)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   170
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   171
(*If it's "undefined", it's zero!*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   172
lemma the_0: "~ (EX! x. P(x)) ==> (THE x. P(x))=0"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   173
apply (unfold the_def)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   174
apply (blast elim!: ReplaceE)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   175
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   176
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   177
(*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
   178
lemma theI2:
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   179
    assumes p1: "~ Q(0) ==> EX! x. P(x)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   180
        and p2: "!!x. P(x) ==> Q(x)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   181
    shows "Q(THE x. P(x))"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   182
apply (rule classical)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   183
apply (rule p2)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   184
apply (rule theI)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   185
apply (rule classical)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   186
apply (rule p1)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   187
apply (erule the_0 [THEN subst], assumption)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   188
done
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   189
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   190
lemma the_eq_trivial [simp]: "(THE x. x = a) = a"
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   191
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   192
13544
895994073bdf various new lemmas for Constructible
paulson
parents: 13357
diff changeset
   193
lemma the_eq_trivial2 [simp]: "(THE x. a = x) = a"
895994073bdf various new lemmas for Constructible
paulson
parents: 13357
diff changeset
   194
by blast
895994073bdf various new lemmas for Constructible
paulson
parents: 13357
diff changeset
   195
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   196
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   197
subsection{*Conditional Terms: @{text "if-then-else"}*}
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   198
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   199
lemma if_true [simp]: "(if True then a else b) = a"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   200
by (unfold if_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   201
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   202
lemma if_false [simp]: "(if False then a else b) = b"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   203
by (unfold if_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   204
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   205
(*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
   206
lemma if_cong:
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   207
    "[| P<->Q;  Q ==> a=c;  ~Q ==> b=d |]  
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   208
     ==> (if P then a else b) = (if Q then c else d)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   209
by (simp add: if_def cong add: conj_cong)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   210
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   211
(*Prevents simplification of x and y: faster and allows the execution
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   212
  of functional programs. NOW THE DEFAULT.*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   213
lemma if_weak_cong: "P<->Q ==> (if P then x else y) = (if Q then x else y)"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   214
by simp
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   215
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   216
(*Not needed for rewriting, since P would rewrite to True anyway*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   217
lemma if_P: "P ==> (if P then a else b) = a"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   218
by (unfold if_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   219
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   220
(*Not needed for rewriting, since P would rewrite to False anyway*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   221
lemma if_not_P: "~P ==> (if P then a else b) = b"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   222
by (unfold if_def, blast)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   223
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   224
lemma split_if [split]:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   225
     "P(if Q then x else y) <-> ((Q --> P(x)) & (~Q --> P(y)))"
14153
76a6ba67bd15 new case_tac
paulson
parents: 13780
diff changeset
   226
by (case_tac Q, simp_all)
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   227
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   228
(** Rewrite rules for boolean case-splitting: faster than 
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   229
	addsplits[split_if]
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
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   232
lemmas split_if_eq1 = split_if [of "%x. x = b", standard]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   233
lemmas split_if_eq2 = split_if [of "%x. a = x", standard]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   234
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   235
lemmas split_if_mem1 = split_if [of "%x. x : b", standard]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   236
lemmas split_if_mem2 = split_if [of "%x. a : x", standard]
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*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   241
lemma if_iff: "a: (if P then x else y) <-> P & a:x | ~P & a: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]:
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   245
    "[| P ==> a: A;  ~P ==> b: A |] ==> (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
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   250
lemma split_if_asm: "P(if Q then x else y) <-> (~((Q & ~P(x)) | (~Q & ~P(y))))"
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
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   256
subsection{*Consequences of Foundation*}
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*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   259
lemma mem_asym: "[| a:b;  ~P ==> b:a |] ==> P"
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*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   266
lemma mem_irrefl: "a:a ==> P"
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
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   272
lemma mem_not_refl: "a ~: a"
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*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   278
lemma mem_imp_not_eq: "a:A ==> a ~= A"
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
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   281
lemma eq_imp_not_mem: "a=A ==> a ~: A"
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
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
   284
subsection{*Rules for Successor*}
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   285
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   286
lemma succ_iff: "i : succ(j) <-> i=j | i:j"
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
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   289
lemma succI1 [simp]: "i : succ(i)"
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
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   292
lemma succI2: "i : j ==> i : succ(j)"
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
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   295
lemma succE [elim!]: 
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   296
    "[| i : succ(j);  i=j ==> P;  i:j ==> P |] ==> P"
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   297
apply (simp add: succ_iff, blast) 
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*)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   301
lemma succCI [intro!]: "(i~:j ==> i=j) ==> i: succ(j)"
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
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   304
lemma succ_not_0 [simp]: "succ(n) ~= 0"
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
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   307
lemmas succ_neq_0 = succ_not_0 [THEN notE, standard, elim!]
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
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   312
(* succ(c) <= B ==> c : B *)
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
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   315
(* succ(b) ~= b *)
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   316
lemmas succ_neq_self = succI1 [THEN mem_imp_not_eq, THEN not_sym, standard]
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   317
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   318
lemma succ_inject_iff [simp]: "succ(m) = succ(n) <-> m=n"
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
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   321
lemmas succ_inject = succ_inject_iff [THEN iffD1, standard, dest!]
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
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   324
subsection{*Miniscoping of the Bounded Universal Quantifier*}
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   327
     "(ALL x:A. P(x) & Q)   <-> (ALL x:A. P(x)) & (A=0 | Q)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   328
     "(ALL x:A. P(x) | Q)   <-> ((ALL x:A. P(x)) | Q)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   329
     "(ALL x:A. P(x) --> Q) <-> ((EX x:A. P(x)) --> Q)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   330
     "(~(ALL x:A. P(x))) <-> (EX x:A. ~P(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   331
     "(ALL x:0.P(x)) <-> True"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   332
     "(ALL x:succ(i).P(x)) <-> P(i) & (ALL x:i. P(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   333
     "(ALL x:cons(a,B).P(x)) <-> P(a) & (ALL x:B. P(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   334
     "(ALL x:RepFun(A,f). P(x)) <-> (ALL y:A. P(f(y)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   335
     "(ALL x:Union(A).P(x)) <-> (ALL y:A. ALL x:y. P(x))" 
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   339
     "(ALL x:A. P & Q(x))   <-> (A=0 | P) & (ALL x:A. Q(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   340
     "(ALL x:A. P | Q(x))   <-> (P | (ALL x:A. Q(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   341
     "(ALL x:A. P --> Q(x)) <-> (P --> (ALL x:A. Q(x)))"
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   345
     "(ALL x:Collect(A,Q).P(x)) <-> (ALL x:A. Q(x) --> P(x))"
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   351
    "(ALL x:A. P(x) & Q(x)) <-> ((ALL x:A. P(x)) & (ALL x:A. Q(x)))"
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
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   355
subsection{*Miniscoping of the Bounded Existential Quantifier*}
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   358
     "(EX x:A. P(x) & Q) <-> ((EX x:A. P(x)) & Q)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   359
     "(EX x:A. P(x) | Q) <-> (EX x:A. P(x)) | (A~=0 & Q)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   360
     "(EX x:A. P(x) --> Q) <-> ((ALL x:A. P(x)) --> (A~=0 & Q))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   361
     "(EX x:0.P(x)) <-> False"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   362
     "(EX x:succ(i).P(x)) <-> P(i) | (EX x:i. P(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   363
     "(EX x:cons(a,B).P(x)) <-> P(a) | (EX x:B. P(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   364
     "(EX x:RepFun(A,f). P(x)) <-> (EX y:A. P(f(y)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   365
     "(EX x:Union(A).P(x)) <-> (EX y:A. EX x:y.  P(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   366
     "(~(EX x:A. P(x))) <-> (ALL x:A. ~P(x))"
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   370
     "(EX x:A. P & Q(x)) <-> (P & (EX x:A. Q(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   371
     "(EX x:A. P | Q(x)) <-> (A~=0 & P) | (EX x:A. Q(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   372
     "(EX x:A. P --> Q(x)) <-> ((A=0 | P) --> (EX x:A. Q(x)))"
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   376
     "(EX x:Collect(A,Q).P(x)) <-> (EX x:A. Q(x) & P(x))"
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   382
    "(EX x:A. P(x) | Q(x)) <-> ((EX x:A. P(x)) | (EX x:A. Q(x)))"
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
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   388
lemma bex_triv_one_point1 [simp]: "(EX x:A. x=a) <-> (a:A)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   389
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   390
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   391
lemma bex_triv_one_point2 [simp]: "(EX x:A. a=x) <-> (a:A)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   392
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   393
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   394
lemma bex_one_point1 [simp]: "(EX x:A. x=a & P(x)) <-> (a:A & P(a))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   395
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   396
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   397
lemma bex_one_point2 [simp]: "(EX x:A. a=x & P(x)) <-> (a:A & P(a))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   398
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   399
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   400
lemma ball_one_point1 [simp]: "(ALL x:A. x=a --> P(x)) <-> (a:A --> P(a))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   401
by blast
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   402
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   403
lemma ball_one_point2 [simp]: "(ALL x:A. a=x --> P(x)) <-> (a:A --> P(a))"
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
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   407
subsection{*Miniscoping of the Replacement Operator*}
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   408
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   409
text{*These cover both @{term Replace} and @{term Collect}*}
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   410
lemma Rep_simps [simp]:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   411
     "{x. y:0, R(x,y)} = 0"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   412
     "{x:0. P(x)} = 0"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   413
     "{x:A. Q} = (if Q then A else 0)"
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
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   420
subsection{*Miniscoping of Unions*}
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   423
     "(UN x:C. cons(a, B(x))) = (if C=0 then 0 else cons(a, UN x:C. B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   424
     "(UN x:C. A(x) Un B')   = (if C=0 then 0 else (UN x:C. A(x)) Un B')"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   425
     "(UN x:C. A' Un B(x))   = (if C=0 then 0 else A' Un (UN x:C. B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   426
     "(UN x:C. A(x) Int B')  = ((UN x:C. A(x)) Int B')"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   427
     "(UN x:C. A' Int B(x))  = (A' Int (UN x:C. B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   428
     "(UN x:C. A(x) - B')    = ((UN x:C. A(x)) - B')"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   429
     "(UN x:C. A' - B(x))    = (if C=0 then 0 else A' - (INT x:C. B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   430
apply (simp_all add: Inter_def) 
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   435
      "(UN x: Union(A). B(x)) = (UN y:A. UN x:y. B(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   436
      "(UN z: (UN x:A. B(x)). C(z)) = (UN  x:A. UN z: B(x). C(z))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   437
      "(UN x: RepFun(A,f). B(x))     = (UN a:A. B(f(a)))"
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
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   442
text{*Opposite of miniscoping: pull the operator out*}
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   445
     "(UN x:C. A(x)) Un B   = (if C=0 then B else (UN x:C. A(x) Un B))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   446
     "((UN x:C. A(x)) Int B) = (UN x:C. A(x) Int B)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   447
     "((UN x:C. A(x)) - B) = (UN x:C. A(x) - B)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   448
apply simp_all 
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   453
     "cons(a, UN x:C. B(x)) = (if C=0 then {a} else (UN x:C. cons(a, B(x))))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   454
     "A Un (UN x:C. B(x))   = (if C=0 then A else (UN x:C. A Un B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   455
     "(A Int (UN x:C. B(x))) = (UN x:C. A Int B(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   456
     "A - (INT x:C. B(x))    = (if C=0 then A else (UN x:C. A - B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   457
     "(UN y:A. UN x:y. B(x)) = (UN x: Union(A). B(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   458
     "(UN a:A. B(f(a))) = (UN x: RepFun(A,f). B(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   459
apply (simp_all add: Inter_def) 
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   464
     "(UN  x:A. UN z: B(x). C(z)) = (UN z: (UN x:A. B(x)). C(z))"
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
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   470
subsection{*Miniscoping of Intersections*}
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   473
     "(INT x:C. A(x) Int B) = (INT x:C. A(x)) Int B"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   474
     "(INT x:C. A(x) - B)   = (INT x:C. A(x)) - B"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   475
     "(INT x:C. A(x) Un B)  = (if C=0 then 0 else (INT x:C. A(x)) Un B)"
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   479
     "(INT x:C. A Int B(x)) = A Int (INT x:C. B(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   480
     "(INT x:C. A - B(x))   = (if C=0 then 0 else A - (UN x:C. B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   481
     "(INT x:C. cons(a, B(x))) = (if C=0 then 0 else cons(a, INT x:C. B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   482
     "(INT x:C. A Un B(x))  = (if C=0 then 0 else A Un (INT x:C. B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   483
apply (simp_all add: Inter_def) 
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
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   489
text{*Opposite of miniscoping: pull the operator out*}
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   493
     "(INT x:C. A(x)) Int B = (INT x:C. A(x) Int B)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   494
     "(INT x:C. A(x)) - B = (INT x:C. A(x) - B)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   495
     "(INT x:C. A(x)) Un B  = (if C=0 then B else (INT x:C. A(x) Un B))"
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:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   500
     "A Int (INT x:C. B(x)) = (INT x:C. A Int B(x))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   501
     "A - (UN x:C. B(x))   = (if C=0 then A else (INT x:C. A - B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   502
     "cons(a, INT x:C. B(x)) = (if C=0 then {a} else (INT x:C. cons(a, B(x))))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   503
     "A Un (INT x:C. B(x))  = (if C=0 then A else (INT x:C. A Un B(x)))"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   504
apply (simp_all add: Inter_def) 
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
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   511
subsection{*Other simprules*}
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]:
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   517
     "0 Un A = A"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   518
     "A Un 0 = A"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   519
     "0 Int A = 0"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   520
     "A Int 0 = 0"
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"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   523
     "Union(0) = 0"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   524
     "Union(cons(b,A)) = b Un Union(A)"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   525
     "Inter({b}) = b"
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   526
by blast+
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   527
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   528
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   529
ML
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   530
{*
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   531
val Upair_iff = thm "Upair_iff";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   532
val UpairI1 = thm "UpairI1";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   533
val UpairI2 = thm "UpairI2";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   534
val UpairE = thm "UpairE";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   535
val Un_iff = thm "Un_iff";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   536
val UnI1 = thm "UnI1";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   537
val UnI2 = thm "UnI2";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   538
val UnE = thm "UnE";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   539
val UnE' = thm "UnE'";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   540
val UnCI = thm "UnCI";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   541
val Int_iff = thm "Int_iff";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   542
val IntI = thm "IntI";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   543
val IntD1 = thm "IntD1";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   544
val IntD2 = thm "IntD2";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   545
val IntE = thm "IntE";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   546
val Diff_iff = thm "Diff_iff";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   547
val DiffI = thm "DiffI";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   548
val DiffD1 = thm "DiffD1";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   549
val DiffD2 = thm "DiffD2";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   550
val DiffE = thm "DiffE";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   551
val cons_iff = thm "cons_iff";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   552
val consI1 = thm "consI1";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   553
val consI2 = thm "consI2";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   554
val consE = thm "consE";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   555
val consE' = thm "consE'";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   556
val consCI = thm "consCI";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   557
val cons_not_0 = thm "cons_not_0";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   558
val cons_neq_0 = thm "cons_neq_0";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   559
val singleton_iff = thm "singleton_iff";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   560
val singletonI = thm "singletonI";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   561
val singletonE = thm "singletonE";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   562
val the_equality = thm "the_equality";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   563
val the_equality2 = thm "the_equality2";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   564
val theI = thm "theI";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   565
val the_0 = thm "the_0";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   566
val theI2 = thm "theI2";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   567
val if_true = thm "if_true";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   568
val if_false = thm "if_false";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   569
val if_cong = thm "if_cong";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   570
val if_weak_cong = thm "if_weak_cong";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   571
val if_P = thm "if_P";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   572
val if_not_P = thm "if_not_P";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   573
val split_if = thm "split_if";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   574
val split_if_eq1 = thm "split_if_eq1";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   575
val split_if_eq2 = thm "split_if_eq2";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   576
val split_if_mem1 = thm "split_if_mem1";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   577
val split_if_mem2 = thm "split_if_mem2";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   578
val if_iff = thm "if_iff";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   579
val if_type = thm "if_type";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   580
val mem_asym = thm "mem_asym";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   581
val mem_irrefl = thm "mem_irrefl";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   582
val mem_not_refl = thm "mem_not_refl";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   583
val mem_imp_not_eq = thm "mem_imp_not_eq";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   584
val succ_iff = thm "succ_iff";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   585
val succI1 = thm "succI1";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   586
val succI2 = thm "succI2";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   587
val succE = thm "succE";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   588
val succCI = thm "succCI";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   589
val succ_not_0 = thm "succ_not_0";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   590
val succ_neq_0 = thm "succ_neq_0";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   591
val succ_subsetD = thm "succ_subsetD";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   592
val succ_neq_self = thm "succ_neq_self";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   593
val succ_inject_iff = thm "succ_inject_iff";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   594
val succ_inject = thm "succ_inject";
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   595
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   596
val split_ifs = thms "split_ifs";
13780
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   597
val ball_simps = thms "ball_simps";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   598
val bex_simps = thms "bex_simps";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   599
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   600
val ball_conj_distrib = thm "ball_conj_distrib";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   601
val bex_disj_distrib = thm "bex_disj_distrib";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   602
val bex_triv_one_point1 = thm "bex_triv_one_point1";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   603
val bex_triv_one_point2 = thm "bex_triv_one_point2";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   604
val bex_one_point1 = thm "bex_one_point1";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   605
val bex_one_point2 = thm "bex_one_point2";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   606
val ball_one_point1 = thm "ball_one_point1";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   607
val ball_one_point2 = thm "ball_one_point2";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   608
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   609
val Rep_simps = thms "Rep_simps";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   610
val misc_simps = thms "misc_simps";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   611
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   612
val UN_simps = thms "UN_simps";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   613
val INT_simps = thms "INT_simps";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   614
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   615
val UN_extend_simps = thms "UN_extend_simps";
af7b79271364 more new-style theories
paulson
parents: 13544
diff changeset
   616
val INT_extend_simps = thms "INT_extend_simps";
13259
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   617
*}
01fa0c8dbc92 conversion of many files to Isar format
paulson
parents: 11770
diff changeset
   618
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 3924
diff changeset
   619
end