diff -r 260a9711f507 -r 0075a8d26a80 src/HOL/AxClasses/Lattice/LatPreInsts.ML --- a/src/HOL/AxClasses/Lattice/LatPreInsts.ML Fri Aug 02 12:16:11 1996 +0200 +++ b/src/HOL/AxClasses/Lattice/LatPreInsts.ML Fri Aug 02 12:25:26 1996 +0200 @@ -22,13 +22,13 @@ 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 (safe_tac (!claset)); 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 (safe_tac (!claset)); by (REPEAT_FIRST (fn i => resolve_tac [sup_ub1, sup_ub2, sup_lb_ubs] i ORELSE atac i)); qed "prod_is_sup"; @@ -36,19 +36,19 @@ (* functions *) goalw thy [is_inf_def, le_fun_def] "is_inf f g (%x. f x && g x)"; - by (safe_tac HOL_cs); + by (safe_tac (!claset)); br inf_lb1 1; br inf_lb2 1; br inf_ub_lbs 1; - by (REPEAT_FIRST (fast_tac HOL_cs)); + by (REPEAT_FIRST (Fast_tac)); 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); + by (safe_tac (!claset)); br sup_ub1 1; br sup_ub2 1; br sup_lb_ubs 1; - by (REPEAT_FIRST (fast_tac HOL_cs)); + by (REPEAT_FIRST (Fast_tac)); qed "fun_is_sup"; @@ -57,7 +57,7 @@ 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); + by (safe_tac (!claset)); br sup_ub1 1; br sup_ub2 1; br sup_lb_ubs 1; @@ -67,7 +67,7 @@ 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); + by (safe_tac (!claset)); br inf_lb1 1; br inf_lb2 1; br inf_ub_lbs 1;