author | wenzelm |
Fri, 06 Oct 2000 17:35:58 +0200 | |
changeset 10168 | 50be659d4222 |
parent 10144 | fe2a4e018dbf |
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 |
|
10007 | 8 |
header {* Wellfoundedness of multiset ordering *} |
7432
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
wenzelm
parents:
diff
changeset
|
9 |
|
10007 | 10 |
theory MultisetOrder = Multiset: |
7432
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}), |
8584 | 15 |
based on a pen-and-paper proof due to Wilfried Buchholz.}\isanewline |
10007 | 16 |
*} |
10144 | 17 |
(*<*)(* FIXME move? *) |
18 |
declare multiset_induct [induct type: multiset] |
|
19 |
declare wf_induct [induct set: wf] |
|
20 |
declare acc_induct [induct set: acc](*>*) |
|
8304 | 21 |
|
10007 | 22 |
subsection {* A technical lemma *} |
7761 | 23 |
|
7527 | 24 |
lemma less_add: "(N, M0 + {#a#}) : mult1 r ==> |
25 |
(EX M. (M, M0) : mult1 r & N = M + {#a#}) | |
|
9659 | 26 |
(EX K. (ALL b. b :# K --> (b, a) : r) & N = M0 + K)" |
10007 | 27 |
(concl is "?case1 (mult1 r) | ?case2") |
28 |
proof (unfold mult1_def) |
|
29 |
let ?r = "\<lambda>K a. ALL b. b :# K --> (b, a) : r" |
|
30 |
let ?R = "\<lambda>N M. EX a M0 K. M = M0 + {#a#} & N = M0 + K & ?r K a" |
|
31 |
let ?case1 = "?case1 {(N, M). ?R N M}" |
|
7527 | 32 |
|
10007 | 33 |
assume "(N, M0 + {#a#}) : {(N, M). ?R N M}" |
7800 | 34 |
hence "EX a' M0' K. |
10007 | 35 |
M0 + {#a#} = M0' + {#a'#} & N = M0' + K & ?r K a'" by simp |
36 |
thus "?case1 | ?case2" |
|
37 |
proof (elim exE conjE) |
|
38 |
fix a' M0' K |
|
39 |
assume N: "N = M0' + K" and r: "?r K a'" |
|
40 |
assume "M0 + {#a#} = M0' + {#a'#}" |
|
7800 | 41 |
hence "M0 = M0' & a = a' | |
10007 | 42 |
(EX K'. M0 = K' + {#a'#} & M0' = K' + {#a#})" |
43 |
by (simp only: add_eq_conv_ex) |
|
44 |
thus ?thesis |
|
45 |
proof (elim disjE conjE exE) |
|
46 |
assume "M0 = M0'" "a = a'" |
|
47 |
with N r have "?r K a & N = M0 + K" by simp |
|
48 |
hence ?case2 .. thus ?thesis .. |
|
49 |
next |
|
50 |
fix K' |
|
51 |
assume "M0' = K' + {#a#}" |
|
52 |
with N have n: "N = K' + K + {#a#}" by (simp add: union_ac) |
|
7527 | 53 |
|
10007 | 54 |
assume "M0 = K' + {#a'#}" |
55 |
with r have "?R (K' + K) M0" by blast |
|
56 |
with n have ?case1 by simp thus ?thesis .. |
|
57 |
qed |
|
58 |
qed |
|
59 |
qed |
|
7527 | 60 |
|
7432
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
wenzelm
parents:
diff
changeset
|
61 |
|
10007 | 62 |
subsection {* The key property *} |
7761 | 63 |
|
10007 | 64 |
lemma all_accessible: "wf r ==> ALL M. M : acc (mult1 r)" |
65 |
proof |
|
66 |
let ?R = "mult1 r" |
|
67 |
let ?W = "acc ?R" |
|
68 |
{ |
|
69 |
fix M M0 a |
|
7800 | 70 |
assume M0: "M0 : ?W" |
71 |
and wf_hyp: "ALL b. (b, a) : r --> (ALL M:?W. M + {#b#} : ?W)" |
|
10007 | 72 |
and acc_hyp: "ALL M. (M, M0) : ?R --> M + {#a#} : ?W" |
73 |
have "M0 + {#a#} : ?W" |
|
74 |
proof (rule accI [of "M0 + {#a#}"]) |
|
75 |
fix N |
|
76 |
assume "(N, M0 + {#a#}) : ?R" |
|
7480 | 77 |
hence "((EX M. (M, M0) : ?R & N = M + {#a#}) | |
10007 | 78 |
(EX K. (ALL b. b :# K --> (b, a) : r) & N = M0 + K))" |
79 |
by (rule less_add) |
|
80 |
thus "N : ?W" |
|
81 |
proof (elim exE disjE conjE) |
|
82 |
fix M assume "(M, M0) : ?R" and N: "N = M + {#a#}" |
|
83 |
from acc_hyp have "(M, M0) : ?R --> M + {#a#} : ?W" .. |
|
84 |
hence "M + {#a#} : ?W" .. |
|
85 |
thus "N : ?W" by (simp only: N) |
|
86 |
next |
|
87 |
fix K |
|
88 |
assume N: "N = M0 + K" |
|
89 |
assume "ALL b. b :# K --> (b, a) : r" |
|
90 |
have "?this --> M0 + K : ?W" (is "?P K") |
|
91 |
proof (induct K) |
|
92 |
from M0 have "M0 + {#} : ?W" by simp |
|
93 |
thus "?P {#}" .. |
|
7442 | 94 |
|
10007 | 95 |
fix K x assume hyp: "?P K" |
96 |
show "?P (K + {#x#})" |
|
97 |
proof |
|
98 |
assume a: "ALL b. b :# (K + {#x#}) --> (b, a) : r" |
|
99 |
hence "(x, a) : r" by simp |
|
100 |
with wf_hyp have b: "ALL M:?W. M + {#x#} : ?W" by blast |
|
7442 | 101 |
|
10007 | 102 |
from a hyp have "M0 + K : ?W" by simp |
103 |
with b have "(M0 + K) + {#x#} : ?W" .. |
|
104 |
thus "M0 + (K + {#x#}) : ?W" by (simp only: union_assoc) |
|
105 |
qed |
|
106 |
qed |
|
107 |
hence "M0 + K : ?W" .. |
|
108 |
thus "N : ?W" by (simp only: N) |
|
109 |
qed |
|
110 |
qed |
|
111 |
} note tedious_reasoning = this |
|
7432
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
wenzelm
parents:
diff
changeset
|
112 |
|
10007 | 113 |
assume wf: "wf r" |
114 |
fix M |
|
115 |
show "M : ?W" |
|
116 |
proof (induct M) |
|
117 |
show "{#} : ?W" |
|
118 |
proof (rule accI) |
|
119 |
fix b assume "(b, {#}) : ?R" |
|
120 |
with not_less_empty show "b : ?W" by contradiction |
|
121 |
qed |
|
7432
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
wenzelm
parents:
diff
changeset
|
122 |
|
10007 | 123 |
fix M a assume "M : ?W" |
124 |
from wf have "ALL M:?W. M + {#a#} : ?W" |
|
125 |
proof induct |
|
126 |
fix a |
|
127 |
assume "ALL b. (b, a) : r --> (ALL M:?W. M + {#b#} : ?W)" |
|
128 |
show "ALL M:?W. M + {#a#} : ?W" |
|
129 |
proof |
|
130 |
fix M assume "M : ?W" |
|
131 |
thus "M + {#a#} : ?W" |
|
132 |
by (rule acc_induct) (rule tedious_reasoning) |
|
133 |
qed |
|
134 |
qed |
|
135 |
thus "M + {#a#} : ?W" .. |
|
136 |
qed |
|
137 |
qed |
|
7432
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
wenzelm
parents:
diff
changeset
|
138 |
|
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
wenzelm
parents:
diff
changeset
|
139 |
|
10007 | 140 |
subsection {* Main result *} |
7761 | 141 |
|
10007 | 142 |
theorem wf_mult1: "wf r ==> wf (mult1 r)" |
143 |
by (rule acc_wfI, rule all_accessible) |
|
7432
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
wenzelm
parents:
diff
changeset
|
144 |
|
10007 | 145 |
theorem wf_mult: "wf r ==> wf (mult r)" |
146 |
by (unfold mult_def, rule wf_trancl, rule wf_mult1) |
|
7432
c32a0fd117a0
Wellfoundedness proof for the multiset order (preliminary version).
wenzelm
parents:
diff
changeset
|
147 |
|
10007 | 148 |
end |