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