| author | paulson | 
| Wed, 08 Dec 1999 13:51:44 +0100 | |
| changeset 8053 | 37ebdaf3bb91 | 
| parent 7982 | d534b897ce39 | 
| child 8281 | 188e2924433e | 
| permissions | -rw-r--r-- | 
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 1 | (* Title: HOL/Isar_examples/MultisetOrder.thy | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 3 | Author: Markus Wenzel | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 4 | |
| 7442 | 5 | Wellfoundedness proof for the multiset order. | 
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 6 | *) | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 7 | |
| 7761 | 8 | header {* Wellfoundedness of multiset ordering *};
 | 
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 9 | |
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 10 | theory MultisetOrder = Multiset:; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 11 | |
| 7968 | 12 | text_raw {*
 | 
| 7982 | 13 |  \footnote{Original tactic script by Tobias Nipkow (see
 | 
| 7968 | 14 |  \url{http://isabelle.in.tum.de/library/HOL/Induct/Multiset.html}),
 | 
| 15 | based on a pen-and-paper proof due to Wilfried Buchholz.} | |
| 16 | *}; | |
| 7527 | 17 | |
| 7761 | 18 | subsection {* A technical lemma *};
 | 
| 19 | ||
| 7527 | 20 | lemma less_add: "(N, M0 + {#a#}) : mult1 r ==>
 | 
| 21 |     (EX M. (M, M0) : mult1 r & N = M + {#a#}) |
 | |
| 22 | (EX K. (ALL b. elem K b --> (b, a) : r) & N = M0 + K)" | |
| 23 | (concl is "?case1 (mult1 r) | ?case2"); | |
| 24 | proof (unfold mult1_def); | |
| 7982 | 25 | let ?r = "\<lambda>K a. ALL b. elem K b --> (b, a) : r"; | 
| 26 |   let ?R = "\<lambda>N M. EX a M0 K. M = M0 + {#a#} & N = M0 + K & ?r K a";
 | |
| 7527 | 27 |   let ?case1 = "?case1 {(N, M). ?R N M}";
 | 
| 28 | ||
| 29 |   assume "(N, M0 + {#a#}) : {(N, M). ?R N M}";
 | |
| 7800 | 30 | hence "EX a' M0' K. | 
| 31 |       M0 + {#a#} = M0' + {#a'#} & N = M0' + K & ?r K a'"; by simp;
 | |
| 7527 | 32 | thus "?case1 | ?case2"; | 
| 33 | proof (elim exE conjE); | |
| 7874 | 34 | fix a' M0' K; | 
| 35 | assume N: "N = M0' + K" and r: "?r K a'"; | |
| 7527 | 36 |     assume "M0 + {#a#} = M0' + {#a'#}";
 | 
| 7800 | 37 | hence "M0 = M0' & a = a' | | 
| 38 |         (EX K'. M0 = K' + {#a'#} & M0' = K' + {#a#})";
 | |
| 7527 | 39 | by (simp only: add_eq_conv_ex); | 
| 40 | thus ?thesis; | |
| 41 | proof (elim disjE conjE exE); | |
| 42 | assume "M0 = M0'" "a = a'"; | |
| 43 | with N r; have "?r K a & N = M0 + K"; by simp; | |
| 44 | hence ?case2; ..; thus ?thesis; ..; | |
| 45 | next; | |
| 46 | fix K'; | |
| 47 |       assume "M0' = K' + {#a#}";
 | |
| 48 |       with N; have n: "N = K' + K + {#a#}"; by (simp add: union_ac);
 | |
| 49 | ||
| 50 |       assume "M0 = K' + {#a'#}";
 | |
| 7565 | 51 | with r; have "?R (K' + K) M0"; by blast; | 
| 7527 | 52 | with n; have ?case1; by simp; thus ?thesis; ..; | 
| 53 | qed; | |
| 54 | qed; | |
| 55 | qed; | |
| 56 | ||
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 57 | |
| 7761 | 58 | subsection {* The key property *};
 | 
| 59 | ||
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 60 | lemma all_accessible: "wf r ==> ALL M. M : acc (mult1 r)"; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 61 | proof; | 
| 7480 | 62 | let ?R = "mult1 r"; | 
| 63 | let ?W = "acc ?R"; | |
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 64 |   {{;
 | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 65 | fix M M0 a; | 
| 7800 | 66 | assume M0: "M0 : ?W" | 
| 67 |       and wf_hyp: "ALL b. (b, a) : r --> (ALL M:?W. M + {#b#} : ?W)"
 | |
| 7480 | 68 |       and acc_hyp: "ALL M. (M, M0) : ?R --> M + {#a#} : ?W";
 | 
| 69 |     have "M0 + {#a#} : ?W";
 | |
| 7442 | 70 |     proof (rule accI [of "M0 + {#a#}"]);
 | 
| 7874 | 71 | fix N; | 
| 72 |       assume "(N, M0 + {#a#}) : ?R";
 | |
| 7480 | 73 |       hence "((EX M. (M, M0) : ?R & N = M + {#a#}) |
 | 
| 7800 | 74 | (EX K. (ALL b. elem K b --> (b, a) : r) & N = M0 + K))"; | 
| 7527 | 75 | by (rule less_add); | 
| 7480 | 76 | thus "N : ?W"; | 
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 77 | proof (elim exE disjE conjE); | 
| 7480 | 78 | 	fix M; assume "(M, M0) : ?R" and N: "N = M + {#a#}";
 | 
| 79 | 	from acc_hyp; have "(M, M0) : ?R --> M + {#a#} : ?W"; ..;
 | |
| 80 | 	hence "M + {#a#} : ?W"; ..;
 | |
| 81 | thus "N : ?W"; by (simp only: N); | |
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 82 | next; | 
| 7451 | 83 | fix K; | 
| 84 | assume N: "N = M0 + K"; | |
| 85 | assume "ALL b. elem K b --> (b, a) : r"; | |
| 7480 | 86 | have "?this --> M0 + K : ?W" (is "?P K"); | 
| 7527 | 87 | proof (induct K rule: multiset_induct); | 
| 7480 | 88 | 	  from M0; have "M0 + {#} : ?W"; by simp;
 | 
| 89 | 	  thus "?P {#}"; ..;
 | |
| 7442 | 90 | |
| 7480 | 91 | fix K x; assume hyp: "?P K"; | 
| 92 | 	  show "?P (K + {#x#})";
 | |
| 7442 | 93 | proof; | 
| 7451 | 94 | 	    assume a: "ALL b. elem (K + {#x#}) b --> (b, a) : r";
 | 
| 7442 | 95 | hence "(x, a) : r"; by simp; | 
| 7800 | 96 | 	    with wf_hyp; have b: "ALL M:?W. M + {#x#} : ?W"; by blast;
 | 
| 7442 | 97 | |
| 7480 | 98 | from a hyp; have "M0 + K : ?W"; by simp; | 
| 99 | 	    with b; have "(M0 + K) + {#x#} : ?W"; ..;
 | |
| 100 | 	    thus "M0 + (K + {#x#}) : ?W"; by (simp only: union_assoc);
 | |
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 101 | qed; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 102 | qed; | 
| 7480 | 103 | hence "M0 + K : ?W"; ..; | 
| 104 | thus "N : ?W"; by (simp only: N); | |
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 105 | qed; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 106 | qed; | 
| 7442 | 107 | }}; note tedious_reasoning = this; | 
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 108 | |
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 109 | assume wf: "wf r"; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 110 | fix M; | 
| 7480 | 111 | show "M : ?W"; | 
| 7527 | 112 | proof (induct M rule: multiset_induct); | 
| 7480 | 113 |     show "{#} : ?W";
 | 
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 114 | proof (rule accI); | 
| 7480 | 115 |       fix b; assume "(b, {#}) : ?R";
 | 
| 116 | with not_less_empty; show "b : ?W"; by contradiction; | |
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 117 | qed; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 118 | |
| 7480 | 119 | fix M a; assume "M : ?W"; | 
| 120 |     from wf; have "ALL M:?W. M + {#a#} : ?W";
 | |
| 7442 | 121 | proof (rule wf_induct [of r]); | 
| 7800 | 122 | fix a; | 
| 123 |       assume "ALL b. (b, a) : r --> (ALL M:?W. M + {#b#} : ?W)";
 | |
| 7480 | 124 |       show "ALL M:?W. M + {#a#} : ?W";
 | 
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 125 | proof; | 
| 7480 | 126 | fix M; assume "M : ?W"; | 
| 7800 | 127 | 	thus "M + {#a#} : ?W";
 | 
| 128 | by (rule acc_induct) (rule tedious_reasoning); | |
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 129 | qed; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 130 | qed; | 
| 7480 | 131 |     thus "M + {#a#} : ?W"; ..;
 | 
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 132 | qed; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 133 | qed; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 134 | |
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 135 | |
| 7761 | 136 | subsection {* Main result *};
 | 
| 137 | ||
| 7432 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 138 | theorem wf_mult1: "wf r ==> wf (mult1 r)"; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 139 | by (rule acc_wfI, rule all_accessible); | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 140 | |
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 141 | theorem wf_mult: "wf r ==> wf (mult r)"; | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 142 | by (unfold mult_def, rule wf_trancl, rule wf_mult1); | 
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 143 | |
| 
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
 wenzelm parents: diff
changeset | 144 | end; |