Skip to content

Commit

Permalink
fixes #113
Browse files Browse the repository at this point in the history
  • Loading branch information
affeldt-aist committed Dec 2, 2024
1 parent a907a9e commit ce0fddd
Show file tree
Hide file tree
Showing 4 changed files with 23 additions and 11 deletions.
11 changes: 10 additions & 1 deletion CHANGELOG.md
Original file line number Diff line number Diff line change
@@ -1,6 +1,15 @@
# Changelog

Last release: [[1.5.2] - 2022-07-06](#152---2022-07-06)
Last release: [[2.1.0] - 2024-01-17](#210---2024-01-17)

## [2.1.0] - 2024-01-17

no documentation

## [2.0.0] - 2023-05-23

- in `finmap.v`:
+ lemmas `bigfcup_imfset`, `fbig_pred1_inj`

## [1.5.2] - 2022-07-06

Expand Down
6 changes: 3 additions & 3 deletions CHANGELOG_UNRELEASED.md
Original file line number Diff line number Diff line change
Expand Up @@ -4,11 +4,11 @@

### Added

- in `finmap.v`:
+ lemmas `bigfcup_imfset`, `fbig_pred1_inj`

### Changed

- in `finmap.v`
+ lemma `partition_disjoint_bigfcup` generalized

### Renamed

### Removed
Expand Down
15 changes: 9 additions & 6 deletions finmap.v
Original file line number Diff line number Diff line change
Expand Up @@ -5,7 +5,7 @@ From HB Require Import structures.
Set Warnings "-notation-incompatible-format".
From mathcomp Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq.
Set Warnings "notation-incompatible-format".
From mathcomp Require Import choice path finset finfun fintype bigop.
From mathcomp Require Import choice finset finfun fintype bigop.

(*****************************************************************************)
(* This file provides representations for finite sets over a choiceType K, *)
Expand Down Expand Up @@ -2607,21 +2607,24 @@ Qed.

Lemma partition_disjoint_bigfcup (f : T -> R) (F : I -> {fset T})
(K : {fset I}) :
(forall i j, i != j -> [disjoint F i & F j])%fset ->
(forall i j, i \in K -> j \in K -> i != j -> [disjoint F i & F j])%fset ->
\big[op/idx]_(i <- \big[fsetU/fset0]_(x <- K) (F x)) f i =
\big[op/idx]_(k <- K) (\big[op/idx]_(i <- F k) f i).
Proof.
move=> disjF; pose P := [fset F i | i in K & F i != fset0]%fset.
have trivP : trivIfset P.
apply/trivIfsetP => _ _ /imfsetP[i _ ->] /imfsetP[j _ ->] neqFij.
by apply: disjF; apply: contraNneq neqFij => ->.
apply/trivIfsetP => _ _ /imfsetP[i iK ->] /imfsetP[j jK ->] neqFij.
move: iK; rewrite !inE/= => /andP[iK Fi0].
move: jK; rewrite !inE/= => /andP[jK Fj0].
by apply: (disjF _ _ iK jK); apply: contraNneq neqFij => ->.
have -> : (\bigcup_(i <- K) F i)%fset = fcover P.
apply/esym; rewrite /P fcover_imfset big_mkcond /=; apply eq_bigr => i _.
by case: ifPn => // /negPn/eqP.
rewrite big_trivIfset // /rhs big_imfset => [|i j _ /andP[jK notFj0] eqFij] /=.
rewrite big_trivIfset // /rhs big_imfset => [|i j iK /andP[jK notFj0] eqFij] /=.
rewrite big_filter big_mkcond; apply eq_bigr => i _.
by case: ifPn => // /negPn /eqP ->; rewrite big_seq_fset0.
by apply: contraNeq (disjF _ _) _; rewrite -fsetI_eq0 eqFij fsetIid.
move: iK; rewrite !inE/= => /andP[iK Fi0].
by apply: contraNeq (disjF _ _ iK jK) _; rewrite -fsetI_eq0 eqFij fsetIid.
Qed.

End FsetBigOps.
Expand Down
2 changes: 1 addition & 1 deletion multiset.v
Original file line number Diff line number Diff line change
Expand Up @@ -7,7 +7,7 @@
Set Warnings "-notation-incompatible-format".
From mathcomp Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq.
Set Warnings "notation-incompatible-format".
From mathcomp Require Import choice path finset finfun fintype bigop tuple.
From mathcomp Require Import choice finset finfun fintype bigop tuple.
Require Import finmap.

(*****************************************************************************)
Expand Down

0 comments on commit ce0fddd

Please sign in to comment.