487 |
487 |
488 |
488 |
489 subsection {* A fold functional for finite sets *} |
489 subsection {* A fold functional for finite sets *} |
490 |
490 |
491 text {* The intended behaviour is |
491 text {* The intended behaviour is |
492 @{text "fold f z {x\<^isub>1, ..., x\<^isub>n} = f x\<^isub>1 (\<dots> (f x\<^isub>n z)\<dots>)"} |
492 @{text "fold f z {x1, ..., xn} = f x1 (\<dots> (f xn z)\<dots>)"} |
493 if @{text f} is ``left-commutative'': |
493 if @{text f} is ``left-commutative'': |
494 *} |
494 *} |
495 |
495 |
496 locale fun_left_comm = |
496 locale fun_left_comm = |
497 fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" |
497 fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" |
2484 |
2484 |
2485 context lattice |
2485 context lattice |
2486 begin |
2486 begin |
2487 |
2487 |
2488 definition |
2488 definition |
2489 Inf_fin :: "'a set \<Rightarrow> 'a" ("\<Sqinter>\<^bsub>fin\<^esub>_" [900] 900) |
2489 Inf_fin :: "'a set \<Rightarrow> 'a" ("\<Sqinter>fin_" [900] 900) |
2490 where |
2490 where |
2491 "Inf_fin = fold1 inf" |
2491 "Inf_fin = fold1 inf" |
2492 |
2492 |
2493 definition |
2493 definition |
2494 Sup_fin :: "'a set \<Rightarrow> 'a" ("\<Squnion>\<^bsub>fin\<^esub>_" [900] 900) |
2494 Sup_fin :: "'a set \<Rightarrow> 'a" ("\<Squnion>fin_" [900] 900) |
2495 where |
2495 where |
2496 "Sup_fin = fold1 sup" |
2496 "Sup_fin = fold1 sup" |
2497 |
2497 |
2498 lemma Inf_le_Sup [simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> \<Sqinter>\<^bsub>fin\<^esub>A \<le> \<Squnion>\<^bsub>fin\<^esub>A" |
2498 lemma Inf_le_Sup [simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> \<Sqinter>finA \<le> \<Squnion>finA" |
2499 apply(unfold Sup_fin_def Inf_fin_def) |
2499 apply(unfold Sup_fin_def Inf_fin_def) |
2500 apply(subgoal_tac "EX a. a:A") |
2500 apply(subgoal_tac "EX a. a:A") |
2501 prefer 2 apply blast |
2501 prefer 2 apply blast |
2502 apply(erule exE) |
2502 apply(erule exE) |
2503 apply(rule order_trans) |
2503 apply(rule order_trans) |
2504 apply(erule (1) fold1_belowI) |
2504 apply(erule (1) fold1_belowI) |
2505 apply(erule (1) lower_semilattice.fold1_belowI [OF dual_lattice]) |
2505 apply(erule (1) lower_semilattice.fold1_belowI [OF dual_lattice]) |
2506 done |
2506 done |
2507 |
2507 |
2508 lemma sup_Inf_absorb [simp]: |
2508 lemma sup_Inf_absorb [simp]: |
2509 "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> sup a (\<Sqinter>\<^bsub>fin\<^esub>A) = a" |
2509 "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> sup a (\<Sqinter>finA) = a" |
2510 apply(subst sup_commute) |
2510 apply(subst sup_commute) |
2511 apply(simp add: Inf_fin_def sup_absorb2 fold1_belowI) |
2511 apply(simp add: Inf_fin_def sup_absorb2 fold1_belowI) |
2512 done |
2512 done |
2513 |
2513 |
2514 lemma inf_Sup_absorb [simp]: |
2514 lemma inf_Sup_absorb [simp]: |
2515 "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> inf a (\<Squnion>\<^bsub>fin\<^esub>A) = a" |
2515 "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> inf a (\<Squnion>finA) = a" |
2516 by (simp add: Sup_fin_def inf_absorb1 |
2516 by (simp add: Sup_fin_def inf_absorb1 |
2517 lower_semilattice.fold1_belowI [OF dual_lattice]) |
2517 lower_semilattice.fold1_belowI [OF dual_lattice]) |
2518 |
2518 |
2519 end |
2519 end |
2520 |
2520 |
2534 (rule arg_cong [where f="fold1 inf"], blast) |
2534 (rule arg_cong [where f="fold1 inf"], blast) |
2535 qed |
2535 qed |
2536 |
2536 |
2537 lemma sup_Inf2_distrib: |
2537 lemma sup_Inf2_distrib: |
2538 assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}" |
2538 assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}" |
2539 shows "sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B) = \<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B}" |
2539 shows "sup (\<Sqinter>finA) (\<Sqinter>finB) = \<Sqinter>fin{sup a b|a b. a \<in> A \<and> b \<in> B}" |
2540 using A proof (induct rule: finite_ne_induct) |
2540 using A proof (induct rule: finite_ne_induct) |
2541 case singleton thus ?case |
2541 case singleton thus ?case |
2542 by (simp add: sup_Inf1_distrib [OF B] fold1_singleton_def [OF Inf_fin_def]) |
2542 by (simp add: sup_Inf1_distrib [OF B] fold1_singleton_def [OF Inf_fin_def]) |
2543 next |
2543 next |
2544 interpret ab_semigroup_idem_mult inf |
2544 interpret ab_semigroup_idem_mult inf |
2551 have "{sup a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {sup a b})" |
2551 have "{sup a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {sup a b})" |
2552 by blast |
2552 by blast |
2553 thus ?thesis by(simp add: insert(1) B(1)) |
2553 thus ?thesis by(simp add: insert(1) B(1)) |
2554 qed |
2554 qed |
2555 have ne: "{sup a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast |
2555 have ne: "{sup a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast |
2556 have "sup (\<Sqinter>\<^bsub>fin\<^esub>(insert x A)) (\<Sqinter>\<^bsub>fin\<^esub>B) = sup (inf x (\<Sqinter>\<^bsub>fin\<^esub>A)) (\<Sqinter>\<^bsub>fin\<^esub>B)" |
2556 have "sup (\<Sqinter>fin(insert x A)) (\<Sqinter>finB) = sup (inf x (\<Sqinter>finA)) (\<Sqinter>finB)" |
2557 using insert by (simp add: fold1_insert_idem_def [OF Inf_fin_def]) |
2557 using insert by (simp add: fold1_insert_idem_def [OF Inf_fin_def]) |
2558 also have "\<dots> = inf (sup x (\<Sqinter>\<^bsub>fin\<^esub>B)) (sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B))" by(rule sup_inf_distrib2) |
2558 also have "\<dots> = inf (sup x (\<Sqinter>finB)) (sup (\<Sqinter>finA) (\<Sqinter>finB))" by(rule sup_inf_distrib2) |
2559 also have "\<dots> = inf (\<Sqinter>\<^bsub>fin\<^esub>{sup x b|b. b \<in> B}) (\<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B})" |
2559 also have "\<dots> = inf (\<Sqinter>fin{sup x b|b. b \<in> B}) (\<Sqinter>fin{sup a b|a b. a \<in> A \<and> b \<in> B})" |
2560 using insert by(simp add:sup_Inf1_distrib[OF B]) |
2560 using insert by(simp add:sup_Inf1_distrib[OF B]) |
2561 also have "\<dots> = \<Sqinter>\<^bsub>fin\<^esub>({sup x b |b. b \<in> B} \<union> {sup a b |a b. a \<in> A \<and> b \<in> B})" |
2561 also have "\<dots> = \<Sqinter>fin({sup x b |b. b \<in> B} \<union> {sup a b |a b. a \<in> A \<and> b \<in> B})" |
2562 (is "_ = \<Sqinter>\<^bsub>fin\<^esub>?M") |
2562 (is "_ = \<Sqinter>fin?M") |
2563 using B insert |
2563 using B insert |
2564 by (simp add: Inf_fin_def fold1_Un2 [OF finB _ finAB ne]) |
2564 by (simp add: Inf_fin_def fold1_Un2 [OF finB _ finAB ne]) |
2565 also have "?M = {sup a b |a b. a \<in> insert x A \<and> b \<in> B}" |
2565 also have "?M = {sup a b |a b. a \<in> insert x A \<and> b \<in> B}" |
2566 by blast |
2566 by blast |
2567 finally show ?case . |
2567 finally show ?case . |
2568 qed |
2568 qed |
2569 |
2569 |
2570 lemma inf_Sup1_distrib: |
2570 lemma inf_Sup1_distrib: |
2571 assumes "finite A" and "A \<noteq> {}" |
2571 assumes "finite A" and "A \<noteq> {}" |
2572 shows "inf x (\<Squnion>\<^bsub>fin\<^esub>A) = \<Squnion>\<^bsub>fin\<^esub>{inf x a|a. a \<in> A}" |
2572 shows "inf x (\<Squnion>finA) = \<Squnion>fin{inf x a|a. a \<in> A}" |
2573 proof - |
2573 proof - |
2574 interpret ab_semigroup_idem_mult sup |
2574 interpret ab_semigroup_idem_mult sup |
2575 by (rule ab_semigroup_idem_mult_sup) |
2575 by (rule ab_semigroup_idem_mult_sup) |
2576 from assms show ?thesis |
2576 from assms show ?thesis |
2577 by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1]) |
2577 by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1]) |
2578 (rule arg_cong [where f="fold1 sup"], blast) |
2578 (rule arg_cong [where f="fold1 sup"], blast) |
2579 qed |
2579 qed |
2580 |
2580 |
2581 lemma inf_Sup2_distrib: |
2581 lemma inf_Sup2_distrib: |
2582 assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}" |
2582 assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}" |
2583 shows "inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B) = \<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B}" |
2583 shows "inf (\<Squnion>finA) (\<Squnion>finB) = \<Squnion>fin{inf a b|a b. a \<in> A \<and> b \<in> B}" |
2584 using A proof (induct rule: finite_ne_induct) |
2584 using A proof (induct rule: finite_ne_induct) |
2585 case singleton thus ?case |
2585 case singleton thus ?case |
2586 by(simp add: inf_Sup1_distrib [OF B] fold1_singleton_def [OF Sup_fin_def]) |
2586 by(simp add: inf_Sup1_distrib [OF B] fold1_singleton_def [OF Sup_fin_def]) |
2587 next |
2587 next |
2588 case (insert x A) |
2588 case (insert x A) |
2595 thus ?thesis by(simp add: insert(1) B(1)) |
2595 thus ?thesis by(simp add: insert(1) B(1)) |
2596 qed |
2596 qed |
2597 have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast |
2597 have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast |
2598 interpret ab_semigroup_idem_mult sup |
2598 interpret ab_semigroup_idem_mult sup |
2599 by (rule ab_semigroup_idem_mult_sup) |
2599 by (rule ab_semigroup_idem_mult_sup) |
2600 have "inf (\<Squnion>\<^bsub>fin\<^esub>(insert x A)) (\<Squnion>\<^bsub>fin\<^esub>B) = inf (sup x (\<Squnion>\<^bsub>fin\<^esub>A)) (\<Squnion>\<^bsub>fin\<^esub>B)" |
2600 have "inf (\<Squnion>fin(insert x A)) (\<Squnion>finB) = inf (sup x (\<Squnion>finA)) (\<Squnion>finB)" |
2601 using insert by (simp add: fold1_insert_idem_def [OF Sup_fin_def]) |
2601 using insert by (simp add: fold1_insert_idem_def [OF Sup_fin_def]) |
2602 also have "\<dots> = sup (inf x (\<Squnion>\<^bsub>fin\<^esub>B)) (inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B))" by(rule inf_sup_distrib2) |
2602 also have "\<dots> = sup (inf x (\<Squnion>finB)) (inf (\<Squnion>finA) (\<Squnion>finB))" by(rule inf_sup_distrib2) |
2603 also have "\<dots> = sup (\<Squnion>\<^bsub>fin\<^esub>{inf x b|b. b \<in> B}) (\<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B})" |
2603 also have "\<dots> = sup (\<Squnion>fin{inf x b|b. b \<in> B}) (\<Squnion>fin{inf a b|a b. a \<in> A \<and> b \<in> B})" |
2604 using insert by(simp add:inf_Sup1_distrib[OF B]) |
2604 using insert by(simp add:inf_Sup1_distrib[OF B]) |
2605 also have "\<dots> = \<Squnion>\<^bsub>fin\<^esub>({inf x b |b. b \<in> B} \<union> {inf a b |a b. a \<in> A \<and> b \<in> B})" |
2605 also have "\<dots> = \<Squnion>fin({inf x b |b. b \<in> B} \<union> {inf a b |a b. a \<in> A \<and> b \<in> B})" |
2606 (is "_ = \<Squnion>\<^bsub>fin\<^esub>?M") |
2606 (is "_ = \<Squnion>fin?M") |
2607 using B insert |
2607 using B insert |
2608 by (simp add: Sup_fin_def fold1_Un2 [OF finB _ finAB ne]) |
2608 by (simp add: Sup_fin_def fold1_Un2 [OF finB _ finAB ne]) |
2609 also have "?M = {inf a b |a b. a \<in> insert x A \<and> b \<in> B}" |
2609 also have "?M = {inf a b |a b. a \<in> insert x A \<and> b \<in> B}" |
2610 by blast |
2610 by blast |
2611 finally show ?case . |
2611 finally show ?case . |
2620 Coincidence on finite sets in complete lattices: |
2620 Coincidence on finite sets in complete lattices: |
2621 *} |
2621 *} |
2622 |
2622 |
2623 lemma Inf_fin_Inf: |
2623 lemma Inf_fin_Inf: |
2624 assumes "finite A" and "A \<noteq> {}" |
2624 assumes "finite A" and "A \<noteq> {}" |
2625 shows "\<Sqinter>\<^bsub>fin\<^esub>A = Inf A" |
2625 shows "\<Sqinter>finA = Inf A" |
2626 proof - |
2626 proof - |
2627 interpret ab_semigroup_idem_mult inf |
2627 interpret ab_semigroup_idem_mult inf |
2628 by (rule ab_semigroup_idem_mult_inf) |
2628 by (rule ab_semigroup_idem_mult_inf) |
2629 from assms show ?thesis |
2629 from assms show ?thesis |
2630 unfolding Inf_fin_def by (induct A set: finite) |
2630 unfolding Inf_fin_def by (induct A set: finite) |
2631 (simp_all add: Inf_insert_simp) |
2631 (simp_all add: Inf_insert_simp) |
2632 qed |
2632 qed |
2633 |
2633 |
2634 lemma Sup_fin_Sup: |
2634 lemma Sup_fin_Sup: |
2635 assumes "finite A" and "A \<noteq> {}" |
2635 assumes "finite A" and "A \<noteq> {}" |
2636 shows "\<Squnion>\<^bsub>fin\<^esub>A = Sup A" |
2636 shows "\<Squnion>finA = Sup A" |
2637 proof - |
2637 proof - |
2638 interpret ab_semigroup_idem_mult sup |
2638 interpret ab_semigroup_idem_mult sup |
2639 by (rule ab_semigroup_idem_mult_sup) |
2639 by (rule ab_semigroup_idem_mult_sup) |
2640 from assms show ?thesis |
2640 from assms show ?thesis |
2641 unfolding Sup_fin_def by (induct A set: finite) |
2641 unfolding Sup_fin_def by (induct A set: finite) |