--- a/src/HOL/MetisExamples/BigO.thy Sat Mar 29 19:14:03 2008 +0100
+++ b/src/HOL/MetisExamples/BigO.thy Sat Mar 29 19:14:05 2008 +0100
@@ -18,7 +18,7 @@
bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set" ("(1O'(_'))")
"O(f::('a => 'b)) == {h. EX c. ALL x. abs (h x) <= c * abs (f x)}"
-ML{*ResAtp.problem_name := "BigO__bigo_pos_const"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_pos_const"*}
lemma bigo_pos_const: "(EX (c::'a::ordered_idom).
ALL x. (abs (h x)) <= (c * (abs (f x))))
= (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
@@ -215,7 +215,7 @@
{h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}"
by (auto simp add: bigo_def bigo_pos_const)
-ML{*ResAtp.problem_name := "BigO__bigo_elt_subset"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_elt_subset"*}
lemma bigo_elt_subset [intro]: "f : O(g) ==> O(f) <= O(g)"
apply (auto simp add: bigo_alt_def)
apply (rule_tac x = "ca * c" in exI)
@@ -233,7 +233,7 @@
done
-ML{*ResAtp.problem_name := "BigO__bigo_refl"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_refl"*}
lemma bigo_refl [intro]: "f : O(f)"
apply(auto simp add: bigo_def)
proof (neg_clausify)
@@ -247,7 +247,7 @@
by (metis 0 2)
qed
-ML{*ResAtp.problem_name := "BigO__bigo_zero"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_zero"*}
lemma bigo_zero: "0 : O(g)"
apply (auto simp add: bigo_def func_zero)
proof (neg_clausify)
@@ -337,7 +337,7 @@
apply (auto del: subsetI simp del: bigo_plus_idemp)
done
-ML{*ResAtp.problem_name := "BigO__bigo_plus_eq"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_plus_eq"*}
lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==>
O(f + g) = O(f) + O(g)"
apply (rule equalityI)
@@ -360,13 +360,13 @@
apply (rule abs_triangle_ineq)
apply (metis add_nonneg_nonneg)
apply (rule add_mono)
-ML{*ResAtp.problem_name := "BigO__bigo_plus_eq_simpler"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_plus_eq_simpler"*}
(*Found by SPASS; SLOW*)
apply (metis le_maxI2 linorder_linear linorder_not_le min_max.less_eq_less_sup.sup_absorb1 mult_le_cancel_right order_trans)
apply (metis le_maxI2 linorder_not_le mult_le_cancel_right order_trans)
done
-ML{*ResAtp.problem_name := "BigO__bigo_bounded_alt"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_bounded_alt"*}
lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==>
f : O(g)"
apply (auto simp add: bigo_def)
@@ -426,7 +426,7 @@
text{*So here is the easier (and more natural) problem using transitivity*}
-ML{*ResAtp.problem_name := "BigO__bigo_bounded_alt_trans"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_bounded_alt_trans"*}
lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)"
apply (auto simp add: bigo_def)
(*Version 1: one-shot proof*)
@@ -434,7 +434,7 @@
done
text{*So here is the easier (and more natural) problem using transitivity*}
-ML{*ResAtp.problem_name := "BigO__bigo_bounded_alt_trans"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_bounded_alt_trans"*}
lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)"
apply (auto simp add: bigo_def)
(*Version 2: single-step proof*)
@@ -473,7 +473,7 @@
apply simp
done
-ML{*ResAtp.problem_name := "BigO__bigo_bounded2"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_bounded2"*}
lemma bigo_bounded2: "ALL x. lb x <= f x ==> ALL x. f x <= lb x + g x ==>
f : lb +o O(g)"
apply (rule set_minus_imp_plus)
@@ -496,7 +496,7 @@
by (metis 4 2 0)
qed
-ML{*ResAtp.problem_name := "BigO__bigo_abs"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_abs"*}
lemma bigo_abs: "(%x. abs(f x)) =o O(f)"
apply (unfold bigo_def)
apply auto
@@ -511,7 +511,7 @@
by (metis 0 2)
qed
-ML{*ResAtp.problem_name := "BigO__bigo_abs2"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_abs2"*}
lemma bigo_abs2: "f =o O(%x. abs(f x))"
apply (unfold bigo_def)
apply auto
@@ -583,7 +583,7 @@
by (simp add: bigo_abs3 [symmetric])
qed
-ML{*ResAtp.problem_name := "BigO__bigo_mult"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_mult"*}
lemma bigo_mult [intro]: "O(f)*O(g) <= O(f * g)"
apply (rule subsetI)
apply (subst bigo_def)
@@ -595,7 +595,7 @@
apply(erule_tac x = x in allE)+
apply(subgoal_tac "c * ca * abs(f x * g x) =
(c * abs(f x)) * (ca * abs(g x))")
-ML{*ResAtp.problem_name := "BigO__bigo_mult_simpler"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_mult_simpler"*}
prefer 2
apply (metis mult_assoc mult_left_commute
OrderedGroup.abs_of_pos OrderedGroup.mult_left_commute
@@ -660,14 +660,14 @@
qed
-ML{*ResAtp.problem_name := "BigO__bigo_mult2"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_mult2"*}
lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)"
apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult)
(*sledgehammer*);
apply (rule_tac x = c in exI)
apply clarify
apply (drule_tac x = x in spec)
-ML{*ResAtp.problem_name := "BigO__bigo_mult2_simpler"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_mult2_simpler"*}
(*sledgehammer [no luck]*);
apply (subgoal_tac "abs(f x) * abs(b x) <= abs(f x) * (c * abs(g x))")
apply (simp add: mult_ac)
@@ -675,11 +675,11 @@
apply (rule abs_ge_zero)
done
-ML{*ResAtp.problem_name:="BigO__bigo_mult3"*}
+ML_command{*ResAtp.problem_name:="BigO__bigo_mult3"*}
lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)"
by (metis bigo_mult set_times_intro subset_iff)
-ML{*ResAtp.problem_name:="BigO__bigo_mult4"*}
+ML_command{*ResAtp.problem_name:="BigO__bigo_mult4"*}
lemma bigo_mult4 [intro]:"f : k +o O(h) ==> g * f : (g * k) +o O(g * h)"
by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib)
@@ -713,13 +713,13 @@
qed
qed
-ML{*ResAtp.problem_name := "BigO__bigo_mult6"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_mult6"*}
lemma bigo_mult6: "ALL x. f x ~= 0 ==>
O(f * g) = (f::'a => ('b::ordered_field)) *o O(g)"
by (metis bigo_mult2 bigo_mult5 order_antisym)
(*proof requires relaxing relevance: 2007-01-25*)
-ML{*ResAtp.problem_name := "BigO__bigo_mult7"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_mult7"*}
declare bigo_mult6 [simp]
lemma bigo_mult7: "ALL x. f x ~= 0 ==>
O(f * g) <= O(f::'a => ('b::ordered_field)) * O(g)"
@@ -731,7 +731,7 @@
done
declare bigo_mult6 [simp del]
-ML{*ResAtp.problem_name := "BigO__bigo_mult8"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_mult8"*}
declare bigo_mult7[intro!]
lemma bigo_mult8: "ALL x. f x ~= 0 ==>
O(f * g) = O(f::'a => ('b::ordered_field)) * O(g)"
@@ -782,7 +782,7 @@
qed
qed
-ML{*ResAtp.problem_name:="BigO__bigo_plus_absorb"*}
+ML_command{*ResAtp.problem_name:="BigO__bigo_plus_absorb"*}
lemma bigo_plus_absorb [simp]: "f : O(g) ==> f +o O(g) = O(g)"
by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff);
@@ -809,7 +809,7 @@
lemma bigo_const1: "(%x. c) : O(%x. 1)"
by (auto simp add: bigo_def mult_ac)
-ML{*ResAtp.problem_name:="BigO__bigo_const2"*}
+ML_command{*ResAtp.problem_name:="BigO__bigo_const2"*}
lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)"
by (metis bigo_const1 bigo_elt_subset);
@@ -832,7 +832,7 @@
apply (rule bigo_const1)
done
-ML{*ResAtp.problem_name := "BigO__bigo_const3"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_const3"*}
lemma bigo_const3: "(c::'a::ordered_field) ~= 0 ==> (%x. 1) : O(%x. c)"
apply (simp add: bigo_def)
proof (neg_clausify)
@@ -856,7 +856,7 @@
O(%x. c) = O(%x. 1)"
by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption)
-ML{*ResAtp.problem_name := "BigO__bigo_const_mult1"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_const_mult1"*}
lemma bigo_const_mult1: "(%x. c * f x) : O(f)"
apply (simp add: bigo_def abs_mult)
proof (neg_clausify)
@@ -872,7 +872,7 @@
lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)"
by (rule bigo_elt_subset, rule bigo_const_mult1)
-ML{*ResAtp.problem_name := "BigO__bigo_const_mult3"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_const_mult3"*}
lemma bigo_const_mult3: "(c::'a::ordered_field) ~= 0 ==> f : O(%x. c * f x)"
apply (simp add: bigo_def)
(*sledgehammer [no luck]*);
@@ -890,7 +890,7 @@
O(%x. c * f x) = O(f)"
by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4)
-ML{*ResAtp.problem_name := "BigO__bigo_const_mult5"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_const_mult5"*}
lemma bigo_const_mult5 [simp]: "(c::'a::ordered_field) ~= 0 ==>
(%x. c) *o O(f) = O(f)"
apply (auto del: subsetI)
@@ -910,7 +910,7 @@
done
-ML{*ResAtp.problem_name := "BigO__bigo_const_mult6"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_const_mult6"*}
lemma bigo_const_mult6 [intro]: "(%x. c) *o O(f) <= O(f)"
apply (auto intro!: subsetI
simp add: bigo_def elt_set_times_def func_times
@@ -967,7 +967,7 @@
apply (blast intro: order_trans mult_right_mono abs_ge_self)
done
-ML{*ResAtp.problem_name := "BigO__bigo_setsum1"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_setsum1"*}
lemma bigo_setsum1: "ALL x y. 0 <= h x y ==>
EX c. ALL x y. abs(f x y) <= c * (h x y) ==>
(%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)"
@@ -984,7 +984,7 @@
(%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)"
by (rule bigo_setsum1, auto)
-ML{*ResAtp.problem_name := "BigO__bigo_setsum3"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_setsum3"*}
lemma bigo_setsum3: "f =o O(h) ==>
(%x. SUM y : A x. (l x y) * f(k x y)) =o
O(%x. SUM y : A x. abs(l x y * h(k x y)))"
@@ -1015,7 +1015,7 @@
apply (erule set_plus_imp_minus)
done
-ML{*ResAtp.problem_name := "BigO__bigo_setsum5"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_setsum5"*}
lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==>
ALL x. 0 <= h x ==>
(%x. SUM y : A x. (l x y) * f(k x y)) =o
@@ -1072,7 +1072,7 @@
apply (simp add: func_times)
done
-ML{*ResAtp.problem_name := "BigO__bigo_fix"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_fix"*}
lemma bigo_fix: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==>
f =o O(h)"
apply (simp add: bigo_alt_def)
@@ -1137,7 +1137,7 @@
apply (erule spec)+
done
-ML{*ResAtp.problem_name:="BigO__bigo_lesso1"*}
+ML_command{*ResAtp.problem_name:="BigO__bigo_lesso1"*}
lemma bigo_lesso1: "ALL x. f x <= g x ==> f <o g =o O(h)"
apply (unfold lesso_def)
apply (subgoal_tac "(%x. max (f x - g x) 0) = 0")
@@ -1149,7 +1149,7 @@
done
-ML{*ResAtp.problem_name := "BigO__bigo_lesso2"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_lesso2"*}
lemma bigo_lesso2: "f =o g +o O(h) ==>
ALL x. 0 <= k x ==> ALL x. k x <= f x ==>
k <o g =o O(h)"
@@ -1186,7 +1186,7 @@
by (metis min_max.less_eq_less_sup.sup_commute min_max.less_eq_less_inf.inf_commute min_max.less_eq_less_inf_sup.sup_inf_absorb min_max.less_eq_less_inf.le_iff_inf 0 max_diff_distrib_left 1 linorder_not_le 8)
qed
-ML{*ResAtp.problem_name := "BigO__bigo_lesso3"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_lesso3"*}
lemma bigo_lesso3: "f =o g +o O(h) ==>
ALL x. 0 <= k x ==> ALL x. g x <= k x ==>
f <o k =o O(h)"
@@ -1203,7 +1203,7 @@
apply (simp del: compare_rls diff_minus);
apply (subst abs_of_nonneg)
apply (drule_tac x = x in spec) back
-ML{*ResAtp.problem_name := "BigO__bigo_lesso3_simpler"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_lesso3_simpler"*}
apply (metis diff_less_0_iff_less linorder_not_le not_leE uminus_add_conv_diff xt1(12) xt1(6))
apply (metis add_minus_cancel diff_le_eq le_diff_eq uminus_add_conv_diff)
apply (metis abs_ge_zero linorder_linear min_max.less_eq_less_sup.sup_absorb1 min_max.less_eq_less_sup.sup_commute)
@@ -1223,7 +1223,7 @@
split: split_max abs_split)
done
-ML{*ResAtp.problem_name := "BigO__bigo_lesso5"*}
+ML_command{*ResAtp.problem_name := "BigO__bigo_lesso5"*}
lemma bigo_lesso5: "f <o g =o O(h) ==>
EX C. ALL x. f x <= g x + C * abs(h x)"
apply (simp only: lesso_def bigo_alt_def)
--- a/src/HOL/MetisExamples/Tarski.thy Sat Mar 29 19:14:03 2008 +0100
+++ b/src/HOL/MetisExamples/Tarski.thy Sat Mar 29 19:14:05 2008 +0100
@@ -408,7 +408,7 @@
(*never proved, 2007-01-22: Tarski__CLF_unnamed_lemma
NOT PROVABLE because of the conjunction used in the definition: we don't
allow reasoning with rules like conjE, which is essential here.*)
-ML{*ResAtp.problem_name:="Tarski__CLF_unnamed_lemma"*}
+ML_command{*ResAtp.problem_name:="Tarski__CLF_unnamed_lemma"*}
lemma (in CLF) [simp]:
"f: pset cl -> pset cl & monotone f (pset cl) (order cl)"
apply (insert f_cl)
@@ -425,7 +425,7 @@
by (simp add: A_def r_def)
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__CLF_CLF_dual"*}
+ML_command{*ResAtp.problem_name:="Tarski__CLF_CLF_dual"*}
declare (in CLF) CLF_def[simp] CL_dualCL[simp] monotone_dual[simp] dualA_iff[simp]
lemma (in CLF) CLF_dual: "(dual cl, f) \<in> CLF"
apply (simp del: dualA_iff)
@@ -461,7 +461,7 @@
-- {* @{text "\<forall>x:H. (x, f (lub H r)) \<in> r"} *}
apply (rule ballI)
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__CLF_lubH_le_flubH_simpler"*}
+ML_command{*ResAtp.problem_name:="Tarski__CLF_lubH_le_flubH_simpler"*}
apply (rule transE)
-- {* instantiates @{text "(x, ?z) \<in> order cl to (x, f x)"}, *}
-- {* because of the def of @{text H} *}
@@ -489,7 +489,7 @@
apply (rule_tac t = "H" in ssubst, assumption)
apply (rule CollectI)
apply (rule conjI)
-ML{*ResAtp.problem_name:="Tarski__CLF_flubH_le_lubH_simpler"*}
+ML_command{*ResAtp.problem_name:="Tarski__CLF_flubH_le_lubH_simpler"*}
(*??no longer terminates, with combinators
apply (metis CO_refl lubH_le_flubH monotone_def monotone_f reflD1 reflD2)
*)
@@ -578,7 +578,7 @@
"H = {x. (x, f x) \<in> r & x \<in> A} ==> lub H cl \<in> fix f A"
apply (simp add: fix_def)
apply (rule conjI)
-ML{*ResAtp.problem_name:="Tarski__CLF_lubH_is_fixp_simpler"*}
+ML_command{*ResAtp.problem_name:="Tarski__CLF_lubH_is_fixp_simpler"*}
apply (metis CO_refl lubH_le_flubH reflD1)
apply (metis antisymE flubH_le_lubH lubH_le_flubH)
done
@@ -613,7 +613,7 @@
apply (rule sym)
apply (simp add: P_def)
apply (rule lubI)
-ML{*ResAtp.problem_name:="Tarski__CLF_T_thm_1_lub_simpler"*}
+ML_command{*ResAtp.problem_name:="Tarski__CLF_T_thm_1_lub_simpler"*}
apply (metis P_def fix_subset)
apply (metis Collect_conj_eq Collect_mem_eq Int_commute Int_lower1 lub_in_lattice vimage_def)
(*??no longer terminates, with combinators
@@ -653,7 +653,7 @@
apply (simp add: glb_dual_lub P_def A_def r_def)
apply (rule dualA_iff [THEN subst])
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__T_thm_1_glb_simpler"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__T_thm_1_glb_simpler"*} (*ALL THEOREMS*)
(*sledgehammer;*)
apply (simp add: CLF.T_thm_1_lub [of _ f, OF dualPO CL_dualCL]
dualPO CL_dualCL CLF_dual dualr_iff)
@@ -812,7 +812,7 @@
apply (rule sublatticeI)
apply (simp add: interval_subset)
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__interval_is_sublattice_simpler"*}
+ML_command{*ResAtp.problem_name:="Tarski__interval_is_sublattice_simpler"*}
(*sledgehammer *)
apply (rule CompleteLatticeI)
apply (simp add: intervalPO)
@@ -831,7 +831,7 @@
lemma (in CLF) Bot_dual_Top: "Bot cl = Top (dual cl)"
by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff)
-ML{*ResAtp.problem_name:="Tarski__Bot_in_lattice"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__Bot_in_lattice"*} (*ALL THEOREMS*)
lemma (in CLF) Bot_in_lattice: "Bot cl \<in> A"
(*sledgehammer; *)
apply (simp add: Bot_def least_def)
@@ -841,12 +841,12 @@
done
(*first proved 2007-01-25 after relaxing relevance*)
-ML{*ResAtp.problem_name:="Tarski__Top_in_lattice"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__Top_in_lattice"*} (*ALL THEOREMS*)
lemma (in CLF) Top_in_lattice: "Top cl \<in> A"
(*sledgehammer;*)
apply (simp add: Top_dual_Bot A_def)
(*first proved 2007-01-25 after relaxing relevance*)
-ML{*ResAtp.problem_name:="Tarski__Top_in_lattice_simpler"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__Top_in_lattice_simpler"*} (*ALL THEOREMS*)
(*sledgehammer*)
apply (rule dualA_iff [THEN subst])
apply (blast intro!: CLF.Bot_in_lattice dualPO CL_dualCL CLF_dual)
@@ -861,7 +861,7 @@
done
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__Bot_prop"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__Bot_prop"*} (*ALL THEOREMS*)
lemma (in CLF) Bot_prop: "x \<in> A ==> (Bot cl, x) \<in> r"
(*sledgehammer*)
apply (simp add: Bot_dual_Top r_def)
@@ -870,12 +870,12 @@
dualA_iff A_def dualPO CL_dualCL CLF_dual)
done
-ML{*ResAtp.problem_name:="Tarski__Bot_in_lattice"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__Bot_in_lattice"*} (*ALL THEOREMS*)
lemma (in CLF) Top_intv_not_empty: "x \<in> A ==> interval r x (Top cl) \<noteq> {}"
apply (metis Top_in_lattice Top_prop empty_iff intervalI reflE)
done
-ML{*ResAtp.problem_name:="Tarski__Bot_intv_not_empty"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__Bot_intv_not_empty"*} (*ALL THEOREMS*)
lemma (in CLF) Bot_intv_not_empty: "x \<in> A ==> interval r (Bot cl) x \<noteq> {}"
apply (metis Bot_prop ex_in_conv intervalI reflE rel_imp_elem)
done
@@ -887,7 +887,7 @@
by (simp add: P_def fix_subset po_subset_po)
(*first proved 2007-01-25 after relaxing relevance*)
-ML{*ResAtp.problem_name:="Tarski__Y_subset_A"*}
+ML_command{*ResAtp.problem_name:="Tarski__Y_subset_A"*}
declare (in Tarski) P_def[simp] Y_ss [simp]
declare fix_subset [intro] subset_trans [intro]
lemma (in Tarski) Y_subset_A: "Y \<subseteq> A"
@@ -903,7 +903,7 @@
by (rule Y_subset_A [THEN lub_in_lattice])
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__lubY_le_flubY"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__lubY_le_flubY"*} (*ALL THEOREMS*)
lemma (in Tarski) lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \<in> r"
(*sledgehammer*)
apply (rule lub_least)
@@ -912,12 +912,12 @@
apply (rule lubY_in_A)
-- {* @{text "Y \<subseteq> P ==> f x = x"} *}
apply (rule ballI)
-ML{*ResAtp.problem_name:="Tarski__lubY_le_flubY_simpler"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__lubY_le_flubY_simpler"*} (*ALL THEOREMS*)
(*sledgehammer *)
apply (rule_tac t = "x" in fix_imp_eq [THEN subst])
apply (erule Y_ss [simplified P_def, THEN subsetD])
-- {* @{text "reduce (f x, f (lub Y cl)) \<in> r to (x, lub Y cl) \<in> r"} by monotonicity *}
-ML{*ResAtp.problem_name:="Tarski__lubY_le_flubY_simplest"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__lubY_le_flubY_simplest"*} (*ALL THEOREMS*)
(*sledgehammer*)
apply (rule_tac f = "f" in monotoneE)
apply (rule monotone_f)
@@ -927,7 +927,7 @@
done
(*first proved 2007-01-25 after relaxing relevance*)
-ML{*ResAtp.problem_name:="Tarski__intY1_subset"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__intY1_subset"*} (*ALL THEOREMS*)
lemma (in Tarski) intY1_subset: "intY1 \<subseteq> A"
(*sledgehammer*)
apply (unfold intY1_def)
@@ -939,7 +939,7 @@
lemmas (in Tarski) intY1_elem = intY1_subset [THEN subsetD]
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__intY1_f_closed"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__intY1_f_closed"*} (*ALL THEOREMS*)
lemma (in Tarski) intY1_f_closed: "x \<in> intY1 \<Longrightarrow> f x \<in> intY1"
(*sledgehammer*)
apply (simp add: intY1_def interval_def)
@@ -947,7 +947,7 @@
apply (rule transE)
apply (rule lubY_le_flubY)
-- {* @{text "(f (lub Y cl), f x) \<in> r"} *}
-ML{*ResAtp.problem_name:="Tarski__intY1_f_closed_simpler"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__intY1_f_closed_simpler"*} (*ALL THEOREMS*)
(*sledgehammer [has been proved before now...]*)
apply (rule_tac f=f in monotoneE)
apply (rule monotone_f)
@@ -960,11 +960,11 @@
apply (simp add: intY1_def interval_def intY1_elem)
done
-ML{*ResAtp.problem_name:="Tarski__intY1_func"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__intY1_func"*} (*ALL THEOREMS*)
lemma (in Tarski) intY1_func: "(%x: intY1. f x) \<in> intY1 -> intY1"
by (metis intY1_f_closed restrict_in_funcset)
-ML{*ResAtp.problem_name:="Tarski__intY1_mono"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__intY1_mono"*} (*ALL THEOREMS*)
lemma (in Tarski) intY1_mono:
"monotone (%x: intY1. f x) intY1 (induced intY1 r)"
(*sledgehammer *)
@@ -973,7 +973,7 @@
done
(*proof requires relaxing relevance: 2007-01-25*)
-ML{*ResAtp.problem_name:="Tarski__intY1_is_cl"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__intY1_is_cl"*} (*ALL THEOREMS*)
lemma (in Tarski) intY1_is_cl:
"(| pset = intY1, order = induced intY1 r |) \<in> CompleteLattice"
(*sledgehammer*)
@@ -986,7 +986,7 @@
done
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__v_in_P"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__v_in_P"*} (*ALL THEOREMS*)
lemma (in Tarski) v_in_P: "v \<in> P"
(*sledgehammer*)
apply (unfold P_def)
@@ -996,7 +996,7 @@
v_def CL_imp_PO intY1_is_cl CLF_def intY1_func intY1_mono)
done
-ML{*ResAtp.problem_name:="Tarski__z_in_interval"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__z_in_interval"*} (*ALL THEOREMS*)
lemma (in Tarski) z_in_interval:
"[| z \<in> P; \<forall>y\<in>Y. (y, z) \<in> induced P r |] ==> z \<in> intY1"
(*sledgehammer *)
@@ -1010,14 +1010,14 @@
apply (simp add: induced_def)
done
-ML{*ResAtp.problem_name:="Tarski__fz_in_int_rel"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__fz_in_int_rel"*} (*ALL THEOREMS*)
lemma (in Tarski) f'z_in_int_rel: "[| z \<in> P; \<forall>y\<in>Y. (y, z) \<in> induced P r |]
==> ((%x: intY1. f x) z, z) \<in> induced intY1 r"
apply (metis P_def acc_def fix_imp_eq fix_subset indI reflE restrict_apply subset_def z_in_interval)
done
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__tarski_full_lemma"*} (*ALL THEOREMS*)
+ML_command{*ResAtp.problem_name:="Tarski__tarski_full_lemma"*} (*ALL THEOREMS*)
lemma (in Tarski) tarski_full_lemma:
"\<exists>L. isLub Y (| pset = P, order = induced P r |) L"
apply (rule_tac x = "v" in exI)
@@ -1047,12 +1047,12 @@
prefer 2 apply (simp add: v_in_P)
apply (unfold v_def)
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__tarski_full_lemma_simpler"*}
+ML_command{*ResAtp.problem_name:="Tarski__tarski_full_lemma_simpler"*}
(*sledgehammer*)
apply (rule indE)
apply (rule_tac [2] intY1_subset)
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__tarski_full_lemma_simplest"*}
+ML_command{*ResAtp.problem_name:="Tarski__tarski_full_lemma_simplest"*}
(*sledgehammer*)
apply (rule CL.glb_lower [OF _ intY1_is_cl, simplified])
apply (simp add: CL_imp_PO intY1_is_cl)
@@ -1070,7 +1070,7 @@
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__Tarski_full"*}
+ML_command{*ResAtp.problem_name:="Tarski__Tarski_full"*}
declare (in CLF) fixf_po[intro] P_def [simp] A_def [simp] r_def [simp]
Tarski.tarski_full_lemma [intro] cl_po [intro] cl_co [intro]
CompleteLatticeI_simp [intro]
@@ -1080,7 +1080,7 @@
apply (rule CompleteLatticeI_simp)
apply (rule fixf_po, clarify)
(*never proved, 2007-01-22*)
-ML{*ResAtp.problem_name:="Tarski__Tarski_full_simpler"*}
+ML_command{*ResAtp.problem_name:="Tarski__Tarski_full_simpler"*}
(*sledgehammer*)
apply (simp add: P_def A_def r_def)
apply (blast intro!: Tarski.tarski_full_lemma cl_po cl_co f_cl)
--- a/src/HOLCF/IOA/NTP/Impl.thy Sat Mar 29 19:14:03 2008 +0100
+++ b/src/HOLCF/IOA/NTP/Impl.thy Sat Mar 29 19:14:05 2008 +0100
@@ -208,7 +208,7 @@
txt {* 10 cases. First 4 are simple, since state doesn't change *}
-ML {* val tac2 = asm_full_simp_tac (ss addsimps [@{thm inv2_def}]) *}
+ML_command {* val tac2 = asm_full_simp_tac (ss addsimps [@{thm inv2_def}]) *}
txt {* 10 - 7 *}
apply (tactic "EVERY1 [tac2,tac2,tac2,tac2]")
@@ -263,7 +263,7 @@
apply (simp (no_asm_simp) add: impl_ioas split del: split_if)
apply (induct_tac "a")
-ML {* val tac3 = asm_full_simp_tac (ss addsimps [@{thm inv3_def}]) *}
+ML_command {* val tac3 = asm_full_simp_tac (ss addsimps [@{thm inv3_def}]) *}
txt {* 10 - 8 *}
@@ -328,7 +328,7 @@
apply (simp (no_asm_simp) add: impl_ioas split del: split_if)
apply (induct_tac "a")
-ML {* val tac4 = asm_full_simp_tac (ss addsimps [@{thm inv4_def}]) *}
+ML_command {* val tac4 = asm_full_simp_tac (ss addsimps [@{thm inv4_def}]) *}
txt {* 10 - 2 *}