src/HOL/Library/Multiset.thy
Mon, 09 Jun 2025 22:14:38 +0200 haftmann more qualified auxiliary operations
Thu, 29 May 2025 14:17:09 +0200 haftmann added lemma
Fri, 02 May 2025 17:24:43 +0200 haftmann executable sorted_list_of_multiset
Tue, 22 Apr 2025 15:41:34 +0200 Manuel Eberl removed possible problematic dependency in HOL-Library.Multiset
Tue, 22 Apr 2025 14:14:00 +0200 Manuel Eberl HOL-Library: multisets of a given size
Tue, 04 Mar 2025 16:58:46 +0100 desharna added attribute "simp" to filter_mset_eq_mempty_iff
Tue, 04 Mar 2025 16:38:21 +0100 desharna removed lemma size_multiset_sum_mset[simp]
Tue, 04 Mar 2025 16:37:14 +0100 desharna added lemma size_mset_sum_mset_conv[simp] (thanks to Manuel Eberl)
Tue, 04 Mar 2025 16:07:55 +0100 desharna added lemma filter_mset_eq_mempty_iff (thanks to Manuel Eberl)
Tue, 04 Mar 2025 15:19:08 +0100 desharna renamed lemma filter_image_mset to filter_mset_image_mset
Tue, 04 Mar 2025 13:10:47 +0100 desharna added lemma sum_mset_image_mset_mono_strong
Tue, 04 Mar 2025 13:10:31 +0100 desharna added lemmas filter_mset_mono_strong, filter_mset_sum_list, set_mset_sum_list[simp] (thanks to Manuel Eberl)
Fri, 06 Dec 2024 20:26:33 +0100 wenzelm clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
Mon, 04 Nov 2024 11:21:19 +0100 wenzelm tuned proofs;
Sun, 03 Nov 2024 22:29:07 +0100 wenzelm tuned proofs;
Fri, 01 Nov 2024 14:10:52 +0000 paulson Library material from Eberl's Parallel_Shear_Sort
Sat, 19 Oct 2024 22:57:43 +0200 wenzelm tuned proofs;
Sat, 19 Oct 2024 19:00:19 +0200 wenzelm more type information;
Fri, 18 Oct 2024 14:20:09 +0200 wenzelm more inner-syntax markup;
Wed, 09 Oct 2024 23:38:29 +0200 wenzelm more inner-syntax markup;
Tue, 01 Oct 2024 20:39:16 +0200 wenzelm drop somewhat pointless 'syntax_consts' declarations;
Mon, 30 Sep 2024 23:32:26 +0200 wenzelm clarified syntax: use outer block (with indent);
Mon, 30 Sep 2024 20:30:59 +0200 wenzelm clarified inner-syntax markup, notably for enumerations: prefer "notation=mixfix" over "entity" via 'syntax_consts' (see also 70076ba563d2);
Fri, 20 Sep 2024 19:51:08 +0200 wenzelm standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
Wed, 28 Aug 2024 22:54:45 +0200 wenzelm more specific "args" syntax, to support more markup for syntax consts;
Sun, 25 Aug 2024 21:10:01 +0200 wenzelm more markup for syntax consts;
Mon, 08 Jul 2024 10:14:22 +0200 desharna added lemma image_mset_diff_if_inj
Mon, 08 Jul 2024 10:08:07 +0200 desharna added lemma minus_add_mset_if_not_in_lhs[simp]
Tue, 11 Jun 2024 10:27:35 +0200 desharna tuned proof
Mon, 10 Jun 2024 21:32:24 +0200 desharna renamed lemmas
Mon, 10 Jun 2024 14:09:55 +0200 desharna renamed theorems
Sat, 08 Jun 2024 14:57:14 +0200 desharna renamed lemmas
Fri, 12 Apr 2024 09:58:32 +0100 paulson Tidying ugly proofs
Fri, 29 Mar 2024 19:28:59 +0100 Manuel Eberl moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Sat, 23 Mar 2024 18:55:38 +0100 desharna redefined wf as an abbreviation for "wf_on UNIV"
Wed, 06 Mar 2024 10:39:45 +0100 blanchet more multiset lemmas
Mon, 05 Feb 2024 10:06:34 +0100 desharna added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
Tue, 23 May 2023 21:43:36 +0200 wenzelm more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
Mon, 08 May 2023 11:24:46 +0200 desharna added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
Wed, 12 Apr 2023 19:56:05 +0200 desharna added lemma multp_image_mset_image_msetI
Mon, 20 Mar 2023 18:33:56 +0100 desharna reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
Fri, 17 Mar 2023 13:56:54 +0100 desharna added lemma multp_repeat_mset_repeat_msetI
Mon, 23 Jan 2023 14:34:07 +0100 desharna added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
Thu, 22 Dec 2022 21:55:51 +0100 desharna merged
Tue, 20 Dec 2022 09:34:37 +0100 desharna used transp_on in assumptions of lemmas Multiset.bex_(least|greatest)_element
Mon, 19 Dec 2022 16:00:49 +0100 desharna strengthened and renamed lemmas preorder.transp_(ge|gr|le|less)
Tue, 20 Dec 2022 18:20:19 +0100 blanchet added lifting_forget as suggested by Peter Lammich
Sun, 18 Dec 2022 14:03:43 +0100 desharna added predicates asym_on and asymp_on and redefined asym and asymp to be abbreviations
Tue, 06 Dec 2022 19:17:05 +0100 desharna Strengthened multiset lemmas w.r.t. irrefl and irreflp
Wed, 07 Dec 2022 10:11:58 +0100 desharna stated goals of some lemmas explicitely to prevent silent changes
Wed, 23 Nov 2022 09:57:59 +0100 desharna strengthened and renamed irreflp_greater[simp] and irreflp_less[simp]
Fri, 28 Oct 2022 15:39:35 +0200 desharna added lemmas multp_mono_strong and mult_mono_strong
Thu, 20 Oct 2022 14:43:29 +0200 desharna tuned proof
Sat, 15 Oct 2022 16:34:19 +0200 desharna added lemma wfP_subset_mset[simp]
Mon, 27 Jun 2022 15:54:18 +0200 traytel strict bounds for BNFs (by Jan van Brügge)
Tue, 21 Jun 2022 14:21:55 +0200 desharna added lemmas monotone{,_on}_multp_multp_image_mset
Mon, 13 Jun 2022 20:02:00 +0200 desharna added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
Sat, 28 May 2022 10:45:45 +0200 desharna added lemmas Multiset.bex_{least,greatest}_element
Mon, 23 May 2022 10:13:08 +0200 desharna added lemma image_mset_filter_mset_swap
Fri, 20 May 2022 11:08:33 +0200 desharna added lemmas filter_mset_cong{0,}
Sun, 28 Nov 2021 19:12:48 +0100 desharna added wfP_less to wellorder and wfP_less_multiset
Sat, 27 Nov 2021 14:45:00 +0100 desharna added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
Sat, 27 Nov 2021 10:46:57 +0100 desharna redefined less_multiset to be based on multp
Sat, 27 Nov 2021 10:28:48 +0100 desharna added lemmas multp_code_eq_multp and multeqp_code_eq_reflclp_multp
Sat, 27 Nov 2021 10:22:42 +0100 desharna added lemmas multp_cancel, multp_cancel_add_mset, and multp_cancel_max
Sat, 27 Nov 2021 10:16:46 +0100 desharna added lemmas multp_implies_one_step, one_step_implies_multp, and subset_implies_multp
Sat, 27 Nov 2021 10:05:59 +0100 desharna added lemma wfP_multp
Sat, 27 Nov 2021 10:01:40 +0100 desharna added lemma mono_multp
Fri, 26 Nov 2021 11:14:10 +0100 desharna added Multiset.multp as predicate equivalent of Multiset.mult
Thu, 25 Nov 2021 14:02:51 +0100 desharna added asymp_{less,greater} to preorder and moved mult1_lessE out
Thu, 25 Nov 2021 12:19:50 +0100 desharna renamed multp_code_iff and multeqp_code_iff
Thu, 25 Nov 2021 12:13:49 +0100 desharna simplified mult_cancel_max and introduced orginal lemma as mult_cancel_max0
Thu, 25 Nov 2021 11:33:38 +0100 desharna renamed Multiset.multp and Multiset.multeqp
Sat, 30 Oct 2021 11:59:41 +0200 wenzelm clarified antiquotations;
Tue, 01 Jun 2021 19:46:34 +0200 nipkow More general fold function for maps
Mon, 17 May 2021 09:07:30 +0000 haftmann mere abbreviation for logical alias
Fri, 23 Apr 2021 09:50:14 +0000 haftmann collecting more lemmas concerning multisets
Mon, 22 Mar 2021 12:18:43 +0000 paulson merged
Mon, 22 Mar 2021 12:18:35 +0000 paulson type class relaxation
Mon, 22 Mar 2021 10:49:51 +0000 haftmann more lemmas
Thu, 18 Mar 2021 06:37:24 +0000 haftmann prefer more direct interpretation
Thu, 11 Mar 2021 07:05:38 +0000 haftmann avoid name clash
Sun, 07 Mar 2021 08:24:24 +0100 haftmann follow corresponding precedence on sets
Sat, 06 Mar 2021 18:42:10 +0000 haftmann consolidated names
Sun, 28 Feb 2021 20:13:07 +0000 haftmann more connections between mset _ = mset _ and permutations
Wed, 24 Feb 2021 13:31:33 +0000 haftmann multiset as equivalence class of permuted lists
Mon, 22 Feb 2021 07:49:48 +0000 haftmann get rid of traditional predicate
Tue, 05 Jan 2021 03:35:46 +0100 Manuel Eberl HOL-Library.Multiset: new notation for prod_mset, consistent with sum_mset
Mon, 04 Jan 2021 19:41:38 +0100 Manuel Eberl HOL-Library: Changed notation for sum_mset
Sun, 15 Nov 2020 07:17:06 +0000 haftmann bundles for reflected term syntax
Thu, 12 Nov 2020 09:06:44 +0100 haftmann bundled syntax for state monad combinators
Wed, 03 Jun 2020 11:44:21 +0200 nipkow should have been copied across from Set.thy as well for better printing
Tue, 21 Jan 2020 11:02:27 +0100 Manuel Eberl Removed multiplicativity assumption from normalization_semidom
Sun, 10 Mar 2019 23:23:03 +0100 wenzelm more formal contributors (with the help of the history);
Sun, 06 Jan 2019 15:04:34 +0100 wenzelm isabelle update -u path_cartouches;
Fri, 04 Jan 2019 23:22:53 +0100 wenzelm isabelle update -u control_cartouches;
Mon, 10 Dec 2018 20:35:08 +0100 wenzelm tuned proofs;
Thu, 08 Nov 2018 09:11:52 +0100 haftmann removed relics of ASCII syntax for indexed big operators
Wed, 03 Oct 2018 09:46:42 +0200 nipkow shuffle -> shuffles
Sun, 23 Sep 2018 12:50:12 +0200 nipkow use standard syntax
Fri, 14 Sep 2018 07:31:55 +0200 nipkow tuned
Thu, 13 Sep 2018 16:15:05 +0200 nipkow removed redundant lemma
Thu, 13 Sep 2018 13:09:03 +0200 nipkow more simp lemmas
Thu, 13 Sep 2018 08:36:51 +0200 nipkow prefer explicit
Thu, 13 Sep 2018 06:36:00 +0200 nipkow typo
Wed, 12 Sep 2018 18:44:31 +0200 nipkow added spaces because otherwise nonatomic arguments look awful: BIGf x -> BIG f x
Sat, 08 Sep 2018 08:08:28 +0000 haftmann left-over rename from 3f9bb52082c4
Thu, 07 Jun 2018 19:36:12 +0200 nipkow utilize 'flip'
Wed, 06 Jun 2018 11:12:37 +0200 nipkow Keep filter input syntax
Tue, 22 May 2018 11:08:37 +0200 nipkow First step to remove nonstandard "[x <- xs. P]" syntax: only input
Mon, 19 Feb 2018 13:56:16 +0100 nipkow added lemma
Wed, 10 Jan 2018 15:21:49 +0100 nipkow Manual updates towards conversion of "op" syntax
Wed, 03 Jan 2018 11:06:13 +0100 blanchet kill old size infrastructure
Sat, 11 Nov 2017 18:41:08 +0000 haftmann dedicated definition for coprimality
Mon, 30 Oct 2017 13:18:44 +0000 haftmann generalized some lemmas on multisets
Thu, 24 Aug 2017 10:47:56 +0200 blanchet tuning (proofs and code)
Tue, 15 Aug 2017 19:47:08 +0200 nipkow added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
Tue, 15 Aug 2017 09:29:35 +0200 nipkow added Min_mset and Max_mset
Thu, 03 Aug 2017 23:03:44 +0200 nipkow tuned
Sat, 15 Jul 2017 14:32:02 +0100 eberlm More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
less more (0) -120 tip