src/HOL/Library/Multiset.thy
changeset 22270 4ccb7e6be929
parent 21417 13c33ad15303
child 22316 f662831459de
equal deleted inserted replaced
22269:7c1e65897693 22270:4ccb7e6be929
   379 subsection {* Multiset orderings *}
   379 subsection {* Multiset orderings *}
   380 
   380 
   381 subsubsection {* Well-foundedness *}
   381 subsubsection {* Well-foundedness *}
   382 
   382 
   383 definition
   383 definition
   384   mult1 :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
   384   mult1 :: "('a => 'a => bool) => 'a multiset => 'a multiset => bool" where
   385   "mult1 r =
   385   "mult1 r =
   386     {(N, M). \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and>
   386     (%N M. \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and>
   387       (\<forall>b. b :# K --> (b, a) \<in> r)}"
   387       (\<forall>b. b :# K --> r b a))"
   388 
   388 
   389 definition
   389 definition
   390   mult :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
   390   mult :: "('a => 'a => bool) => 'a multiset => 'a multiset => bool" where
   391   "mult r = (mult1 r)\<^sup>+"
   391   "mult r = (mult1 r)\<^sup>+\<^sup>+"
   392 
   392 
   393 lemma not_less_empty [iff]: "(M, {#}) \<notin> mult1 r"
   393 lemma not_less_empty [iff]: "\<not> mult1 r M {#}"
   394   by (simp add: mult1_def)
   394   by (simp add: mult1_def)
   395 
   395 
   396 lemma less_add: "(N, M0 + {#a#}) \<in> mult1 r ==>
   396 lemma less_add: "mult1 r N (M0 + {#a#})==>
   397     (\<exists>M. (M, M0) \<in> mult1 r \<and> N = M + {#a#}) \<or>
   397     (\<exists>M. mult1 r M M0 \<and> N = M + {#a#}) \<or>
   398     (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K)"
   398     (\<exists>K. (\<forall>b. b :# K --> r b a) \<and> N = M0 + K)"
   399   (is "_ \<Longrightarrow> ?case1 (mult1 r) \<or> ?case2")
   399   (is "_ \<Longrightarrow> ?case1 (mult1 r) \<or> ?case2")
   400 proof (unfold mult1_def)
   400 proof (unfold mult1_def)
   401   let ?r = "\<lambda>K a. \<forall>b. b :# K --> (b, a) \<in> r"
   401   let ?r = "\<lambda>K a. \<forall>b. b :# K --> r b a"
   402   let ?R = "\<lambda>N M. \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and> ?r K a"
   402   let ?R = "\<lambda>N M. \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and> ?r K a"
   403   let ?case1 = "?case1 {(N, M). ?R N M}"
   403   let ?case1 = "?case1 ?R"
   404 
   404 
   405   assume "(N, M0 + {#a#}) \<in> {(N, M). ?R N M}"
   405   assume "?R N (M0 + {#a#})"
   406   then have "\<exists>a' M0' K.
   406   then have "\<exists>a' M0' K.
   407       M0 + {#a#} = M0' + {#a'#} \<and> N = M0' + K \<and> ?r K a'" by simp
   407       M0 + {#a#} = M0' + {#a'#} \<and> N = M0' + K \<and> ?r K a'" by simp
   408   then show "?case1 \<or> ?case2"
   408   then show "?case1 \<or> ?case2"
   409   proof (elim exE conjE)
   409   proof (elim exE conjE)
   410     fix a' M0' K
   410     fix a' M0' K
   428       with n have ?case1 by simp then show ?thesis ..
   428       with n have ?case1 by simp then show ?thesis ..
   429     qed
   429     qed
   430   qed
   430   qed
   431 qed
   431 qed
   432 
   432 
   433 lemma all_accessible: "wf r ==> \<forall>M. M \<in> acc (mult1 r)"
   433 lemma all_accessible: "wfP r ==> \<forall>M. acc (mult1 r) M"
   434 proof
   434 proof
   435   let ?R = "mult1 r"
   435   let ?R = "mult1 r"
   436   let ?W = "acc ?R"
   436   let ?W = "acc ?R"
   437   {
   437   {
   438     fix M M0 a
   438     fix M M0 a
   439     assume M0: "M0 \<in> ?W"
   439     assume M0: "?W M0"
   440       and wf_hyp: "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
   440       and wf_hyp: "!!b. r b a ==> \<forall>M \<triangleright> ?W. ?W (M + {#b#})"
   441       and acc_hyp: "\<forall>M. (M, M0) \<in> ?R --> M + {#a#} \<in> ?W"
   441       and acc_hyp: "\<forall>M. ?R M M0 --> ?W (M + {#a#})"
   442     have "M0 + {#a#} \<in> ?W"
   442     have "?W (M0 + {#a#})"
   443     proof (rule accI [of "M0 + {#a#}"])
   443     proof (rule accI [of _ "M0 + {#a#}"])
   444       fix N
   444       fix N
   445       assume "(N, M0 + {#a#}) \<in> ?R"
   445       assume "?R N (M0 + {#a#})"
   446       then have "((\<exists>M. (M, M0) \<in> ?R \<and> N = M + {#a#}) \<or>
   446       then have "((\<exists>M. ?R M M0 \<and> N = M + {#a#}) \<or>
   447           (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K))"
   447           (\<exists>K. (\<forall>b. b :# K --> r b a) \<and> N = M0 + K))"
   448         by (rule less_add)
   448         by (rule less_add)
   449       then show "N \<in> ?W"
   449       then show "?W N"
   450       proof (elim exE disjE conjE)
   450       proof (elim exE disjE conjE)
   451         fix M assume "(M, M0) \<in> ?R" and N: "N = M + {#a#}"
   451         fix M assume "?R M M0" and N: "N = M + {#a#}"
   452         from acc_hyp have "(M, M0) \<in> ?R --> M + {#a#} \<in> ?W" ..
   452         from acc_hyp have "?R M M0 --> ?W (M + {#a#})" ..
   453         then have "M + {#a#} \<in> ?W" ..
   453         then have "?W (M + {#a#})" ..
   454         then show "N \<in> ?W" by (simp only: N)
   454         then show "?W N" by (simp only: N)
   455       next
   455       next
   456         fix K
   456         fix K
   457         assume N: "N = M0 + K"
   457         assume N: "N = M0 + K"
   458         assume "\<forall>b. b :# K --> (b, a) \<in> r"
   458         assume "\<forall>b. b :# K --> r b a"
   459         then have "M0 + K \<in> ?W"
   459         then have "?W (M0 + K)"
   460         proof (induct K)
   460         proof (induct K)
   461           case empty
   461           case empty
   462           from M0 show "M0 + {#} \<in> ?W" by simp
   462           from M0 show "?W (M0 + {#})" by simp
   463         next
   463         next
   464           case (add K x)
   464           case (add K x)
   465           from add.prems have "(x, a) \<in> r" by simp
   465           from add.prems have "r x a" by simp
   466           with wf_hyp have "\<forall>M \<in> ?W. M + {#x#} \<in> ?W" by blast
   466           with wf_hyp have "\<forall>M \<triangleright> ?W. ?W (M + {#x#})" by blast
   467           moreover from add have "M0 + K \<in> ?W" by simp
   467           moreover from add have "?W (M0 + K)" by simp
   468           ultimately have "(M0 + K) + {#x#} \<in> ?W" ..
   468           ultimately have "?W ((M0 + K) + {#x#})" ..
   469           then show "M0 + (K + {#x#}) \<in> ?W" by (simp only: union_assoc)
   469           then show "?W (M0 + (K + {#x#}))" by (simp only: union_assoc)
   470         qed
   470         qed
   471         then show "N \<in> ?W" by (simp only: N)
   471         then show "?W N" by (simp only: N)
   472       qed
   472       qed
   473     qed
   473     qed
   474   } note tedious_reasoning = this
   474   } note tedious_reasoning = this
   475 
   475 
   476   assume wf: "wf r"
   476   assume wf: "wfP r"
   477   fix M
   477   fix M
   478   show "M \<in> ?W"
   478   show "?W M"
   479   proof (induct M)
   479   proof (induct M)
   480     show "{#} \<in> ?W"
   480     show "?W {#}"
   481     proof (rule accI)
   481     proof (rule accI)
   482       fix b assume "(b, {#}) \<in> ?R"
   482       fix b assume "?R b {#}"
   483       with not_less_empty show "b \<in> ?W" by contradiction
   483       with not_less_empty show "?W b" by contradiction
   484     qed
   484     qed
   485 
   485 
   486     fix M a assume "M \<in> ?W"
   486     fix M a assume "?W M"
   487     from wf have "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
   487     from wf have "\<forall>M \<triangleright> ?W. ?W (M + {#a#})"
   488     proof induct
   488     proof induct
   489       fix a
   489       fix a
   490       assume "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
   490       assume "!!b. r b a ==> \<forall>M \<triangleright> ?W. ?W (M + {#b#})"
   491       show "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
   491       show "\<forall>M \<triangleright> ?W. ?W (M + {#a#})"
   492       proof
   492       proof
   493         fix M assume "M \<in> ?W"
   493         fix M assume "?W M"
   494         then show "M + {#a#} \<in> ?W"
   494         then show "?W (M + {#a#})"
   495           by (rule acc_induct) (rule tedious_reasoning)
   495           by (rule acc_induct) (rule tedious_reasoning)
   496       qed
   496       qed
   497     qed
   497     qed
   498     then show "M + {#a#} \<in> ?W" ..
   498     then show "?W (M + {#a#})" ..
   499   qed
   499   qed
   500 qed
   500 qed
   501 
   501 
   502 theorem wf_mult1: "wf r ==> wf (mult1 r)"
   502 theorem wf_mult1: "wfP r ==> wfP (mult1 r)"
   503   by (rule acc_wfI, rule all_accessible)
   503   by (rule acc_wfI, rule all_accessible)
   504 
   504 
   505 theorem wf_mult: "wf r ==> wf (mult r)"
   505 theorem wf_mult: "wfP r ==> wfP (mult r)"
   506   by (unfold mult_def, rule wf_trancl, rule wf_mult1)
   506   by (unfold mult_def, rule wfP_trancl, rule wf_mult1)
   507 
   507 
   508 
   508 
   509 subsubsection {* Closure-free presentation *}
   509 subsubsection {* Closure-free presentation *}
   510 
   510 
   511 (*Badly needed: a linear arithmetic procedure for multisets*)
   511 (*Badly needed: a linear arithmetic procedure for multisets*)
   514 by (simp add: multiset_eq_conv_count_eq)
   514 by (simp add: multiset_eq_conv_count_eq)
   515 
   515 
   516 text {* One direction. *}
   516 text {* One direction. *}
   517 
   517 
   518 lemma mult_implies_one_step:
   518 lemma mult_implies_one_step:
   519   "trans r ==> (M, N) \<in> mult r ==>
   519   "transP r ==> mult r M N ==>
   520     \<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and>
   520     \<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and>
   521     (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r)"
   521     (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. r k j)"
   522   apply (unfold mult_def mult1_def set_of_def)
   522   apply (unfold mult_def mult1_def set_of_def)
   523   apply (erule converse_trancl_induct, clarify)
   523   apply (erule converse_trancl_induct', clarify)
   524    apply (rule_tac x = M0 in exI, simp, clarify)
   524    apply (rule_tac x = M0 in exI, simp, clarify)
   525   apply (case_tac "a :# K")
   525   apply (case_tac "a :# Ka")
   526    apply (rule_tac x = I in exI)
   526    apply (rule_tac x = I in exI)
   527    apply (simp (no_asm))
   527    apply (simp (no_asm))
   528    apply (rule_tac x = "(K - {#a#}) + Ka" in exI)
   528    apply (rule_tac x = "(Ka - {#a#}) + K" in exI)
   529    apply (simp (no_asm_simp) add: union_assoc [symmetric])
   529    apply (simp (no_asm_simp) add: union_assoc [symmetric])
   530    apply (drule_tac f = "\<lambda>M. M - {#a#}" in arg_cong)
   530    apply (drule_tac f = "\<lambda>M. M - {#a#}" in arg_cong)
   531    apply (simp add: diff_union_single_conv)
   531    apply (simp add: diff_union_single_conv)
   532    apply (simp (no_asm_use) add: trans_def)
   532    apply (simp (no_asm_use) add: trans_def)
   533    apply blast
   533    apply blast
   554   apply (erule size_eq_Suc_imp_elem [THEN exE])
   554   apply (erule size_eq_Suc_imp_elem [THEN exE])
   555   apply (drule elem_imp_eq_diff_union, auto)
   555   apply (drule elem_imp_eq_diff_union, auto)
   556   done
   556   done
   557 
   557 
   558 lemma one_step_implies_mult_aux:
   558 lemma one_step_implies_mult_aux:
   559   "trans r ==>
   559   "\<forall>I J K. (size J = n \<and> J \<noteq> {#} \<and> (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. r k j))
   560     \<forall>I J K. (size J = n \<and> J \<noteq> {#} \<and> (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r))
   560     --> mult r (I + K) (I + J)"
   561       --> (I + K, I + J) \<in> mult r"
       
   562   apply (induct_tac n, auto)
   561   apply (induct_tac n, auto)
   563   apply (frule size_eq_Suc_imp_eq_union, clarify)
   562   apply (frule size_eq_Suc_imp_eq_union, clarify)
   564   apply (rename_tac "J'", simp)
   563   apply (rename_tac "J'", simp)
   565   apply (erule notE, auto)
   564   apply (erule notE, auto)
   566   apply (case_tac "J' = {#}")
   565   apply (case_tac "J' = {#}")
   567    apply (simp add: mult_def)
   566    apply (simp add: mult_def)
   568    apply (rule r_into_trancl)
   567    apply (rule trancl.r_into_trancl)
   569    apply (simp add: mult1_def set_of_def, blast)
   568    apply (simp add: mult1_def set_of_def, blast)
   570   txt {* Now we know @{term "J' \<noteq> {#}"}. *}
   569   txt {* Now we know @{term "J' \<noteq> {#}"}. *}
   571   apply (cut_tac M = K and P = "\<lambda>x. (x, a) \<in> r" in multiset_partition)
   570   apply (cut_tac M = K and P = "\<lambda>x. r x a" in multiset_partition)
   572   apply (erule_tac P = "\<forall>k \<in> set_of K. ?P k" in rev_mp)
   571   apply (erule_tac P = "\<forall>k \<in> set_of K. ?P k" in rev_mp)
   573   apply (erule ssubst)
   572   apply (erule ssubst)
   574   apply (simp add: Ball_def, auto)
   573   apply (simp add: Ball_def, auto)
   575   apply (subgoal_tac
   574   apply (subgoal_tac
   576     "((I + {# x : K. (x, a) \<in> r #}) + {# x : K. (x, a) \<notin> r #},
   575     "mult r ((I + {# x : K. r x a #}) + {# x : K. \<not> r x a #})
   577       (I + {# x : K. (x, a) \<in> r #}) + J') \<in> mult r")
   576       ((I + {# x : K. r x a #}) + J')")
   578    prefer 2
   577    prefer 2
   579    apply force
   578    apply force
   580   apply (simp (no_asm_use) add: union_assoc [symmetric] mult_def)
   579   apply (simp (no_asm_use) add: union_assoc [symmetric] mult_def)
   581   apply (erule trancl_trans)
   580   apply (erule trancl_trans')
   582   apply (rule r_into_trancl)
   581   apply (rule trancl.r_into_trancl)
   583   apply (simp add: mult1_def set_of_def)
   582   apply (simp add: mult1_def set_of_def)
   584   apply (rule_tac x = a in exI)
   583   apply (rule_tac x = a in exI)
   585   apply (rule_tac x = "I + J'" in exI)
   584   apply (rule_tac x = "I + J'" in exI)
   586   apply (simp add: union_ac)
   585   apply (simp add: union_ac)
   587   done
   586   done
   588 
   587 
   589 lemma one_step_implies_mult:
   588 lemma one_step_implies_mult:
   590   "trans r ==> J \<noteq> {#} ==> \<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r
   589   "J \<noteq> {#} ==> \<forall>k \<in> set_of K. \<exists>j \<in> set_of J. r k j
   591     ==> (I + K, I + J) \<in> mult r"
   590     ==> mult r (I + K) (I + J)"
   592   apply (insert one_step_implies_mult_aux, blast)
   591   apply (insert one_step_implies_mult_aux, blast)
   593   done
   592   done
   594 
   593 
   595 
   594 
   596 subsubsection {* Partial-order properties *}
   595 subsubsection {* Partial-order properties *}
   597 
   596 
   598 instance multiset :: (type) ord ..
   597 instance multiset :: (type) ord ..
   599 
   598 
   600 defs (overloaded)
   599 defs (overloaded)
   601   less_multiset_def: "M' < M == (M', M) \<in> mult {(x', x). x' < x}"
   600   less_multiset_def: "op < == mult op <"
   602   le_multiset_def: "M' <= M == M' = M \<or> M' < (M::'a multiset)"
   601   le_multiset_def: "M' <= M == M' = M \<or> M' < (M::'a multiset)"
   603 
   602 
   604 lemma trans_base_order: "trans {(x', x). x' < (x::'a::order)}"
   603 lemma trans_base_order: "transP (op < :: 'a::order => 'a => bool)"
   605   unfolding trans_def by (blast intro: order_less_trans)
   604   unfolding trans_def by (blast intro: order_less_trans)
   606 
   605 
   607 text {*
   606 text {*
   608  \medskip Irreflexivity.
   607  \medskip Irreflexivity.
   609 *}
   608 *}
   627 
   626 
   628 text {* Transitivity. *}
   627 text {* Transitivity. *}
   629 
   628 
   630 theorem mult_less_trans: "K < M ==> M < N ==> K < (N::'a::order multiset)"
   629 theorem mult_less_trans: "K < M ==> M < N ==> K < (N::'a::order multiset)"
   631   apply (unfold less_multiset_def mult_def)
   630   apply (unfold less_multiset_def mult_def)
   632   apply (blast intro: trancl_trans)
   631   apply (blast intro: trancl_trans')
   633   done
   632   done
   634 
   633 
   635 text {* Asymmetry. *}
   634 text {* Asymmetry. *}
   636 
   635 
   637 theorem mult_less_not_sym: "M < N ==> \<not> N < (M::'a::order multiset)"
   636 theorem mult_less_not_sym: "M < N ==> \<not> N < (M::'a::order multiset)"
   674 
   673 
   675 
   674 
   676 subsubsection {* Monotonicity of multiset union *}
   675 subsubsection {* Monotonicity of multiset union *}
   677 
   676 
   678 lemma mult1_union:
   677 lemma mult1_union:
   679     "(B, D) \<in> mult1 r ==> trans r ==> (C + B, C + D) \<in> mult1 r"
   678     "mult1 r B D ==> mult1 r (C + B) (C + D)"
   680   apply (unfold mult1_def, auto)
   679   apply (unfold mult1_def, auto)
   681   apply (rule_tac x = a in exI)
   680   apply (rule_tac x = a in exI)
   682   apply (rule_tac x = "C + M0" in exI)
   681   apply (rule_tac x = "C + M0" in exI)
   683   apply (simp add: union_assoc)
   682   apply (simp add: union_assoc)
   684   done
   683   done
   685 
   684 
   686 lemma union_less_mono2: "B < D ==> C + B < C + (D::'a::order multiset)"
   685 lemma union_less_mono2: "B < D ==> C + B < C + (D::'a::order multiset)"
   687   apply (unfold less_multiset_def mult_def)
   686   apply (unfold less_multiset_def mult_def)
   688   apply (erule trancl_induct)
   687   apply (erule trancl_induct')
   689    apply (blast intro: mult1_union transI order_less_trans r_into_trancl)
   688    apply (blast intro: mult1_union)
   690   apply (blast intro: mult1_union transI order_less_trans r_into_trancl trancl_trans)
   689   apply (blast intro: mult1_union trancl.r_into_trancl trancl_trans')
   691   done
   690   done
   692 
   691 
   693 lemma union_less_mono1: "B < D ==> B + C < D + (C::'a::order multiset)"
   692 lemma union_less_mono1: "B < D ==> B + C < D + (C::'a::order multiset)"
   694   apply (subst union_commute [of B C])
   693   apply (subst union_commute [of B C])
   695   apply (subst union_commute [of D C])
   694   apply (subst union_commute [of D C])
   708 
   707 
   709 lemma empty_leI [iff]: "{#} <= (M::'a::order multiset)"
   708 lemma empty_leI [iff]: "{#} <= (M::'a::order multiset)"
   710   apply (unfold le_multiset_def less_multiset_def)
   709   apply (unfold le_multiset_def less_multiset_def)
   711   apply (case_tac "M = {#}")
   710   apply (case_tac "M = {#}")
   712    prefer 2
   711    prefer 2
   713    apply (subgoal_tac "({#} + {#}, {#} + M) \<in> mult (Collect (split op <))")
   712    apply (subgoal_tac "mult op < ({#} + {#}) ({#} + M)")
   714     prefer 2
   713     prefer 2
   715     apply (rule one_step_implies_mult)
   714     apply (rule one_step_implies_mult)
   716       apply (simp only: trans_def, auto)
   715       apply auto
   717   done
   716   done
   718 
   717 
   719 lemma union_upper1: "A <= A + (B::'a::order multiset)"
   718 lemma union_upper1: "A <= A + (B::'a::order multiset)"
   720 proof -
   719 proof -
   721   have "A + {#} <= A + B" by (blast intro: union_le_mono)
   720   have "A + {#} <= A + B" by (blast intro: union_le_mono)