| author | nipkow | 
| Sun, 15 Feb 2009 16:25:16 +0100 | |
| changeset 29923 | 24f56736c56f | 
| parent 23750 | a1db5f819d00 | 
| child 36862 | 952b2b102a0a | 
| permissions | -rw-r--r-- | 
| 5261 | 1  | 
(* Title: HOL/Lambda/ListOrder.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Tobias Nipkow  | 
|
4  | 
Copyright 1998 TU Muenchen  | 
|
5  | 
*)  | 
|
6  | 
||
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
9771 
diff
changeset
 | 
7  | 
header {* Lifting an order to lists of elements *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
9771 
diff
changeset
 | 
8  | 
|
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents: 
19086 
diff
changeset
 | 
9  | 
theory ListOrder imports Main begin  | 
| 5261 | 10  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
9771 
diff
changeset
 | 
11  | 
text {*
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
9771 
diff
changeset
 | 
12  | 
Lifting an order to lists of elements, relating exactly one  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
9771 
diff
changeset
 | 
13  | 
element.  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
9771 
diff
changeset
 | 
14  | 
*}  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
9771 
diff
changeset
 | 
15  | 
|
| 19086 | 16  | 
definition  | 
| 22270 | 17  | 
  step1 :: "('a => 'a => bool) => 'a list => 'a list => bool" where
 | 
| 19086 | 18  | 
"step1 r =  | 
| 22270 | 19  | 
(\<lambda>ys xs. \<exists>us z z' vs. xs = us @ z # vs \<and> r z' z \<and> ys =  | 
20  | 
us @ z' # vs)"  | 
|
| 9771 | 21  | 
|
22  | 
||
| 22270 | 23  | 
lemma step1_converse [simp]: "step1 (r^--1) = (step1 r)^--1"  | 
24  | 
apply (unfold step1_def)  | 
|
25  | 
apply (blast intro!: order_antisym)  | 
|
26  | 
done  | 
|
27  | 
||
28  | 
lemma in_step1_converse [iff]: "(step1 (r^--1) x y) = ((step1 r)^--1 x y)"  | 
|
29  | 
apply auto  | 
|
30  | 
done  | 
|
31  | 
||
32  | 
lemma not_Nil_step1 [iff]: "\<not> step1 r [] xs"  | 
|
| 9771 | 33  | 
apply (unfold step1_def)  | 
34  | 
apply blast  | 
|
35  | 
done  | 
|
36  | 
||
| 22270 | 37  | 
lemma not_step1_Nil [iff]: "\<not> step1 r xs []"  | 
| 9771 | 38  | 
apply (unfold step1_def)  | 
39  | 
apply blast  | 
|
40  | 
done  | 
|
41  | 
||
42  | 
lemma Cons_step1_Cons [iff]:  | 
|
| 22270 | 43  | 
"(step1 r (y # ys) (x # xs)) =  | 
44  | 
(r y x \<and> xs = ys \<or> x = y \<and> step1 r ys xs)"  | 
|
| 9771 | 45  | 
apply (unfold step1_def)  | 
46  | 
apply (rule iffI)  | 
|
47  | 
apply (erule exE)  | 
|
48  | 
apply (rename_tac ts)  | 
|
49  | 
apply (case_tac ts)  | 
|
| 13601 | 50  | 
apply fastsimp  | 
| 9771 | 51  | 
apply force  | 
52  | 
apply (erule disjE)  | 
|
53  | 
apply blast  | 
|
54  | 
apply (blast intro: Cons_eq_appendI)  | 
|
55  | 
done  | 
|
56  | 
||
57  | 
lemma append_step1I:  | 
|
| 22270 | 58  | 
"step1 r ys xs \<and> vs = us \<or> ys = xs \<and> step1 r vs us  | 
59  | 
==> step1 r (ys @ vs) (xs @ us)"  | 
|
| 9771 | 60  | 
apply (unfold step1_def)  | 
61  | 
apply auto  | 
|
62  | 
apply blast  | 
|
63  | 
apply (blast intro: append_eq_appendI)  | 
|
64  | 
done  | 
|
65  | 
||
| 18257 | 66  | 
lemma Cons_step1E [elim!]:  | 
| 22270 | 67  | 
assumes "step1 r ys (x # xs)"  | 
68  | 
and "!!y. ys = y # xs \<Longrightarrow> r y x \<Longrightarrow> R"  | 
|
69  | 
and "!!zs. ys = x # zs \<Longrightarrow> step1 r zs xs \<Longrightarrow> R"  | 
|
| 18257 | 70  | 
shows R  | 
| 23464 | 71  | 
using assms  | 
| 18257 | 72  | 
apply (cases ys)  | 
| 9771 | 73  | 
apply (simp add: step1_def)  | 
74  | 
apply blast  | 
|
75  | 
done  | 
|
76  | 
||
77  | 
lemma Snoc_step1_SnocD:  | 
|
| 22270 | 78  | 
"step1 r (ys @ [y]) (xs @ [x])  | 
79  | 
==> (step1 r ys xs \<and> y = x \<or> ys = xs \<and> r y x)"  | 
|
| 9771 | 80  | 
apply (unfold step1_def)  | 
81  | 
apply (clarify del: disjCI)  | 
|
82  | 
apply (rename_tac vs)  | 
|
83  | 
apply (rule_tac xs = vs in rev_exhaust)  | 
|
84  | 
apply force  | 
|
85  | 
apply simp  | 
|
86  | 
apply blast  | 
|
87  | 
done  | 
|
88  | 
||
| 18241 | 89  | 
lemma Cons_acc_step1I [intro!]:  | 
| 23750 | 90  | 
"accp r x ==> accp (step1 r) xs \<Longrightarrow> accp (step1 r) (x # xs)"  | 
91  | 
apply (induct arbitrary: xs set: accp)  | 
|
| 9771 | 92  | 
apply (erule thin_rl)  | 
| 23750 | 93  | 
apply (erule accp_induct)  | 
94  | 
apply (rule accp.accI)  | 
|
| 9771 | 95  | 
apply blast  | 
96  | 
done  | 
|
97  | 
||
| 23750 | 98  | 
lemma lists_accD: "listsp (accp r) xs ==> accp (step1 r) xs"  | 
| 22270 | 99  | 
apply (induct set: listsp)  | 
| 23750 | 100  | 
apply (rule accp.accI)  | 
| 9771 | 101  | 
apply simp  | 
| 23750 | 102  | 
apply (rule accp.accI)  | 
103  | 
apply (fast dest: accp_downward)  | 
|
| 9771 | 104  | 
done  | 
105  | 
||
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
9771 
diff
changeset
 | 
106  | 
lemma ex_step1I:  | 
| 22270 | 107  | 
"[| x \<in> set xs; r y x |]  | 
108  | 
==> \<exists>ys. step1 r ys xs \<and> y \<in> set ys"  | 
|
| 9771 | 109  | 
apply (unfold step1_def)  | 
110  | 
apply (drule in_set_conv_decomp [THEN iffD1])  | 
|
111  | 
apply force  | 
|
112  | 
done  | 
|
113  | 
||
| 23750 | 114  | 
lemma lists_accI: "accp (step1 r) xs ==> listsp (accp r) xs"  | 
115  | 
apply (induct set: accp)  | 
|
| 9771 | 116  | 
apply clarify  | 
| 23750 | 117  | 
apply (rule accp.accI)  | 
| 22270 | 118  | 
apply (drule_tac r=r in ex_step1I, assumption)  | 
| 9771 | 119  | 
apply blast  | 
120  | 
done  | 
|
| 5261 | 121  | 
|
| 11639 | 122  | 
end  |