src/HOL/Probability/Interval_Integral.thy
changeset 62390 842917225d56
parent 62083 7582b39f51ed
child 63040 eb4ddd18d635
equal deleted inserted replaced
62380:29800666e526 62390:842917225d56
   300 lemma interval_integral_reflect:
   300 lemma interval_integral_reflect:
   301   "(LBINT x=a..b. f x) = (LBINT x=-b..-a. f (-x))"
   301   "(LBINT x=a..b. f x) = (LBINT x=-b..-a. f (-x))"
   302 proof (induct a b rule: linorder_wlog)
   302 proof (induct a b rule: linorder_wlog)
   303   case (sym a b) then show ?case
   303   case (sym a b) then show ?case
   304     by (auto simp add: interval_lebesgue_integral_def interval_integrable_endpoints_reverse
   304     by (auto simp add: interval_lebesgue_integral_def interval_integrable_endpoints_reverse
   305              split: split_if_asm)
   305              split: if_split_asm)
   306 next
   306 next
   307   case (le a b) then show ?case
   307   case (le a b) then show ?case
   308     unfolding interval_lebesgue_integral_def
   308     unfolding interval_lebesgue_integral_def
   309     by (subst set_integral_reflect)
   309     by (subst set_integral_reflect)
   310        (auto simp: interval_lebesgue_integrable_def einterval_iff
   310        (auto simp: interval_lebesgue_integrable_def einterval_iff