# HG changeset patch # User immler # Date 1438101932 -7200 # Node ID 53d8c91c86ec304f965849ebc83ca34c6d9de97f # Parent 9372f29acd47b2854f6e193535881029a6719e5c# Parent db6430b18964a025e7b04d670d93e3fb1a9ad923 merged diff -r db6430b18964 -r 53d8c91c86ec Admin/isatest/isatest-makedist --- a/Admin/isatest/isatest-makedist Tue Jul 28 17:46:35 2015 +0200 +++ b/Admin/isatest/isatest-makedist Tue Jul 28 18:45:32 2015 +0200 @@ -117,7 +117,7 @@ $MAKEALL $HOME/settings/mac-poly-M8-skip_proofs; $MAKEALL $HOME/settings/mac-poly-M8-quick_and_dirty" sleep 15 -$SSH macbroy6 "$MAKEALL $HOME/settings/mac-poly64-M2" +$SSH macbroy6 "$MAKEALL -x HOL-Mirabelle-ex $HOME/settings/mac-poly64-M2" sleep 15 $SSH macbroy30 "$MAKEALL $HOME/settings/mac-poly-M2" sleep 15 diff -r db6430b18964 -r 53d8c91c86ec src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Tue Jul 28 17:46:35 2015 +0200 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Jul 28 18:45:32 2015 +0200 @@ -10,17 +10,18 @@ "~~/src/HOL/Library/Indicator_Function" begin -lemma cSup_abs_le: (* TODO: is this really needed? *) +lemma cSup_abs_le: (* TODO: move to Conditionally_Complete_Lattices.thy? *) fixes S :: "real set" - shows "S \ {} \ (\x\S. \x\ \ a) \ \Sup S\ \ a" + shows "S \ {} \ (\x. x\S \ \x\ \ a) \ \Sup S\ \ a" by (auto simp add: abs_le_interval_iff intro: cSup_least) (metis cSup_upper2 bdd_aboveI) -lemma cInf_abs_ge: (* TODO: is this really needed? *) +lemma cInf_abs_ge: fixes S :: "real set" - shows "S \ {} \ (\x\S. \x\ \ a) \ \Inf S\ \ a" - by (simp add: Inf_real_def) (insert cSup_abs_le [of "uminus ` S"], auto) - -lemma cSup_asclose: (* TODO: is this really needed? *) + shows "S \ {} \ (\x. x\S \ \x\ \ a) \ \Inf S\ \ a" + using cSup_abs_le [of "uminus ` S"] + by (fastforce simp add: Inf_real_def) + +lemma cSup_asclose: fixes S :: "real set" assumes S: "S \ {}" and b: "\x\S. \x - l\ \ e" @@ -34,7 +35,7 @@ unfolding th by (auto intro!: cSup_upper2 cSup_least) qed -lemma cInf_asclose: (* TODO: is this really needed? *) +lemma cInf_asclose: fixes S :: "real set" assumes S: "S \ {}" and b: "\x\S. \x - l\ \ e" @@ -4931,11 +4932,10 @@ from p(4)[unfolded xk, OF insertI1] guess u v by (elim exE) note uv=this have "finite {k. \x. (x, k) \ p}" - apply (rule finite_subset[of _ "snd ` p"],rule) - unfolding subset_eq image_iff mem_Collect_eq - apply (erule exE) - apply (rule_tac x="(xa,x)" in bexI) + apply (rule finite_subset[of _ "snd ` p"]) using p + apply safe + apply (metis image_iff snd_conv) apply auto done then have int: "interior (cbox u v) \ interior (\{k. \x. (x, k) \ p}) = {}" @@ -4978,7 +4978,7 @@ apply rule apply rule apply (erule insertE) - defer + apply (simp add: uv xk) apply (rule UnI2) apply (drule q1(3)[rule_format]) unfolding xk uv @@ -5002,7 +5002,7 @@ apply rule apply (erule insertE) apply (rule UnI2) - defer + apply (simp add: False uv xk) apply (drule q1(3)[rule_format]) using False unfolding xk uv @@ -11771,7 +11771,9 @@ prefer 5 unfolding real_norm_def apply rule apply (rule cSup_abs_le) - prefer 5 + using assms + apply (force simp add: ) + prefer 4 apply rule apply (rule_tac g=h in absolutely_integrable_integrable_bound) using assms diff -r db6430b18964 -r 53d8c91c86ec src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Tue Jul 28 17:46:35 2015 +0200 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Tue Jul 28 18:45:32 2015 +0200 @@ -1203,7 +1203,7 @@ show ?lhs using i False apply (auto simp add: dependent_def) - by (metis in_span_insert insert_Diff insert_Diff_if insert_iff) + by (metis in_span_insert insert_Diff_if insert_Diff_single insert_absorb) qed qed