author | chunhan |
Mon, 30 Dec 2013 23:41:58 +0800 | |
changeset 86 | 690636b7b6f1 |
parent 74 | 271e9818b6f6 |
permissions | -rwxr-xr-x |
74
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
1 |
(*<*) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
2 |
theory My_list_prefix |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
3 |
imports List_Prefix |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
4 |
begin |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
5 |
(*>*) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
6 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
7 |
(* cmp:: 1:complete equal; 2:less; 3:greater; 4: len equal,but ele no equal *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
8 |
fun cmp :: "'a list \<Rightarrow> 'a list \<Rightarrow> nat" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
9 |
where |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
10 |
"cmp [] [] = 1" | |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
11 |
"cmp [] (e#es) = 2" | |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
12 |
"cmp (e#es) [] = 3" | |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
13 |
"cmp (e#es) (a#as) = (let r = cmp es as in |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
14 |
if (e = a) then r else 4)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
15 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
16 |
(* list_com:: fetch the same ele of the same left order into a new list*) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
17 |
fun list_com :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
18 |
where |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
19 |
"list_com [] ys = []" | |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
20 |
"list_com xs [] = []" | |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
21 |
"list_com (x#xs) (y#ys) = (if x = y then x#(list_com xs ys) else [])" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
22 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
23 |
(* list_com_rev:: by the right order of list_com *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
24 |
definition list_com_rev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "\<bullet>" 50) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
25 |
where |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
26 |
"xs \<bullet> ys \<equiv> rev (list_com (rev xs) (rev ys))" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
27 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
28 |
(* list_diff:: list substract, once different return tailer *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
29 |
fun list_diff :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
30 |
where |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
31 |
"list_diff [] xs = []" | |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
32 |
"list_diff (x#xs) [] = x#xs" | |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
33 |
"list_diff (x#xs) (y#ys) = (if x = y then list_diff xs ys else (x#xs))" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
34 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
35 |
(* list_diff_rev:: list substract with rev order*) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
36 |
definition list_diff_rev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "\<setminus>" 51) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
37 |
where |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
38 |
"xs \<setminus> ys \<equiv> rev (list_diff (rev xs) (rev ys))" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
39 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
40 |
(* xs <= ys:: \<exists>zs. ys = xs @ zs *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
41 |
(* no_junior:: xs is ys' tail,or equal *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
42 |
definition no_junior :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<preceq>" 50) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
43 |
where |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
44 |
"xs \<preceq> ys \<equiv> rev xs \<le> rev ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
45 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
46 |
(* < :: xs <= ys \<and> xs \<noteq> ys *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
47 |
(* is_ancestor:: xs is ys' tail, but no equal *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
48 |
definition is_ancestor :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<prec>" 50) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
49 |
where |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
50 |
"xs \<prec> ys \<equiv> rev xs < rev ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
51 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
52 |
lemma list_com_diff [simp]: "(list_com xs ys) @ (list_diff xs ys) = xs" (is "?P xs ys") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
53 |
by (rule_tac P = ?P in cmp.induct, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
54 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
55 |
lemma list_com_diff_rev [simp]: "(xs \<setminus> ys) @ (xs \<bullet> ys) = xs" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
56 |
apply (simp only:list_com_rev_def list_diff_rev_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
57 |
by (fold rev_append, simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
58 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
59 |
lemma list_com_commute: "list_com xs ys = list_com ys xs" (is "?P xs ys") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
60 |
by (rule_tac P = ?P in cmp.induct, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
61 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
62 |
lemma list_com_ido: "xs \<le> ys \<longrightarrow> list_com xs ys = xs" (is "?P xs ys") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
63 |
by (rule_tac P = ?P in cmp.induct, auto+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
64 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
65 |
lemma list_com_rev_ido [simp]: "xs \<preceq> ys \<Longrightarrow> xs \<bullet> ys = xs" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
66 |
by (cut_tac list_com_ido, auto simp: no_junior_def list_com_rev_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
67 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
68 |
lemma list_com_rev_commute [iff]: "(xs \<bullet> ys) = (ys \<bullet> xs)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
69 |
by (simp only:list_com_rev_def list_com_commute) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
70 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
71 |
lemma list_com_rev_ido1 [simp]: "xs \<preceq> ys \<Longrightarrow> ys \<bullet> xs = xs" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
72 |
by simp |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
73 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
74 |
lemma list_diff_le: "(list_diff xs ys = []) = (xs \<le> ys)" (is "?P xs ys") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
75 |
by (rule_tac P = ?P in cmp.induct, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
76 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
77 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
78 |
lemma list_diff_rev_le: "(xs \<setminus> ys = []) = (xs \<preceq> ys)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
79 |
by (auto simp:list_diff_rev_def no_junior_def list_diff_le) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
80 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
81 |
lemma list_diff_lt: "(list_diff xs ys = [] \<and> list_diff ys xs \<noteq> []) = (xs < ys)" (is "?P xs ys") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
82 |
by (rule_tac P = ?P in cmp.induct, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
83 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
84 |
lemma list_diff_rev_lt: "(xs \<setminus> ys = [] \<and> ys \<setminus> xs \<noteq> []) = (xs \<prec> ys)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
85 |
by (auto simp: list_diff_rev_def list_diff_lt is_ancestor_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
86 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
87 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
88 |
(* xs diff ys result not [] \<Longrightarrow> \<exists> e \<in> xs. a \<in> ys. e \<noteq> a *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
89 |
lemma list_diff_neq: |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
90 |
"\<forall> e es a as. list_diff xs ys = (e#es) \<and> list_diff ys xs = (a#as) \<longrightarrow> e \<noteq> a" (is "?P xs ys") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
91 |
by (rule_tac P = ?P in cmp.induct, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
92 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
93 |
lemma list_diff_rev_neq_pre: "\<forall> e es a as. xs \<setminus> ys = rev (e#es) \<and> ys \<setminus> xs = rev (a#as) \<longrightarrow> e \<noteq> a" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
94 |
apply (simp only:list_diff_rev_def, clarify) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
95 |
apply (insert list_diff_neq, atomize) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
96 |
by (erule_tac x = "rev xs" in allE, erule_tac x = "rev ys" in allE, blast) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
97 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
98 |
lemma list_diff_rev_neq: "\<forall> e es a as. xs \<setminus> ys = es @ [e] \<and> ys \<setminus> xs = as @ [a] \<longrightarrow> e \<noteq> a" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
99 |
apply (rule_tac allI)+ |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
100 |
apply (insert list_diff_rev_neq_pre, atomize) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
101 |
apply (erule_tac x = "xs" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
102 |
apply (erule_tac x = "ys" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
103 |
apply (erule_tac x = "e" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
104 |
apply (erule_tac x = "rev es" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
105 |
apply (erule_tac x = "a" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
106 |
apply (erule_tac x = "rev as" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
107 |
by auto |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
108 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
109 |
lemma list_com_self [simp]: "list_com zs zs = zs" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
110 |
by (induct_tac zs, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
111 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
112 |
lemma list_com_rev_self [simp]: "zs \<bullet> zs = zs" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
113 |
by (simp add:list_com_rev_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
114 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
115 |
lemma list_com_append [simp]: "(list_com (zs @ xs) (zs @ ys)) = (zs @ (list_com xs ys))" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
116 |
by (induct_tac zs, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
117 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
118 |
lemma list_inter_append [simp]: "((xs @ zs) \<bullet> (ys @ zs)) = ((xs \<bullet> ys) @ zs)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
119 |
by (simp add:list_com_rev_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
120 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
121 |
lemma list_diff_djoin_pre: |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
122 |
"\<forall> e es a as. list_diff xs ys = e#es \<and> list_diff ys xs = a#as \<longrightarrow> (\<forall> zs zs'. (list_diff (xs @ zs) (ys @ zs') = [e]@es@zs))" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
123 |
(is "?P xs ys") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
124 |
by (rule_tac P = ?P in cmp.induct, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
125 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
126 |
lemma list_diff_djoin_rev_pre: |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
127 |
"\<forall> e es a as. xs \<setminus> ys = rev (e#es) \<and> ys \<setminus> xs = rev (a#as) \<longrightarrow> (\<forall> zs zs'. ((zs @ xs) \<setminus> (zs' @ ys) = rev ([e]@es@rev zs)))" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
128 |
apply (simp only: list_diff_rev_def, clarify) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
129 |
apply (insert list_diff_djoin_pre, atomize) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
130 |
apply (erule_tac x = "rev xs" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
131 |
apply (erule_tac x = "rev ys" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
132 |
apply (erule_tac x = "e" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
133 |
apply (erule_tac x = "es" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
134 |
apply (erule_tac x = "a" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
135 |
apply (erule_tac x = "as" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
136 |
by simp |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
137 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
138 |
lemma list_diff_djoin_rev: |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
139 |
"xs \<setminus> ys = es @ [e] \<and> ys \<setminus> xs = as @ [a] \<Longrightarrow> zs @ xs \<setminus> zs' @ ys = zs @ es @ [e]" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
140 |
apply (insert list_diff_djoin_rev_pre [rule_format, simplified]) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
141 |
apply (clarsimp, atomize) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
142 |
apply (erule_tac x = "xs" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
143 |
apply (erule_tac x = "ys" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
144 |
apply (erule_tac x = "rev es" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
145 |
apply (erule_tac x = "e" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
146 |
apply (erule_tac x = "rev as" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
147 |
apply (erule_tac x = "a" in allE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
148 |
by auto |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
149 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
150 |
lemmas list_diff_djoin_rev_simplified = conjI [THEN list_diff_djoin_rev, simp] |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
151 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
152 |
lemmas list_diff_djoin = conjI [THEN list_diff_djoin_pre [rule_format], simp] |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
153 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
154 |
lemma list_diff_ext_left [simp]: "(list_diff (zs @ xs) (zs @ ys) = (list_diff xs ys))" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
155 |
by (induct_tac zs, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
156 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
157 |
lemma list_diff_rev_ext_left [simp]: "((xs @ zs \<setminus> ys @ zs) = (xs \<setminus> ys))" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
158 |
by (auto simp: list_diff_rev_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
159 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
160 |
declare no_junior_def [simp] |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
161 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
162 |
lemma no_juniorE: "\<lbrakk>xs \<preceq> ys; \<And> zs. ys = zs @ xs \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
163 |
proof - |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
164 |
assume h: "xs \<preceq> ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
165 |
and h1: "\<And> zs. ys = zs @ xs \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
166 |
show "R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
167 |
proof - |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
168 |
from h have "rev xs \<le> rev ys" by (simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
169 |
from this obtain zs where eq_rev: "rev ys = rev xs @ zs" by (auto simp:prefix_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
170 |
show R |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
171 |
proof(rule h1 [where zs = "rev zs"]) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
172 |
from rev_rev_ident and eq_rev have "rev (rev (ys)) = rev zs @ rev (rev xs)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
173 |
by simp |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
174 |
thus "ys = rev zs @ xs" by simp |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
175 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
176 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
177 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
178 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
179 |
lemma no_juniorI: "\<lbrakk>ys = zs @ xs\<rbrakk> \<Longrightarrow> xs \<preceq> ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
180 |
by simp |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
181 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
182 |
lemma no_junior_ident [simp]: "xs \<preceq> xs" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
183 |
by simp |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
184 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
185 |
lemma no_junior_expand: "xs \<preceq> ys = ((xs \<prec> ys) \<or> xs = ys)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
186 |
by (simp only:no_junior_def is_ancestor_def strict_prefix_def, blast) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
187 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
188 |
lemma no_junior_same_prefix: " e # \<tau> \<preceq> e' # \<tau>' \<Longrightarrow> \<tau> \<preceq> \<tau>'" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
189 |
apply (simp add:no_junior_def ) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
190 |
apply (erule disjE, simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
191 |
apply (simp only:prefix_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
192 |
by (erule exE, rule_tac x = "[e] @ zs" in exI, auto) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
193 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
194 |
lemma no_junior_noteq: "\<lbrakk>\<tau> \<preceq> a # \<tau>'; \<tau> \<noteq> a # \<tau>'\<rbrakk> \<Longrightarrow> \<tau> \<preceq> \<tau>'" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
195 |
apply (erule no_juniorE) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
196 |
by (case_tac zs, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
197 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
198 |
lemma is_ancestor_app [simp]: "xs \<prec> ys \<Longrightarrow> xs \<prec> zs @ ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
199 |
by (auto simp:is_ancestor_def strict_prefix_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
200 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
201 |
lemma is_ancestor_cons [simp]: "xs \<prec> ys \<Longrightarrow> xs \<prec> a # ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
202 |
by (auto simp:is_ancestor_def strict_prefix_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
203 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
204 |
lemma no_junior_app [simp]: "xs \<preceq> ys \<Longrightarrow> xs \<preceq> zs @ ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
205 |
by simp |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
206 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
207 |
lemma is_ancestor_no_junior [simp]: "xs \<prec> ys \<Longrightarrow> xs \<preceq> ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
208 |
by (simp add:is_ancestor_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
209 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
210 |
lemma is_ancestor_y [simp]: "ys \<prec> y#ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
211 |
by (simp add:is_ancestor_def strict_prefix_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
212 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
213 |
lemma no_junior_cons [simp]: "xs \<preceq> ys \<Longrightarrow> xs \<prec> (y#ys)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
214 |
by (unfold no_junior_expand, auto) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
215 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
216 |
lemma no_junior_anti_sym: "\<lbrakk>xs \<preceq> ys; ys \<preceq> xs\<rbrakk> \<Longrightarrow> xs = ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
217 |
by simp |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
218 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
219 |
declare no_junior_def [simp del] |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
220 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
221 |
(* djoin:: xs and ys is not the other's tail, not equal either *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
222 |
definition djoin :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<asymp>" 50) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
223 |
where |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
224 |
"xs \<asymp> ys \<equiv> \<not> (xs \<preceq> ys \<or> ys \<preceq> xs)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
225 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
226 |
(* dinj:: function f's returning list is not tailing when paras not equal *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
227 |
definition dinj :: "('a \<Rightarrow> 'b list) \<Rightarrow> bool" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
228 |
where |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
229 |
"dinj f \<equiv> (\<forall> a b. a \<noteq> b \<longrightarrow> f a \<asymp> f b)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
230 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
231 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
232 |
(* list_cmp:: list comparison: one is other's prefix or no equal at some position *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
233 |
lemma list_cmp: "xs \<le> ys \<or> ys \<le> xs \<or> (\<exists> zs x y a b. xs = zs @ [a] @ x \<and> ys = zs @ [b] @ y \<and> a \<noteq> b)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
234 |
proof(cases "list_diff xs ys") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
235 |
assume " list_diff xs ys = []" with list_diff_le show ?thesis by blast |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
236 |
next |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
237 |
fix e es |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
238 |
assume h: "list_diff xs ys = e # es" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
239 |
show ?thesis |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
240 |
proof(cases "list_diff ys xs") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
241 |
assume " list_diff ys xs = []" with list_diff_le show ?thesis by blast |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
242 |
next |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
243 |
fix a as assume h1: "list_diff ys xs = (a # as)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
244 |
have "xs = (list_com xs ys) @ [e] @ es \<and> ys = (list_com xs ys) @ [a] @ as \<and> e \<noteq> a" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
245 |
apply (simp, fold h1, fold h) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
246 |
apply (simp,subst list_com_commute, simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
247 |
apply (rule_tac list_diff_neq[rule_format]) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
248 |
by (insert h1, insert h, blast) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
249 |
thus ?thesis by blast |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
250 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
251 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
252 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
253 |
(* In fact, this is a case split *) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
254 |
lemma list_diff_ind: "\<lbrakk>list_diff xs ys = [] \<Longrightarrow> R; list_diff ys xs = [] \<Longrightarrow> R; |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
255 |
\<And> e es a as. \<lbrakk>list_diff xs ys = e#es; list_diff ys xs = a#as; e \<noteq> a\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
256 |
proof - |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
257 |
assume h1: "list_diff xs ys = [] \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
258 |
and h2: "list_diff ys xs = [] \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
259 |
and h3: "\<And> e es a as. \<lbrakk>list_diff xs ys = e#es; list_diff ys xs = a#as; e \<noteq> a\<rbrakk> \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
260 |
show R |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
261 |
proof(cases "list_diff xs ys") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
262 |
assume "list_diff xs ys = []" from h1 [OF this] show R . |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
263 |
next |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
264 |
fix e es |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
265 |
assume he: "list_diff xs ys = e#es" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
266 |
show R |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
267 |
proof(cases "list_diff ys xs") |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
268 |
assume "list_diff ys xs = []" from h2 [OF this] show R . |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
269 |
next |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
270 |
fix a as |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
271 |
assume ha: "list_diff ys xs = a#as" show R |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
272 |
proof(rule h3 [OF he ha]) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
273 |
from list_diff_neq [rule_format, OF conjI [OF he ha ]] |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
274 |
show "e \<noteq> a" . |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
275 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
276 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
277 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
278 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
279 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
280 |
lemma list_diff_rev_ind: |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
281 |
"\<lbrakk>xs \<setminus> ys = [] \<Longrightarrow> R; ys \<setminus> xs = [] \<Longrightarrow> R; \<And> e es a as. \<lbrakk>xs \<setminus> ys = es@[e]; ys \<setminus> xs = as@[a]; e \<noteq> a\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
282 |
proof - |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
283 |
fix xs ys R |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
284 |
assume h1: "xs \<setminus> ys = [] \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
285 |
and h2: "ys \<setminus> xs = [] \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
286 |
and h3: "\<And> e es a as. \<lbrakk>xs \<setminus> ys = es@[e]; ys \<setminus> xs = as@[a]; e \<noteq> a\<rbrakk> \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
287 |
show R |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
288 |
proof (rule list_diff_ind [where xs = "rev xs" and ys = "rev ys"]) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
289 |
assume "list_diff (rev xs) (rev ys) = []" thus R by (auto intro:h1 simp:list_diff_rev_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
290 |
next |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
291 |
assume "list_diff (rev ys) (rev xs) = []" thus R by (auto intro:h2 simp:list_diff_rev_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
292 |
next |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
293 |
fix e es a as |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
294 |
assume "list_diff (rev xs) (rev ys) = e # es" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
295 |
and "list_diff (rev ys) (rev xs) = a # as" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
296 |
and " e \<noteq> a" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
297 |
thus R by (auto intro:h3 simp:list_diff_rev_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
298 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
299 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
300 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
301 |
lemma djoin_diff_iff: "(xs \<asymp> ys) = (\<exists> e es a as. list_diff (rev xs) (rev ys) = e#es \<and> list_diff (rev ys) (rev xs) = a#as \<and> a \<noteq> e)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
302 |
proof (rule list_diff_ind [where xs = "rev xs" and ys = "rev ys"]) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
303 |
assume "list_diff (rev xs) (rev ys) = []" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
304 |
hence "xs \<preceq> ys" by (unfold no_junior_def, simp add:list_diff_le) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
305 |
thus ?thesis |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
306 |
apply (auto simp:djoin_def no_junior_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
307 |
by (fold list_diff_le, simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
308 |
next |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
309 |
assume "list_diff (rev ys) (rev xs) = []" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
310 |
hence "ys \<preceq> xs" by (unfold no_junior_def, simp add:list_diff_le) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
311 |
thus ?thesis |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
312 |
apply (auto simp:djoin_def no_junior_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
313 |
by (fold list_diff_le, simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
314 |
next |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
315 |
fix e es a as |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
316 |
assume he: "list_diff (rev xs) (rev ys) = e # es" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
317 |
and ha: "list_diff (rev ys) (rev xs) = a # as" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
318 |
and hn: "e \<noteq> a" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
319 |
show ?thesis |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
320 |
proof |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
321 |
from he ha hn |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
322 |
show |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
323 |
"\<exists>e es a as. list_diff (rev xs) (rev ys) = e # es \<and> list_diff (rev ys) (rev xs) = a # as \<and> a \<noteq> e" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
324 |
by blast |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
325 |
next |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
326 |
from he ha hn |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
327 |
show "xs \<asymp> ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
328 |
by (auto simp:djoin_def no_junior_def, fold list_diff_le, simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
329 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
330 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
331 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
332 |
lemma djoin_diff_rev_iff: "(xs \<asymp> ys) = (\<exists> e es a as. xs \<setminus> ys = es@[e] \<and> ys \<setminus> xs = as@[a] \<and> a \<noteq> e)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
333 |
apply (auto simp:djoin_diff_iff list_diff_rev_def) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
334 |
apply (rule_tac x = e in exI, safe) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
335 |
apply (rule_tac x = "rev es" in exI) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
336 |
apply (rule_tac injD[where f = rev], simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
337 |
apply (rule_tac x = "a" in exI, safe) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
338 |
apply (rule_tac x = "rev as" in exI) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
339 |
apply (rule_tac injD[where f = rev], simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
340 |
done |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
341 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
342 |
lemma djoin_revE: "\<lbrakk>xs \<asymp> ys; \<And>e es a as. \<lbrakk>xs \<setminus> ys = es@[e]; ys \<setminus> xs = as@[a]; a \<noteq> e\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
343 |
by (unfold djoin_diff_rev_iff, blast) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
344 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
345 |
lemma djoin_append_left[simp, intro]: "xs \<asymp> ys \<Longrightarrow> (zs' @ xs) \<asymp> (zs @ ys)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
346 |
by (auto simp:djoin_diff_iff intro:list_diff_djoin[simplified]) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
347 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
348 |
lemma djoin_cons_left[simp]: "xs \<asymp> ys \<Longrightarrow> (e # xs) \<asymp> (a # ys)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
349 |
by (drule_tac zs' = "[e]" and zs = "[a]" in djoin_append_left, simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
350 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
351 |
lemma djoin_simp_1 [simp]: "xs \<asymp> ys \<Longrightarrow> xs \<asymp> (zs @ ys)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
352 |
by (drule_tac djoin_append_left [where zs' = "[]"], simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
353 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
354 |
lemma djoin_simp_2 [simp]: "xs \<asymp> ys \<Longrightarrow> (zs' @ xs) \<asymp> ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
355 |
by (drule_tac djoin_append_left [where zs = "[]"], simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
356 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
357 |
lemma djoin_append_right[simp]: "xs \<asymp> ys \<Longrightarrow> (xs @ zs) \<asymp> (ys @ zs)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
358 |
by (simp add:djoin_diff_iff) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
359 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
360 |
lemma djoin_cons_append[simp]: "xs \<asymp> ys \<Longrightarrow> (x # xs) \<asymp> (zs @ ys)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
361 |
by (subgoal_tac "[x] @ xs \<asymp> zs @ ys", simp, blast) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
362 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
363 |
lemma djoin_append_cons[simp]: "xs \<asymp> ys \<Longrightarrow> (zs @ xs) \<asymp> (y # ys)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
364 |
by (subgoal_tac "zs @ xs \<asymp> [y] @ ys", simp, blast) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
365 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
366 |
lemma djoin_neq [simp]: "xs \<asymp> ys \<Longrightarrow> xs \<noteq> ys" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
367 |
by (simp only:djoin_diff_iff, clarsimp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
368 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
369 |
lemma djoin_cons [simp]: "e \<noteq> a \<Longrightarrow> e # xs \<asymp> a # xs" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
370 |
by (unfold djoin_diff_iff, simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
371 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
372 |
lemma djoin_append_e [simp]: "e \<noteq> a \<Longrightarrow> (zs @ [e] @ xs) \<asymp> (zs' @ [a] @ xs)" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
373 |
by (unfold djoin_diff_iff, simp) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
374 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
375 |
lemma djoin_mono [simp]: "\<lbrakk>xs \<asymp> ys; xs \<preceq> xs'; ys \<preceq> ys'\<rbrakk> \<Longrightarrow> xs' \<asymp> ys'" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
376 |
proof(erule_tac djoin_revE,unfold djoin_diff_rev_iff) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
377 |
fix e es a as |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
378 |
assume hx: "xs \<preceq> xs'" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
379 |
and hy: "ys \<preceq> ys'" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
380 |
and hmx: "xs \<setminus> ys = es @ [e]" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
381 |
and hmy: "ys \<setminus> xs = as @ [a]" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
382 |
and neq: "a \<noteq> e" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
383 |
have "xs' \<setminus> ys' = ((xs' \<setminus> xs) @ es) @ [e] \<and> ys' \<setminus> xs' = ((ys' \<setminus> ys) @ as) @ [a] \<and> a \<noteq> e" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
384 |
proof - |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
385 |
from hx have heqx: "(xs' \<setminus> xs) @ xs = xs'" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
386 |
by (cut_tac list_com_diff_rev [of xs' xs], subgoal_tac "xs' \<bullet> xs = xs", simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
387 |
moreover from hy have heqy: "(ys' \<setminus> ys) @ ys = ys'" |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
388 |
by (cut_tac list_com_diff_rev [of ys' ys], subgoal_tac "ys' \<bullet> ys = ys", simp+) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
389 |
moreover from list_diff_djoin_rev_simplified [OF hmx hmy] |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
390 |
have "((xs' \<setminus> xs) @ xs) \<setminus> ((ys' \<setminus> ys) @ ys) = (xs' \<setminus> xs) @ es @ [e]" by simp |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
391 |
moreover from list_diff_djoin_rev_simplified [OF hmy hmx] |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
392 |
have "((ys' \<setminus> ys) @ ys) \<setminus> ((xs' \<setminus> xs) @ xs) = (ys' \<setminus> ys) @ as @ [a]" by simp |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
393 |
ultimately show ?thesis by (simp add:neq) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
394 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
395 |
thus "\<exists>e es a as. xs' \<setminus> ys' = es @ [e] \<and> ys' \<setminus> xs' = as @ [a] \<and> a \<noteq> e" by blast |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
396 |
qed |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
397 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
398 |
lemmas djoin_append_e_simplified [simp] = djoin_append_e [simplified] |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
399 |
|
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
400 |
(*<*) |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
401 |
end |
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff
changeset
|
402 |
(*>*) |