mono.ML
author clasohm
Wed, 02 Mar 1994 12:26:55 +0100
changeset 48 21291189b51e
parent 0 7949f97df77a
child 128 89669c58e506
permissions -rw-r--r--
changed "." to "$" and Cons to infix "#" to eliminate ambiguity

(*  Title: 	HOL/mono
    ID:         $Id$
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1991  University of Cambridge

Monotonicity of various operations
*)

val [prem] = goal Set.thy
    "[| !!x. P(x) ==> Q(x) |] ==> Collect(P) <= Collect(Q)";
by (fast_tac (set_cs addIs [prem]) 1);
val Collect_mono = result();

goal Set.thy "!!A B. A<=B ==> f``A <= f``B";
by (fast_tac set_cs 1);
val image_mono = result();

goal Set.thy "!!A B. A<=B ==> Union(A) <= Union(B)";
by (fast_tac set_cs 1);
val Union_mono = result();

goal Set.thy "!!A B. B<=A ==> Inter(A) <= Inter(B)";
by (fast_tac set_cs 1);
val Inter_anti_mono = result();

val prems = goal Set.thy
    "[| A<=B;  !!x. x:A ==> f(x)<=g(x) |] ==> \
\    (UN x:A. f(x)) <= (UN x:B. g(x))";
by (fast_tac (set_cs addIs (prems RL [subsetD])) 1);
val UN_mono = result();

val [prem] = goal Set.thy
    "[| !!x. f(x)<=g(x) |] ==> (UN x. f(x)) <= (UN x. g(x))";
by (fast_tac (set_cs addIs [prem RS subsetD]) 1);
val UN1_mono = result();

val prems = goal Set.thy
    "[| B<=A;  !!x. x:A ==> f(x)<=g(x) |] ==> \
\    (INT x:A. f(x)) <= (INT x:A. g(x))";
by (fast_tac (set_cs addIs (prems RL [subsetD])) 1);
val INT_anti_mono = result();

(*The inclusion is POSITIVE! *)
val [prem] = goal Set.thy
    "[| !!x. f(x)<=g(x) |] ==> (INT x. f(x)) <= (INT x. g(x))";
by (fast_tac (set_cs addIs [prem RS subsetD]) 1);
val INT1_mono = result();

goal Set.thy "!!A B. [| A<=C;  B<=D |] ==> A Un B <= C Un D";
by (fast_tac set_cs 1);
val Un_mono = result();

goal Set.thy "!!A B. [| A<=C;  B<=D |] ==> A Int B <= C Int D";
by (fast_tac set_cs 1);
val Int_mono = result();

goal Set.thy "!!A::'a set. [| A<=C;  D<=B |] ==> A-B <= C-D";
by (fast_tac set_cs 1);
val Diff_mono = result();

goal Set.thy "!!A B. A<=B ==> Compl(B) <= Compl(A)";
by (fast_tac set_cs 1);
val Compl_anti_mono = result();

val prems = goal Prod.thy
    "[| A<=C;  !!x. x:A ==> B<=D |] ==> Sigma(A,%x.B) <= Sigma(C,%x.D)";
by (cut_facts_tac prems 1);
by (fast_tac (set_cs addIs (prems RL [subsetD]) 
                     addSIs [SigmaI] 
                     addSEs [SigmaE]) 1);
val Sigma_mono = result();