thys/Sulzmann.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 26 Jun 2017 17:43:28 +0100
changeset 248 b90ff5abb437
parent 247 f35753951058
child 254 7c89d3f6923e
permissions -rw-r--r--
added a proof that Positional ordering is equivalent to direct posix definition
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
148
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
   
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
theory Sulzmann
204
cd9e40280784 added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
     3
  imports "Lexer" "~~/src/HOL/Library/Multiset"
148
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
begin
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
section {* Sulzmann's "Ordering" of Values *}
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
     9
fun 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    10
  size :: "val \<Rightarrow> nat"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    11
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    12
  "size (Void) = 0"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    13
| "size (Char c) = 0"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    14
| "size (Left v) = 1 + size v"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    15
| "size (Right v) = 1 + size v"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    16
| "size (Seq v1 v2) = 1 + (size v1) + (size v2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    17
| "size (Stars []) = 0"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    18
| "size (Stars (v#vs)) = 1 + (size v) + (size (Stars vs))" 
148
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    20
lemma Star_size [simp]:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    21
  "\<lbrakk>n < length vs; 0 < length vs\<rbrakk> \<Longrightarrow> size (nth vs n) < size (Stars vs)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    22
apply(induct vs arbitrary: n)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    23
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    24
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    25
by (metis One_nat_def Suc_pred less_Suc0 less_Suc_eq list.size(3) not_add_less1 not_less_eq nth_Cons' trans_less_add2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    26
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    27
lemma Star_size0 [simp]:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    28
  "0 < length vs \<Longrightarrow> 0 < size (Stars vs)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    29
apply(induct vs)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    30
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    31
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    32
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    33
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    34
fun 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    35
  at :: "val \<Rightarrow> nat list \<Rightarrow> val"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    36
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    37
  "at v [] = v"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    38
| "at (Left v) (0#ps)= at v ps"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    39
| "at (Right v) (Suc 0#ps)= at v ps"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    40
| "at (Seq v1 v2) (0#ps)= at v1 ps"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    41
| "at (Seq v1 v2) (Suc 0#ps)= at v2 ps"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    42
| "at (Stars vs) (n#ps)= at (nth vs n) ps"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    43
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    44
fun 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    45
  ato :: "val \<Rightarrow> nat list \<Rightarrow> val option"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    46
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    47
  "ato v [] = Some v"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    48
| "ato (Left v) (0#ps)= ato v ps"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    49
| "ato (Right v) (Suc 0#ps)= ato v ps"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    50
| "ato (Seq v1 v2) (0#ps)= ato v1 ps"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    51
| "ato (Seq v1 v2) (Suc 0#ps)= ato v2 ps"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    52
| "ato (Stars vs) (n#ps)= (if (n < length vs) then ato (nth vs n) ps else None)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    53
| "ato v p = None"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    54
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    55
fun Pos :: "val \<Rightarrow> (nat list) set"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    56
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    57
  "Pos (Void) = {[]}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    58
| "Pos (Char c) = {[]}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    59
| "Pos (Left v) = {[]} \<union> {0#ps | ps. ps \<in> Pos v}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    60
| "Pos (Right v) = {[]} \<union> {1#ps | ps. ps \<in> Pos v}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    61
| "Pos (Seq v1 v2) = {[]} \<union> {0#ps | ps. ps \<in> Pos v1} \<union> {1#ps | ps. ps \<in> Pos v2}" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    62
| "Pos (Stars []) = {[]}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    63
| "Pos (Stars (v#vs)) = {[]} \<union> {0#ps | ps. ps \<in> Pos v} \<union> {(Suc n)#ps | n ps. n#ps \<in> Pos (Stars vs)}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    64
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    65
lemma Pos_empty:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    66
  shows "[] \<in> Pos v"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    67
apply(induct v rule: Pos.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    68
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    69
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    70
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    71
lemma Pos_finite_aux:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    72
  assumes "\<forall>v \<in> set vs. finite (Pos v)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    73
  shows "finite (Pos (Stars vs))"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    74
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    75
apply(induct vs)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    76
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    77
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    78
apply(subgoal_tac "finite (Pos (Stars vs) - {[]})")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    79
apply(rule_tac f="\<lambda>l. Suc (hd l) # tl l" in finite_surj)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    80
apply(assumption)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    81
back
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    82
apply(auto simp add: image_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    83
apply(rule_tac x="n#ps" in bexI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    84
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    85
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    86
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    87
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    88
lemma Pos_finite:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    89
  shows "finite (Pos v)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    90
apply(induct v rule: val.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    91
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    92
apply(simp add: Pos_finite_aux)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    93
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    94
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    95
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    96
lemma ato_test:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    97
  assumes "p \<in> Pos v"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    98
  shows "\<exists>v'. ato v p = Some v'"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
    99
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   100
apply(induct v arbitrary: p rule: Pos.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   101
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   102
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   103
by (metis ato.simps(6) option.distinct(1))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   104
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   105
definition pflat :: "val \<Rightarrow> nat list => string option"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   106
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   107
  "pflat v p \<equiv> (if p \<in> Pos v then Some (flat (at v p)) else None)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   108
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   109
fun intlen :: "'a list \<Rightarrow> int"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   110
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   111
  "intlen [] = 0"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   112
| "intlen (x#xs) = 1 + intlen xs"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   113
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   114
lemma inlen_bigger:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   115
  shows "0 \<le> intlen xs"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   116
apply(induct xs)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   117
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   118
done 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   119
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   120
lemma intlen_append:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   121
  shows "intlen (xs @ ys) = intlen xs + intlen ys"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   122
apply(induct xs arbitrary: ys)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   123
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   124
done 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   125
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   126
lemma intlen_length:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   127
  assumes "length xs < length ys"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   128
  shows "intlen xs < intlen ys"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   129
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   130
apply(induct xs arbitrary: ys)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   131
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   132
apply(case_tac ys)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   133
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   134
apply (smt inlen_bigger)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   135
by (smt Suc_lessE intlen.simps(2) length_Suc_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   136
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   137
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   138
definition pflat_len :: "val \<Rightarrow> nat list => int"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   139
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   140
  "pflat_len v p \<equiv> (if p \<in> Pos v then intlen (flat (at v p)) else -1)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   141
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   142
lemma pflat_len_simps:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   143
  shows "pflat_len (Seq v1 v2) (0#p) = pflat_len v1 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   144
  and   "pflat_len (Seq v1 v2) (Suc 0#p) = pflat_len v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   145
  and   "pflat_len (Left v) (0#p) = pflat_len v p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   146
  and   "pflat_len (Left v) (Suc 0#p) = -1"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   147
  and   "pflat_len (Right v) (Suc 0#p) = pflat_len v p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   148
  and   "pflat_len (Right v) (0#p) = -1"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   149
  and   "pflat_len v [] = intlen (flat v)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   150
apply(auto simp add: pflat_len_def Pos_empty)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   151
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   152
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   153
lemma pflat_len_Stars_simps:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   154
  assumes "n < length vs"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   155
  shows "pflat_len (Stars vs) (n#p) = pflat_len (vs!n) p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   156
using assms 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   157
apply(induct vs arbitrary: n p)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   158
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   159
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   160
apply(simp add: pflat_len_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   161
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   162
apply (metis at.simps(6))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   163
apply (metis Suc_less_eq Suc_pred)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   164
by (metis diff_Suc_1 less_Suc_eq_0_disj nth_Cons')
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   165
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   166
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   167
lemma pflat_len_Stars_simps2:
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   168
  shows "pflat_len (Stars (v#vs)) (Suc n # p) = pflat_len (Stars vs) (n#p)"
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   169
  and   "pflat_len (Stars (v#vs)) (0 # p) = pflat_len v p"
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   170
using assms 
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   171
apply(simp_all add: pflat_len_def)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   172
done
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   173
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   174
lemma Two_to_Three_aux:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   175
  assumes "p \<in> Pos v1 \<union> Pos v2" "pflat_len v1 p = pflat_len v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   176
  shows "p \<in> Pos v1 \<inter> Pos v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   177
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   178
apply(simp add: pflat_len_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   179
apply(auto split: if_splits)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   180
apply (smt inlen_bigger)+
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   181
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   182
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   183
lemma Two_to_Three:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   184
  assumes "\<forall>p \<in> Pos v1 \<union> Pos v2. pflat v1 p = pflat v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   185
  shows "Pos v1 = Pos v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   186
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   187
by (metis Un_iff option.distinct(1) pflat_def subsetI subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   188
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   189
lemma Two_to_Three_orig:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   190
  assumes "\<forall>p \<in> Pos v1 \<union> Pos v2. pflat_len v1 p = pflat_len v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   191
  shows "Pos v1 = Pos v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   192
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   193
by (metis Un_iff inlen_bigger neg_0_le_iff_le not_one_le_zero pflat_len_def subsetI subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   194
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   195
lemma set_eq1:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   196
  assumes "insert [] A = insert [] B" "[] \<notin> A"  "[] \<notin> B"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   197
  shows "A = B"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   198
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   199
by (simp add: insert_ident)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   200
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   201
lemma set_eq2:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   202
  assumes "A \<union> B = A \<union> C"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   203
  and "A \<inter> B = {}" "A \<inter> C = {}"  
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   204
  shows "B = C"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   205
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   206
using Un_Int_distrib sup_bot.left_neutral sup_commute by blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   207
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   208
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   209
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   210
lemma Three_to_One:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   211
  assumes "\<turnstile> v1 : r" "\<turnstile> v2 : r" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   212
  and "Pos v1 = Pos v2" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   213
  shows "v1 = v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   214
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   215
apply(induct v1 arbitrary: r v2 rule: Pos.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   216
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   217
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   218
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   219
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   220
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   221
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   222
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   223
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   224
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   225
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   226
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   227
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   228
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   229
apply(simp add: insert_ident)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   230
apply(drule_tac x="r1a" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   231
apply(drule_tac x="v1a" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   232
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   233
apply(drule_tac meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   234
thm subset_antisym
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   235
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   236
apply(auto)[3]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   237
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   238
apply(simp add: insert_ident)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   239
using Pos_empty apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   240
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   241
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   242
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   243
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   244
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   245
apply(simp add: insert_ident)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   246
using Pos_empty apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   247
apply(simp add: insert_ident)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   248
apply(drule_tac x="r2a" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   249
apply(drule_tac x="v2b" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   250
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   251
apply(drule_tac meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   252
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   253
apply(auto)[3]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   254
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   255
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   256
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   257
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   258
apply(simp add: insert_ident)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   259
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   260
apply(drule_tac x="r1a" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   261
apply(drule_tac x="r2a" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   262
apply(drule_tac x="v1b" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   263
apply(drule_tac x="v2c" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   264
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   265
apply(drule_tac meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   266
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   267
apply(rule subsetI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   268
apply(subgoal_tac "0 # x \<in> {0 # ps |ps. ps \<in> Pos v1a}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   269
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   270
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   271
apply(subgoal_tac "0 # x \<in> {0 # ps |ps. ps \<in> Pos v1b}  \<union> {Suc 0 # ps |ps. ps \<in> Pos v2c}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   272
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   273
apply (metis (no_types, lifting) Un_iff)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   274
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   275
apply(rule subsetI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   276
apply(subgoal_tac "0 # x \<in> {0 # ps |ps. ps \<in> Pos v1b}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   277
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   278
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   279
apply(subgoal_tac "0 # x \<in> {0 # ps |ps. ps \<in> Pos v1a}  \<union> {Suc 0 # ps |ps. ps \<in> Pos v2b}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   280
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   281
apply (metis (no_types, lifting) Un_iff)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   282
apply(simp (no_asm_use))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   283
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   284
apply(drule_tac meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   285
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   286
apply(rule subsetI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   287
apply(subgoal_tac "Suc 0 # x \<in> {Suc 0 # ps |ps. ps \<in> Pos v2b}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   288
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   289
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   290
apply(subgoal_tac "Suc 0 # x \<in> {0 # ps |ps. ps \<in> Pos v1b}  \<union> {Suc 0 # ps |ps. ps \<in> Pos v2c}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   291
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   292
apply (metis (no_types, lifting) Un_iff)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   293
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   294
apply(rule subsetI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   295
apply(subgoal_tac "Suc 0 # x \<in> {Suc 0 # ps |ps. ps \<in> Pos v2c}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   296
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   297
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   298
apply(subgoal_tac "Suc 0 # x \<in> {0 # ps |ps. ps \<in> Pos v1b}  \<union> {Suc 0 # ps |ps. ps \<in> Pos v2b}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   299
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   300
apply (metis (no_types, lifting) Un_iff)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   301
apply(simp (no_asm_use))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   302
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   303
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   304
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   305
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   306
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   307
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   308
using Pos_empty apply fastforce
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   309
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   310
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   311
apply(erule Prf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   312
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   313
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   314
using Pos_empty apply fastforce
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   315
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   316
apply(simp add: insert_ident)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   317
apply(drule_tac x="rb" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   318
apply(drule_tac x="STAR rb" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   319
apply(drule_tac x="vb" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   320
apply(drule_tac x="Stars vsb" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   321
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   322
apply(drule_tac meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   323
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   324
apply(rule subsetI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   325
apply(subgoal_tac "0 # x \<in> {0 # ps |ps. ps \<in> Pos va}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   326
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   327
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   328
apply(subgoal_tac "0 # x \<in> {0 # ps |ps. ps \<in> Pos vb} \<union> {Suc n # ps |n ps. n # ps \<in> Pos (Stars vsb)}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   329
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   330
apply (metis (no_types, lifting) Un_iff)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   331
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   332
apply(rule subsetI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   333
apply(subgoal_tac "0 # x \<in> {0 # ps |ps. ps \<in> Pos vb}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   334
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   335
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   336
apply(subgoal_tac "0 # x \<in> {0 # ps |ps. ps \<in> Pos va} \<union> {Suc n # ps |n ps. n # ps \<in> Pos (Stars vsa)}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   337
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   338
apply (metis (no_types, lifting) Un_iff)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   339
apply(simp (no_asm_use))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   340
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   341
apply(drule_tac meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   342
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   343
apply(rule subsetI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   344
apply(case_tac vsa)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   345
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   346
apply (simp add: Pos_empty)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   347
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   348
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   349
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   350
apply (simp add: Pos_empty)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   351
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   352
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   353
apply(subgoal_tac 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   354
  "Suc 0 # ps \<in> {Suc n # ps |n ps. n = 0 \<and> ps \<in> Pos a \<or> (\<exists>na. n = Suc na \<and> na # ps \<in> Pos (Stars list))}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   355
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   356
apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   357
apply(subgoal_tac "Suc 0 # ps \<in> {0 # ps |ps. ps \<in> Pos vb} \<union>  {Suc n # ps |n ps. n # ps \<in> Pos (Stars vsb)}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   358
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   359
apply (metis (no_types, lifting) Un_iff)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   360
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   361
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   362
apply(subgoal_tac 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   363
  "Suc (Suc n) # ps \<in> {Suc n # ps |n ps. n = 0 \<and> ps \<in> Pos a \<or> (\<exists>na. n = Suc na \<and> na # ps \<in> Pos (Stars list))}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   364
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   365
apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   366
apply(subgoal_tac "Suc (Suc n) # ps \<in> {0 # ps |ps. ps \<in> Pos vb} \<union> {Suc n # ps |n ps. n # ps \<in> Pos (Stars vsb)}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   367
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   368
apply (metis (no_types, lifting) Un_iff)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   369
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   370
apply(rule subsetI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   371
apply(case_tac vsb)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   372
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   373
apply (simp add: Pos_empty)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   374
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   375
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   376
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   377
apply (simp add: Pos_empty)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   378
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   379
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   380
apply(subgoal_tac 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   381
  "Suc 0 # ps \<in> {Suc n # ps |n ps. n = 0 \<and> ps \<in> Pos a \<or> (\<exists>na. n = Suc na \<and> na # ps \<in> Pos (Stars list))}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   382
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   383
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   384
apply(subgoal_tac "Suc 0 # ps \<in> {Suc n # ps |n ps. n # ps \<in> Pos (Stars vsa)}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   385
apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   386
using list.inject apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   387
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   388
apply(subgoal_tac 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   389
  "Suc (Suc n) # ps \<in> {Suc n # ps |n ps. n = 0 \<and> ps \<in> Pos a \<or> (\<exists>na. n = Suc na \<and> na # ps \<in> Pos (Stars list))}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   390
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   391
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   392
apply(subgoal_tac "Suc (Suc n) # ps \<in> {0 # ps |ps. ps \<in> Pos vb} \<union> {Suc n # ps |n ps. n # ps \<in> Pos (Stars vsa)}")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   393
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   394
apply (metis (no_types, lifting) Un_iff)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   395
apply(simp (no_asm_use))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   396
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   397
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   398
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   399
definition prefix_list:: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<sqsubseteq>pre _")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   400
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   401
  "ps1 \<sqsubseteq>pre ps2 \<equiv> (\<exists>ps'. ps1 @ps' = ps2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   402
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   403
definition sprefix_list:: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<sqsubset>spre _")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   404
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   405
  "ps1 \<sqsubset>spre ps2 \<equiv> (ps1 \<sqsubseteq>pre ps2 \<and> ps1 \<noteq> ps2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   406
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   407
inductive lex_lists :: "nat list \<Rightarrow> nat list \<Rightarrow> bool" ("_ \<sqsubset>lex _")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   408
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   409
  "[] \<sqsubset>lex p#ps"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   410
| "ps1 \<sqsubset>lex ps2 \<Longrightarrow> (p#ps1) \<sqsubset>lex (p#ps2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   411
| "p1 < p2 \<Longrightarrow> (p1#ps1) \<sqsubset>lex (p2#ps2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   412
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   413
lemma lex_irrfl:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   414
  fixes ps1 ps2 :: "nat list"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   415
  assumes "ps1 \<sqsubset>lex ps2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   416
  shows "ps1 \<noteq> ps2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   417
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   418
apply(induct rule: lex_lists.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   419
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   420
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   421
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   422
lemma lex_append:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   423
  assumes "ps2 \<noteq> []"  
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   424
  shows "ps \<sqsubset>lex ps @ ps2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   425
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   426
apply(induct ps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   427
apply(auto intro: lex_lists.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   428
apply(case_tac ps2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   429
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   430
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   431
apply(auto intro: lex_lists.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   432
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   433
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   434
lemma lexordp_simps [simp]:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   435
  fixes xs ys :: "nat list"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   436
  shows "[] \<sqsubset>lex ys = (ys \<noteq> [])"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   437
  and   "xs \<sqsubset>lex [] = False"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   438
  and   "(x # xs) \<sqsubset>lex (y # ys) \<longleftrightarrow> (x < y \<or> (\<not> y < x \<and> xs \<sqsubset>lex ys))"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   439
apply -
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   440
apply (metis lex_append lex_lists.simps list.simps(3))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   441
using lex_lists.cases apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   442
using lex_lists.cases lex_lists.intros(2) lex_lists.intros(3) not_less_iff_gr_or_eq by fastforce
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   443
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   444
lemma lex_append_cancel [simp]:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   445
  fixes ps ps1 ps2 :: "nat list"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   446
  shows "ps @ ps1 \<sqsubset>lex ps @ ps2 \<longleftrightarrow> ps1 \<sqsubset>lex ps2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   447
apply(induct ps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   448
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   449
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   450
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   451
lemma lex_trans:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   452
  fixes ps1 ps2 ps3 :: "nat list"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   453
  assumes "ps1 \<sqsubset>lex ps2" "ps2 \<sqsubset>lex ps3"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   454
  shows "ps1 \<sqsubset>lex ps3"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   455
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   456
apply(induct arbitrary: ps3 rule: lex_lists.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   457
apply(erule lex_lists.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   458
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   459
apply(rotate_tac 2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   460
apply(erule lex_lists.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   461
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   462
apply(erule lex_lists.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   463
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   464
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   465
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   466
lemma trichotomous_aux:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   467
  fixes p q :: "nat list"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   468
  assumes "p \<sqsubset>lex q" "p \<noteq> q"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   469
  shows "\<not>(q \<sqsubset>lex p)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   470
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   471
apply(induct rule: lex_lists.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   472
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   473
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   474
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   475
lemma trichotomous_aux2:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   476
  fixes p q :: "nat list"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   477
  assumes "p \<sqsubset>lex q" "q \<sqsubset>lex p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   478
  shows "False"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   479
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   480
apply(induct rule: lex_lists.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   481
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   482
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   483
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   484
lemma trichotomous:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   485
  fixes p q :: "nat list"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   486
  shows "p = q \<or> p \<sqsubset>lex q \<or> q \<sqsubset>lex p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   487
apply(induct p arbitrary: q)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   488
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   489
apply(case_tac q)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   490
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   491
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   492
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   493
definition dpos :: "val \<Rightarrow> val \<Rightarrow> nat list \<Rightarrow> bool" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   494
  where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   495
  "dpos v1 v2 p \<equiv> (p \<in> Pos v1 \<union> Pos v2) \<and> (p \<notin> Pos v1 \<inter> Pos v2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   496
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   497
definition
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   498
  "DPos v1 v2 \<equiv> {p. dpos v1 v2 p}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   499
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   500
lemma outside_lemma:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   501
  assumes "p \<notin> Pos v1 \<union> Pos v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   502
  shows "pflat_len v1 p = pflat_len v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   503
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   504
apply(auto simp add: pflat_len_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   505
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   506
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   507
lemma dpos_lemma_aux:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   508
  assumes "p \<in> Pos v1 \<union> Pos v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   509
  and "pflat_len v1 p = pflat_len v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   510
  shows "p \<in> Pos v1 \<inter> Pos v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   511
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   512
apply(auto simp add: pflat_len_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   513
apply (smt inlen_bigger)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   514
apply (smt inlen_bigger)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   515
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   516
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   517
lemma dpos_lemma:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   518
  assumes "p \<in> Pos v1 \<union> Pos v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   519
  and "pflat_len v1 p = pflat_len v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   520
  shows "\<not>dpos v1 v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   521
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   522
apply(auto simp add: dpos_def dpos_lemma_aux)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   523
using dpos_lemma_aux apply auto[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   524
using dpos_lemma_aux apply auto[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   525
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   526
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   527
lemma dpos_lemma2:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   528
  assumes "p \<in> Pos v1 \<union> Pos v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   529
  and "dpos v1 v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   530
  shows "pflat_len v1 p \<noteq> pflat_len v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   531
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   532
using dpos_lemma by blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   533
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   534
lemma DPos_lemma:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   535
  assumes "p \<in> DPos v1 v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   536
  shows "pflat_len v1 p \<noteq> pflat_len v2 p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   537
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   538
unfolding DPos_def 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   539
apply(auto simp add: pflat_len_def dpos_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   540
apply (smt inlen_bigger)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   541
by (smt inlen_bigger)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   542
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   543
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   544
definition val_ord:: "val \<Rightarrow> nat list \<Rightarrow> val \<Rightarrow> bool" ("_ \<sqsubset>val _ _")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   545
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   546
  "v1 \<sqsubset>val p v2 \<equiv> (p \<in> Pos v1 \<and> pflat_len v1 p > pflat_len v2 p \<and>
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   547
                   (\<forall>q \<in> Pos v1 \<union> Pos v2. q \<sqsubset>lex p \<longrightarrow> pflat_len v1 q = pflat_len v2 q))"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   548
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   549
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   550
definition val_ord_ex:: "val \<Rightarrow> val \<Rightarrow> bool" ("_ :\<sqsubset>val _")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   551
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   552
  "v1 :\<sqsubset>val v2 \<equiv> (\<exists>p. v1 \<sqsubset>val p v2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   553
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   554
definition val_ord_ex1:: "val \<Rightarrow> val \<Rightarrow> bool" ("_ :\<sqsubseteq>val _")
148
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   555
where
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   556
  "v1 :\<sqsubseteq>val v2 \<equiv> v1 :\<sqsubset>val v2 \<or> v1 = v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   557
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   558
lemma val_ord_shorterI:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   559
  assumes "length (flat v') < length (flat v)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   560
  shows "v :\<sqsubset>val v'" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   561
using assms(1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   562
apply(subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   563
apply(rule_tac x="[]" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   564
apply(subst val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   565
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   566
apply (simp add: Pos_empty)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   567
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   568
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   569
apply (simp add: intlen_length)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   570
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   571
done
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   572
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   573
lemma val_ord_spre:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   574
  assumes "(flat v') \<sqsubset>spre (flat v)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   575
  shows "v :\<sqsubset>val v'" 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   576
using assms(1)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   577
apply(rule_tac val_ord_shorterI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   578
apply(simp add: sprefix_list_def prefix_list_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   579
apply(auto)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   580
apply(case_tac ps')
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   581
apply(auto)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   582
by (metis append_eq_conv_conj drop_all le_less_linear neq_Nil_conv)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   583
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   584
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   585
lemma val_ord_ALTI:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   586
  assumes "v \<sqsubset>val p v'" "flat v = flat v'"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   587
  shows "(Left v) \<sqsubset>val (0#p) (Left v')" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   588
using assms(1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   589
apply(subst (asm) val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   590
apply(erule conjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   591
apply(subst val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   592
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   593
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   594
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   595
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   596
apply(rule ballI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   597
apply(rule impI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   598
apply(simp only: Pos.simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   599
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   600
using assms(2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   601
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   602
apply(auto simp add: pflat_len_simps)[2]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   603
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   604
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   605
lemma val_ord_ALTI2:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   606
  assumes "v \<sqsubset>val p v'" "flat v = flat v'"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   607
  shows "(Right v) \<sqsubset>val (1#p) (Right v')" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   608
using assms(1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   609
apply(subst (asm) val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   610
apply(erule conjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   611
apply(subst val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   612
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   613
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   614
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   615
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   616
apply(rule ballI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   617
apply(rule impI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   618
apply(simp only: Pos.simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   619
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   620
using assms(2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   621
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   622
apply(auto simp add: pflat_len_simps)[2]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   623
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   624
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   625
lemma val_ord_ALTE:
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   626
  assumes "(Left v1) \<sqsubset>val (p # ps) (Left v2)"
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   627
  shows "p = 0 \<and> v1 \<sqsubset>val ps v2"
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   628
using assms(1)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   629
apply(simp add: val_ord_def)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   630
apply(auto simp add: pflat_len_simps)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   631
apply (metis (no_types, hide_lams) assms lex_lists.intros(2) outside_lemma pflat_len_simps(3) val_ord_def)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   632
by (metis (no_types, hide_lams) assms lex_lists.intros(2) outside_lemma pflat_len_simps(3) val_ord_def)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   633
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   634
lemma val_ord_ALTE2:
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   635
  assumes "(Right v1) \<sqsubset>val (p # ps) (Right v2)"
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   636
  shows "p = 1 \<and> v1 \<sqsubset>val ps v2"
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   637
using assms(1)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   638
apply(simp add: val_ord_def)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   639
apply(auto simp add: pflat_len_simps)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   640
apply (metis (no_types, hide_lams) assms lex_lists.intros(2) outside_lemma pflat_len_simps(5) val_ord_def)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   641
by (metis (no_types, hide_lams) assms lex_lists.intros(2) outside_lemma pflat_len_simps(5) val_ord_def)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   642
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   643
lemma val_ord_STARI:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   644
  assumes "v1 \<sqsubset>val p v2" "flat (Stars (v1#vs1)) = flat (Stars (v2#vs2))"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   645
  shows "(Stars (v1#vs1)) \<sqsubset>val (0#p) (Stars (v2#vs2))" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   646
using assms(1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   647
apply(subst (asm) val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   648
apply(erule conjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   649
apply(subst val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   650
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   651
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   652
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   653
apply(subst pflat_len_Stars_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   654
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   655
apply(subst pflat_len_Stars_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   656
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   657
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   658
apply(rule ballI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   659
apply(rule impI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   660
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   661
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   662
using assms(2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   663
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   664
apply(auto simp add: pflat_len_Stars_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   665
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   666
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   667
lemma val_ord_STARI2:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   668
  assumes "(Stars vs1) \<sqsubset>val n#p (Stars vs2)" "flat (Stars vs1) = flat (Stars vs2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   669
  shows "(Stars (v#vs1)) \<sqsubset>val (Suc n#p) (Stars (v#vs2))" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   670
using assms(1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   671
apply(subst (asm) val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   672
apply(erule conjE)+
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   673
apply(subst val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   674
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   675
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   676
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   677
apply(case_tac vs1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   678
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   679
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   680
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   681
apply(case_tac vs2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   682
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   683
apply (simp add: pflat_len_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   684
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   685
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   686
apply (simp add: pflat_len_Stars_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   687
using pflat_len_def apply auto[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   688
apply(rule ballI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   689
apply(rule impI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   690
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   691
using assms(2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   692
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   693
apply (simp add: pflat_len_simps(7))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   694
apply (simp add: pflat_len_Stars_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   695
using assms(2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   696
apply(auto simp add: pflat_len_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   697
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   698
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   699
apply(auto simp add: pflat_len_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   700
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   701
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   702
apply(auto simp add: pflat_len_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   703
apply(auto simp add: pflat_len_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   704
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   705
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   706
apply(auto simp add: pflat_len_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   707
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   708
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   709
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   710
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   711
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   712
lemma val_ord_SEQI:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   713
  assumes "v1 \<sqsubset>val p v1'" "flat (Seq v1 v2) = flat (Seq v1' v2')"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   714
  shows "(Seq v1 v2) \<sqsubset>val (0#p) (Seq v1' v2')" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   715
using assms(1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   716
apply(subst (asm) val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   717
apply(erule conjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   718
apply(subst val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   719
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   720
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   721
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   722
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   723
apply(rule ballI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   724
apply(rule impI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   725
apply(simp only: Pos.simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   726
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   727
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   728
using assms(2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   729
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   730
apply(auto simp add: pflat_len_simps)[2]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   731
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   732
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   733
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   734
lemma val_ord_SEQI2:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   735
  assumes "v2 \<sqsubset>val p v2'" "flat v2 = flat v2'"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   736
  shows "(Seq v v2) \<sqsubset>val (1#p) (Seq v v2')" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   737
using assms(1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   738
apply(subst (asm) val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   739
apply(erule conjE)+
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   740
apply(subst val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   741
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   742
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   743
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   744
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   745
apply(rule ballI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   746
apply(rule impI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   747
apply(simp only: Pos.simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   748
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   749
apply(auto simp add: pflat_len_def intlen_append)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   750
apply(auto simp add: assms(2))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   751
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   752
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   753
lemma val_ord_SEQE_0:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   754
  assumes "(Seq v1 v2) \<sqsubset>val 0#p (Seq v1' v2')" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   755
  shows "v1 \<sqsubset>val p v1'"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   756
using assms(1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   757
apply(simp add: val_ord_def val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   758
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   759
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   760
apply(simp add: val_ord_def pflat_len_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   761
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   762
apply(drule_tac x="0#q" in bspec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   763
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   764
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   765
apply(drule_tac x="0#q" in bspec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   766
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   767
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   768
apply(drule_tac x="0#q" in bspec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   769
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   770
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   771
apply(simp add: val_ord_def pflat_len_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   772
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   773
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   774
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   775
lemma val_ord_SEQE_1:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   776
  assumes "(Seq v1 v2) \<sqsubset>val (Suc 0)#p (Seq v1' v2')" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   777
  shows "v2 \<sqsubset>val p v2'"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   778
using assms(1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   779
apply(simp add: val_ord_def pflat_len_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   780
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   781
apply(drule_tac x="1#q" in bspec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   782
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   783
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   784
apply(drule_tac x="1#q" in bspec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   785
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   786
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   787
apply(drule_tac x="1#q" in bspec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   788
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   789
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   790
apply(drule_tac x="1#q" in bspec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   791
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   792
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   793
apply(simp add: intlen_append)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   794
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   795
apply(simp add: intlen_append)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   796
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   797
apply(simp add: intlen_append)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   798
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   799
apply(simp add: intlen_append)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   800
apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   801
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   802
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   803
lemma val_ord_SEQE_2:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   804
  assumes "(Seq v1 v2) \<sqsubset>val (Suc 0)#p (Seq v1' v2')" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   805
  and "\<turnstile> v1 : r" "\<turnstile> v1' : r" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   806
  shows "v1 = v1'"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   807
proof -
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   808
  have "\<forall>q \<in> Pos v1 \<union> Pos v1'. 0 # q \<sqsubset>lex 1#p \<longrightarrow> pflat_len v1 q = pflat_len v1' q"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   809
  using assms(1) 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   810
  apply(simp add: val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   811
  apply(rule ballI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   812
  apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   813
  apply(drule_tac x="0#q" in bspec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   814
  apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   815
  apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   816
  done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   817
  then have "Pos v1 = Pos v1'"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   818
  apply(rule_tac Two_to_Three_orig)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   819
  apply(rule ballI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   820
  apply(drule_tac x="pa" in bspec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   821
  apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   822
  apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   823
  done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   824
  then show "v1 = v1'" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   825
  apply(rule_tac Three_to_One)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   826
  apply(rule assms)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   827
  apply(rule assms)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   828
  apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   829
  done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   830
qed
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   831
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   832
lemma val_ord_SEQ:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   833
  assumes "(Seq v1 v2) :\<sqsubset>val (Seq v1' v2')" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   834
  and "flat (Seq v1 v2) = flat (Seq v1' v2')"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   835
  and "\<turnstile> v1 : r" "\<turnstile> v1' : r" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   836
  shows "(v1 :\<sqsubset>val v1') \<or> (v1 = v1' \<and> (v2 :\<sqsubset>val v2'))"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   837
using assms(1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   838
apply(subst (asm) val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   839
apply(erule exE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   840
apply(simp only: val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   841
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   842
apply(erule conjE)+
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   843
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   844
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   845
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   846
apply(erule exE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   847
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   848
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   849
apply(subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   850
apply(rule_tac x="ps" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   851
apply(rule val_ord_SEQE_0)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   852
apply(simp add: val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   853
apply(erule exE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   854
apply(rule disjI2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   855
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   856
thm val_ord_SEQE_1
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   857
apply(rule_tac val_ord_SEQE_2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   858
apply(auto simp add: val_ord_def)[3]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   859
apply(rule assms(3))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   860
apply(rule assms(4))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   861
apply(subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   862
apply(rule_tac x="ps" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   863
apply(rule_tac val_ord_SEQE_1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   864
apply(auto simp add: val_ord_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   865
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   866
using assms(2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   867
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   868
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   869
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
   870
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   871
lemma val_ord_ex_trans:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   872
  assumes "v1 :\<sqsubset>val v2" "v2 :\<sqsubset>val v3"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   873
  shows "v1 :\<sqsubset>val v3"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   874
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   875
unfolding val_ord_ex_def
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   876
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   877
apply(subgoal_tac "p = pa \<or> p \<sqsubset>lex pa \<or> pa \<sqsubset>lex p")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   878
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   879
apply(rule trichotomous)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   880
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   881
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   882
apply(rule_tac x="pa" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   883
apply(subst val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   884
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   885
apply(simp add: val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   886
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   887
apply(simp add: val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   888
apply(simp add: val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   889
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   890
using outside_lemma apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   891
apply(simp add: val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   892
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   893
using outside_lemma apply force
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   894
apply auto[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   895
apply(simp add: val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   896
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   897
apply (metis (no_types, hide_lams) lex_trans outside_lemma)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   898
apply(simp add: val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   899
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   900
by (metis IntE Two_to_Three_aux UnCI lex_trans outside_lemma)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   901
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   902
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   903
definition fdpos :: "val \<Rightarrow> val \<Rightarrow> nat list \<Rightarrow> bool" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   904
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   905
  "fdpos v1 v2 p \<equiv>  ({q. q \<sqsubset>lex p} \<inter> DPos v1 v2 = {})"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   906
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   907
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   908
lemma pos_append:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   909
  assumes "p @ q \<in> Pos v"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   910
  shows "q \<in> Pos (at v p)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   911
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   912
apply(induct arbitrary: p q rule: Pos.induct) 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   913
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   914
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   915
apply(simp add: append_eq_Cons_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   916
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   917
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   918
apply(simp add: append_eq_Cons_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   919
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   920
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   921
apply(simp add: append_eq_Cons_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   922
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   923
apply(simp add: append_eq_Cons_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   924
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   925
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   926
apply(simp add: append_eq_Cons_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   927
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   928
apply(simp add: append_eq_Cons_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   929
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   930
by (metis append_Cons at.simps(6))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   931
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   932
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   933
lemma Pos_pre:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   934
  assumes "p \<in> Pos v" "q \<sqsubseteq>pre p"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   935
  shows "q \<in> Pos v"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   936
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   937
apply(induct v arbitrary: p q rule: Pos.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   938
apply(simp_all add: prefix_list_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   939
apply (meson append_eq_Cons_conv append_is_Nil_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   940
apply (meson append_eq_Cons_conv append_is_Nil_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   941
apply (metis (no_types, lifting) Cons_eq_append_conv append_is_Nil_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   942
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   943
apply (meson append_eq_Cons_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   944
apply(simp add: append_eq_Cons_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   945
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   946
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   947
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   948
lemma lex_lists_order:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   949
  assumes "q' \<sqsubset>lex q" "\<not>(q' \<sqsubseteq>pre q)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   950
  shows "\<not>(q \<sqsubset>lex q')"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   951
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   952
apply(induct rule: lex_lists.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   953
apply(simp add: prefix_list_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   954
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   955
using trichotomous_aux2 by auto
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   956
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   957
lemma lex_appendL:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   958
  assumes "q \<sqsubset>lex p" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   959
  shows "q \<sqsubset>lex p @ q'"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   960
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   961
apply(induct arbitrary: q' rule: lex_lists.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   962
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   963
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   964
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   965
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   966
inductive 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   967
  CPrf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<Turnstile> _ : _" [100, 100] 100)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   968
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   969
 "\<lbrakk>\<Turnstile> v1 : r1; \<Turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<Turnstile>  Seq v1 v2 : SEQ r1 r2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   970
| "\<Turnstile> v1 : r1 \<Longrightarrow> \<Turnstile> Left v1 : ALT r1 r2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   971
| "\<Turnstile> v2 : r2 \<Longrightarrow> \<Turnstile> Right v2 : ALT r1 r2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   972
| "\<Turnstile> Void : ONE"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   973
| "\<Turnstile> Char c : CHAR c"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   974
| "\<Turnstile> Stars [] : STAR r"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   975
| "\<lbrakk>\<Turnstile> v : r; flat v \<noteq> []; \<Turnstile> Stars vs : STAR r\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (v # vs) : STAR r"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   976
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   977
lemma Prf_CPrf:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   978
  assumes "\<Turnstile> v : r"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   979
  shows "\<turnstile> v : r"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   980
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   981
apply(induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   982
apply(auto intro: Prf.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   983
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   984
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   985
definition
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   986
  "CPT r s = {v. flat v = s \<and> \<Turnstile> v : r}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   987
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   988
definition
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   989
  "CPTpre r s = {v. \<exists>s'. flat v @ s' = s \<and> \<Turnstile> v : r}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   990
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   991
lemma CPT_CPTpre_subset:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   992
  shows "CPT r s \<subseteq> CPTpre r s"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   993
apply(auto simp add: CPT_def CPTpre_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   994
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   995
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   996
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   997
lemma CPTpre_subsets:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   998
  "CPTpre ZERO s = {}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   999
  "CPTpre ONE s \<subseteq> {Void}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1000
  "CPTpre (CHAR c) s \<subseteq> {Char c}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1001
  "CPTpre (ALT r1 r2) s \<subseteq> Left ` CPTpre r1 s \<union> Right ` CPTpre r2 s"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1002
  "CPTpre (SEQ r1 r2) s \<subseteq> {Seq v1 v2 | v1 v2. v1 \<in> CPTpre r1 s \<and> v2 \<in> CPTpre r2 (drop (length (flat v1)) s)}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1003
  "CPTpre (STAR r) s \<subseteq> {Stars []} \<union>
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1004
    {Stars (v#vs) | v vs. v \<in> CPTpre r s \<and> flat v \<noteq> [] \<and> Stars vs \<in> CPTpre (STAR r) (drop (length (flat v)) s)}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1005
  "CPTpre (STAR r) [] = {Stars []}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1006
apply(auto simp add: CPTpre_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1007
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1008
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1009
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1010
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1011
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1012
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1013
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1014
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1015
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1016
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1017
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1018
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1019
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1020
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1021
apply(rule CPrf.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1022
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1023
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1024
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1025
lemma CPTpre_simps:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1026
  shows "CPTpre ONE s = {Void}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1027
  and   "CPTpre (CHAR c) (d#s) = (if c = d then {Char c} else {})"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1028
  and   "CPTpre (ALT r1 r2) s = Left ` CPTpre r1 s \<union> Right ` CPTpre r2 s"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1029
  and   "CPTpre (SEQ r1 r2) s = 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1030
          {Seq v1 v2 | v1 v2. v1 \<in> CPTpre r1 s \<and> v2 \<in> CPTpre r2 (drop (length (flat v1)) s)}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1031
apply -
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1032
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1033
apply(rule CPTpre_subsets)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1034
apply(auto simp add: CPTpre_def intro: "CPrf.intros")[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1035
apply(case_tac "c = d")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1036
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1037
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1038
apply(rule CPTpre_subsets)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1039
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1040
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1041
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1042
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1043
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1044
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1045
apply(rule CPTpre_subsets)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1046
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1047
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1048
apply(rule CPTpre_subsets)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1049
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1050
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1051
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1052
lemma CPT_simps:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1053
  shows "CPT ONE s = (if s = [] then {Void} else {})"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1054
  and   "CPT (CHAR c) [d] = (if c = d then {Char c} else {})"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1055
  and   "CPT (ALT r1 r2) s = Left ` CPT r1 s \<union> Right ` CPT r2 s"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1056
  and   "CPT (SEQ r1 r2) s = 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1057
          {Seq v1 v2 | v1 v2 s1 s2. s1 @ s2 = s \<and> v1 \<in> CPT r1 s1 \<and> v2 \<in> CPT r2 s2}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1058
apply -
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1059
apply(rule subset_antisym)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1060
apply(auto simp add: CPT_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1061
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1062
apply(simp_all)[7]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1063
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1064
apply(simp_all)[7]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1065
apply(auto simp add: CPT_def intro: CPrf.intros)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1066
apply(auto simp add: CPT_def intro: CPrf.intros)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1067
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1068
apply(simp_all)[7]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1069
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1070
apply(simp_all)[7]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1071
apply(auto simp add: CPT_def image_def intro: CPrf.intros)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1072
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1073
apply(simp_all)[7]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1074
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1075
apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1076
apply(auto simp add: CPT_def image_def intro: CPrf.intros)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1077
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1078
apply(simp_all)[7]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1079
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1080
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1081
lemma CPTpre_SEQ:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1082
  assumes "v \<in> CPTpre (SEQ r1 r2) s"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1083
  shows "\<exists>s'. flat v = s' \<and> (s' \<sqsubseteq>pre s) \<and> s' \<in> L (SEQ r1 r2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1084
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1085
apply(simp add: CPTpre_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1086
apply(auto simp add: CPTpre_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1087
apply (simp add: prefix_list_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1088
by (metis L.simps(4) L_flat_Prf1 Prf.intros(1) Prf_CPrf flat.simps(5))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1089
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1090
lemma Cond_prefix:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1091
  assumes "\<forall>s\<^sub>3. s1 @ s\<^sub>3 \<in> L r1 \<longrightarrow> s\<^sub>3 = [] \<or> (\<forall>s\<^sub>4. s1 @ s\<^sub>3 @ s\<^sub>4 \<sqsubseteq>pre s1 @ s2 \<longrightarrow> s\<^sub>4 \<notin> L r2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1092
  and "t1 \<in> L r1" "t2 \<in> L r2" "t1 @ t2 \<sqsubseteq>pre s1 @ s2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1093
  shows "t1 \<sqsubseteq>pre s1"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1094
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1095
apply(auto simp add: Sequ_def prefix_list_def append_eq_append_conv2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1096
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1097
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1098
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1099
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1100
lemma test: 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1101
  assumes "finite A"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1102
  shows "finite {vs. Stars vs \<in> A}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1103
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1104
apply(induct A)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1105
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1106
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1107
apply(case_tac x)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1108
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1109
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1110
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1111
lemma CPTpre_STAR_finite:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1112
  assumes "\<And>s. finite (CPTpre r s)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1113
  shows "finite (CPTpre (STAR r) s)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1114
apply(induct s rule: length_induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1115
apply(case_tac xs)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1116
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1117
apply(simp add: CPTpre_subsets)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1118
apply(rule finite_subset)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1119
apply(rule CPTpre_subsets)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1120
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1121
apply(rule_tac B="(\<lambda>(v, vs). Stars (v#vs)) ` {(v, vs). v \<in> CPTpre r (a#list) \<and> flat v \<noteq> [] \<and> Stars vs \<in> CPTpre (STAR r) (drop (length (flat v)) (a#list))}" in finite_subset)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1122
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1123
apply(rule finite_imageI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1124
apply(simp add: Collect_case_prod_Sigma)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1125
apply(rule finite_SigmaI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1126
apply(rule assms)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1127
apply(case_tac "flat v = []")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1128
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1129
apply(drule_tac x="drop (length (flat v)) (a # list)" in spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1130
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1131
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1132
apply(rule test)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1133
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1134
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1135
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1136
lemma CPTpre_finite:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1137
  shows "finite (CPTpre r s)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1138
apply(induct r arbitrary: s)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1139
apply(simp add: CPTpre_subsets)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1140
apply(rule finite_subset)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1141
apply(rule CPTpre_subsets)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1142
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1143
apply(rule finite_subset)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1144
apply(rule CPTpre_subsets)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1145
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1146
sorry
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1147
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1148
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1149
lemma CPT_finite:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1150
  shows "finite (CPT r s)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1151
apply(rule finite_subset)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1152
apply(rule CPT_CPTpre_subset)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1153
apply(rule CPTpre_finite)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1154
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1155
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1156
lemma Posix_CPT:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1157
  assumes "s \<in> r \<rightarrow> v"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1158
  shows "v \<in> CPT r s"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1159
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1160
apply(induct rule: Posix.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1161
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1162
apply(rule CPrf.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1163
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1164
apply(rule CPrf.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1165
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1166
apply(rule CPrf.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1167
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1168
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1169
apply(rule CPrf.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1170
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1171
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1172
apply(rule CPrf.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1173
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1174
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1175
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1176
apply(rule CPrf.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1177
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1178
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1179
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1180
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1181
apply(rule CPrf.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1182
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1183
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1184
lemma Posix_val_ord:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1185
  assumes "s \<in> r \<rightarrow> v1" "v2 \<in> CPTpre r s" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1186
  shows "v1 :\<sqsubseteq>val v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1187
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1188
apply(induct arbitrary: v2 rule: Posix.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1189
apply(simp add: CPTpre_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1190
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1191
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1192
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1193
apply(simp add: val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1194
apply(simp add: CPTpre_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1195
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1196
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1197
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1198
apply(simp add: val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1199
(* ALT1 *)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1200
prefer 3
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1201
(* SEQ case *)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1202
apply(subst (asm) (3) CPTpre_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1203
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1204
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1205
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1206
apply(case_tac "s' = []")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1207
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1208
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1209
apply(simp add: val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1210
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1211
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1212
apply(simp add: val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1213
apply(simp (no_asm) add: val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1214
apply(rule_tac x="[]" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1215
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1216
apply(rule intlen_length)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1217
apply (metis Posix1(2) append_assoc append_eq_conv_conj drop_eq_Nil not_le)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1218
apply(subgoal_tac "length (flat v1a) \<le> length s1")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1219
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1220
apply (metis L_flat_Prf1 Prf_CPrf append_eq_append_conv_if append_take_drop_id drop_eq_Nil)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1221
apply(subst (asm) append_eq_append_conv_if)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1222
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1223
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1224
apply(drule_tac x="v1a" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1225
apply(drule meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1226
apply(auto simp add: CPTpre_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1227
using append_eq_conv_conj apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1228
apply(subst (asm) (2)val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1229
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1230
apply(subst (asm) val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1231
apply(erule exE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1232
apply(subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1233
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1234
apply(subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1235
apply(rule_tac x="0#p" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1236
apply(rule val_ord_SEQI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1237
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1238
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1239
apply (metis Posix1(2) append_assoc append_take_drop_id)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1240
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1241
apply(drule_tac x="v2b" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1242
apply(drule meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1243
apply(auto simp add: CPTpre_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1244
apply (simp add: Posix1(2))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1245
apply(subst (asm) val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1246
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1247
apply(subst (asm) val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1248
apply(erule exE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1249
apply(subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1250
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1251
apply(subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1252
apply(rule_tac x="1#p" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1253
apply(rule val_ord_SEQI2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1254
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1255
apply (simp add: Posix1(2))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1256
apply(subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1257
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1258
(* ALT *)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1259
apply(subst (asm) (2) CPTpre_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1260
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1261
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1262
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1263
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1264
apply(case_tac "s' = []")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1265
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1266
apply(drule_tac x="v1" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1267
apply(drule meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1268
apply(auto simp add: CPTpre_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1269
apply(subst (asm) val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1270
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1271
apply(subst (asm) val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1272
apply(erule exE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1273
apply(subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1274
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1275
apply(subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1276
apply(rule_tac x="0#p" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1277
apply(rule val_ord_ALTI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1278
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1279
using Posix1(2) apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1280
using val_ord_ex1_def apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1281
apply(subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1282
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1283
apply (simp add: Posix1(2) val_ord_shorterI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1284
apply(subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1285
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1286
apply(case_tac "s' = []")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1287
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1288
apply(subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1289
apply(rule_tac x="[0]" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1290
apply(subst val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1291
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1292
apply(simp add: Pos_empty)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1293
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1294
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1295
apply (smt inlen_bigger)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1296
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1297
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1298
apply(simp add: pflat_len_simps)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1299
using Posix1(2) apply auto[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1300
apply(rule ballI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1301
apply(rule impI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1302
apply(case_tac "q = []")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1303
using Posix1(2) apply auto[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1304
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1305
apply(rule val_ord_shorterI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1306
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1307
apply (simp add: Posix1(2))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1308
(* ALT RIGHT *)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1309
apply(subst (asm) (2) CPTpre_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1310
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1311
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1312
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1313
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1314
apply(case_tac "s' = []")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1315
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1316
apply (simp add: L_flat_Prf1 Prf_CPrf)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1317
apply(subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1318
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1319
apply(rule val_ord_shorterI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1320
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1321
apply (simp add: Posix1(2))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1322
apply(case_tac "s' = []")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1323
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1324
apply(drule_tac x="v2a" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1325
apply(drule meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1326
apply(auto simp add: CPTpre_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1327
apply(subst (asm) val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1328
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1329
apply(subst (asm) val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1330
apply(erule exE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1331
apply(subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1332
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1333
apply(subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1334
apply(rule_tac x="1#p" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1335
apply(rule val_ord_ALTI2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1336
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1337
using Posix1(2) apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1338
apply (simp add: val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1339
apply(subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1340
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1341
apply(rule val_ord_shorterI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1342
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1343
apply (simp add: Posix1(2))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1344
(* STAR empty case *)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1345
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1346
apply(subst (asm) CPTpre_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1347
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1348
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1349
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1350
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1351
apply (simp add: val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1352
(* STAR non-empty case *)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1353
apply(subst (asm) (3) CPTpre_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1354
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1355
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1356
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1357
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1358
apply (simp add: val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1359
apply(rule val_ord_shorterI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1360
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1361
apply(case_tac "s' = []")
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1362
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1363
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1364
apply (simp add: val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1365
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1366
apply(rule val_ord_shorterI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1367
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1368
apply (metis Posix1(2) append_assoc append_eq_conv_conj drop_all flat.simps(7) flat_Stars length_append not_less)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1369
apply(drule_tac x="va" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1370
apply(drule meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1371
apply(auto simp add: CPTpre_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1372
apply (smt L.simps(6) L_flat_Prf1 Prf_CPrf append_eq_append_conv2 flat_Stars self_append_conv)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1373
apply (subst (asm) (2) val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1374
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1375
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1376
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1377
apply(drule_tac x="Stars vsa" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1378
apply(drule meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1379
apply(auto simp add: CPTpre_def)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1380
apply (simp add: Posix1(2))
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1381
apply (subst (asm) val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1382
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1383
apply (subst (asm) val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1384
apply(erule exE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1385
apply (subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1386
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1387
apply (subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1388
apply(case_tac p)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1389
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1390
apply (metis Posix1(2) flat_Stars not_less_iff_gr_or_eq pflat_len_simps(7) same_append_eq val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1391
using Posix1(2) val_ord_STARI2 apply fastforce
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1392
apply(simp add: val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1393
apply (subst (asm) val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1394
apply(erule exE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1395
apply (subst val_ord_ex1_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1396
apply(rule disjI1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1397
apply (subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1398
by (metis Posix1(2) flat.simps(7) flat_Stars val_ord_STARI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1399
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1400
lemma Posix_val_ord_stronger:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1401
  assumes "s \<in> r \<rightarrow> v1" "v2 \<in> CPT r s" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1402
  shows "v1 :\<sqsubseteq>val v2"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1403
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1404
apply(rule_tac Posix_val_ord)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1405
apply(assumption)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1406
apply(simp add: CPTpre_def CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1407
done
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1408
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1409
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1410
lemma STAR_val_ord:
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1411
  assumes "Stars (v1 # vs1) \<sqsubset>val (Suc p # ps) Stars (v2 # vs2)" "flat v1 = flat v2"
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1412
  shows "(Stars vs1) \<sqsubset>val (p # ps) (Stars vs2)"
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1413
using assms(1,2)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1414
apply -
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1415
apply(simp(no_asm)  add: val_ord_def)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1416
apply(rule conjI)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1417
apply(simp add: val_ord_def)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1418
apply(rule conjI)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1419
apply(simp add: val_ord_def)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1420
apply(auto simp add: pflat_len_simps pflat_len_Stars_simps2)[1]
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1421
apply(rule ballI)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1422
apply(rule impI)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1423
apply(simp add: val_ord_def pflat_len_simps pflat_len_Stars_simps2 intlen_append)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1424
apply(clarify)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1425
apply(case_tac q)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1426
apply(simp add: val_ord_def pflat_len_simps pflat_len_Stars_simps2 intlen_append)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1427
apply(clarify)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1428
apply(erule disjE)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1429
prefer 2
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1430
apply(drule_tac x="Suc a # list" in bspec)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1431
apply(simp)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1432
apply(drule mp)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1433
apply(simp)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1434
apply(simp add: val_ord_def pflat_len_simps pflat_len_Stars_simps2 intlen_append)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1435
apply(drule_tac x="Suc a # list" in bspec)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1436
apply(simp)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1437
apply(drule mp)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1438
apply(simp)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1439
apply(simp add: val_ord_def pflat_len_simps pflat_len_Stars_simps2 intlen_append)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1440
done
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1441
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1442
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1443
lemma Posix_val_ord_reverse:
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1444
  assumes "s \<in> r \<rightarrow> v1" 
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1445
  shows "\<not>(\<exists>v2 \<in> CPT r s. v2 :\<sqsubset>val v1)"
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1446
using assms
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1447
by (metis Posix_val_ord_stronger less_irrefl val_ord_def 
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1448
    val_ord_ex1_def val_ord_ex_def val_ord_ex_trans)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1449
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1450
thm Posix.intros(6)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1451
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1452
inductive Prop :: "rexp \<Rightarrow> val list \<Rightarrow> bool"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1453
where
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1454
  "Prop r []"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1455
| "\<lbrakk>Prop r vs; 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1456
    \<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = concat (map flat vs) \<and> flat v @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))\<rbrakk> 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1457
   \<Longrightarrow> Prop r (v # vs)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1458
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1459
lemma STAR_val_ord_ex:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1460
  assumes "Stars (v # vs1) :\<sqsubset>val Stars (v # vs2)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1461
  shows "Stars vs1 :\<sqsubset>val Stars vs2"
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1462
using assms
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1463
apply(subst (asm) val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1464
apply(erule exE)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1465
apply(case_tac p)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1466
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1467
apply(simp add: val_ord_def pflat_len_simps intlen_append)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1468
apply(subst val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1469
apply(rule_tac x="[]" in exI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1470
apply(simp add: val_ord_def pflat_len_simps Pos_empty)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1471
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1472
apply(case_tac a)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1473
apply(clarify)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1474
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1475
using STAR_val_ord val_ord_ex_def apply blast
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1476
apply(auto simp add: pflat_len_Stars_simps2 val_ord_def pflat_len_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1477
done
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1478
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1479
lemma STAR_val_ord_exI:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1480
  assumes "Stars vs1 :\<sqsubset>val Stars vs2" "flat (Stars vs1) = flat (Stars vs2)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1481
  shows "Stars (vs @ vs1) :\<sqsubset>val Stars (vs @ vs2)"
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1482
using assms
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1483
apply(induct vs)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1484
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1485
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1486
apply(simp add: val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1487
apply(erule exE)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1488
apply(case_tac p)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1489
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1490
apply(rule_tac x="[]" in exI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1491
apply(simp add: val_ord_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1492
apply(auto simp add: pflat_len_simps intlen_append)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1493
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1494
apply(rule_tac x="Suc aa#list" in exI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1495
apply(rule val_ord_STARI2)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1496
apply(simp)
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1497
apply(simp)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1498
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1499
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1500
lemma STAR_val_ord_ex_append:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1501
  assumes "Stars (vs @ vs1) :\<sqsubset>val Stars (vs @ vs2)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1502
  shows "Stars vs1 :\<sqsubset>val Stars vs2"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1503
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1504
apply(induct vs)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1505
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1506
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1507
apply(drule STAR_val_ord_ex)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1508
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1509
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1510
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1511
lemma STAR_val_ord_ex_append_eq:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1512
  assumes "flat (Stars vs1) = flat (Stars vs2)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1513
  shows "Stars (vs @ vs1) :\<sqsubset>val Stars (vs @ vs2) \<longleftrightarrow> Stars vs1 :\<sqsubset>val Stars vs2"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1514
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1515
apply(rule_tac iffI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1516
apply(erule STAR_val_ord_ex_append)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1517
apply(rule STAR_val_ord_exI)
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1518
apply(auto)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1519
done
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1520
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1521
lemma Posix_STARI:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1522
  assumes "\<forall>v \<in> set vs. flat v \<noteq> [] \<and> (flat v) \<in> r \<rightarrow> v" "Prop r vs"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1523
  shows "flat (Stars vs) \<in> STAR r \<rightarrow> Stars vs"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1524
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1525
apply(induct vs arbitrary: r)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1526
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1527
apply(rule Posix.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1528
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1529
apply(rule Posix.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1530
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1531
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1532
apply(erule Prop.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1533
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1534
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1535
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1536
apply(erule Prop.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1537
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1538
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1539
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1540
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1541
lemma CPrf_stars:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1542
  assumes "\<Turnstile> Stars vs : STAR r"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1543
  shows "\<forall>v \<in> set vs. flat v \<noteq> [] \<and> \<Turnstile> v : r"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1544
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1545
apply(induct vs)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1546
apply(auto)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1547
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1548
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1549
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1550
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1551
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1552
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1553
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1554
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1555
done
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1556
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1557
definition PT :: "rexp \<Rightarrow> string \<Rightarrow> val set"
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1558
where "PT r s \<equiv> {v. flat v = s \<and> \<turnstile> v : r}"
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1559
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1560
lemma exists:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1561
  assumes "s \<in> (L r)\<star>"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1562
  shows "\<exists>vs. concat (map flat vs) = s \<and> \<turnstile> Stars vs : STAR r"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1563
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1564
apply(drule_tac Star_string)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1565
apply(auto)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1566
by (metis L_flat_Prf2 Prf_Stars Star_val)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1567
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1568
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1569
lemma val_ord_Posix_Stars:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1570
  assumes "(Stars vs) \<in> CPT (STAR r) (flat (Stars vs))" "\<forall>v \<in> set vs. flat v \<in> r \<rightarrow> v"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1571
  and "\<not>(\<exists>vs2 \<in> PT (STAR r) (flat (Stars vs)). vs2 :\<sqsubset>val (Stars vs))"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1572
  shows "(flat (Stars vs)) \<in> (STAR r) \<rightarrow> Stars vs" 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1573
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1574
apply(induct vs)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1575
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1576
apply(rule Posix.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1577
apply(simp (no_asm))
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1578
apply(rule Posix.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1579
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1580
apply(auto simp add: CPT_def PT_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1581
defer
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1582
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1583
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1584
apply(auto simp add: CPT_def PT_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1585
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1586
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1587
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1588
apply(auto simp add: CPT_def PT_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1589
apply(erule Prf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1590
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1591
apply (metis CPrf_stars Cons_eq_map_conv Posix_CPT Posix_STAR2 Posix_val_ord_reverse list.exhaust list.set_intros(2) map_idI val.distinct(25))
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1592
apply(clarify)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1593
apply(drule_tac x="Stars (a#v#vsa)" in spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1594
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1595
apply(drule mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1596
apply (meson CPrf_stars Prf.intros(7) Prf_CPrf list.set_intros(1))
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1597
apply(subst (asm) (2) val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1598
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1599
apply (metis flat.simps(7) flat_Stars intlen.cases not_less_iff_gr_or_eq pflat_len_simps(7) val_ord_STARI2 val_ord_def val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1600
apply(auto simp add: CPT_def PT_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1601
using CPrf_stars apply auto[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1602
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1603
apply(auto simp add: CPT_def PT_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1604
apply(subgoal_tac "\<exists>vA. flat vA = flat a @ s\<^sub>3 \<and> \<turnstile> vA : r")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1605
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1606
apply (meson L_flat_Prf2)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1607
apply(subgoal_tac "\<exists>vB. flat (Stars vB) = s\<^sub>4 \<and> \<turnstile> (Stars vB) : (STAR r)")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1608
apply(clarify)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1609
apply(drule_tac x="Stars (vA # vB)" in spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1610
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1611
apply(drule mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1612
using Prf.intros(7) apply blast
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1613
apply(subst (asm) (2) val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1614
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1615
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1616
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1617
using exists apply blast
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1618
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1619
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1620
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1621
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1622
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1623
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1624
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1625
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1626
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1627
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1628
apply(clarify)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1629
apply(rotate_tac 3)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1630
apply(erule Prf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1631
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1632
apply (metis CPrf_stars intlen.cases less_irrefl list.set_intros(1) val_ord_def val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1633
apply(drule_tac x="Stars (v#va#vsb)" in spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1634
apply(drule mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1635
apply (simp add: Posix1a Prf.intros(7))
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1636
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1637
apply(subst (asm) (2) val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1638
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1639
apply (metis flat.simps(7) flat_Stars intlen.cases not_less_iff_gr_or_eq pflat_len_simps(7) val_ord_STARI2 val_ord_def val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1640
proof -
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1641
  fix a :: val and vsa :: "val list" and s\<^sub>3 :: "char list" and vA :: val and vB :: "val list"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1642
  assume a1: "s\<^sub>3 \<noteq> []"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1643
  assume a2: "s\<^sub>3 @ concat (map flat vB) = concat (map flat vsa)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1644
  assume a3: "flat vA = flat a @ s\<^sub>3"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1645
  assume a4: "\<forall>p. \<not> Stars (vA # vB) \<sqsubset>val p Stars (a # vsa)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1646
  have f5: "\<And>n cs. drop n (cs::char list) = [] \<or> n < length cs"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1647
    by (meson drop_eq_Nil not_less)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1648
  have f6: "\<not> length (flat vA) \<le> length (flat a)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1649
    using a3 a1 by simp
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1650
  have "flat (Stars (a # vsa)) = flat (Stars (vA # vB))"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1651
    using a3 a2 by simp
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1652
  then show False
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1653
    using f6 f5 a4 by (metis (full_types) drop_eq_Nil val_ord_STARI val_ord_ex_def val_ord_shorterI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1654
qed
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1655
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1656
lemma Prf_Stars_append:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1657
  assumes "\<turnstile> Stars vs1 : STAR r" "\<turnstile> Stars vs2 : STAR r"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1658
  shows "\<turnstile> Stars (vs1 @ vs2) : STAR r"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1659
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1660
apply(induct vs1 arbitrary: vs2)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1661
apply(auto intro: Prf.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1662
apply(erule Prf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1663
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1664
apply(auto intro: Prf.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1665
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1666
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1667
lemma CPrf_Stars_appendE:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1668
  assumes "\<Turnstile> Stars (vs1 @ vs2) : STAR r"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1669
  shows "\<Turnstile> Stars vs1 : STAR r \<and> \<Turnstile> Stars vs2 : STAR r" 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1670
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1671
apply(induct vs1 arbitrary: vs2)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1672
apply(auto intro: CPrf.intros)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1673
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1674
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1675
apply(auto intro: CPrf.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1676
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1677
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1678
lemma val_ord_Posix:
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1679
  assumes "v1 \<in> CPT r s" "\<not>(\<exists>v2 \<in> PT r s. v2 :\<sqsubset>val v1)"
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1680
  shows "s \<in> r \<rightarrow> v1" 
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1681
using assms
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1682
apply(induct r arbitrary: s v1)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1683
apply(auto simp add: CPT_def PT_def)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1684
apply(erule CPrf.cases)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1685
apply(simp_all)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1686
(* ONE *)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1687
apply(auto simp add: CPT_def)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1688
apply(erule CPrf.cases)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1689
apply(simp_all)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1690
apply(rule Posix.intros)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1691
(* CHAR *)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1692
apply(auto simp add: CPT_def)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1693
apply(erule CPrf.cases)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1694
apply(simp_all)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1695
apply(rule Posix.intros)
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1696
prefer 2
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1697
(* ALT *)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1698
apply(auto simp add: CPT_def PT_def)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1699
apply(erule CPrf.cases)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1700
apply(simp_all)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1701
apply(rule Posix.intros)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1702
apply(drule_tac x="flat v1a" in meta_spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1703
apply(drule_tac x="v1a" in meta_spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1704
apply(drule meta_mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1705
apply(simp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1706
apply(drule meta_mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1707
apply(auto)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1708
apply(drule_tac x="Left v2" in spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1709
apply(simp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1710
apply(drule mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1711
apply(rule Prf.intros)
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1712
apply(simp)
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1713
apply (meson val_ord_ALTI val_ord_ex_def)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1714
apply(assumption)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1715
(* ALT Right *)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1716
apply(auto simp add: CPT_def)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1717
apply(rule Posix.intros)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1718
apply(rotate_tac 1)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1719
apply(drule_tac x="flat v2" in meta_spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1720
apply(drule_tac x="v2" in meta_spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1721
apply(drule meta_mp)
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1722
apply(simp)
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1723
apply(drule meta_mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1724
apply(auto)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1725
apply(drule_tac x="Right v2a" in spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1726
apply(simp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1727
apply(drule mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1728
apply(rule Prf.intros)
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1729
apply(simp)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1730
apply(subst (asm) (2) val_ord_ex_def)
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1731
apply(erule exE)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1732
apply(drule val_ord_ALTI2)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1733
apply(assumption)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1734
apply(auto simp add: val_ord_ex_def)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1735
apply(assumption)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1736
apply(auto)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1737
apply(subgoal_tac "\<exists>v2'. flat v2' = flat v2 \<and> \<turnstile> v2' : r1a")
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1738
apply(clarify)
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1739
apply(drule_tac x="Left v2'" in spec)
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1740
apply(simp)
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1741
apply(drule mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1742
apply(rule Prf.intros)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1743
apply(assumption)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1744
apply(simp add: val_ord_ex_def)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1745
apply(subst (asm) (3) val_ord_def)
246
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1746
apply(simp)
23657fad2017 updated
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
  1747
apply(simp add: pflat_len_simps)
247
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1748
apply(drule_tac x="[0]" in spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1749
apply(simp add: pflat_len_simps Pos_empty)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1750
apply(drule mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1751
apply (smt inlen_bigger)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1752
apply(erule disjE)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1753
apply blast
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1754
apply auto[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1755
apply (meson L_flat_Prf2)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1756
(* SEQ *)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1757
apply(auto simp add: CPT_def)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1758
apply(erule CPrf.cases)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1759
apply(simp_all)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1760
apply(rule Posix.intros)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1761
apply(drule_tac x="flat v1a" in meta_spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1762
apply(drule_tac x="v1a" in meta_spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1763
apply(drule meta_mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1764
apply(simp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1765
apply(drule meta_mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1766
apply(auto)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1767
apply(auto simp add: PT_def)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1768
apply(drule_tac x="Seq v2a v2" in spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1769
apply(simp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1770
apply(drule mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1771
apply (simp add: Prf.intros(1) Prf_CPrf)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1772
using val_ord_SEQI val_ord_ex_def apply fastforce
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1773
apply(assumption)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1774
apply(rotate_tac 1)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1775
apply(drule_tac x="flat v2" in meta_spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1776
apply(drule_tac x="v2" in meta_spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1777
apply(simp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1778
apply(auto)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1779
apply(drule meta_mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1780
apply(auto)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1781
apply(auto simp add: PT_def)[1]
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1782
apply(drule_tac x="Seq v1a v2a" in spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1783
apply(simp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1784
apply(drule mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1785
apply (simp add: Prf.intros(1) Prf_CPrf)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1786
apply (meson val_ord_SEQI2 val_ord_ex_def)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1787
apply(assumption)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1788
(* SEQ side condition *)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1789
apply(auto simp add: PT_def)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1790
apply(subgoal_tac "\<exists>vA. flat vA = flat v1a @ s\<^sub>3 \<and> \<turnstile> vA : r1a")
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1791
prefer 2
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1792
apply (meson L_flat_Prf2)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1793
apply(subgoal_tac "\<exists>vB. flat vB = s\<^sub>4 \<and> \<turnstile> vB : r2a")
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1794
prefer 2
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1795
apply (meson L_flat_Prf2)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1796
apply(clarify)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1797
apply(drule_tac x="Seq vA vB" in spec)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1798
apply(simp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1799
apply(drule mp)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1800
apply (simp add: Prf.intros(1))
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1801
apply(subst (asm) (3) val_ord_ex_def)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1802
apply (metis append_Nil2 append_assoc append_eq_conv_conj flat.simps(5) length_append not_add_less1 not_less_iff_gr_or_eq val_ord_SEQI val_ord_ex_def val_ord_shorterI)
f35753951058 updated
Christian Urban <urbanc@in.tum.de>
parents: 246
diff changeset
  1803
(* STAR *)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1804
apply(auto simp add: CPT_def)[1]
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1805
apply(erule CPrf.cases)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1806
apply(simp_all)[6]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1807
using Posix_STAR2 apply blast
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1808
apply(clarify)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1809
apply(rule val_ord_Posix_Stars)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1810
apply(auto simp add: CPT_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1811
apply (simp add: CPrf.intros(7))
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1812
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1813
apply(drule_tac x="flat v" in meta_spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1814
apply(drule_tac x="v" in meta_spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1815
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1816
apply(drule meta_mp)
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1817
apply(auto)[1]
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1818
apply(drule_tac x="Stars (v2#vs)" in spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1819
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1820
apply(drule mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1821
using Prf.intros(7) Prf_CPrf apply blast
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1822
apply(subst (asm) (2) val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1823
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1824
using val_ord_STARI val_ord_ex_def apply fastforce
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1825
apply(assumption)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1826
apply(drule_tac x="flat va" in meta_spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1827
apply(drule_tac x="va" in meta_spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1828
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1829
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1830
using CPrf_stars apply blast
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1831
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1832
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1833
apply(subgoal_tac "\<exists>pre post. vs = pre @ [va] @ post")
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1834
prefer 2
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1835
apply (metis append_Cons append_Nil in_set_conv_decomp_first)
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1836
apply(clarify)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1837
apply(drule_tac x="Stars (v#(pre @ [v2] @ post))" in spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1838
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1839
apply(drule mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1840
apply(rule Prf.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1841
apply (simp add: Prf_CPrf)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1842
apply(rule Prf_Stars_append)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1843
apply(drule CPrf_Stars_appendE)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1844
apply(auto simp add: Prf_CPrf)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1845
apply (metis CPrf_Stars_appendE CPrf_stars Prf_CPrf Prf_Stars list.set_intros(2) set_ConsD)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1846
apply(subgoal_tac "\<not> Stars ([v] @ pre @ v2 # post) :\<sqsubset>val Stars ([v] @ pre @ va # post)")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1847
apply(subst (asm) STAR_val_ord_ex_append_eq)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1848
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1849
apply(subst (asm) STAR_val_ord_ex_append_eq)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1850
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1851
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1852
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1853
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1854
apply(simp)
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1855
apply(simp add: val_ord_ex_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1856
apply(erule exE)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1857
apply(rotate_tac 6)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1858
apply(drule_tac x="0#p" in spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1859
apply (simp add: val_ord_STARI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1860
apply(auto simp add: PT_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1861
done
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1862
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1863
inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<preceq>_ _" [100, 100, 100] 100)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1864
where
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1865
  C2: "v1 \<preceq>r1 v1' \<Longrightarrow> (Seq v1 v2) \<preceq>(SEQ r1 r2) (Seq v1' v2')" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1866
| C1: "v2 \<preceq>r2 v2' \<Longrightarrow> (Seq v1 v2) \<preceq>(SEQ r1 r2) (Seq v1 v2')" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1867
| A1: "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<preceq>(ALT r1 r2) (Left v1)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1868
| A2: "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<preceq>(ALT r1 r2) (Right v2)"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1869
| A3: "v2 \<preceq>r2 v2' \<Longrightarrow> (Right v2) \<preceq>(ALT r1 r2) (Right v2')"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1870
| A4: "v1 \<preceq>r1 v1' \<Longrightarrow> (Left v1) \<preceq>(ALT r1 r2) (Left v1')"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1871
| K1: "flat (Stars (v # vs)) = [] \<Longrightarrow> (Stars []) \<preceq>(STAR r) (Stars (v # vs))"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1872
| K2: "flat (Stars (v # vs)) \<noteq> [] \<Longrightarrow> (Stars (v # vs)) \<preceq>(STAR r) (Stars [])"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1873
| K3: "v1 \<preceq>r v2 \<Longrightarrow> (Stars (v1 # vs1)) \<preceq>(STAR r) (Stars (v2 # vs2))"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1874
| K4: "(Stars vs1) \<preceq>(STAR r) (Stars vs2) \<Longrightarrow> (Stars (v # vs1)) \<preceq>(STAR r) (Stars (v # vs2))"
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1875
(*| MY1: "Void \<preceq>ONE Void" 
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1876
| MY2: "(Char c) \<preceq>(CHAR c) (Char c)" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1877
| MY3: "(Stars []) \<preceq>(STAR r) (Stars [])" 
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1878
*)
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1879
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1880
(*
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1881
lemma ValOrd_refl: 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1882
  assumes "\<turnstile> v : r" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1883
  shows "v \<preceq>r v"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1884
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1885
apply(induct r rule: Prf.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1886
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1887
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1888
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1889
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1890
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1891
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1892
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1893
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1894
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1895
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1896
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1897
done
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1898
*)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1899
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1900
lemma ValOrd_irrefl: 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1901
  assumes "\<turnstile> v : r"  "v \<preceq>r v" 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1902
  shows "False"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1903
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1904
apply(induct v r rule: Prf.induct)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1905
apply(erule ValOrd.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1906
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1907
apply(erule ValOrd.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1908
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1909
apply(erule ValOrd.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1910
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1911
apply(erule ValOrd.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1912
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1913
apply(erule ValOrd.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1914
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1915
apply(erule ValOrd.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1916
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1917
apply(erule ValOrd.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1918
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1919
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1920
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1921
lemma prefix_sprefix:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1922
  shows "xs \<sqsubseteq>pre ys \<longleftrightarrow> (xs = ys \<or> xs \<sqsubset>spre ys)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1923
apply(auto simp add: sprefix_list_def prefix_list_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1924
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1925
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1926
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1927
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1928
lemma Posix_CPT2:
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1929
  assumes "v1 \<preceq>r v2" "flat v2 \<sqsubseteq>pre flat v1"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1930
  shows "v1 :\<sqsubset>val v2" 
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  1931
using assms
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1932
apply(induct v1 r v2 arbitrary: rule: ValOrd.induct)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1933
prefer 3
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1934
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1935
apply(auto simp add: prefix_sprefix)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1936
apply(rule val_ord_spre)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1937
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1938
prefer 3
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1939
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1940
apply(auto simp add: prefix_sprefix)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1941
apply(auto simp add: val_ord_ex_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1942
apply(rule_tac x="[0]" in exI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1943
apply(auto simp add: val_ord_def Pos_empty pflat_len_simps)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1944
apply (smt inlen_bigger)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1945
apply(rule val_ord_spre)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1946
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1947
prefer 3
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1948
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1949
apply(auto simp add: prefix_sprefix)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1950
using val_ord_ALTI2 val_ord_ex_def apply fastforce
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1951
apply(rule val_ord_spre)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1952
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1953
prefer 3
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1954
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1955
apply(auto simp add: prefix_sprefix)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1956
using val_ord_ALTI val_ord_ex_def apply fastforce
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1957
apply(rule val_ord_spre)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1958
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1959
(* SEQ case *)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1960
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1961
apply(auto simp add: prefix_sprefix)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1962
apply(auto simp add: append_eq_append_conv2)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1963
apply(case_tac "us = []")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1964
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1965
apply(auto simp add: val_ord_ex1_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1966
apply (metis flat.simps(5) val_ord_SEQI val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1967
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1968
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1969
apply(case_tac "us = []")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1970
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1971
apply (metis flat.simps(5) val_ord_SEQI val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1972
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1973
apply(rule disjI2)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1974
apply (metis append_self_conv prefix_list_def sprefix_list_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1975
apply(auto simp add: val_ord_ex_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1976
apply (metis append_assoc flat.simps(5) val_ord_SEQI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1977
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1978
apply(sugoal_ tac "")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1979
thm val_ord_SEQI
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1980
apply(rule val_ord_SEQI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1981
thm val_ord_SEQI
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1982
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1983
apply(case_tac "us
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1984
apply(case_tac "v1 = v1'")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1985
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1986
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1987
apply(auto simp add: val_ord_ex1_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1988
apply(auto simp add: val_ord_ex_def)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1989
apply(rule_tac x="[0]" in exI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1990
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1991
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1992
apply(rule val_ord_spre)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1993
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1994
prefer 3
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1995
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1996
using val_ord_ex1_def val_ord_spre apply auto[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1997
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1998
apply(erule disjE)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  1999
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2000
apply(subst val_ord_ex1_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2001
apply(rule disjI1)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2002
apply(rule val_ord_spre)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2003
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2004
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2005
apply(simp add: append_eq_append_conv2)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2006
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2007
apply(case_tac "us = []")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2008
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2009
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2010
apply(auto simp add: prefix_sprefix)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2011
apply(subst (asm) val_ord_ex1_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2012
apply(erule disjE)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2013
apply(subst val_ord_ex1_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2014
apply(rule disjI1)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2015
apply (metis flat.simps(5) val_ord_SEQI val_ord_ex_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2016
apply(clarify)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2017
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2018
apply(subgoal_tac "flat v1' \<sqsubset>spre flat v1")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2019
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2020
apply (simp add: prefix_list_def sprefix_list_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2021
apply(drule val_ord_spre)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2022
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
  2023
apply(rule val_ord_sprefixI)
245
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2024
apply(erule ValOrd.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2025
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2026
apply(erule ValOrd.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2027
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2028
apply(erule ValOrd.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2029
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2030
apply(erule ValOrd.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2031
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2032
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2033
(* HERE *)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2034
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2035
apply(subst val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2036
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2037
apply(drule_tac x="v2a" in  meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2038
apply(rotate_tac 5)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2039
apply(drule_tac x="v2'" in  meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2040
apply(rule_tac x="0#p" in exI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2041
apply(rule val_ord_SEQI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2042
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2043
apply(drule_tac r="r1a" in  val_ord_SEQ)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2044
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2045
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2046
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2047
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2048
lemma Posix_CPT:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2049
  assumes "v1 :\<sqsubset>val v2" "v1 \<in> CPT r s" "v2 \<in> CPT r s"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2050
  shows "v1 \<preceq>r v2" 
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2051
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2052
apply(induct r arbitrary: v1 v2 s rule: rexp.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2053
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2054
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2055
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2056
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2057
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2058
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2059
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2060
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2061
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2062
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2063
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2064
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2065
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2066
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2067
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2068
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2069
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2070
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2071
(*SEQ case *)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2072
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2073
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2074
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2075
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2076
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2077
apply(erule CPrf.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2078
apply(simp_all)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2079
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2080
thm val_ord_SEQ
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2081
apply(drule_tac r="r1a" in  val_ord_SEQ)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2082
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2083
using Prf_CPrf apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2084
using Prf_CPrf apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2085
apply(erule disjE)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2086
apply(rule C2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2087
prefer 2
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2088
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2089
apply(rule C1)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2090
apply blast
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2091
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2092
apply(simp add: append_eq_append_conv2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2093
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2094
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2095
apply(drule_tac x="v1a" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2096
apply(rotate_tac 8)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2097
apply(drule_tac x="v1b" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2098
apply(rotate_tac 8)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2099
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2100
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2101
(* HERE *)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2102
apply(subst (asm) (3) val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2103
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2104
apply(subst (asm) val_ord_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2105
apply(clarify)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2106
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2107
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2108
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2109
apply(simp add: val_ord_ex_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2110
oops
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2111
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2112
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2113
lemma ValOrd_trans:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2114
  assumes "x \<preceq>r y" "y \<preceq>r z"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2115
  and "x \<in> CPT r s" "y \<in> CPT r s" "z \<in> CPT r s"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2116
  shows "x \<preceq>r z"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2117
using assms
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2118
apply(induct x r y arbitrary: s z rule: ValOrd.induct)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2119
apply(rotate_tac 2)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2120
apply(erule ValOrd.cases)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2121
apply(simp_all)[13]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2122
apply(rule ValOrd.intros)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2123
apply(drule_tac x="s" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2124
apply(drule_tac x="v1'a" in meta_spec)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2125
apply(drule_tac meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2126
apply(simp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2127
apply(drule_tac meta_mp)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2128
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2129
oops
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2130
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2131
lemma ValOrd_preorder:
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2132
  "preorder_on (CPT r s) {(v1, v2). v1 \<preceq>r v2 \<and> v1 \<in> (CPT r s) \<and> v2 \<in> (CPT r s)}"
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2133
apply(simp add: preorder_on_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2134
apply(rule conjI)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2135
apply(simp add: refl_on_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2136
apply(auto)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2137
apply(rule ValOrd_refl)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2138
apply(simp add: CPT_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2139
apply(rule Prf_CPrf)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2140
apply(auto)[1]
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2141
apply(simp add: trans_def)
b16702bb6242 updated
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
  2142
apply(auto)
148
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2143
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2144
definition ValOrdEq :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<ge>_ _" [100, 100, 100] 100)
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2145
where 
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2146
  "v\<^sub>1 \<ge>r v\<^sub>2 \<equiv> v\<^sub>1 = v\<^sub>2 \<or> (v\<^sub>1 >r v\<^sub>2 \<and> flat v\<^sub>1 = flat v\<^sub>2)"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2147
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2148
(*
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2149
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2150
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2151
inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100)
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2152
where
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2153
  "v2 \<succ>r2 v2' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1 v2')" 
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2154
| "\<lbrakk>v1 \<succ>r1 v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2155
| "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2156
| "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2157
| "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2158
| "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2159
| "Void \<succ>EMPTY Void"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2160
| "(Char c) \<succ>(CHAR c) (Char c)"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2161
| "flat (Stars (v # vs)) = [] \<Longrightarrow> (Stars []) \<succ>(STAR r) (Stars (v # vs))"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2162
| "flat (Stars (v # vs)) \<noteq> [] \<Longrightarrow> (Stars (v # vs)) \<succ>(STAR r) (Stars [])"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2163
| "\<lbrakk>v1 \<succ>r v2; v1 \<noteq> v2\<rbrakk> \<Longrightarrow> (Stars (v1 # vs1)) \<succ>(STAR r) (Stars (v2 # vs2))"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2164
| "(Stars vs1) \<succ>(STAR r) (Stars vs2) \<Longrightarrow> (Stars (v # vs1)) \<succ>(STAR r) (Stars (v # vs2))"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2165
| "(Stars []) \<succ>(STAR r) (Stars [])"
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2166
*)
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2167
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2168
154
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2169
section {* Bit-Encodings *}
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2170
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2171
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2172
fun 
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2173
  code :: "val \<Rightarrow> rexp \<Rightarrow> bool list"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2174
where
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2175
  "code Void ONE = []"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2176
| "code (Char c) (CHAR d) = []"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2177
| "code (Left v) (ALT r1 r2) = False # (code v r1)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2178
| "code (Right v) (ALT r1 r2) = True # (code v r2)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2179
| "code (Seq v1 v2) (SEQ r1 r2) = (code v1 r1) @ (code v2 r2)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2180
| "code (Stars []) (STAR r) = [True]"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2181
| "code (Stars (v # vs)) (STAR r) =  False # (code v r) @ code (Stars vs) (STAR r)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2182
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2183
fun 
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2184
  Stars_add :: "val \<Rightarrow> val \<Rightarrow> val"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2185
where
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2186
  "Stars_add v (Stars vs) = Stars (v # vs)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2187
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2188
function
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2189
  decode' :: "bool list \<Rightarrow> rexp \<Rightarrow> (val * bool list)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2190
where
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2191
  "decode' ds ZERO = (Void, [])"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2192
| "decode' ds ONE = (Void, ds)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2193
| "decode' ds (CHAR d) = (Char d, ds)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2194
| "decode' [] (ALT r1 r2) = (Void, [])"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2195
| "decode' (False # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r1 in (Left v, ds'))"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2196
| "decode' (True # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r2 in (Right v, ds'))"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2197
| "decode' ds (SEQ r1 r2) = (let (v1, ds') = decode' ds r1 in
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2198
                             let (v2, ds'') = decode' ds' r2 in (Seq v1 v2, ds''))"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2199
| "decode' [] (STAR r) = (Void, [])"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2200
| "decode' (True # ds) (STAR r) = (Stars [], ds)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2201
| "decode' (False # ds) (STAR r) = (let (v, ds') = decode' ds r in
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2202
                                    let (vs, ds'') = decode' ds' (STAR r) 
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2203
                                    in (Stars_add v vs, ds''))"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2204
by pat_completeness auto
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2205
204
cd9e40280784 added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
  2206
termination
cd9e40280784 added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
  2207
apply(size_change)
cd9e40280784 added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
  2208
oops
cd9e40280784 added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
  2209
154
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2210
term "inv_image (measure(%cs. size cs) <*lex*> measure(%s. size s)) (%(ds,r). (r,ds))"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2211
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2212
lemma decode'_smaller:
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2213
  assumes "decode'_dom (ds, r)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2214
  shows "length (snd (decode' ds r)) \<le> length ds"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2215
using assms
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2216
apply(induct ds r)
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2217
apply(auto simp add: decode'.psimps split: prod.split)
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2218
using dual_order.trans apply blast
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2219
by (meson dual_order.trans le_SucI)
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2220
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2221
termination "decode'"  
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2222
apply(relation "inv_image (measure(%cs. size cs) <*lex*> measure(%s. size s)) (%(ds,r). (r,ds))") 
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2223
apply(auto dest!: decode'_smaller)
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2224
by (metis less_Suc_eq_le snd_conv)
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2225
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2226
fun 
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2227
  decode :: "bool list \<Rightarrow> rexp \<Rightarrow> val option"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2228
where
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2229
  "decode ds r = (let (v, ds') = decode' ds r 
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2230
                  in (if ds' = [] then Some v else None))"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2231
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2232
lemma decode'_code:
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2233
  assumes "\<turnstile> v : r"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2234
  shows "decode' ((code v r) @ ds) r = (v, ds)"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2235
using assms
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2236
by (induct v r arbitrary: ds) (auto)
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2237
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2238
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2239
lemma decode_code:
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2240
  assumes "\<turnstile> v : r"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2241
  shows "decode (code v r) r = Some v"
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2242
using assms decode'_code[of _ _ "[]"]
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2243
by auto
2de3cf684ba0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 148
diff changeset
  2244
159
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2245
datatype arexp =
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2246
  AZERO
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2247
| AONE "bool list"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2248
| ACHAR "bool list" char
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2249
| ASEQ "bool list" arexp arexp
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2250
| AALT "bool list" arexp arexp
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2251
| ASTAR "bool list" arexp
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2252
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2253
fun fuse :: "bool list \<Rightarrow> arexp \<Rightarrow> arexp" where
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2254
  "fuse bs AZERO = AZERO"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2255
| "fuse bs (AONE cs) = AONE (bs @ cs)" 
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2256
| "fuse bs (ACHAR cs c) = ACHAR (bs @ cs) c"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2257
| "fuse bs (AALT cs r1 r2) = AALT (bs @ cs) r1 r2"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2258
| "fuse bs (ASEQ cs r1 r2) = ASEQ (bs @ cs) r1 r2"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2259
| "fuse bs (ASTAR cs r) = ASTAR (bs @ cs) r"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2260
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2261
fun internalise :: "rexp \<Rightarrow> arexp" where
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2262
  "internalise ZERO = AZERO"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2263
| "internalise ONE = AONE []"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2264
| "internalise (CHAR c) = ACHAR [] c"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2265
| "internalise (ALT r1 r2) = AALT [] (fuse [False] (internalise r1)) 
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2266
                                     (fuse [True]  (internalise r2))"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2267
| "internalise (SEQ r1 r2) = ASEQ [] (internalise r1) (internalise r2)"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2268
| "internalise (STAR r) = ASTAR [] (internalise r)"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2269
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2270
fun retrieve :: "arexp \<Rightarrow> val \<Rightarrow> bool list" where
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2271
  "retrieve (AONE bs) Void = bs"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2272
| "retrieve (ACHAR bs c) (Char d) = bs"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2273
| "retrieve (AALT bs r1 r2) (Left v) = bs @ retrieve r1 v"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2274
| "retrieve (AALT bs r1 r2) (Right v) = bs @ retrieve r2 v"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2275
| "retrieve (ASEQ bs r1 r2) (Seq v1 v2) = bs @ retrieve r1 v1 @ retrieve r2 v2"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2276
| "retrieve (ASTAR bs r) (Stars []) = bs @ [True]"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2277
| "retrieve (ASTAR bs r) (Stars (v#vs)) = 
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2278
     bs @ [False] @ retrieve r v @ retrieve (ASTAR [] r) (Stars vs)"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2279
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2280
fun
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2281
 anullable :: "arexp \<Rightarrow> bool"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2282
where
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2283
  "anullable (AZERO) = False"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2284
| "anullable (AONE bs) = True"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2285
| "anullable (ACHAR bs c) = False"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2286
| "anullable (AALT bs r1 r2) = (anullable r1 \<or> anullable r2)"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2287
| "anullable (ASEQ bs r1 r2) = (anullable r1 \<and> anullable r2)"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2288
| "anullable (ASTAR bs r) = True"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2289
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2290
fun 
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2291
  amkeps :: "arexp \<Rightarrow> bool list"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2292
where
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2293
  "amkeps(AONE bs) = bs"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2294
| "amkeps(ASEQ bs r1 r2) = bs @ (amkeps r1) @ (amkeps r2)"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2295
| "amkeps(AALT bs r1 r2) = (if anullable(r1) then bs @ (amkeps r1) else bs @ (amkeps r2))"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2296
| "amkeps(ASTAR bs r) = bs @ [True]"
148
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2297
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2298
159
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2299
fun
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2300
 ader :: "char \<Rightarrow> arexp \<Rightarrow> arexp"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2301
where
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2302
  "ader c (AZERO) = AZERO"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2303
| "ader c (AONE bs) = AZERO"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2304
| "ader c (ACHAR bs d) = (if c = d then AONE bs else AZERO)"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2305
| "ader c (AALT bs r1 r2) = AALT bs (ader c r1) (ader c r2)"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2306
| "ader c (ASEQ bs r1 r2) = 
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2307
     (if anullable r1
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2308
      then AALT bs (ASEQ [] (ader c r1) r2) (fuse (amkeps r1) (ader c r2))
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2309
      else ASEQ bs (ader c r1) r2)"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2310
| "ader c (ASTAR bs r) = ASEQ bs (fuse [False] (ader c r)) (ASTAR [] r)"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2311
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2312
lemma
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2313
  assumes "\<turnstile> v : der c r"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2314
  shows "Some (injval r c v) = decode (retrieve (ader c (internalise r)) v) r"
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2315
using assms
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2316
apply(induct c r arbitrary: v rule: der.induct)
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2317
apply(simp_all)
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2318
apply(erule Prf_elims)
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2319
apply(erule Prf_elims)
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2320
apply(case_tac "c = d")
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2321
apply(simp)
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2322
apply(erule Prf_elims)
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2323
apply(simp)
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2324
apply(simp)
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2325
apply(erule Prf_elims)
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2326
apply(auto split: prod.splits)[1]
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2327
oops
940530087f30 updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  2328
148
702ed601349b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2329
end