| author | wenzelm | 
| Tue, 01 Aug 2017 22:19:37 +0200 | |
| changeset 66305 | 7454317f883c | 
| parent 63167 | 0909deb8059b | 
| child 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 
54545
 
483131676087
effectively reverted d25fc4c0ff62, to avoid quasi-cyclic dependencies with HOL-Cardinals and minimize BNF dependencies
 
blanchet 
parents: 
54482 
diff
changeset
 | 
1  | 
(* Title: HOL/Cardinals/Order_Union.thy  | 
| 52184 | 2  | 
Author: Andrei Popescu, TU Muenchen  | 
3  | 
||
| 52199 | 4  | 
The ordinal-like sum of two orders with disjoint fields  | 
| 52184 | 5  | 
*)  | 
6  | 
||
| 63167 | 7  | 
section \<open>Order Union\<close>  | 
| 52184 | 8  | 
|
9  | 
theory Order_Union  | 
|
| 55027 | 10  | 
imports Order_Relation  | 
| 52184 | 11  | 
begin  | 
12  | 
||
13  | 
definition Osum :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" (infix "Osum" 60) where  | 
|
14  | 
  "r Osum r' = r \<union> r' \<union> {(a, a'). a \<in> Field r \<and> a' \<in> Field r'}"
 | 
|
15  | 
||
| 52191 | 16  | 
notation Osum (infix "\<union>o" 60)  | 
| 52184 | 17  | 
|
18  | 
lemma Field_Osum: "Field (r \<union>o r') = Field r \<union> Field r'"  | 
|
19  | 
unfolding Osum_def Field_def by blast  | 
|
20  | 
||
21  | 
lemma Osum_wf:  | 
|
22  | 
assumes FLD: "Field r Int Field r' = {}" and
 | 
|
23  | 
WF: "wf r" and WF': "wf r'"  | 
|
24  | 
shows "wf (r Osum r')"  | 
|
25  | 
unfolding wf_eq_minimal2 unfolding Field_Osum  | 
|
26  | 
proof(intro allI impI, elim conjE)  | 
|
27  | 
  fix A assume *: "A \<subseteq> Field r \<union> Field r'" and **: "A \<noteq> {}"
 | 
|
28  | 
obtain B where B_def: "B = A Int Field r" by blast  | 
|
29  | 
show "\<exists>a\<in>A. \<forall>a'\<in>A. (a', a) \<notin> r \<union>o r'"  | 
|
30  | 
  proof(cases "B = {}")
 | 
|
31  | 
    assume Case1: "B \<noteq> {}"
 | 
|
32  | 
    hence "B \<noteq> {} \<and> B \<le> Field r" using B_def by auto
 | 
|
33  | 
then obtain a where 1: "a \<in> B" and 2: "\<forall>a1 \<in> B. (a1,a) \<notin> r"  | 
|
| 
55021
 
e40090032de9
compile (importing 'Metis' or 'Main' would have been an alternative)
 
blanchet 
parents: 
54545 
diff
changeset
 | 
34  | 
using WF unfolding wf_eq_minimal2 by blast  | 
| 52184 | 35  | 
hence 3: "a \<in> Field r \<and> a \<notin> Field r'" using B_def FLD by auto  | 
36  | 
(* *)  | 
|
37  | 
have "\<forall>a1 \<in> A. (a1,a) \<notin> r Osum r'"  | 
|
38  | 
proof(intro ballI)  | 
|
39  | 
fix a1 assume **: "a1 \<in> A"  | 
|
40  | 
      {assume Case11: "a1 \<in> Field r"
 | 
|
41  | 
hence "(a1,a) \<notin> r" using B_def ** 2 by auto  | 
|
42  | 
moreover  | 
|
43  | 
have "(a1,a) \<notin> r'" using 3 by (auto simp add: Field_def)  | 
|
44  | 
ultimately have "(a1,a) \<notin> r Osum r'"  | 
|
45  | 
using 3 unfolding Osum_def by auto  | 
|
46  | 
}  | 
|
47  | 
moreover  | 
|
48  | 
      {assume Case12: "a1 \<notin> Field r"
 | 
|
49  | 
hence "(a1,a) \<notin> r" unfolding Field_def by auto  | 
|
50  | 
moreover  | 
|
51  | 
have "(a1,a) \<notin> r'" using 3 unfolding Field_def by auto  | 
|
52  | 
ultimately have "(a1,a) \<notin> r Osum r'"  | 
|
53  | 
using 3 unfolding Osum_def by auto  | 
|
54  | 
}  | 
|
55  | 
ultimately show "(a1,a) \<notin> r Osum r'" by blast  | 
|
56  | 
qed  | 
|
57  | 
thus ?thesis using 1 B_def by auto  | 
|
58  | 
next  | 
|
59  | 
    assume Case2: "B = {}"
 | 
|
60  | 
    hence 1: "A \<noteq> {} \<and> A \<le> Field r'" using * ** B_def by auto
 | 
|
61  | 
then obtain a' where 2: "a' \<in> A" and 3: "\<forall>a1' \<in> A. (a1',a') \<notin> r'"  | 
|
| 
55021
 
e40090032de9
compile (importing 'Metis' or 'Main' would have been an alternative)
 
blanchet 
parents: 
54545 
diff
changeset
 | 
62  | 
using WF' unfolding wf_eq_minimal2 by blast  | 
| 52184 | 63  | 
hence 4: "a' \<in> Field r' \<and> a' \<notin> Field r" using 1 FLD by blast  | 
64  | 
(* *)  | 
|
65  | 
have "\<forall>a1' \<in> A. (a1',a') \<notin> r Osum r'"  | 
|
66  | 
proof(unfold Osum_def, auto simp add: 3)  | 
|
67  | 
fix a1' assume "(a1', a') \<in> r"  | 
|
68  | 
thus False using 4 unfolding Field_def by blast  | 
|
69  | 
next  | 
|
70  | 
fix a1' assume "a1' \<in> A" and "a1' \<in> Field r"  | 
|
71  | 
thus False using Case2 B_def by auto  | 
|
72  | 
qed  | 
|
73  | 
thus ?thesis using 2 by blast  | 
|
74  | 
qed  | 
|
75  | 
qed  | 
|
76  | 
||
77  | 
lemma Osum_Refl:  | 
|
78  | 
assumes FLD: "Field r Int Field r' = {}" and
 | 
|
79  | 
REFL: "Refl r" and REFL': "Refl r'"  | 
|
80  | 
shows "Refl (r Osum r')"  | 
|
| 
58127
 
b7cab82f488e
renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
 
blanchet 
parents: 
55027 
diff
changeset
 | 
81  | 
using assms  | 
| 52184 | 82  | 
unfolding refl_on_def Field_Osum unfolding Osum_def by blast  | 
83  | 
||
84  | 
lemma Osum_trans:  | 
|
85  | 
assumes FLD: "Field r Int Field r' = {}" and
 | 
|
86  | 
TRANS: "trans r" and TRANS': "trans r'"  | 
|
87  | 
shows "trans (r Osum r')"  | 
|
88  | 
proof(unfold trans_def, auto)  | 
|
89  | 
fix x y z assume *: "(x, y) \<in> r \<union>o r'" and **: "(y, z) \<in> r \<union>o r'"  | 
|
90  | 
show "(x, z) \<in> r \<union>o r'"  | 
|
91  | 
proof-  | 
|
92  | 
    {assume Case1: "(x,y) \<in> r"
 | 
|
93  | 
hence 1: "x \<in> Field r \<and> y \<in> Field r" unfolding Field_def by auto  | 
|
94  | 
have ?thesis  | 
|
95  | 
proof-  | 
|
96  | 
       {assume Case11: "(y,z) \<in> r"
 | 
|
97  | 
hence "(x,z) \<in> r" using Case1 TRANS trans_def[of r] by blast  | 
|
98  | 
hence ?thesis unfolding Osum_def by auto  | 
|
99  | 
}  | 
|
100  | 
moreover  | 
|
101  | 
       {assume Case12: "(y,z) \<in> r'"
 | 
|
102  | 
hence "y \<in> Field r'" unfolding Field_def by auto  | 
|
103  | 
hence False using FLD 1 by auto  | 
|
104  | 
}  | 
|
105  | 
moreover  | 
|
106  | 
       {assume Case13: "z \<in> Field r'"
 | 
|
107  | 
hence ?thesis using 1 unfolding Osum_def by auto  | 
|
108  | 
}  | 
|
109  | 
ultimately show ?thesis using ** unfolding Osum_def by blast  | 
|
110  | 
qed  | 
|
111  | 
}  | 
|
112  | 
moreover  | 
|
113  | 
    {assume Case2: "(x,y) \<in> r'"
 | 
|
114  | 
hence 2: "x \<in> Field r' \<and> y \<in> Field r'" unfolding Field_def by auto  | 
|
115  | 
have ?thesis  | 
|
116  | 
proof-  | 
|
117  | 
       {assume Case21: "(y,z) \<in> r"
 | 
|
118  | 
hence "y \<in> Field r" unfolding Field_def by auto  | 
|
119  | 
hence False using FLD 2 by auto  | 
|
120  | 
}  | 
|
121  | 
moreover  | 
|
122  | 
       {assume Case22: "(y,z) \<in> r'"
 | 
|
123  | 
hence "(x,z) \<in> r'" using Case2 TRANS' trans_def[of r'] by blast  | 
|
124  | 
hence ?thesis unfolding Osum_def by auto  | 
|
125  | 
}  | 
|
126  | 
moreover  | 
|
127  | 
       {assume Case23: "y \<in> Field r"
 | 
|
128  | 
hence False using FLD 2 by auto  | 
|
129  | 
}  | 
|
130  | 
ultimately show ?thesis using ** unfolding Osum_def by blast  | 
|
131  | 
qed  | 
|
132  | 
}  | 
|
133  | 
moreover  | 
|
134  | 
    {assume Case3: "x \<in> Field r \<and> y \<in> Field r'"
 | 
|
135  | 
have ?thesis  | 
|
136  | 
proof-  | 
|
137  | 
       {assume Case31: "(y,z) \<in> r"
 | 
|
138  | 
hence "y \<in> Field r" unfolding Field_def by auto  | 
|
139  | 
hence False using FLD Case3 by auto  | 
|
140  | 
}  | 
|
141  | 
moreover  | 
|
142  | 
       {assume Case32: "(y,z) \<in> r'"
 | 
|
143  | 
hence "z \<in> Field r'" unfolding Field_def by blast  | 
|
144  | 
hence ?thesis unfolding Osum_def using Case3 by auto  | 
|
145  | 
}  | 
|
146  | 
moreover  | 
|
147  | 
       {assume Case33: "y \<in> Field r"
 | 
|
148  | 
hence False using FLD Case3 by auto  | 
|
149  | 
}  | 
|
150  | 
ultimately show ?thesis using ** unfolding Osum_def by blast  | 
|
151  | 
qed  | 
|
152  | 
}  | 
|
153  | 
ultimately show ?thesis using * unfolding Osum_def by blast  | 
|
154  | 
qed  | 
|
155  | 
qed  | 
|
156  | 
||
157  | 
lemma Osum_Preorder:  | 
|
158  | 
"\<lbrakk>Field r Int Field r' = {}; Preorder r; Preorder r'\<rbrakk> \<Longrightarrow> Preorder (r Osum r')"
 | 
|
159  | 
unfolding preorder_on_def using Osum_Refl Osum_trans by blast  | 
|
160  | 
||
161  | 
lemma Osum_antisym:  | 
|
162  | 
assumes FLD: "Field r Int Field r' = {}" and
 | 
|
163  | 
AN: "antisym r" and AN': "antisym r'"  | 
|
164  | 
shows "antisym (r Osum r')"  | 
|
165  | 
proof(unfold antisym_def, auto)  | 
|
166  | 
fix x y assume *: "(x, y) \<in> r \<union>o r'" and **: "(y, x) \<in> r \<union>o r'"  | 
|
167  | 
show "x = y"  | 
|
168  | 
proof-  | 
|
169  | 
    {assume Case1: "(x,y) \<in> r"
 | 
|
170  | 
hence 1: "x \<in> Field r \<and> y \<in> Field r" unfolding Field_def by auto  | 
|
171  | 
have ?thesis  | 
|
172  | 
proof-  | 
|
173  | 
have "(y,x) \<in> r \<Longrightarrow> ?thesis"  | 
|
174  | 
using Case1 AN antisym_def[of r] by blast  | 
|
175  | 
moreover  | 
|
176  | 
       {assume "(y,x) \<in> r'"
 | 
|
177  | 
hence "y \<in> Field r'" unfolding Field_def by auto  | 
|
178  | 
hence False using FLD 1 by auto  | 
|
179  | 
}  | 
|
180  | 
moreover  | 
|
181  | 
have "x \<in> Field r' \<Longrightarrow> False" using FLD 1 by auto  | 
|
182  | 
ultimately show ?thesis using ** unfolding Osum_def by blast  | 
|
183  | 
qed  | 
|
184  | 
}  | 
|
185  | 
moreover  | 
|
186  | 
    {assume Case2: "(x,y) \<in> r'"
 | 
|
187  | 
hence 2: "x \<in> Field r' \<and> y \<in> Field r'" unfolding Field_def by auto  | 
|
188  | 
have ?thesis  | 
|
189  | 
proof-  | 
|
190  | 
       {assume "(y,x) \<in> r"
 | 
|
191  | 
hence "y \<in> Field r" unfolding Field_def by auto  | 
|
192  | 
hence False using FLD 2 by auto  | 
|
193  | 
}  | 
|
194  | 
moreover  | 
|
195  | 
have "(y,x) \<in> r' \<Longrightarrow> ?thesis"  | 
|
196  | 
using Case2 AN' antisym_def[of r'] by blast  | 
|
197  | 
moreover  | 
|
198  | 
       {assume "y \<in> Field r"
 | 
|
199  | 
hence False using FLD 2 by auto  | 
|
200  | 
}  | 
|
201  | 
ultimately show ?thesis using ** unfolding Osum_def by blast  | 
|
202  | 
qed  | 
|
203  | 
}  | 
|
204  | 
moreover  | 
|
205  | 
    {assume Case3: "x \<in> Field r \<and> y \<in> Field r'"
 | 
|
206  | 
have ?thesis  | 
|
207  | 
proof-  | 
|
208  | 
       {assume "(y,x) \<in> r"
 | 
|
209  | 
hence "y \<in> Field r" unfolding Field_def by auto  | 
|
210  | 
hence False using FLD Case3 by auto  | 
|
211  | 
}  | 
|
212  | 
moreover  | 
|
213  | 
       {assume Case32: "(y,x) \<in> r'"
 | 
|
214  | 
hence "x \<in> Field r'" unfolding Field_def by blast  | 
|
215  | 
hence False using FLD Case3 by auto  | 
|
216  | 
}  | 
|
217  | 
moreover  | 
|
218  | 
have "\<not> y \<in> Field r" using FLD Case3 by auto  | 
|
219  | 
ultimately show ?thesis using ** unfolding Osum_def by blast  | 
|
220  | 
qed  | 
|
221  | 
}  | 
|
222  | 
ultimately show ?thesis using * unfolding Osum_def by blast  | 
|
223  | 
qed  | 
|
224  | 
qed  | 
|
225  | 
||
226  | 
lemma Osum_Partial_order:  | 
|
227  | 
"\<lbrakk>Field r Int Field r' = {}; Partial_order r; Partial_order r'\<rbrakk> \<Longrightarrow>
 | 
|
228  | 
Partial_order (r Osum r')"  | 
|
229  | 
unfolding partial_order_on_def using Osum_Preorder Osum_antisym by blast  | 
|
230  | 
||
231  | 
lemma Osum_Total:  | 
|
232  | 
assumes FLD: "Field r Int Field r' = {}" and
 | 
|
233  | 
TOT: "Total r" and TOT': "Total r'"  | 
|
234  | 
shows "Total (r Osum r')"  | 
|
235  | 
using assms  | 
|
236  | 
unfolding total_on_def Field_Osum unfolding Osum_def by blast  | 
|
237  | 
||
238  | 
lemma Osum_Linear_order:  | 
|
239  | 
"\<lbrakk>Field r Int Field r' = {}; Linear_order r; Linear_order r'\<rbrakk> \<Longrightarrow>
 | 
|
240  | 
Linear_order (r Osum r')"  | 
|
241  | 
unfolding linear_order_on_def using Osum_Partial_order Osum_Total by blast  | 
|
242  | 
||
243  | 
lemma Osum_minus_Id1:  | 
|
244  | 
assumes "r \<le> Id"  | 
|
245  | 
shows "(r Osum r') - Id \<le> (r' - Id) \<union> (Field r \<times> Field r')"  | 
|
246  | 
proof-  | 
|
247  | 
let ?Left = "(r Osum r') - Id"  | 
|
248  | 
let ?Right = "(r' - Id) \<union> (Field r \<times> Field r')"  | 
|
249  | 
  {fix a::'a and b assume *: "(a,b) \<notin> Id"
 | 
|
250  | 
   {assume "(a,b) \<in> r"
 | 
|
251  | 
with * have False using assms by auto  | 
|
252  | 
}  | 
|
253  | 
moreover  | 
|
254  | 
   {assume "(a,b) \<in> r'"
 | 
|
255  | 
with * have "(a,b) \<in> r' - Id" by auto  | 
|
256  | 
}  | 
|
257  | 
ultimately  | 
|
258  | 
have "(a,b) \<in> ?Left \<Longrightarrow> (a,b) \<in> ?Right"  | 
|
259  | 
unfolding Osum_def by auto  | 
|
260  | 
}  | 
|
261  | 
thus ?thesis by auto  | 
|
262  | 
qed  | 
|
263  | 
||
264  | 
lemma Osum_minus_Id2:  | 
|
265  | 
assumes "r' \<le> Id"  | 
|
266  | 
shows "(r Osum r') - Id \<le> (r - Id) \<union> (Field r \<times> Field r')"  | 
|
267  | 
proof-  | 
|
268  | 
let ?Left = "(r Osum r') - Id"  | 
|
269  | 
let ?Right = "(r - Id) \<union> (Field r \<times> Field r')"  | 
|
270  | 
  {fix a::'a and b assume *: "(a,b) \<notin> Id"
 | 
|
271  | 
   {assume "(a,b) \<in> r'"
 | 
|
272  | 
with * have False using assms by auto  | 
|
273  | 
}  | 
|
274  | 
moreover  | 
|
275  | 
   {assume "(a,b) \<in> r"
 | 
|
276  | 
with * have "(a,b) \<in> r - Id" by auto  | 
|
277  | 
}  | 
|
278  | 
ultimately  | 
|
279  | 
have "(a,b) \<in> ?Left \<Longrightarrow> (a,b) \<in> ?Right"  | 
|
280  | 
unfolding Osum_def by auto  | 
|
281  | 
}  | 
|
282  | 
thus ?thesis by auto  | 
|
283  | 
qed  | 
|
284  | 
||
285  | 
lemma Osum_minus_Id:  | 
|
286  | 
assumes TOT: "Total r" and TOT': "Total r'" and  | 
|
287  | 
NID: "\<not> (r \<le> Id)" and NID': "\<not> (r' \<le> Id)"  | 
|
288  | 
shows "(r Osum r') - Id \<le> (r - Id) Osum (r' - Id)"  | 
|
289  | 
proof-  | 
|
290  | 
  {fix a a' assume *: "(a,a') \<in> (r Osum r')" and **: "a \<noteq> a'"
 | 
|
291  | 
have "(a,a') \<in> (r - Id) Osum (r' - Id)"  | 
|
292  | 
proof-  | 
|
293  | 
     {assume "(a,a') \<in> r \<or> (a,a') \<in> r'"
 | 
|
294  | 
with ** have ?thesis unfolding Osum_def by auto  | 
|
295  | 
}  | 
|
296  | 
moreover  | 
|
297  | 
     {assume "a \<in> Field r \<and> a' \<in> Field r'"
 | 
|
298  | 
hence "a \<in> Field(r - Id) \<and> a' \<in> Field (r' - Id)"  | 
|
299  | 
using assms Total_Id_Field by blast  | 
|
300  | 
hence ?thesis unfolding Osum_def by auto  | 
|
301  | 
}  | 
|
| 54482 | 302  | 
ultimately show ?thesis using * unfolding Osum_def by fast  | 
| 52184 | 303  | 
qed  | 
304  | 
}  | 
|
305  | 
thus ?thesis by(auto simp add: Osum_def)  | 
|
306  | 
qed  | 
|
307  | 
||
308  | 
lemma wf_Int_Times:  | 
|
309  | 
assumes "A Int B = {}"
 | 
|
310  | 
shows "wf(A \<times> B)"  | 
|
| 54482 | 311  | 
unfolding wf_def using assms by blast  | 
| 52184 | 312  | 
|
313  | 
lemma Osum_wf_Id:  | 
|
314  | 
assumes TOT: "Total r" and TOT': "Total r'" and  | 
|
315  | 
        FLD: "Field r Int Field r' = {}" and
 | 
|
316  | 
WF: "wf(r - Id)" and WF': "wf(r' - Id)"  | 
|
317  | 
shows "wf ((r Osum r') - Id)"  | 
|
318  | 
proof(cases "r \<le> Id \<or> r' \<le> Id")  | 
|
319  | 
assume Case1: "\<not>(r \<le> Id \<or> r' \<le> Id)"  | 
|
320  | 
  have "Field(r - Id) Int Field(r' - Id) = {}"
 | 
|
321  | 
using FLD mono_Field[of "r - Id" r] mono_Field[of "r' - Id" r']  | 
|
322  | 
Diff_subset[of r Id] Diff_subset[of r' Id] by blast  | 
|
323  | 
thus ?thesis  | 
|
324  | 
using Case1 Osum_minus_Id[of r r'] assms Osum_wf[of "r - Id" "r' - Id"]  | 
|
325  | 
wf_subset[of "(r - Id) \<union>o (r' - Id)" "(r Osum r') - Id"] by auto  | 
|
326  | 
next  | 
|
327  | 
have 1: "wf(Field r \<times> Field r')"  | 
|
328  | 
using FLD by (auto simp add: wf_Int_Times)  | 
|
329  | 
assume Case2: "r \<le> Id \<or> r' \<le> Id"  | 
|
330  | 
moreover  | 
|
331  | 
  {assume Case21: "r \<le> Id"
 | 
|
332  | 
hence "(r Osum r') - Id \<le> (r' - Id) \<union> (Field r \<times> Field r')"  | 
|
333  | 
using Osum_minus_Id1[of r r'] by simp  | 
|
334  | 
moreover  | 
|
335  | 
   {have "Domain(Field r \<times> Field r') Int Range(r' - Id) = {}"
 | 
|
336  | 
using FLD unfolding Field_def by blast  | 
|
337  | 
hence "wf((r' - Id) \<union> (Field r \<times> Field r'))"  | 
|
338  | 
using 1 WF' wf_Un[of "Field r \<times> Field r'" "r' - Id"]  | 
|
339  | 
by (auto simp add: Un_commute)  | 
|
340  | 
}  | 
|
| 
55021
 
e40090032de9
compile (importing 'Metis' or 'Main' would have been an alternative)
 
blanchet 
parents: 
54545 
diff
changeset
 | 
341  | 
ultimately have ?thesis using wf_subset by blast  | 
| 52184 | 342  | 
}  | 
343  | 
moreover  | 
|
344  | 
  {assume Case22: "r' \<le> Id"
 | 
|
345  | 
hence "(r Osum r') - Id \<le> (r - Id) \<union> (Field r \<times> Field r')"  | 
|
346  | 
using Osum_minus_Id2[of r' r] by simp  | 
|
347  | 
moreover  | 
|
348  | 
   {have "Range(Field r \<times> Field r') Int Domain(r - Id) = {}"
 | 
|
349  | 
using FLD unfolding Field_def by blast  | 
|
350  | 
hence "wf((r - Id) \<union> (Field r \<times> Field r'))"  | 
|
351  | 
using 1 WF wf_Un[of "r - Id" "Field r \<times> Field r'"]  | 
|
352  | 
by (auto simp add: Un_commute)  | 
|
353  | 
}  | 
|
| 
55021
 
e40090032de9
compile (importing 'Metis' or 'Main' would have been an alternative)
 
blanchet 
parents: 
54545 
diff
changeset
 | 
354  | 
ultimately have ?thesis using wf_subset by blast  | 
| 52184 | 355  | 
}  | 
356  | 
ultimately show ?thesis by blast  | 
|
357  | 
qed  | 
|
358  | 
||
359  | 
lemma Osum_Well_order:  | 
|
360  | 
assumes FLD: "Field r Int Field r' = {}" and
 | 
|
361  | 
WELL: "Well_order r" and WELL': "Well_order r'"  | 
|
362  | 
shows "Well_order (r Osum r')"  | 
|
363  | 
proof-  | 
|
364  | 
have "Total r \<and> Total r'" using WELL WELL'  | 
|
365  | 
by (auto simp add: order_on_defs)  | 
|
366  | 
thus ?thesis using assms unfolding well_order_on_def  | 
|
367  | 
using Osum_Linear_order Osum_wf_Id by blast  | 
|
368  | 
qed  | 
|
369  | 
||
370  | 
end  |