src/HOL/AxClasses/Lattice/LatMorph.ML
author paulson
Wed, 23 Jul 1997 11:52:22 +0200
changeset 3564 f886dbd91ee5
parent 1899 0075a8d26a80
child 4091 771b1f6422a8
permissions -rw-r--r--
Now Datatype.occs_in_prems prints the necessary warning ITSELF. It is also easier to invoke and even works if the induction variable is a parameter (rather than a free variable).


open LatMorph;


(** monotone functions vs. "&&"- / "||"-semi-morphisms **)

goalw thy [is_mono_def] "is_mono f = (ALL x y. f (x && y) [= f x && f y)";
  by (safe_tac (!claset));
  (*==> (level 1)*)
    by (stac le_inf_eq 1);
    br conjI 1;
    by (Step_tac 1);
    by (Step_tac 1);
    be mp 1;
    br inf_lb1 1;
    by (Step_tac 1);
    by (Step_tac 1);
    be mp 1;
    br inf_lb2 1;
  (*==> (level 11)*)
    br (conjI RS (le_trans RS mp)) 1;
    br inf_lb2 2;
    by (subgoal_tac "x && y = x" 1);
    be subst 1;
    by (Fast_tac 1);
    by (stac inf_connect 1);
    ba 1;
qed "mono_inf_eq";

goalw thy [is_mono_def] "is_mono f = (ALL x y. f x || f y [= f (x || y))";
  by (safe_tac (!claset));
  (*==> (level 1)*)
    by (stac ge_sup_eq 1);
    br conjI 1;
    by (Step_tac 1);
    by (Step_tac 1);
    be mp 1;
    br sup_ub1 1;
    by (Step_tac 1);
    by (Step_tac 1);
    be mp 1;
    br sup_ub2 1;
  (*==> (level 11)*)
    br (conjI RS (le_trans RS mp)) 1;
    br sup_ub1 1;
    by (subgoal_tac "x || y = y" 1);
    be subst 1;
    by (Fast_tac 1);
    by (stac sup_connect 1);
    ba 1;
qed "mono_sup_eq";