Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
split_weak_cong, nat_case_weak_cong, nat_rec_weak_cong. Used in llist.ML
to make simplifications faster.
HOL/gfp: re-ordered premises to put mono(f) early (first or right after
A==gfp(f) in the def_ rules). Renamed some variables in rules, A to X and
h to A. Renamed coinduct to weak_coinduct, coinduct2 to coinduct.
Strengthened coinduct as suggested by j. Frost, to have the premise X <= f(X
Un gfp(f)).
HOL/llist: used stronger coinduct rule to strengthen LList_coinduct,
LList_equalityI, llist_equalityI, llist_fun_equalityI and to delete the "2"
form of those rules. Proved List_Fun_LList_I, LListD_Fun_diag_I and
llistD_Fun_range_I to help use the new coinduction rules; most proofs
involving them required small changes. Proved prod_fun_range_eq_diag as
lemma for llist_equalityI.
(* 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();