src/ZF/mono.ML
author wenzelm
Thu, 14 Oct 1999 15:04:36 +0200
changeset 7866 3ccaa11b6df9
parent 6053 8a1059aa01f0
child 9907 473a6604da94
permissions -rw-r--r--
pdf: generate thumbnails if ISABELLE_THUMBPDF set;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 811
diff changeset
     1
(*  Title:      ZF/mono
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 811
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Monotonicity of various operations (for lattice properties see subset.ML)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
(** Replacement, in its various formulations **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
(*Not easy to express monotonicity in P, since any "bigger" predicate
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
  would have to be single-valued*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    13
Goal "A<=B ==> Replace(A,P) <= Replace(B,P)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3890
diff changeset
    14
by (blast_tac (claset() addSEs [ReplaceE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    15
qed "Replace_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    17
Goal "A<=B ==> {f(x). x:A} <= {f(x). x:B}";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    18
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    19
qed "RepFun_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    21
Goal "A<=B ==> Pow(A) <= Pow(B)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    22
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    23
qed "Pow_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    25
Goal "A<=B ==> Union(A) <= Union(B)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    26
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    27
qed "Union_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
    29
val prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
    "[| A<=C;  !!x. x:A ==> B(x)<=D(x) \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
\    |] ==> (UN x:A. B(x)) <= (UN x:C. D(x))";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3890
diff changeset
    32
by (blast_tac (claset() addIs (prems RL [subsetD])) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    33
qed "UN_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
(*Intersection is ANTI-monotonic.  There are TWO premises! *)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    36
Goal "[| A<=B;  a:A |] ==> Inter(B) <= Inter(A)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    37
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    38
qed "Inter_anti_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    40
Goal "C<=D ==> cons(a,C) <= cons(a,D)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    41
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    42
qed "cons_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    44
Goal "[| A<=C;  B<=D |] ==> A Un B <= C Un D";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    45
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    46
qed "Un_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    48
Goal "[| A<=C;  B<=D |] ==> A Int B <= C Int D";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    49
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    50
qed "Int_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    52
Goal "[| A<=C;  D<=B |] ==> A-B <= C-D";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    53
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    54
qed "Diff_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    55
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
(** Standard products, sums and function spaces **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
    58
Goal "[| A<=C;  ALL x:A. B(x) <= D(x) |] ==> Sigma(A,B) <= Sigma(C,D)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    59
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    60
qed "Sigma_mono_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
val Sigma_mono = ballI RSN (2,Sigma_mono_lemma);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    63
Goalw sum_defs "[| A<=C;  B<=D |] ==> A+B <= C+D";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
by (REPEAT (ares_tac [subset_refl,Un_mono,Sigma_mono] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    65
qed "sum_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
(*Note that B->A and C->A are typically disjoint!*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    68
Goal "B<=C ==> A->B <= A->C";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3890
diff changeset
    69
by (blast_tac (claset() addIs [lam_type] addEs [Pi_lamE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    70
qed "Pi_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    72
Goalw [lam_def] "A<=B ==> Lambda(A,c) <= Lambda(B,c)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
by (etac RepFun_mono 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    74
qed "lam_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
(** Quine-inspired ordered pairs, products, injections and sums **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    78
Goalw [QPair_def] "[| a<=c;  b<=d |] ==> <a;b> <= <c;d>";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
by (REPEAT (ares_tac [sum_mono] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    80
qed "QPair_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    82
Goal "[| A<=C;  ALL x:A. B(x) <= D(x) |] ==>  \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
\                          QSigma(A,B) <= QSigma(C,D)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    84
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    85
qed "QSigma_mono_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
val QSigma_mono = ballI RSN (2,QSigma_mono_lemma);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    88
Goalw [QInl_def] "a<=b ==> QInl(a) <= QInl(b)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    90
qed "QInl_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    92
Goalw [QInr_def] "a<=b ==> QInr(a) <= QInr(b)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    94
qed "QInr_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    96
Goal "[| A<=C;  B<=D |] ==> A <+> B <= C <+> D";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
    97
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    98
qed "qsum_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
(** Converse, domain, range, field **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   103
Goal "r<=s ==> converse(r) <= converse(s)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
   104
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   105
qed "converse_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   107
Goal "r<=s ==> domain(r)<=domain(s)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
   108
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   109
qed "domain_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
811
9bac814082e4 Used bind_thm to store domain_rel_subset and range_rel_subset
lcp
parents: 760
diff changeset
   111
bind_thm ("domain_rel_subset", [domain_mono, domain_subset] MRS subset_trans);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   112
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   113
Goal "r<=s ==> range(r)<=range(s)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
   114
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   115
qed "range_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   116
811
9bac814082e4 Used bind_thm to store domain_rel_subset and range_rel_subset
lcp
parents: 760
diff changeset
   117
bind_thm ("range_rel_subset", [range_mono, range_subset] MRS subset_trans);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   119
Goal "r<=s ==> field(r)<=field(s)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
   120
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   121
qed "field_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   123
Goal "r <= A*A ==> field(r) <= A";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
by (etac (field_mono RS subset_trans) 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
   125
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   126
qed "field_rel_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   127
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   128
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   129
(** Images **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   131
val [prem1,prem2] = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
    "[| !! x y. <x,y>:r ==> <x,y>:s;  A<=B |] ==> r``A <= s``B";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3890
diff changeset
   133
by (blast_tac (claset() addIs [prem1, prem2 RS subsetD]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   134
qed "image_pair_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   136
val [prem1,prem2] = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   137
    "[| !! x y. <x,y>:r ==> <x,y>:s;  A<=B |] ==> r-``A <= s-``B";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3890
diff changeset
   138
by (blast_tac (claset() addIs [prem1, prem2 RS subsetD]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   139
qed "vimage_pair_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   140
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   141
Goal "[| r<=s;  A<=B |] ==> r``A <= s``B";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
   142
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   143
qed "image_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   145
Goal "[| r<=s;  A<=B |] ==> r-``A <= s-``B";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
   146
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   147
qed "vimage_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   148
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   149
val [sub,PQimp] = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   150
    "[| A<=B;  !!x. x:A ==> P(x) --> Q(x) |] ==> Collect(A,P) <= Collect(B,Q)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3890
diff changeset
   151
by (blast_tac (claset() addIs [sub RS subsetD, PQimp RS mp]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   152
qed "Collect_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   153
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   154
(** Monotonicity of implications -- some could go to FOL **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   155
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   156
Goal "A<=B ==> x:A --> x:B";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2803
diff changeset
   157
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   158
qed "in_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   159
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   160
goal IFOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)";
2602
5ac837d98a85 Renamed structure Int (intuitionistic prover) to IntPr
paulson
parents: 2469
diff changeset
   161
by (IntPr.fast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   162
qed "conj_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   163
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
goal IFOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1|P2) --> (Q1|Q2)";
2602
5ac837d98a85 Renamed structure Int (intuitionistic prover) to IntPr
paulson
parents: 2469
diff changeset
   165
by (IntPr.fast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   166
qed "disj_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   167
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   168
goal IFOL.thy "!!P1 P2 Q1 Q2.[| Q1-->P1; P2-->Q2 |] ==> (P1-->P2)-->(Q1-->Q2)";
2602
5ac837d98a85 Renamed structure Int (intuitionistic prover) to IntPr
paulson
parents: 2469
diff changeset
   169
by (IntPr.fast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   170
qed "imp_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   171
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   172
goal IFOL.thy "P-->P";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
by (rtac impI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   174
by (assume_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   175
qed "imp_refl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   177
val [PQimp] = goal IFOL.thy
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3015
diff changeset
   178
    "[| !!x. P(x) --> Q(x) |] ==> (EX x. P(x)) --> (EX x. Q(x))";
2803
734fc343ec2a Conducted the IFOL proofs using intuitionistic tools
paulson
parents: 2602
diff changeset
   179
by IntPr.safe_tac;
3015
65778b9d865f Ran expandshort
paulson
parents: 2925
diff changeset
   180
by (etac (PQimp RS mp RS exI) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   181
qed "ex_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   182
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   183
val [PQimp] = goal IFOL.thy
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3015
diff changeset
   184
    "[| !!x. P(x) --> Q(x) |] ==> (ALL x. P(x)) --> (ALL x. Q(x))";
2803
734fc343ec2a Conducted the IFOL proofs using intuitionistic tools
paulson
parents: 2602
diff changeset
   185
by IntPr.safe_tac;
3015
65778b9d865f Ran expandshort
paulson
parents: 2925
diff changeset
   186
by (etac (spec RS (PQimp RS mp)) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   187
qed "all_mono";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
   188
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
   189
(*Used in intr_elim.ML and in individual datatype definitions*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
   190
val basic_monos = [subset_refl, imp_refl, disj_mono, conj_mono, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 811
diff changeset
   191
                   ex_mono, Collect_mono, Part_mono, in_mono];
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
   192