src/HOL/AxClasses/Lattice/LatMorph.ML
author wenzelm
Mon, 22 Jun 1998 17:26:46 +0200
changeset 5069 3ea049f7979d
parent 4153 e534c4c32d54
child 5711 5a1cd4b4b20e
permissions -rw-r--r--
isatool fixgoal;


open LatMorph;


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

Goalw [is_mono_def] "is_mono f = (ALL x y. f (x && y) [= f x && f y)";
  by Safe_tac;
  (*==> (level 1)*)
    by (stac le_inf_eq 1);
    by (rtac conjI 1);
    by (Step_tac 1);
    by (Step_tac 1);
    by (etac mp 1);
    by (rtac inf_lb1 1);
    by (Step_tac 1);
    by (Step_tac 1);
    by (etac mp 1);
    by (rtac inf_lb2 1);
  (*==> (level 11)*)
    by (rtac (conjI RS (le_trans RS mp)) 1);
    by (rtac inf_lb2 2);
    by (subgoal_tac "x && y = x" 1);
    by (etac subst 1);
    by (Fast_tac 1);
    by (stac inf_connect 1);
    by (assume_tac 1);
qed "mono_inf_eq";

Goalw [is_mono_def] "is_mono f = (ALL x y. f x || f y [= f (x || y))";
  by Safe_tac;
  (*==> (level 1)*)
    by (stac ge_sup_eq 1);
    by (rtac conjI 1);
    by (Step_tac 1);
    by (Step_tac 1);
    by (etac mp 1);
    by (rtac sup_ub1 1);
    by (Step_tac 1);
    by (Step_tac 1);
    by (etac mp 1);
    by (rtac sup_ub2 1);
  (*==> (level 11)*)
    by (rtac (conjI RS (le_trans RS mp)) 1);
    by (rtac sup_ub1 1);
    by (subgoal_tac "x || y = y" 1);
    by (etac subst 1);
    by (Fast_tac 1);
    by (stac sup_connect 1);
    by (assume_tac 1);
qed "mono_sup_eq";