diff -r f552152d7747 -r 46cfc4b8112e src/HOL/Decision_Procs/Dense_Linear_Order.thy --- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy Wed Mar 17 19:37:44 2010 +0100 +++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy Thu Mar 18 12:58:52 2010 +0100 @@ -19,9 +19,9 @@ context linorder begin -lemma less_not_permute[noatp]: "\ (x < y \ y < x)" by (simp add: not_less linear) +lemma less_not_permute[no_atp]: "\ (x < y \ y < x)" by (simp add: not_less linear) -lemma gather_simps[noatp]: +lemma gather_simps[no_atp]: shows "(\x. (\y \ L. y < x) \ (\y \ U. x < y) \ x < u \ P x) \ (\x. (\y \ L. y < x) \ (\y \ (insert u U). x < y) \ P x)" and "(\x. (\y \ L. y < x) \ (\y \ U. x < y) \ l < x \ P x) \ (\x. (\y \ (insert l L). y < x) \ (\y \ U. x < y) \ P x)" @@ -29,61 +29,61 @@ and "(\x. (\y \ L. y < x) \ (\y \ U. x < y) \ l < x) \ (\x. (\y \ (insert l L). y < x) \ (\y \ U. x < y))" by auto lemma - gather_start[noatp]: "(\x. P x) \ (\x. (\y \ {}. y < x) \ (\y\ {}. x < y) \ P x)" + gather_start[no_atp]: "(\x. P x) \ (\x. (\y \ {}. y < x) \ (\y\ {}. x < y) \ P x)" by simp text{* Theorems for @{text "\z. \x. x < z \ (P x \ P\<^bsub>-\\<^esub>)"}*} -lemma minf_lt[noatp]: "\z . \x. x < z \ (x < t \ True)" by auto -lemma minf_gt[noatp]: "\z . \x. x < z \ (t < x \ False)" +lemma minf_lt[no_atp]: "\z . \x. x < z \ (x < t \ True)" by auto +lemma minf_gt[no_atp]: "\z . \x. x < z \ (t < x \ False)" by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le) -lemma minf_le[noatp]: "\z. \x. x < z \ (x \ t \ True)" by (auto simp add: less_le) -lemma minf_ge[noatp]: "\z. \x. x < z \ (t \ x \ False)" +lemma minf_le[no_atp]: "\z. \x. x < z \ (x \ t \ True)" by (auto simp add: less_le) +lemma minf_ge[no_atp]: "\z. \x. x < z \ (t \ x \ False)" by (auto simp add: less_le not_less not_le) -lemma minf_eq[noatp]: "\z. \x. x < z \ (x = t \ False)" by auto -lemma minf_neq[noatp]: "\z. \x. x < z \ (x \ t \ True)" by auto -lemma minf_P[noatp]: "\z. \x. x < z \ (P \ P)" by blast +lemma minf_eq[no_atp]: "\z. \x. x < z \ (x = t \ False)" by auto +lemma minf_neq[no_atp]: "\z. \x. x < z \ (x \ t \ True)" by auto +lemma minf_P[no_atp]: "\z. \x. x < z \ (P \ P)" by blast text{* Theorems for @{text "\z. \x. x < z \ (P x \ P\<^bsub>+\\<^esub>)"}*} -lemma pinf_gt[noatp]: "\z . \x. z < x \ (t < x \ True)" by auto -lemma pinf_lt[noatp]: "\z . \x. z < x \ (x < t \ False)" +lemma pinf_gt[no_atp]: "\z . \x. z < x \ (t < x \ True)" by auto +lemma pinf_lt[no_atp]: "\z . \x. z < x \ (x < t \ False)" by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le) -lemma pinf_ge[noatp]: "\z. \x. z < x \ (t \ x \ True)" by (auto simp add: less_le) -lemma pinf_le[noatp]: "\z. \x. z < x \ (x \ t \ False)" +lemma pinf_ge[no_atp]: "\z. \x. z < x \ (t \ x \ True)" by (auto simp add: less_le) +lemma pinf_le[no_atp]: "\z. \x. z < x \ (x \ t \ False)" by (auto simp add: less_le not_less not_le) -lemma pinf_eq[noatp]: "\z. \x. z < x \ (x = t \ False)" by auto -lemma pinf_neq[noatp]: "\z. \x. z < x \ (x \ t \ True)" by auto -lemma pinf_P[noatp]: "\z. \x. z < x \ (P \ P)" by blast +lemma pinf_eq[no_atp]: "\z. \x. z < x \ (x = t \ False)" by auto +lemma pinf_neq[no_atp]: "\z. \x. z < x \ (x \ t \ True)" by auto +lemma pinf_P[no_atp]: "\z. \x. z < x \ (P \ P)" by blast -lemma nmi_lt[noatp]: "t \ U \ \x. \True \ x < t \ (\ u\ U. u \ x)" by auto -lemma nmi_gt[noatp]: "t \ U \ \x. \False \ t < x \ (\ u\ U. u \ x)" +lemma nmi_lt[no_atp]: "t \ U \ \x. \True \ x < t \ (\ u\ U. u \ x)" by auto +lemma nmi_gt[no_atp]: "t \ U \ \x. \False \ t < x \ (\ u\ U. u \ x)" by (auto simp add: le_less) -lemma nmi_le[noatp]: "t \ U \ \x. \True \ x\ t \ (\ u\ U. u \ x)" by auto -lemma nmi_ge[noatp]: "t \ U \ \x. \False \ t\ x \ (\ u\ U. u \ x)" by auto -lemma nmi_eq[noatp]: "t \ U \ \x. \False \ x = t \ (\ u\ U. u \ x)" by auto -lemma nmi_neq[noatp]: "t \ U \\x. \True \ x \ t \ (\ u\ U. u \ x)" by auto -lemma nmi_P[noatp]: "\ x. ~P \ P \ (\ u\ U. u \ x)" by auto -lemma nmi_conj[noatp]: "\\x. \P1' \ P1 x \ (\ u\ U. u \ x) ; +lemma nmi_le[no_atp]: "t \ U \ \x. \True \ x\ t \ (\ u\ U. u \ x)" by auto +lemma nmi_ge[no_atp]: "t \ U \ \x. \False \ t\ x \ (\ u\ U. u \ x)" by auto +lemma nmi_eq[no_atp]: "t \ U \ \x. \False \ x = t \ (\ u\ U. u \ x)" by auto +lemma nmi_neq[no_atp]: "t \ U \\x. \True \ x \ t \ (\ u\ U. u \ x)" by auto +lemma nmi_P[no_atp]: "\ x. ~P \ P \ (\ u\ U. u \ x)" by auto +lemma nmi_conj[no_atp]: "\\x. \P1' \ P1 x \ (\ u\ U. u \ x) ; \x. \P2' \ P2 x \ (\ u\ U. u \ x)\ \ \x. \(P1' \ P2') \ (P1 x \ P2 x) \ (\ u\ U. u \ x)" by auto -lemma nmi_disj[noatp]: "\\x. \P1' \ P1 x \ (\ u\ U. u \ x) ; +lemma nmi_disj[no_atp]: "\\x. \P1' \ P1 x \ (\ u\ U. u \ x) ; \x. \P2' \ P2 x \ (\ u\ U. u \ x)\ \ \x. \(P1' \ P2') \ (P1 x \ P2 x) \ (\ u\ U. u \ x)" by auto -lemma npi_lt[noatp]: "t \ U \ \x. \False \ x < t \ (\ u\ U. x \ u)" by (auto simp add: le_less) -lemma npi_gt[noatp]: "t \ U \ \x. \True \ t < x \ (\ u\ U. x \ u)" by auto -lemma npi_le[noatp]: "t \ U \ \x. \False \ x \ t \ (\ u\ U. x \ u)" by auto -lemma npi_ge[noatp]: "t \ U \ \x. \True \ t \ x \ (\ u\ U. x \ u)" by auto -lemma npi_eq[noatp]: "t \ U \ \x. \False \ x = t \ (\ u\ U. x \ u)" by auto -lemma npi_neq[noatp]: "t \ U \ \x. \True \ x \ t \ (\ u\ U. x \ u )" by auto -lemma npi_P[noatp]: "\ x. ~P \ P \ (\ u\ U. x \ u)" by auto -lemma npi_conj[noatp]: "\\x. \P1' \ P1 x \ (\ u\ U. x \ u) ; \x. \P2' \ P2 x \ (\ u\ U. x \ u)\ +lemma npi_lt[no_atp]: "t \ U \ \x. \False \ x < t \ (\ u\ U. x \ u)" by (auto simp add: le_less) +lemma npi_gt[no_atp]: "t \ U \ \x. \True \ t < x \ (\ u\ U. x \ u)" by auto +lemma npi_le[no_atp]: "t \ U \ \x. \False \ x \ t \ (\ u\ U. x \ u)" by auto +lemma npi_ge[no_atp]: "t \ U \ \x. \True \ t \ x \ (\ u\ U. x \ u)" by auto +lemma npi_eq[no_atp]: "t \ U \ \x. \False \ x = t \ (\ u\ U. x \ u)" by auto +lemma npi_neq[no_atp]: "t \ U \ \x. \True \ x \ t \ (\ u\ U. x \ u )" by auto +lemma npi_P[no_atp]: "\ x. ~P \ P \ (\ u\ U. x \ u)" by auto +lemma npi_conj[no_atp]: "\\x. \P1' \ P1 x \ (\ u\ U. x \ u) ; \x. \P2' \ P2 x \ (\ u\ U. x \ u)\ \ \x. \(P1' \ P2') \ (P1 x \ P2 x) \ (\ u\ U. x \ u)" by auto -lemma npi_disj[noatp]: "\\x. \P1' \ P1 x \ (\ u\ U. x \ u) ; \x. \P2' \ P2 x \ (\ u\ U. x \ u)\ +lemma npi_disj[no_atp]: "\\x. \P1' \ P1 x \ (\ u\ U. x \ u) ; \x. \P2' \ P2 x \ (\ u\ U. x \ u)\ \ \x. \(P1' \ P2') \ (P1 x \ P2 x) \ (\ u\ U. x \ u)" by auto -lemma lin_dense_lt[noatp]: "t \ U \ \x l u. (\ t. l < t \ t < u \ t \ U) \ l< x \ x < u \ x < t \ (\ y. l < y \ y < u \ y < t)" +lemma lin_dense_lt[no_atp]: "t \ U \ \x l u. (\ t. l < t \ t < u \ t \ U) \ l< x \ x < u \ x < t \ (\ y. l < y \ y < u \ y < t)" proof(clarsimp) fix x l u y assume tU: "t \ U" and noU: "\t. l < t \ t < u \ t \ U" and lx: "l < x" and xu: "x U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l < x \ x < u \ t < x \ (\ y. l < y \ y < u \ t < y)" +lemma lin_dense_gt[no_atp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l < x \ x < u \ t < x \ (\ y. l < y \ y < u \ t < y)" proof(clarsimp) fix x l u y assume tU: "t \ U" and noU: "\t. l < t \ t < u \ t \ U" and lx: "l < x" and xu: "x U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ x \ t \ (\ y. l < y \ y < u \ y\ t)" +lemma lin_dense_le[no_atp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ x \ t \ (\ y. l < y \ y < u \ y\ t)" proof(clarsimp) fix x l u y assume tU: "t \ U" and noU: "\t. l < t \ t < u \ t \ U" and lx: "l < x" and xu: "x t < y" by auto thus "y \ t" by (simp add: not_less) qed -lemma lin_dense_ge[noatp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ t \ x \ (\ y. l < y \ y < u \ t \ y)" +lemma lin_dense_ge[no_atp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ t \ x \ (\ y. l < y \ y < u \ t \ y)" proof(clarsimp) fix x l u y assume tU: "t \ U" and noU: "\t. l < t \ t < u \ t \ U" and lx: "l < x" and xu: "x y y" by (simp add: not_less) qed -lemma lin_dense_eq[noatp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ x = t \ (\ y. l < y \ y < u \ y= t)" by auto -lemma lin_dense_neq[noatp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ x \ t \ (\ y. l < y \ y < u \ y\ t)" by auto -lemma lin_dense_P[noatp]: "\x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P \ (\ y. l < y \ y < u \ P)" by auto +lemma lin_dense_eq[no_atp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ x = t \ (\ y. l < y \ y < u \ y= t)" by auto +lemma lin_dense_neq[no_atp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ x \ t \ (\ y. l < y \ y < u \ y\ t)" by auto +lemma lin_dense_P[no_atp]: "\x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P \ (\ y. l < y \ y < u \ P)" by auto -lemma lin_dense_conj[noatp]: +lemma lin_dense_conj[no_atp]: "\\x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P1 x \ (\ y. l < y \ y < u \ P1 y) ; \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P2 x @@ -146,7 +146,7 @@ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ (P1 x \ P2 x) \ (\ y. l < y \ y < u \ (P1 y \ P2 y))" by blast -lemma lin_dense_disj[noatp]: +lemma lin_dense_disj[no_atp]: "\\x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P1 x \ (\ y. l < y \ y < u \ P1 y) ; \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P2 x @@ -155,11 +155,11 @@ \ (\ y. l < y \ y < u \ (P1 y \ P2 y))" by blast -lemma npmibnd[noatp]: "\\x. \ MP \ P x \ (\ u\ U. u \ x); \x. \PP \ P x \ (\ u\ U. x \ u)\ +lemma npmibnd[no_atp]: "\\x. \ MP \ P x \ (\ u\ U. u \ x); \x. \PP \ P x \ (\ u\ U. x \ u)\ \ \x. \ MP \ \PP \ P x \ (\ u\ U. \ u' \ U. u \ x \ x \ u')" by auto -lemma finite_set_intervals[noatp]: +lemma finite_set_intervals[no_atp]: assumes px: "P x" and lx: "l \ x" and xu: "x \ u" and linS: "l\ S" and uinS: "u \ S" and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" shows "\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a \ x \ x \ b \ P x" @@ -191,7 +191,7 @@ from ainS binS noy ax xb px show ?thesis by blast qed -lemma finite_set_intervals2[noatp]: +lemma finite_set_intervals2[no_atp]: assumes px: "P x" and lx: "l \ x" and xu: "x \ u" and linS: "l\ S" and uinS: "u \ S" and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" shows "(\ s\ S. P s) \ (\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a < x \ x < b \ P x)" @@ -215,7 +215,7 @@ "{y. x < y \ y < z} = {} \ \ x < z" by (auto dest: dense) -lemma dlo_qe_bnds[noatp]: +lemma dlo_qe_bnds[no_atp]: assumes ne: "L \ {}" and neU: "U \ {}" and fL: "finite L" and fU: "finite U" shows "(\x. (\y \ L. y < x) \ (\y \ U. x < y)) \ (\ l \ L. \u \ U. l < u)" proof (simp only: atomize_eq, rule iffI) @@ -239,7 +239,7 @@ ultimately show "\x. (\y\L. y < x) \ (\y\U. x < y)" by auto qed -lemma dlo_qe_noub[noatp]: +lemma dlo_qe_noub[no_atp]: assumes ne: "L \ {}" and fL: "finite L" shows "(\x. (\y \ L. y < x) \ (\y \ {}. x < y)) \ True" proof(simp add: atomize_eq) @@ -249,7 +249,7 @@ thus "\x. \y\L. y < x" by blast qed -lemma dlo_qe_nolb[noatp]: +lemma dlo_qe_nolb[no_atp]: assumes ne: "U \ {}" and fU: "finite U" shows "(\x. (\y \ {}. y < x) \ (\y \ U. x < y)) \ True" proof(simp add: atomize_eq) @@ -259,14 +259,14 @@ thus "\x. \y\U. x < y" by blast qed -lemma exists_neq[noatp]: "\(x::'a). x \ t" "\(x::'a). t \ x" +lemma exists_neq[no_atp]: "\(x::'a). x \ t" "\(x::'a). t \ x" using gt_ex[of t] by auto -lemmas dlo_simps[noatp] = order_refl less_irrefl not_less not_le exists_neq +lemmas dlo_simps[no_atp] = order_refl less_irrefl not_less not_le exists_neq le_less neq_iff linear less_not_permute -lemma axiom[noatp]: "dense_linorder (op \) (op <)" by (rule dense_linorder_axioms) -lemma atoms[noatp]: +lemma axiom[no_atp]: "dense_linorder (op \) (op <)" by (rule dense_linorder_axioms) +lemma atoms[no_atp]: shows "TERM (less :: 'a \ _)" and "TERM (less_eq :: 'a \ _)" and "TERM (op = :: 'a \ _)" . @@ -277,21 +277,21 @@ end (* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *) -lemma dnf[noatp]: +lemma dnf[no_atp]: "(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))" by blast+ -lemmas weak_dnf_simps[noatp] = simp_thms dnf +lemmas weak_dnf_simps[no_atp] = simp_thms dnf -lemma nnf_simps[noatp]: +lemma nnf_simps[no_atp]: "(\(P \ Q)) = (\P \ \Q)" "(\(P \ Q)) = (\P \ \Q)" "(P \ Q) = (\P \ Q)" "(P = Q) = ((P \ Q) \ (\P \ \ Q))" "(\ \(P)) = P" by blast+ -lemma ex_distrib[noatp]: "(\x. P x \ Q x) \ ((\x. P x) \ (\x. Q x))" by blast +lemma ex_distrib[no_atp]: "(\x. P x \ Q x) \ ((\x. P x) \ (\x. Q x))" by blast -lemmas dnf_simps[noatp] = weak_dnf_simps nnf_simps ex_distrib +lemmas dnf_simps[no_atp] = weak_dnf_simps nnf_simps ex_distrib use "~~/src/HOL/Tools/Qelim/langford.ML" method_setup dlo = {* @@ -316,10 +316,10 @@ locale linorder_no_ub = linorder_stupid_syntax + assumes gt_ex: "\y. less x y" begin -lemma ge_ex[noatp]: "\y. x \ y" using gt_ex by auto +lemma ge_ex[no_atp]: "\y. x \ y" using gt_ex by auto text {* Theorems for @{text "\z. \x. z \ x \ (P x \ P\<^bsub>+\\<^esub>)"} *} -lemma pinf_conj[noatp]: +lemma pinf_conj[no_atp]: assumes ex1: "\z1. \x. z1 \ x \ (P1 x \ P1')" and ex2: "\z2. \x. z2 \ x \ (P2 x \ P2')" shows "\z. \x. z \ x \ ((P1 x \ P2 x) \ (P1' \ P2'))" @@ -335,7 +335,7 @@ thus ?thesis by blast qed -lemma pinf_disj[noatp]: +lemma pinf_disj[no_atp]: assumes ex1: "\z1. \x. z1 \ x \ (P1 x \ P1')" and ex2: "\z2. \x. z2 \ x \ (P2 x \ P2')" shows "\z. \x. z \ x \ ((P1 x \ P2 x) \ (P1' \ P2'))" @@ -351,7 +351,7 @@ thus ?thesis by blast qed -lemma pinf_ex[noatp]: assumes ex:"\z. \x. z \ x \ (P x \ P1)" and p1: P1 shows "\ x. P x" +lemma pinf_ex[no_atp]: assumes ex:"\z. \x. z \ x \ (P x \ P1)" and p1: P1 shows "\ x. P x" proof- from ex obtain z where z: "\x. z \ x \ (P x \ P1)" by blast from gt_ex obtain x where x: "z \ x" by blast @@ -365,11 +365,11 @@ locale linorder_no_lb = linorder_stupid_syntax + assumes lt_ex: "\y. less y x" begin -lemma le_ex[noatp]: "\y. y \ x" using lt_ex by auto +lemma le_ex[no_atp]: "\y. y \ x" using lt_ex by auto text {* Theorems for @{text "\z. \x. x \ z \ (P x \ P\<^bsub>-\\<^esub>)"} *} -lemma minf_conj[noatp]: +lemma minf_conj[no_atp]: assumes ex1: "\z1. \x. x \ z1 \ (P1 x \ P1')" and ex2: "\z2. \x. x \ z2 \ (P2 x \ P2')" shows "\z. \x. x \ z \ ((P1 x \ P2 x) \ (P1' \ P2'))" @@ -384,7 +384,7 @@ thus ?thesis by blast qed -lemma minf_disj[noatp]: +lemma minf_disj[no_atp]: assumes ex1: "\z1. \x. x \ z1 \ (P1 x \ P1')" and ex2: "\z2. \x. x \ z2 \ (P2 x \ P2')" shows "\z. \x. x \ z \ ((P1 x \ P2 x) \ (P1' \ P2'))" @@ -399,7 +399,7 @@ thus ?thesis by blast qed -lemma minf_ex[noatp]: assumes ex:"\z. \x. x \ z \ (P x \ P1)" and p1: P1 shows "\ x. P x" +lemma minf_ex[no_atp]: assumes ex:"\z. \x. x \ z \ (P x \ P1)" and p1: P1 shows "\ x. P x" proof- from ex obtain z where z: "\x. x \ z \ (P x \ P1)" by blast from lt_ex obtain x where x: "x \ z" by blast @@ -422,7 +422,7 @@ context constr_dense_linorder begin -lemma rinf_U[noatp]: +lemma rinf_U[no_atp]: assumes fU: "finite U" and lin_dense: "\x l u. (\ t. l \ t \ t\ u \ t \ U) \ l\ x \ x \ u \ P x \ (\ y. l \ y \ y \ u \ P y )" @@ -465,7 +465,7 @@ ultimately show ?thesis by blast qed -theorem fr_eq[noatp]: +theorem fr_eq[no_atp]: assumes fU: "finite U" and lin_dense: "\x l u. (\ t. l \ t \ t\ u \ t \ U) \ l\ x \ x \ u \ P x \ (\ y. l \ y \ y \ u \ P y )" @@ -493,16 +493,16 @@ ultimately have "?E = ?D" by blast thus "?E \ ?D" by simp qed -lemmas minf_thms[noatp] = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P -lemmas pinf_thms[noatp] = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P +lemmas minf_thms[no_atp] = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P +lemmas pinf_thms[no_atp] = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P -lemmas nmi_thms[noatp] = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P -lemmas npi_thms[noatp] = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P -lemmas lin_dense_thms[noatp] = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P +lemmas nmi_thms[no_atp] = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P +lemmas npi_thms[no_atp] = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P +lemmas lin_dense_thms[no_atp] = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P -lemma ferrack_axiom[noatp]: "constr_dense_linorder less_eq less between" +lemma ferrack_axiom[no_atp]: "constr_dense_linorder less_eq less between" by (rule constr_dense_linorder_axioms) -lemma atoms[noatp]: +lemma atoms[no_atp]: shows "TERM (less :: 'a \ _)" and "TERM (less_eq :: 'a \ _)" and "TERM (op = :: 'a \ _)" .