src/HOL/AxClasses/Lattice/LatPreInsts.ML
author paulson
Tue, 16 Jul 1996 15:49:46 +0200
changeset 1868 836950047d85
parent 1573 6d66b59f94a9
child 1899 0075a8d26a80
permissions -rw-r--r--
Put in minimal simpset to avoid excessive simplification, just as in revision 1.9 of HOL/indrule.ML


open LatPreInsts;


(** complete lattices **)

goal thy "is_inf x y (Inf {x, y})";
  br (bin_is_Inf_eq RS subst) 1;
  br Inf_is_Inf 1;
qed "Inf_is_inf";

goal thy "is_sup x y (Sup {x, y})";
  br (bin_is_Sup_eq RS subst) 1;
  br Sup_is_Sup 1;
qed "Sup_is_sup";



(** product lattices **)

(* pairs *)

goalw thy [is_inf_def, le_prod_def] "is_inf p q (fst p && fst q, snd p && snd q)";
  by (Simp_tac 1);
  by (safe_tac HOL_cs);
  by (REPEAT_FIRST (fn i => resolve_tac [inf_lb1, inf_lb2, inf_ub_lbs] i ORELSE atac i));
qed "prod_is_inf";

goalw thy [is_sup_def, le_prod_def] "is_sup p q (fst p || fst q, snd p || snd q)";
  by (Simp_tac 1);
  by (safe_tac HOL_cs);
  by (REPEAT_FIRST (fn i => resolve_tac [sup_ub1, sup_ub2, sup_lb_ubs] i ORELSE atac i));
qed "prod_is_sup";


(* functions *)

goalw thy [is_inf_def, le_fun_def] "is_inf f g (%x. f x && g x)";
  by (safe_tac HOL_cs);
  br inf_lb1 1;
  br inf_lb2 1;
  br inf_ub_lbs 1;
  by (REPEAT_FIRST (fast_tac HOL_cs));
qed "fun_is_inf";

goalw thy [is_sup_def, le_fun_def] "is_sup f g (%x. f x || g x)";
  by (safe_tac HOL_cs);
  br sup_ub1 1;
  br sup_ub2 1;
  br sup_lb_ubs 1;
  by (REPEAT_FIRST (fast_tac HOL_cs));
qed "fun_is_sup";



(** dual lattices **)

goalw thy [is_inf_def, le_dual_def] "is_inf x y (Abs_dual (Rep_dual x || Rep_dual y))";
  by (stac Abs_dual_inverse' 1);
  by (safe_tac HOL_cs);
  br sup_ub1 1;
  br sup_ub2 1;
  br sup_lb_ubs 1;
  ba 1;
  ba 1;
qed "dual_is_inf";

goalw thy [is_sup_def, le_dual_def] "is_sup x y (Abs_dual (Rep_dual x && Rep_dual y))";
  by (stac Abs_dual_inverse' 1);
  by (safe_tac HOL_cs);
  br inf_lb1 1;
  br inf_lb2 1;
  br inf_ub_lbs 1;
  ba 1;
  ba 1;
qed "dual_is_sup";