src/HOL/Int.thy
changeset 30843 3419ca741dbf
parent 30839 bf99ceb7d015
child 30960 fec1a04b7220
equal deleted inserted replaced
30840:98809b3f5e3c 30843:3419ca741dbf
  1380 qed
  1380 qed
  1381 
  1381 
  1382 
  1382 
  1383 subsection {* @{term setsum} and @{term setprod} *}
  1383 subsection {* @{term setsum} and @{term setprod} *}
  1384 
  1384 
  1385 text {*By Jeremy Avigad*}
       
  1386 
       
  1387 lemma of_nat_setsum: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))"
  1385 lemma of_nat_setsum: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))"
  1388   apply (cases "finite A")
  1386   apply (cases "finite A")
  1389   apply (erule finite_induct, auto)
  1387   apply (erule finite_induct, auto)
  1390   done
  1388   done
  1391 
  1389 
  1401 
  1399 
  1402 lemma of_int_setprod: "of_int (setprod f A) = (\<Prod>x\<in>A. of_int(f x))"
  1400 lemma of_int_setprod: "of_int (setprod f A) = (\<Prod>x\<in>A. of_int(f x))"
  1403   apply (cases "finite A")
  1401   apply (cases "finite A")
  1404   apply (erule finite_induct, auto)
  1402   apply (erule finite_induct, auto)
  1405   done
  1403   done
  1406 
       
  1407 lemma setprod_nonzero_nat:
       
  1408     "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::nat)) ==> setprod f A \<noteq> 0"
       
  1409   by (rule setprod_nonzero, auto)
       
  1410 
       
  1411 lemma setprod_pos_nat:
       
  1412   "finite S ==> (ALL x : S. f x > (0::nat)) ==> setprod f S > 0"
       
  1413 using setprod_nonzero_nat by simp
       
  1414 
       
  1415 lemma setprod_zero_eq_nat:
       
  1416     "finite A ==> (setprod f A = (0::nat)) = (\<exists>x \<in> A. f x = 0)"
       
  1417   by (rule setprod_zero_eq, auto)
       
  1418 
       
  1419 lemma setprod_nonzero_int:
       
  1420     "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::int)) ==> setprod f A \<noteq> 0"
       
  1421   by (rule setprod_nonzero, auto)
       
  1422 
       
  1423 lemma setprod_zero_eq_int:
       
  1424     "finite A ==> (setprod f A = (0::int)) = (\<exists>x \<in> A. f x = 0)"
       
  1425   by (rule setprod_zero_eq, auto)
       
  1426 
  1404 
  1427 lemmas int_setsum = of_nat_setsum [where 'a=int]
  1405 lemmas int_setsum = of_nat_setsum [where 'a=int]
  1428 lemmas int_setprod = of_nat_setprod [where 'a=int]
  1406 lemmas int_setprod = of_nat_setprod [where 'a=int]
  1429 
  1407 
  1430 
  1408